From e2bb6d8d530fb36fac3712149620fee8d14449de Mon Sep 17 00:00:00 2001 From: Ralf Jung <jung@mpi-sws.org> Date: Fri, 4 May 2018 10:23:55 +0200 Subject: [PATCH] all SBI must be a COFE --- theories/bi/derived_laws_sbi.v | 6 ++---- theories/bi/interface.v | 3 +++ theories/proofmode/coq_tactics.v | 2 +- theories/proofmode/ltac_tactics.v | 3 +-- theories/tests/proofmode.v | 8 ++++---- 5 files changed, 11 insertions(+), 11 deletions(-) diff --git a/theories/bi/derived_laws_sbi.v b/theories/bi/derived_laws_sbi.v index 2acb046bb..43c87d15f 100644 --- a/theories/bi/derived_laws_sbi.v +++ b/theories/bi/derived_laws_sbi.v @@ -210,8 +210,6 @@ Section löb. 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). @@ -412,7 +410,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 `{Cofe PROP} P Q : Timeless Q → Timeless (P → Q). +Global Instance impl_timeless P Q : Timeless Q → Timeless (P → Q). Proof. rewrite /Timeless=> HQ. rewrite later_false_em. apply or_mono, impl_intro_l; first done. @@ -425,7 +423,7 @@ Proof. intros; rewrite /Timeless except_0_sep later_sep; auto using sep_mono. Qed. -Global Instance wand_timeless `{Cofe PROP} P Q : Timeless Q → Timeless (P -∗ Q). +Global Instance wand_timeless 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 82c3ae2dd..71ebd6a09 100644 --- a/theories/bi/interface.v +++ b/theories/bi/interface.v @@ -229,6 +229,7 @@ Structure sbi := Sbi { sbi_internal_eq : ∀ A : ofeT, A → A → sbi_car; sbi_later : sbi_car → sbi_car; sbi_ofe_mixin : OfeMixin sbi_car; + sbi_cofe : Cofe (OfeT sbi_car sbi_ofe_mixin); sbi_bi_mixin : BiMixin sbi_entails sbi_emp sbi_pure sbi_and sbi_or sbi_impl sbi_forall sbi_exist sbi_sep sbi_wand sbi_persistently; sbi_sbi_mixin : SbiMixin sbi_entails sbi_pure sbi_or sbi_impl @@ -247,6 +248,8 @@ Canonical Structure sbi_ofeC. Coercion sbi_bi (PROP : sbi) : bi := {| bi_ofe_mixin := sbi_ofe_mixin PROP; bi_bi_mixin := sbi_bi_mixin PROP |}. Canonical Structure sbi_bi. +Global Instance sbi_cofe' (PROP : sbi) : Cofe PROP. +Proof. apply sbi_cofe. Qed. Arguments sbi_car : simpl never. Arguments sbi_dist : simpl never. diff --git a/theories/proofmode/coq_tactics.v b/theories/proofmode/coq_tactics.v index 4d35abc14..d8b0f59fd 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 `{Cofe PROP} Δ Δ' i Q : +Lemma tac_löb Δ Δ' 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 dae0fef3d..49999e5b6 100644 --- a/theories/proofmode/ltac_tactics.v +++ b/theories/proofmode/ltac_tactics.v @@ -1659,8 +1659,7 @@ 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 _ _ _ _); - [iSolveTC || fail "iLöb: PROP is not a Cofe" - |reflexivity || fail "iLöb: spatial context not empty, this should not happen" + [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 8e9eecbdd..1a8abb751 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} `{Cofe PROP}. +Context {PROP : sbi}. 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 (Hn). - iEval (rewrite (Nat.add_comm x y) // Hn). + iIntros (H). + iEval (rewrite (Nat.add_comm x y) // H). done. Qed. -- GitLab