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

more consistent naming

parent b7a767a7
No related branches found
No related tags found
No related merge requests found
...@@ -206,11 +206,11 @@ Lemma ofe_fun_validI {A} {B : A → ucmraT} (g : ofe_fun B) : ✓ g ⊣⊢ ∀ i ...@@ -206,11 +206,11 @@ Lemma ofe_fun_validI {A} {B : A → ucmraT} (g : ofe_fun B) : ✓ g ⊣⊢ ∀ i
Proof. exact: uPred_primitive.ofe_fun_validI. Qed. Proof. exact: uPred_primitive.ofe_fun_validI. Qed.
(** Consistency/soundness statement *) (** Consistency/soundness statement *)
Lemma soundness_pure φ : bi_emp_valid (PROP:=uPredI M) φ φ. Lemma pure_soundness φ : bi_emp_valid (PROP:=uPredI M) φ φ.
Proof. apply soundness_pure. Qed. Proof. apply pure_soundness. Qed.
Lemma soundness_later P : bi_emp_valid ( P) bi_emp_valid P. Lemma later_soundness P : bi_emp_valid ( P) bi_emp_valid P.
Proof. apply soundness_later. Qed. Proof. apply later_soundness. Qed.
End restate. End restate.
(** See [derived.v] for the version for basic updates. *) (** See [derived.v] for the version for basic updates. *)
......
...@@ -92,7 +92,7 @@ Global Instance uPred_ownM_sep_homomorphism : ...@@ -92,7 +92,7 @@ Global Instance uPred_ownM_sep_homomorphism :
Proof. split; [split; try apply _|]. apply ownM_op. apply ownM_unit'. Qed. Proof. split; [split; try apply _|]. apply ownM_op. apply ownM_unit'. Qed.
(** Consistency/soundness statement *) (** Consistency/soundness statement *)
Lemma soundness_bupd_plain P `{!Plain P} : bi_emp_valid (|==> P) bi_emp_valid P. Lemma bupd_plain_soundness P `{!Plain P} : bi_emp_valid (|==> P) bi_emp_valid P.
Proof. Proof.
eapply bi_emp_valid_mono. etrans; last exact: bupd_plainly. apply bupd_mono'. eapply bi_emp_valid_mono. etrans; last exact: bupd_plainly. apply bupd_mono'.
apply: plain. apply: plain.
...@@ -101,8 +101,8 @@ Qed. ...@@ -101,8 +101,8 @@ Qed.
Corollary soundness φ n : (▷^n φ : uPred M)%I φ. Corollary soundness φ n : (▷^n φ : uPred M)%I φ.
Proof. Proof.
induction n as [|n IH]=> /=. induction n as [|n IH]=> /=.
- apply soundness_pure. - apply pure_soundness.
- intros H. by apply IH, soundness_later. - intros H. by apply IH, later_soundness.
Qed. Qed.
Corollary consistency_modal n : ¬ (▷^n False : uPred M)%I. Corollary consistency_modal n : ¬ (▷^n False : uPred M)%I.
......
...@@ -62,7 +62,7 @@ Qed. ...@@ -62,7 +62,7 @@ Qed.
Lemma fupd_plain_soundness `{!invPreG Σ} E1 E2 (P: iProp Σ) `{!Plain P}: Lemma fupd_plain_soundness `{!invPreG Σ} E1 E2 (P: iProp Σ) `{!Plain P}:
( `{Hinv: !invG Σ}, bi_emp_valid (|={E1,E2}=> P)) bi_emp_valid P. ( `{Hinv: !invG Σ}, bi_emp_valid (|={E1,E2}=> P)) bi_emp_valid P.
Proof. Proof.
iIntros (Hfupd). apply soundness_later. iMod wsat_alloc as (Hinv) "[Hw HE]". iIntros (Hfupd). apply later_soundness. iMod wsat_alloc as (Hinv) "[Hw HE]".
iAssert (|={,E2}=> P)%I as "H". iAssert (|={,E2}=> P)%I as "H".
{ iMod fupd_intro_mask'; last iApply Hfupd. done. } { iMod fupd_intro_mask'; last iApply Hfupd. done. }
rewrite uPred_fupd_eq /uPred_fupd_def. rewrite uPred_fupd_eq /uPred_fupd_def.
......
...@@ -801,10 +801,10 @@ Lemma ofe_fun_validI {A} {B : A → ucmraT} (g : ofe_fun B) : ✓ g ⊣⊢ ∀ i ...@@ -801,10 +801,10 @@ Lemma ofe_fun_validI {A} {B : A → ucmraT} (g : ofe_fun B) : ✓ g ⊣⊢ ∀ i
Proof. by unseal. Qed. Proof. by unseal. Qed.
(** Consistency/soundness statement *) (** Consistency/soundness statement *)
Lemma soundness_pure φ : (True φ ) φ. Lemma pure_soundness φ : (True φ ) φ.
Proof. unseal=> -[H]. by apply (H 0 ε); eauto using ucmra_unit_validN. Qed. Proof. unseal=> -[H]. by apply (H 0 ε); eauto using ucmra_unit_validN. Qed.
Lemma soundness_later P : (True P) (True P). Lemma later_soundness P : (True P) (True P).
Proof. Proof.
unseal=> -[HP]; split=> n x Hx _. unseal=> -[HP]; split=> n x Hx _.
apply uPred_mono with n ε; eauto using ucmra_unit_leastN. apply uPred_mono with n ε; eauto using ucmra_unit_leastN.
......
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