diff --git a/theories/base_logic/upred.v b/theories/base_logic/upred.v index 7f5e40560a456112729483e01c9278f46f5dce68..22c3b7f9aa688ebb08ad1a8265638dfb548db174 100644 --- a/theories/base_logic/upred.v +++ b/theories/base_logic/upred.v @@ -508,9 +508,9 @@ Proof. - (* (P ⊢ Q) → â–· P ⊢ â–· Q *) intros P Q. unseal=> HP; split=>-[|n] x ??; [done|apply HP; eauto using cmra_validN_S]. - - (* (â–· P → P) ⊢ P *) - intros P. unseal; split=> n x ? HP; induction n as [|n IH]; [by apply HP|]. - apply HP, IH, uPred_mono with (S n) x; eauto using cmra_validN_S. + - (* P ⊢ â–· P *) + intros P. unseal; split=> n x ? HP. destruct n; first done. simpl. + apply uPred_mono with (S n) x; eauto using cmra_validN_S. - (* (∀ a, â–· Φ a) ⊢ â–· ∀ a, Φ a *) intros A Φ. unseal; by split=> -[|n] x. - (* (â–· ∃ a, Φ a) ⊢ â–· False ∨ (∃ a, â–· Φ a) *) diff --git a/theories/bi/derived_laws_bi.v b/theories/bi/derived_laws_bi.v index ff5a64e6ff49c498804b07d5f1d451ac00b20f83..7b8a79d49fc874cc604a195e0627b505666f6bbe 100644 --- a/theories/bi/derived_laws_bi.v +++ b/theories/bi/derived_laws_bi.v @@ -122,6 +122,12 @@ Proof. + apply and_intro; first done. by apply pure_intro. + rewrite -EQ impl_elim_r. done. Qed. +Lemma entails_impl_True P Q : (P ⊢ Q) ↔ (True ⊢ (P → Q)). +Proof. + rewrite entails_eq_True. split. + - by intros <-. + - intros. apply (anti_symm (⊢)); last done. apply True_intro. +Qed. Lemma and_mono P P' Q Q' : (P ⊢ Q) → (P' ⊢ Q') → P ∧ P' ⊢ Q ∧ Q'. Proof. auto. Qed. diff --git a/theories/bi/derived_laws_sbi.v b/theories/bi/derived_laws_sbi.v index 147a2f7397a2f27020d323b553c3a675d13c27ed..2acb046bba01cbe73c1dbb42f7d954000f611501 100644 --- a/theories/bi/derived_laws_sbi.v +++ b/theories/bi/derived_laws_sbi.v @@ -158,12 +158,6 @@ Global Instance later_flip_mono' : Proper (flip (⊢) ==> flip (⊢)) (@sbi_later PROP). Proof. intros P Q; apply later_mono. Qed. -Lemma later_intro P : P ⊢ â–· P. -Proof. - rewrite -(and_elim_l (â–· P)%I P) -(löb (â–· P ∧ P)%I). - apply impl_intro_l. by rewrite {1}(and_elim_r (â–· P)%I). -Qed. - Lemma later_True : â–· True ⊣⊢ True. Proof. apply (anti_symm (⊢)); auto using later_intro. Qed. Lemma later_emp `{!BiAffine PROP} : â–· emp ⊣⊢ emp. @@ -209,6 +203,45 @@ Proof. intros. by rewrite /Persistent -later_persistently {1}(persistent P). Qed Global Instance later_absorbing P : Absorbing P → Absorbing (â–· P). Proof. intros ?. by rewrite /Absorbing -later_absorbingly absorbing. Qed. +Section löb. + (* Proof following https://en.wikipedia.org/wiki/L%C3%B6b's_theorem#Proof_of_L%C3%B6b's_theorem *) + Definition flöb_pre (P Ψ : PROP) : PROP := (â–· Ψ → P)%I. + + Local Instance flöb_pre_contractive P : Contractive (flöb_pre P). + Proof. solve_contractive. Qed. + + Context `{Cofe PROP}. + + Definition flöb (P : PROP) := fixpoint (flöb_pre P). + + Lemma weak_löb P : (â–· P ⊢ P) → (True ⊢ P). + Proof. + set (Ψ := flöb P). assert (Ψ ⊣⊢ (â–· Ψ → P)) as HΨ. + { exact: fixpoint_unfold. } + intros HP. rewrite -HP. + assert (Ψ ⊢ (â–· Ψ → P)) as HΨ'%entails_impl_True by by rewrite -HΨ. + rewrite ->(later_intro (Ψ → _))%I in HΨ'. + rewrite ->later_impl in HΨ'. + rewrite ->later_impl in HΨ'. + assert (â–· Ψ ⊢ P) as HΨP. + { rewrite -HP. rewrite -(idemp (∧) (â–· Ψ))%I {2}(later_intro (â–· Ψ))%I. + apply impl_elim_l', entails_impl_True. done. } + rewrite -HΨP HΨ -2!later_intro. + apply (entails_impl_True _ P). done. + Qed. + + Lemma löb P : (â–· P → P) ⊢ P. + Proof. + apply entails_impl_True, weak_löb. apply impl_intro_r. + rewrite -{2}(idemp (∧) (â–· P → P))%I. + rewrite {2}(later_intro (â–· P → P))%I. + rewrite later_impl. + rewrite assoc impl_elim_l. + rewrite impl_elim_r. done. + Qed. + +End löb. + (* Iterated later modality *) Global Instance laterN_ne m : NonExpansive (@sbi_laterN PROP m). Proof. induction m; simpl. by intros ???. solve_proper. Qed. @@ -379,7 +412,7 @@ Proof. intros; rewrite /Timeless except_0_and later_and; auto. Qed. Global Instance or_timeless P Q : Timeless P → Timeless Q → Timeless (P ∨ Q). Proof. intros; rewrite /Timeless except_0_or later_or; auto. Qed. -Global Instance impl_timeless P Q : Timeless Q → Timeless (P → Q). +Global Instance impl_timeless `{Cofe PROP} P Q : Timeless Q → Timeless (P → Q). Proof. rewrite /Timeless=> HQ. rewrite later_false_em. apply or_mono, impl_intro_l; first done. @@ -392,7 +425,7 @@ Proof. intros; rewrite /Timeless except_0_sep later_sep; auto using sep_mono. Qed. -Global Instance wand_timeless P Q : Timeless Q → Timeless (P -∗ Q). +Global Instance wand_timeless `{Cofe PROP} P Q : Timeless Q → Timeless (P -∗ Q). Proof. rewrite /Timeless=> HQ. rewrite later_false_em. apply or_mono, wand_intro_l; first done. diff --git a/theories/bi/interface.v b/theories/bi/interface.v index b90fdfd99dd4835864b8968ca058b45a277574f2..82c3ae2ddc2973a299167de28722dbe2a59b81af 100644 --- a/theories/bi/interface.v +++ b/theories/bi/interface.v @@ -146,7 +146,7 @@ Section bi_mixin. sbi_mixin_later_eq_2 {A : ofeT} (x y : A) : â–· (x ≡ y) ⊢ Next x ≡ Next y; sbi_mixin_later_mono P Q : (P ⊢ Q) → â–· P ⊢ â–· Q; - sbi_mixin_löb P : (â–· P → P) ⊢ P; + sbi_mixin_later_intro P : P ⊢ â–· P; sbi_mixin_later_forall_2 {A} (Φ : A → PROP) : (∀ a, â–· Φ a) ⊢ â–· ∀ a, Φ a; sbi_mixin_later_exist_false {A} (Φ : A → PROP) : @@ -454,8 +454,8 @@ Proof. eapply sbi_mixin_later_eq_2, sbi_sbi_mixin. Qed. Lemma later_mono P Q : (P ⊢ Q) → â–· P ⊢ â–· Q. Proof. eapply sbi_mixin_later_mono, sbi_sbi_mixin. Qed. -Lemma löb P : (â–· P → P) ⊢ P. -Proof. eapply sbi_mixin_löb, sbi_sbi_mixin. Qed. +Lemma later_intro P : P ⊢ â–· P. +Proof. eapply sbi_mixin_later_intro, sbi_sbi_mixin. Qed. Lemma later_forall_2 {A} (Φ : A → PROP) : (∀ a, â–· Φ a) ⊢ â–· ∀ a, Φ a. Proof. eapply sbi_mixin_later_forall_2, sbi_sbi_mixin. Qed. diff --git a/theories/bi/monpred.v b/theories/bi/monpred.v index 828bf4eb0e0d8d4580c752ed1e01b3ced00ecddf..86a74749ea28d0edcf3b1cc9595e72e9e2c14522 100644 --- a/theories/bi/monpred.v +++ b/theories/bi/monpred.v @@ -336,8 +336,7 @@ Proof. - intros A x y. split=> i. by apply bi.later_eq_1. - intros A x y. split=> i. by apply bi.later_eq_2. - intros P Q [?]. split=> i. by apply bi.later_mono. - - intros P. split=> i /=. - setoid_rewrite bi.pure_impl_forall. rewrite /= !bi.forall_elim //. by apply bi.löb. + - intros P. split=> i /=. by apply bi.later_intro. - intros A Ψ. split=> i. by apply bi.later_forall_2. - intros A Ψ. split=> i. by apply bi.later_exist_false. - intros P Q. split=> i. by apply bi.later_sep_1. diff --git a/theories/proofmode/coq_tactics.v b/theories/proofmode/coq_tactics.v index d8b0f59fd72593a90e12ce242c26f9774306d5d0..4d35abc1414cc4d9442283af21860638aec78d84 100644 --- a/theories/proofmode/coq_tactics.v +++ b/theories/proofmode/coq_tactics.v @@ -1396,7 +1396,7 @@ Proof. - by rewrite Hs //= right_id. Qed. -Lemma tac_löb Δ Δ' i Q : +Lemma tac_löb `{Cofe PROP} Δ Δ' i Q : env_spatial_is_nil Δ = true → envs_app true (Esnoc Enil i (â–· Q)%I) Δ = Some Δ' → envs_entails Δ' Q → envs_entails Δ Q. diff --git a/theories/proofmode/ltac_tactics.v b/theories/proofmode/ltac_tactics.v index 49999e5b61006220a477eabb52f4430a30eb3ee8..dae0fef3dee0d2d9d3869ec620a636585085055e 100644 --- a/theories/proofmode/ltac_tactics.v +++ b/theories/proofmode/ltac_tactics.v @@ -1659,7 +1659,8 @@ Tactic Notation "iLöbCore" "as" constr (IH) := refine should use the other unification algorithm, which should not have this issue. *) notypeclasses refine (tac_löb _ _ IH _ _ _ _); - [reflexivity || fail "iLöb: spatial context not empty, this should not happen" + [iSolveTC || fail "iLöb: PROP is not a Cofe" + |reflexivity || fail "iLöb: spatial context not empty, this should not happen" |env_reflexivity || fail "iLöb:" IH "not fresh"|]. Tactic Notation "iLöbRevert" constr(Hs) "with" tactic(tac) := diff --git a/theories/tests/proofmode.v b/theories/tests/proofmode.v index 1a8abb751472fa40f1af592237ecb4793b9199a4..8e9eecbdd482c9be182dc814525386089dc76046 100644 --- a/theories/tests/proofmode.v +++ b/theories/tests/proofmode.v @@ -1,9 +1,9 @@ From iris.proofmode Require Import tactics intro_patterns. From stdpp Require Import gmap hlist. -Set Default Proof Using "Type". +Set Default Proof Using "Type*". Section tests. -Context {PROP : sbi}. +Context {PROP : sbi} `{Cofe PROP}. Implicit Types P Q R : PROP. Lemma demo_0 P Q : â–¡ (P ∨ Q) -∗ (∀ x, ⌜x = 0⌠∨ ⌜x = 1âŒ) → (Q ∨ P). @@ -381,8 +381,8 @@ Proof. iIntros (Hφ). iAssert ⌜φâŒ%I with "[%]" as "$"; auto with iFrame. Qe Lemma test_iEval x y : ⌜ (y + x)%nat = 1 ⌠-∗ ⌜ S (x + y) = 2%nat ⌠: PROP. Proof. - iIntros (H). - iEval (rewrite (Nat.add_comm x y) // H). + iIntros (Hn). + iEval (rewrite (Nat.add_comm x y) // Hn). done. Qed.