diff --git a/theories/base_logic/bi.v b/theories/base_logic/bi.v index df9d498f107866292e4578c23506a031da790650..3b93a6e656c40babf0959b839bacbf404c0b456a 100644 --- a/theories/base_logic/bi.v +++ b/theories/base_logic/bi.v @@ -204,13 +204,13 @@ Lemma discrete_fun_validI {A} {B : A → ucmraT} (g : discrete_fun B) : ✓ g Proof. exact: uPred_primitive.discrete_fun_validI. Qed. (** Consistency/soundness statement *) -Lemma pure_soundness φ : bi_emp_valid (PROP:=uPredI M) ⌜ φ ⌠→ φ. +Lemma pure_soundness φ : (⊢@{uPredI M} ⌜ φ âŒ) → φ. Proof. apply pure_soundness. Qed. -Lemma internal_eq_soundness {A : ofeT} (x y : A) : (True ⊢ x ≡ y) → x ≡ y. +Lemma internal_eq_soundness {A : ofeT} (x y : A) : (⊢@{uPredI M} x ≡ y) → x ≡ y. Proof. apply internal_eq_soundness. Qed. -Lemma later_soundness P : bi_emp_valid (â–· P) → bi_emp_valid P. +Lemma later_soundness P : (⊢ â–· P) → ⊢ P. Proof. apply later_soundness. Qed. (** See [derived.v] for a similar soundness result for basic updates. *) End restate. diff --git a/theories/base_logic/derived.v b/theories/base_logic/derived.v index e300eb8e527d15d8efa45d0a4488a4b64c9e14fb..01dff355c205c86e21914865c8f0bc720b3442af 100644 --- a/theories/base_logic/derived.v +++ b/theories/base_logic/derived.v @@ -84,7 +84,7 @@ 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. +Lemma bupd_plain_soundness P `{!Plain P} : (⊢ |==> P) → ⊢ P. Proof. eapply bi_emp_valid_mono. etrans; last exact: bupd_plainly. apply bupd_mono'. apply: plain. diff --git a/theories/base_logic/lib/fancy_updates.v b/theories/base_logic/lib/fancy_updates.v index 3be5ca307c88c23f0b9a2209284818a3d5568bdd..6317a748799cc64d71e14594b65af22834a507d3 100644 --- a/theories/base_logic/lib/fancy_updates.v +++ b/theories/base_logic/lib/fancy_updates.v @@ -59,8 +59,8 @@ Proof. by iFrame. Qed. -Lemma fupd_plain_soundness `{!invPreG Σ} E1 E2 (P: iProp Σ) `{!Plain P}: - (∀ `{Hinv: !invG Σ}, bi_emp_valid (|={E1,E2}=> P)) → bi_emp_valid P. +Lemma fupd_plain_soundness `{!invPreG Σ} E1 E2 (P: iProp Σ) `{!Plain P} : + (∀ `{Hinv: !invG Σ}, ⊢ |={E1,E2}=> P) → ⊢ P. Proof. iIntros (Hfupd). apply later_soundness. iMod wsat_alloc as (Hinv) "[Hw HE]". iAssert (|={⊤,E2}=> P)%I as "H". diff --git a/theories/bi/embedding.v b/theories/bi/embedding.v index 3eaa91784100ffdc0e859c0b2e8fffe90df61878..bbf87f95dfbc71b2f9277ad45185f87e5776cf6d 100644 --- a/theories/bi/embedding.v +++ b/theories/bi/embedding.v @@ -15,7 +15,7 @@ Record BiEmbedMixin (PROP1 PROP2 : bi) `(Embed PROP1 PROP2) := { bi_embed_mixin_ne : NonExpansive (embed (A:=PROP1) (B:=PROP2)); bi_embed_mixin_mono : Proper ((⊢) ==> (⊢)) (embed (A:=PROP1) (B:=PROP2)); bi_embed_mixin_emp_valid_inj (P : PROP1) : - bi_emp_valid (PROP:=PROP2) ⎡P⎤ → bi_emp_valid P; + (⊢@{PROP2} ⎡P⎤) → ⊢ P; bi_embed_mixin_emp_2 : emp ⊢@{PROP2} ⎡emp⎤; bi_embed_mixin_impl_2 (P Q : PROP1) : (⎡P⎤ → ⎡Q⎤) ⊢@{PROP2} ⎡P → Q⎤; diff --git a/theories/proofmode/class_instances_bi.v b/theories/proofmode/class_instances_bi.v index f5adfc0d3aab58ff2ef5166ade45dbb60028c11c..fd6b4d4d530ac5c1a7f65d6f757068fce9fffc11 100644 --- a/theories/proofmode/class_instances_bi.v +++ b/theories/proofmode/class_instances_bi.v @@ -31,7 +31,7 @@ Implicit Types P Q R : PROP. Implicit Types mP : option PROP. (** AsEmpValid *) -Global Instance as_emp_valid_emp_valid P : AsEmpValid0 (bi_emp_valid P) P | 0. +Global Instance as_emp_valid_emp_valid P : AsEmpValid0 (⊢ P) P | 0. Proof. by rewrite /AsEmpValid. Qed. Global Instance as_emp_valid_entails P Q : AsEmpValid0 (P ⊢ Q) (P -∗ Q). Proof. split. apply bi.entails_wand. apply bi.wand_entails. Qed. @@ -986,10 +986,10 @@ Qed. (* These instances must be used only after [into_forall_wand_pure] and [into_forall_wand_pure] above. *) Global Instance into_forall_wand P Q : - IntoForall (P -∗ Q) (λ _ : bi_emp_valid P, Q) | 10. + IntoForall (P -∗ Q) (λ _ : ⊢ P, Q) | 10. Proof. rewrite /IntoForall. apply forall_intro=><-. rewrite emp_wand //. Qed. Global Instance into_forall_impl `{!BiAffine PROP} P Q : - IntoForall (P → Q) (λ _ : bi_emp_valid P, Q) | 10. + IntoForall (P → Q) (λ _ : ⊢ P, Q) | 10. Proof. rewrite /IntoForall. apply forall_intro=><-. rewrite -True_emp True_impl //. Qed. (** FromForall *) diff --git a/theories/proofmode/classes.v b/theories/proofmode/classes.v index a71e3527760550471e291c74be94512c03aaf9ec..d9603ff412841d81aeb129c68a314f3285f4287b 100644 --- a/theories/proofmode/classes.v +++ b/theories/proofmode/classes.v @@ -492,7 +492,7 @@ Hint Mode IntoEmbed + + + ! - : typeclass_instances. [Hint Mode - ! -], but the fact that Coq ignores primitive projections for hints modes would make this fail.*) Class AsEmpValid {PROP : bi} (φ : Prop) (P : PROP) := - as_emp_valid : φ ↔ bi_emp_valid P. + as_emp_valid : φ ↔ ⊢ P. Arguments AsEmpValid {_} _%type _%I. Class AsEmpValid0 {PROP : bi} (φ : Prop) (P : PROP) := as_emp_valid_0 : AsEmpValid φ P. @@ -500,10 +500,10 @@ Arguments AsEmpValid0 {_} _%type _%I. Existing Instance as_emp_valid_0 | 0. Lemma as_emp_valid_1 (φ : Prop) {PROP : bi} (P : PROP) `{!AsEmpValid φ P} : - φ → bi_emp_valid P. + φ → ⊢ P. Proof. by apply as_emp_valid. Qed. Lemma as_emp_valid_2 (φ : Prop) {PROP : bi} (P : PROP) `{!AsEmpValid φ P} : - bi_emp_valid P → φ. + (⊢ P) → φ. Proof. by apply as_emp_valid. Qed. (* Input: [P]; Outputs: [N], diff --git a/theories/proofmode/coq_tactics.v b/theories/proofmode/coq_tactics.v index ccf395f4d6d4aa34e8c5a22d3988171d5b5106d0..8d38f7f972ed3fbd6f1e3f00a4f89d653ff62792 100644 --- a/theories/proofmode/coq_tactics.v +++ b/theories/proofmode/coq_tactics.v @@ -15,7 +15,7 @@ Implicit Types Δ : envs PROP. Implicit Types P Q : PROP. (** * Starting and stopping the proof mode *) -Lemma tac_start P : envs_entails (Envs Enil Enil 1) P → bi_emp_valid P. +Lemma tac_start P : envs_entails (Envs Enil Enil 1) P → ⊢ P. Proof. rewrite envs_entails_eq !of_envs_eq /=. rewrite intuitionistically_True_emp left_id=><-. @@ -477,7 +477,7 @@ Proof. by rewrite -(entails_wand P) // intuitionistically_emp emp_wand. Qed. -Definition IntoEmpValid (φ : Type) (P : PROP) := φ → bi_emp_valid P. +Definition IntoEmpValid (φ : Type) (P : PROP) := φ → ⊢ P. (** These lemmas are [Defined] because the guardedness checker must see through them. See https://gitlab.mpi-sws.org/iris/iris/issues/274. For the same reason, their bodies use as little automation as possible. *) @@ -489,7 +489,7 @@ Proof. rewrite /IntoEmpValid => Hφ Hi1 Hi2. apply Hi1, Hi2, Hφ. Defined. Lemma into_emp_valid_forall {A} (φ : A → Type) P x : IntoEmpValid (φ x) P → IntoEmpValid (∀ x : A, φ x) P. Proof. rewrite /IntoEmpValid => Hi1 Hi2. apply Hi1, Hi2. Defined. -Lemma into_emp_valid_proj φ P : IntoEmpValid φ P → φ → bi_emp_valid P. +Lemma into_emp_valid_proj φ P : IntoEmpValid φ P → φ → ⊢ P. Proof. intros HP. apply HP. Defined. (** When called by the proof mode, the proof of [P] is produced by calling diff --git a/theories/proofmode/ltac_tactics.v b/theories/proofmode/ltac_tactics.v index 493d2ebbbbda70e1602ff4b408f23912b3cdf4da..418f25c0db18669065a873a724773286758e2099 100644 --- a/theories/proofmode/ltac_tactics.v +++ b/theories/proofmode/ltac_tactics.v @@ -747,7 +747,7 @@ Tactic Notation "iPoseProofCoreHyp" constr(H) "as" constr(Hnew) := end. (* The tactic [iIntoEmpValid] tactic "imports a Coq lemma into the proofmode", -i.e. it solves a goal [IntoEmpValid ψ ?Q]. The argument [ψ] must be of the +i.e., it solves a goal [IntoEmpValid ψ ?Q]. The argument [ψ] must be of the following shape: [∀ (x_1 : A_1) .. (x_n : A_n), φ] @@ -756,14 +756,14 @@ for which we have an instance [AsEmpValid φ ?Q]. Examples of such [φ]s are -- [bi_emp_valid P], in which case [?Q] is unified with [P]. +- [⊢ P], in which case [?Q] is unified with [P]. - [P1 ⊢ P2], in which case [?Q] is unified with [P1 -∗ P2]. - [P1 ⊣⊢ P2], in which case [?Q] is unified with [P1 ↔ P2]. The tactic instantiates each dependent argument [x_i : A_i] with an evar, and generates a goal [A_i] for each non-dependent argument [x_i : A_i]. -For example, if goal is [bi_emp_valid (∀ x, P x → R1 x ⊢ R2 x) ?Q], then the +For example, if goal is [IntoEmpValid (∀ x, P x → R1 x ⊢ R2 x) ?Q], then the [iIntoEmpValid] tactic generates an evar [?x], a subgoal [P ?x], and unifies [?Q] with [R1 ?x -∗ R2 ?x]. *) Ltac iIntoEmpValid_go := first @@ -772,7 +772,7 @@ Ltac iIntoEmpValid_go := first [(*goal for [φ] *)|iIntoEmpValid_go] |(* Case [∀ x : A, φ] *) notypeclasses refine (into_emp_valid_forall _ _ _ _); iIntoEmpValid_go - |(* Case [P ⊢ Q], [P ⊣⊢ Q], [bi_emp_valid P] *) + |(* Case [P ⊢ Q], [P ⊣⊢ Q], [⊢ P] *) notypeclasses refine (into_emp_valid_here _ _ _)]. Ltac iIntoEmpValid := @@ -3191,7 +3191,7 @@ Tactic Notation "iAccu" := (** Automation *) Hint Extern 0 (_ ⊢ _) => iStartProof : core. -Hint Extern 0 (bi_emp_valid _) => iStartProof : core. +Hint Extern 0 (⊢ _) => iStartProof : core. (* Make sure that by and done solve trivial things in proof mode *) Hint Extern 0 (envs_entails _ _) => iPureIntro; try done : core. diff --git a/theories/si_logic/bi.v b/theories/si_logic/bi.v index d7d9f5aa3d353a0afd7c99652c536bf35b625b93..d613e3f0ec9e4095a1094c005620d3566bb94757 100644 --- a/theories/si_logic/bi.v +++ b/theories/si_logic/bi.v @@ -159,13 +159,13 @@ Proof. done. Qed. Module siProp. Section restate. -Lemma pure_soundness φ : bi_emp_valid (PROP:=siPropI) ⌜ φ ⌠→ φ. +Lemma pure_soundness φ : (⊢@{siPropI} ⌜ φ âŒ) → φ. Proof. apply pure_soundness. Qed. -Lemma internal_eq_soundness {A : ofeT} (x y : A) : (True ⊢@{siPropI} x ≡ y) → x ≡ y. +Lemma internal_eq_soundness {A : ofeT} (x y : A) : (⊢@{siPropI} x ≡ y) → x ≡ y. Proof. apply internal_eq_soundness. Qed. -Lemma later_soundness P : bi_emp_valid (PROP:=siPropI) (â–· P) → bi_emp_valid P. +Lemma later_soundness (P : siProp) : (⊢ â–· P) → ⊢ P. Proof. apply later_soundness. Qed. End restate. End siProp.