Skip to content
Snippets Groups Projects
Commit d9498942 authored by Ralf Jung's avatar Ralf Jung
Browse files

uPred entailment commutes over limits

parent 9c55bc2c
No related branches found
No related tags found
No related merge requests found
......@@ -68,7 +68,7 @@ End ofe_mixin.
Hint Extern 1 (_ {_} _) => apply equiv_dist; assumption.
(** Discrete COFEs and Timeless elements *)
(** Discrete OFEs and Timeless elements *)
(* TODO: On paper, We called these "discrete elements". I think that makes
more sense. *)
Class Timeless `{Equiv A, Dist A} (x : A) := timeless y : x {0} y x y.
......
......@@ -172,5 +172,13 @@ Lemma entails_equiv_l (P Q R : uPred M) : (P ⊣⊢ Q) → (Q ⊢ R) → (P ⊢
Proof. by intros ->. Qed.
Lemma entails_equiv_r (P Q R : uPred M) : (P Q) (Q ⊣⊢ R) (P R).
Proof. by intros ? <-. Qed.
Lemma entails_lim (P Q : chain (uPredC M)) :
( n, P n Q n) compl P compl Q.
Proof.
intros Hlim. split. intros n m Hval HP.
eapply uPred_holds_ne, Hlim, HP; eauto using conv_compl.
Qed.
End entails.
End uPred.
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment