diff --git a/theories/base_logic/bi.v b/theories/base_logic/bi.v index b6d0650377ccafa68e80b88602b3e5c1a0af2fac..72c3ea3e31e2fe3fc2e5b651e1dc8f8b4deee30c 100644 --- a/theories/base_logic/bi.v +++ b/theories/base_logic/bi.v @@ -206,13 +206,15 @@ Lemma ofe_fun_validI {A} {B : A → ucmraT} (g : ofe_fun B) : ✓ g ⊣⊢ ∀ i Proof. exact: uPred_primitive.ofe_fun_validI. Qed. (** Consistency/soundness statement *) -Lemma soundness_pure φ : bi_emp_valid (PROP:=uPredI M) ⌜ φ ⌠→ φ. -Proof. apply soundness_pure. Qed. +Lemma pure_soundness φ : bi_emp_valid (PROP:=uPredI M) ⌜ φ ⌠→ φ. +Proof. apply pure_soundness. Qed. -Lemma soundness_later P : bi_emp_valid (â–· P) → bi_emp_valid P. -Proof. apply soundness_later. Qed. +Lemma later_soundness P : bi_emp_valid (â–· P) → bi_emp_valid P. +Proof. apply later_soundness. Qed. +(** See [derived.v] for a similar soundness result for basic updates. *) End restate. + (** New unseal tactic that also unfolds the BI layer. This is used by [base_logic.double_negation]. TODO: Can we get rid of this? *) diff --git a/theories/base_logic/derived.v b/theories/base_logic/derived.v index e5b4462bd614ca47ec2c3469c6d2f7bf3dc3c8ce..0adcc1a5c714ad87056f98e924db35bacaec7877 100644 --- a/theories/base_logic/derived.v +++ b/theories/base_logic/derived.v @@ -92,11 +92,17 @@ Global Instance uPred_ownM_sep_homomorphism : Proof. split; [split; try apply _|]. apply ownM_op. apply ownM_unit'. Qed. (** Consistency/soundness statement *) +Lemma bupd_plain_soundness P `{!Plain P} : bi_emp_valid (|==> P) → bi_emp_valid P. +Proof. + eapply bi_emp_valid_mono. etrans; last exact: bupd_plainly. apply bupd_mono'. + apply: plain. +Qed. + Corollary soundness φ n : (â–·^n ⌜ φ ⌠: uPred M)%I → φ. Proof. induction n as [|n IH]=> /=. - - apply soundness_pure. - - intros H. by apply IH, soundness_later. + - apply pure_soundness. + - intros H. by apply IH, later_soundness. Qed. Corollary consistency_modal n : ¬ (â–·^n False : uPred M)%I. diff --git a/theories/base_logic/lib/fancy_updates.v b/theories/base_logic/lib/fancy_updates.v index 0889c404dea27917a7a67c3585cee8af77f89d0b..6dfaa0d0c3a4fa48a50287bc15751e36b22f12db 100644 --- a/theories/base_logic/lib/fancy_updates.v +++ b/theories/base_logic/lib/fancy_updates.v @@ -59,11 +59,12 @@ Proof. by iFrame. Qed. -Lemma fupd_plain_soundness `{!invPreG Σ} E (P: iProp Σ) `{!Plain P}: - (∀ `{Hinv: !invG Σ}, (|={⊤,E}=> P)%I) → (â–· P)%I. +Lemma fupd_plain_soundness `{!invPreG Σ} E1 E2 (P: iProp Σ) `{!Plain P}: + (∀ `{Hinv: !invG Σ}, bi_emp_valid (|={E1,E2}=> P)) → bi_emp_valid P. Proof. - iIntros (Hfupd). iMod wsat_alloc as (Hinv) "[Hw HE]". - iPoseProof (Hfupd Hinv) as "H". + iIntros (Hfupd). apply later_soundness. iMod wsat_alloc as (Hinv) "[Hw HE]". + iAssert (|={⊤,E2}=> P)%I as "H". + { iMod fupd_intro_mask'; last iApply Hfupd. done. } rewrite uPred_fupd_eq /uPred_fupd_def. iMod ("H" with "[$]") as "[Hw [HE >H']]"; iFrame. Qed. @@ -73,8 +74,8 @@ Lemma step_fupdN_soundness `{!invPreG Σ} φ n : φ. Proof. intros Hiter. - apply (soundness (M:=iResUR Σ) _ (S (S n))); simpl. - apply (fupd_plain_soundness ⊤ _)=> Hinv. + apply (soundness (M:=iResUR Σ) _ (S n)); simpl. + apply (fupd_plain_soundness ⊤ ⊤ _)=> Hinv. iPoseProof (Hiter Hinv) as "H". clear Hiter. destruct n as [|n]. - iApply fupd_plainly_mask_empty. iMod "H" as %?; auto. diff --git a/theories/base_logic/upred.v b/theories/base_logic/upred.v index d9c180eb0b6669536a0de8c947159d2cb48fbfa3..b9c90e9708a24cc6dca41d2c45f6a5786ffde16a 100644 --- a/theories/base_logic/upred.v +++ b/theories/base_logic/upred.v @@ -801,10 +801,10 @@ Lemma ofe_fun_validI {A} {B : A → ucmraT} (g : ofe_fun B) : ✓ g ⊣⊢ ∀ i Proof. by unseal. Qed. (** Consistency/soundness statement *) -Lemma soundness_pure φ : (True ⊢ ⌜ φ âŒ) → φ. +Lemma pure_soundness φ : (True ⊢ ⌜ φ âŒ) → φ. 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. unseal=> -[HP]; split=> n x Hx _. apply uPred_mono with n ε; eauto using ucmra_unit_leastN. diff --git a/theories/program_logic/total_adequacy.v b/theories/program_logic/total_adequacy.v index 9e8ba8f6dbfe6f2be17f8b305b95d6167248ebe3..7d5512c4cd4fb38836bd811af0fd20b2f08e3ef4 100644 --- a/theories/program_logic/total_adequacy.v +++ b/theories/program_logic/total_adequacy.v @@ -123,8 +123,8 @@ Theorem twp_total Σ Λ `{!invPreG Σ} s e σ Φ : stateI σ [] 0 ∗ WP e @ s; ⊤ [{ Φ }])%I) → sn erased_step ([e], σ). (* i.e. ([e], σ) is strongly normalizing *) Proof. - intros Hwp. apply (soundness (M:=iResUR Σ) _ 2); simpl. - apply (fupd_plain_soundness ⊤ _)=> Hinv. + intros Hwp. apply (soundness (M:=iResUR Σ) _ 1); simpl. + apply (fupd_plain_soundness ⊤ ⊤ _)=> Hinv. iMod (Hwp) as (stateI fork_post) "[Hσ H]". iApply (@twptp_total _ _ (IrisG _ _ Hinv stateI fork_post) with "Hσ"). by iApply (@twp_twptp _ _ (IrisG _ _ Hinv stateI fork_post)).