diff --git a/theories/base_logic/bi.v b/theories/base_logic/bi.v index 9db6cf68316ea7e64f869080976387b5bb38bf1f..b6d0650377ccafa68e80b88602b3e5c1a0af2fac 100644 --- a/theories/base_logic/bi.v +++ b/theories/base_logic/bi.v @@ -206,9 +206,11 @@ 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_iter φ n : Nat.iter n sbi_later (⌜ φ ⌠: uPred M)%I → φ. -Proof. exact: uPred_primitive.soundness. Qed. +Lemma soundness_pure φ : bi_emp_valid (PROP:=uPredI M) ⌜ φ ⌠→ φ. +Proof. apply soundness_pure. Qed. +Lemma soundness_later P : bi_emp_valid (â–· P) → bi_emp_valid P. +Proof. apply soundness_later. Qed. End restate. (** New unseal tactic that also unfolds the BI layer. diff --git a/theories/base_logic/derived.v b/theories/base_logic/derived.v index aa2e11301236d328ee97c8216fa81e8a4bacf3b7..e5b4462bd614ca47ec2c3469c6d2f7bf3dc3c8ce 100644 --- a/theories/base_logic/derived.v +++ b/theories/base_logic/derived.v @@ -92,15 +92,18 @@ Global Instance uPred_ownM_sep_homomorphism : Proof. split; [split; try apply _|]. apply ownM_op. apply ownM_unit'. Qed. (** Consistency/soundness statement *) -Lemma soundness φ n : (â–·^n ⌜ φ ⌠: uPred M)%I → φ. -Proof. rewrite laterN_iter. apply soundness_iter. Qed. +Corollary soundness φ n : (â–·^n ⌜ φ ⌠: uPred M)%I → φ. +Proof. + induction n as [|n IH]=> /=. + - apply soundness_pure. + - intros H. by apply IH, soundness_later. +Qed. Corollary consistency_modal n : ¬ (â–·^n False : uPred M)%I. Proof. exact (soundness False n). Qed. Corollary consistency : ¬(False : uPred M)%I. Proof. exact (consistency_modal 0). Qed. - End derived. End uPred. diff --git a/theories/base_logic/upred.v b/theories/base_logic/upred.v index f6052c634c3bf2af8d0d23e21874199c7ff465f4..d9c180eb0b6669536a0de8c947159d2cb48fbfa3 100644 --- a/theories/base_logic/upred.v +++ b/theories/base_logic/upred.v @@ -801,12 +801,14 @@ Lemma ofe_fun_validI {A} {B : A → ucmraT} (g : ofe_fun B) : ✓ g ⊣⊢ ∀ i Proof. by unseal. Qed. (** Consistency/soundness statement *) -Lemma soundness φ n : (True ⊢ iter n uPred_later (⌜ φ âŒ)%I) → φ. +Lemma soundness_pure φ : (True ⊢ ⌜ φ âŒ) → φ. +Proof. unseal=> -[H]. by apply (H 0 ε); eauto using ucmra_unit_validN. Qed. + +Lemma soundness_later P : (True ⊢ â–· P) → (True ⊢ P). Proof. - cut (iter n (@uPred_later M) (⌜ φ âŒ)%I n ε → φ). - { intros help H. eapply help, H; eauto using ucmra_unit_validN. by unseal. } - unseal. induction n as [|n IH]=> H; auto. + unseal=> -[HP]; split=> n x Hx _. + apply uPred_mono with n ε; eauto using ucmra_unit_leastN. + by apply (HP (S n)); eauto using ucmra_unit_validN. Qed. - End primitive. End uPred_primitive.