From d1e5eb2503867e12f3a770196ea4cb828155868d Mon Sep 17 00:00:00 2001 From: Robbert Krebbers <mail@robbertkrebbers.nl> Date: Fri, 5 May 2023 21:11:21 +0200 Subject: [PATCH] Show that for non-step indexed BIs, <pers> can trivially be inhabited. --- iris/base_logic/bi.v | 16 +++++++--- iris/bi/interface.v | 72 +++++++++++++++++++++++++++++++++++--------- iris/bi/monpred.v | 16 ++++++++-- iris/si_logic/bi.v | 17 ++++++++--- tests/heapprop.v | 37 +++++++++++------------ 5 files changed, 113 insertions(+), 45 deletions(-) diff --git a/iris/base_logic/bi.v b/iris/base_logic/bi.v index 5f8257722..9077f0feb 100644 --- a/iris/base_logic/bi.v +++ b/iris/base_logic/bi.v @@ -13,8 +13,7 @@ Local Existing Instance entails_po. Lemma uPred_bi_mixin (M : ucmra) : BiMixin uPred_entails uPred_emp uPred_pure uPred_and uPred_or uPred_impl - (@uPred_forall M) (@uPred_exist M) uPred_sep uPred_wand - uPred_persistently. + (@uPred_forall M) (@uPred_exist M) uPred_sep uPred_wand. Proof. split. - exact: entails_po. @@ -27,7 +26,6 @@ Proof. - exact: exist_ne. - exact: sep_ne. - exact: wand_ne. - - exact: persistently_ne. - exact: pure_intro. - exact: pure_elim'. - exact: and_elim_l. @@ -49,6 +47,15 @@ Proof. - exact: sep_assoc'. - exact: wand_intro_r. - exact: wand_elim_l'. +Qed. + +Lemma uPred_bi_persistently_mixin (M : ucmra) : + BiPersistentlyMixin + uPred_entails uPred_emp uPred_and + (@uPred_exist M) uPred_sep uPred_persistently. +Proof. + split. + - exact: persistently_ne. - exact: persistently_mono. - exact: persistently_idemp_2. - (* emp ⊢ <pers> emp (ADMISSIBLE) *) @@ -96,7 +103,8 @@ Qed. Canonical Structure uPredI (M : ucmra) : bi := {| bi_ofe_mixin := ofe_mixin_of (uPred M); bi_bi_mixin := uPred_bi_mixin M; - bi_bi_later_mixin := uPred_bi_later_mixin M |}. + bi_bi_later_mixin := uPred_bi_later_mixin M; + bi_bi_persistently_mixin := uPred_bi_persistently_mixin M |}. Lemma uPred_internal_eq_mixin M : BiInternalEqMixin (uPredI M) (@uPred_internal_eq M). Proof. diff --git a/iris/bi/interface.v b/iris/bi/interface.v index c751e3016..74cba5df6 100644 --- a/iris/bi/interface.v +++ b/iris/bi/interface.v @@ -19,7 +19,6 @@ Section bi_mixin. Context (bi_exist : ∀ A, (A → PROP) → PROP). Context (bi_sep : PROP → PROP → PROP). Context (bi_wand : PROP → PROP → PROP). - Context (bi_persistently : PROP → PROP). Bind Scope bi_scope with PROP. Local Infix "⊢" := bi_entails. @@ -36,7 +35,6 @@ Section bi_mixin. (bi_exist _ (λ x, .. (bi_exist _ (λ y, P%I)) ..)) : bi_scope. Local Infix "∗" := bi_sep : bi_scope. Local Infix "-∗" := bi_wand : bi_scope. - Local Notation "'<pers>' P" := (bi_persistently P) : bi_scope. (** * Axioms for a general BI (logic of bunched implications) *) @@ -64,7 +62,6 @@ Section bi_mixin. Proper (pointwise_relation _ (dist n) ==> dist n) (bi_exist A); bi_mixin_sep_ne : NonExpansive2 bi_sep; bi_mixin_wand_ne : NonExpansive2 bi_wand; - bi_mixin_persistently_ne : NonExpansive bi_persistently; (** Higher-order logic *) bi_mixin_pure_intro (φ : Prop) P : φ → P ⊢ ⌜ φ âŒ; @@ -95,8 +92,14 @@ Section bi_mixin. bi_mixin_sep_assoc' P Q R : (P ∗ Q) ∗ R ⊢ P ∗ (Q ∗ R); bi_mixin_wand_intro_r P Q R : (P ∗ Q ⊢ R) → P ⊢ Q -∗ R; bi_mixin_wand_elim_l' P Q R : (P ⊢ Q -∗ R) → P ∗ Q ⊢ R; + }. + + Context (bi_persistently : PROP → PROP). + Local Notation "'<pers>' P" := (bi_persistently P) : bi_scope. + + Record BiPersistentlyMixin := { + bi_mixin_persistently_ne : NonExpansive bi_persistently; - (** Persistently *) (* In the ordered RA model: Holds without further assumptions. *) bi_mixin_persistently_mono P Q : (P ⊢ Q) → <pers> P ⊢ <pers> Q; (* In the ordered RA model: `core` is idempotent *) @@ -121,6 +124,41 @@ Section bi_mixin. bi_mixin_persistently_and_sep_elim P Q : <pers> P ∧ Q ⊢ P ∗ Q; }. + Lemma bi_persistently_mixin_discrete : + (∀ n (P Q : PROP), P ≡{n}≡ Q → P ≡ Q) → + (∀ {A} (Φ : A → PROP), (emp ⊢ ∃ x, Φ x) → ∃ x, emp ⊢ Φ x) → + (∀ P : PROP, (<pers> P)%I = ⌜ emp ⊢ P âŒ%I) → + BiMixin → BiPersistentlyMixin. + Proof. + intros Hdiscrete Hex Hpers Hbi. pose proof (bi_mixin_entails_po Hbi). + split. + - intros n P Q [HPQ HQP]%Hdiscrete%(bi_mixin_equiv_entails Hbi). + rewrite !Hpers. apply (bi_mixin_pure_ne Hbi). split=> ?; by etrans. + - intros P Q HPQ. rewrite !Hpers. apply (bi_mixin_pure_elim' Hbi)=> ?. + apply (bi_mixin_pure_intro Hbi). by trans P. + - intros P. rewrite !Hpers. apply (bi_mixin_pure_elim' Hbi)=> ?. + by do 2 apply (bi_mixin_pure_intro Hbi). + - rewrite Hpers. by apply (bi_mixin_pure_intro Hbi). + - intros P Q. rewrite !Hpers. + apply (bi_mixin_impl_elim_l' Hbi). apply (bi_mixin_pure_elim' Hbi)=> ?. + apply (bi_mixin_impl_intro_r Hbi). + etrans; [apply (bi_mixin_and_elim_r Hbi)|]. + apply (bi_mixin_pure_elim' Hbi)=> ?. + apply (bi_mixin_pure_intro Hbi). by apply (bi_mixin_and_intro Hbi). + - intros A Φ. rewrite !Hpers. apply (bi_mixin_pure_elim' Hbi)=> /Hex [x ?]. + etrans; [|apply (bi_mixin_exist_intro Hbi x)]; simpl. + rewrite Hpers. by apply (bi_mixin_pure_intro Hbi). + - intros P Q. rewrite !Hpers. + apply (bi_mixin_wand_elim_l' Hbi). apply (bi_mixin_pure_elim' Hbi)=> ?. + apply (bi_mixin_wand_intro_r Hbi). by apply (bi_mixin_pure_intro Hbi). + - intros P Q. rewrite !Hpers. + apply (bi_mixin_impl_elim_l' Hbi). apply (bi_mixin_pure_elim' Hbi)=> ?. + apply (bi_mixin_impl_intro_r Hbi). + etrans; [apply (bi_mixin_and_elim_r Hbi)|]. + etrans; [apply (bi_mixin_emp_sep_1 Hbi)|]. + by apply (bi_mixin_sep_mono Hbi). + Qed. + (** We equip any BI with a later modality. This avoids an additional layer in the BI hierachy and improves performance significantly (see Iris issue #303). @@ -157,7 +195,7 @@ Section bi_mixin. BiMixin → BiLaterMixin. Proof. intros Hlater Hbi. pose proof (bi_mixin_entails_po Hbi). - split; repeat intro; rewrite ?Hlater ?Hequiv //. + split; repeat intro; rewrite ?Hlater //. - apply (bi_mixin_forall_intro Hbi)=> a. etrans; [apply (bi_mixin_forall_elim Hbi a)|]. by rewrite Hlater. - etrans; [|apply (bi_mixin_or_intro_r Hbi)]. @@ -194,7 +232,9 @@ Structure bi := Bi { bi_ofe_mixin : OfeMixin bi_car; bi_cofe_aux : Cofe (Ofe bi_car bi_ofe_mixin); bi_bi_mixin : BiMixin bi_entails bi_emp bi_pure bi_and bi_or bi_impl bi_forall - bi_exist bi_sep bi_wand bi_persistently; + bi_exist bi_sep bi_wand; + bi_bi_persistently_mixin : + BiPersistentlyMixin bi_entails bi_emp bi_and bi_exist bi_sep bi_persistently; bi_bi_later_mixin : BiLaterMixin bi_entails bi_pure bi_or bi_impl bi_forall bi_exist bi_sep bi_persistently bi_later; }. @@ -326,7 +366,7 @@ Proof. eapply bi_mixin_sep_ne, bi_bi_mixin. Qed. Global Instance wand_ne : NonExpansive2 (@bi_wand PROP). Proof. eapply bi_mixin_wand_ne, bi_bi_mixin. Qed. Global Instance persistently_ne : NonExpansive (@bi_persistently PROP). -Proof. eapply bi_mixin_persistently_ne, bi_bi_mixin. Qed. +Proof. eapply bi_mixin_persistently_ne, bi_bi_persistently_mixin. Qed. (* Higher-order logic *) Lemma pure_intro (φ : Prop) P : φ → P ⊢ ⌜ φ âŒ. @@ -381,24 +421,28 @@ Proof. eapply bi_mixin_wand_elim_l', bi_bi_mixin. Qed. (* Persistently *) Lemma persistently_mono P Q : (P ⊢ Q) → <pers> P ⊢ <pers> Q. -Proof. eapply bi_mixin_persistently_mono, bi_bi_mixin. Qed. +Proof. eapply bi_mixin_persistently_mono, bi_bi_persistently_mixin. Qed. Lemma persistently_idemp_2 P : <pers> P ⊢ <pers> <pers> P. -Proof. eapply bi_mixin_persistently_idemp_2, bi_bi_mixin. Qed. +Proof. eapply bi_mixin_persistently_idemp_2, bi_bi_persistently_mixin. Qed. Lemma persistently_emp_2 : emp ⊢@{PROP} <pers> emp. -Proof. eapply bi_mixin_persistently_emp_2, bi_bi_mixin. Qed. +Proof. eapply bi_mixin_persistently_emp_2, bi_bi_persistently_mixin. Qed. Lemma persistently_and_2 (P Q : PROP) : ((<pers> P) ∧ (<pers> Q)) ⊢ <pers> (P ∧ Q). -Proof. eapply bi_mixin_persistently_and_2, bi_bi_mixin. Qed. +Proof. eapply bi_mixin_persistently_and_2, bi_bi_persistently_mixin. Qed. Lemma persistently_exist_1 {A} (Ψ : A → PROP) : <pers> (∃ a, Ψ a) ⊢ ∃ a, <pers> (Ψ a). -Proof. eapply bi_mixin_persistently_exist_1, bi_bi_mixin. Qed. +Proof. eapply bi_mixin_persistently_exist_1, bi_bi_persistently_mixin. Qed. Lemma persistently_absorbing P Q : <pers> P ∗ Q ⊢ <pers> P. -Proof. eapply (bi_mixin_persistently_absorbing bi_entails), bi_bi_mixin. Qed. +Proof. + eapply (bi_mixin_persistently_absorbing bi_entails), bi_bi_persistently_mixin. +Qed. Lemma persistently_and_sep_elim P Q : <pers> P ∧ Q ⊢ P ∗ Q. -Proof. eapply (bi_mixin_persistently_and_sep_elim bi_entails), bi_bi_mixin. Qed. +Proof. + eapply (bi_mixin_persistently_and_sep_elim bi_entails), bi_bi_persistently_mixin. +Qed. (* Later *) Global Instance later_ne : NonExpansive (@bi_later PROP). diff --git a/iris/bi/monpred.v b/iris/bi/monpred.v index 978c53876..2272e36ee 100644 --- a/iris/bi/monpred.v +++ b/iris/bi/monpred.v @@ -315,8 +315,7 @@ Section instances. Lemma monPred_bi_mixin : BiMixin (PROP:=monPred I PROP) monPred_entails monPred_emp monPred_pure monPred_and monPred_or - monPred_impl monPred_forall monPred_exist monPred_sep monPred_wand - monPred_persistently. + monPred_impl monPred_forall monPred_exist monPred_sep monPred_wand. Proof. split; rewrite ?monPred_defs.monPred_unseal; try by (split=> ? /=; repeat f_equiv). @@ -354,6 +353,15 @@ Section instances. apply bi.wand_intro_r. by rewrite -HR /= !Hij. - intros P Q R [HP]. split=> i. apply bi.wand_elim_l'. rewrite HP /= bi.forall_elim bi.pure_impl_forall bi.forall_elim //. + Qed. + + Lemma monPred_bi_persistently_mixin : + BiPersistentlyMixin (PROP:=monPred I PROP) + monPred_entails monPred_emp monPred_and + monPred_exist monPred_sep monPred_persistently. + Proof. + split; rewrite ?monPred_defs.monPred_unseal; + try by (split=> ? /=; repeat f_equiv). - intros P Q [?]. split=> i /=. by f_equiv. - intros P. split=> i. by apply bi.persistently_idemp_2. - split=> i. by apply bi.persistently_emp_intro. @@ -386,7 +394,9 @@ Section instances. Qed. Canonical Structure monPredI : bi := - {| bi_ofe_mixin := monPred_ofe_mixin; bi_bi_mixin := monPred_bi_mixin; + {| bi_ofe_mixin := monPred_ofe_mixin; + bi_bi_mixin := monPred_bi_mixin; + bi_bi_persistently_mixin := monPred_bi_persistently_mixin; bi_bi_later_mixin := monPred_bi_later_mixin |}. (** We restate the unsealing lemmas so that they also unfold the BI layer. The diff --git a/iris/si_logic/bi.v b/iris/si_logic/bi.v index 0b4d63892..e366744e3 100644 --- a/iris/si_logic/bi.v +++ b/iris/si_logic/bi.v @@ -20,8 +20,7 @@ Local Existing Instance entails_po. Lemma siProp_bi_mixin : BiMixin siProp_entails siProp_emp siProp_pure siProp_and siProp_or siProp_impl - (@siProp_forall) (@siProp_exist) siProp_sep siProp_wand - siProp_persistently. + (@siProp_forall) (@siProp_exist) siProp_sep siProp_wand. Proof. split. - exact: entails_po. @@ -34,7 +33,6 @@ Proof. - exact: exist_ne. - exact: and_ne. - exact: impl_ne. - - solve_proper. - exact: pure_intro. - exact: pure_elim'. - exact: and_elim_l. @@ -68,6 +66,15 @@ Proof. apply impl_intro_r. - (* (P ⊢ Q -∗ R) → P ∗ Q ⊢ R *) apply impl_elim_l'. +Qed. + +Lemma siProp_bi_persistently_mixin : + BiPersistentlyMixin + siProp_entails siProp_emp siProp_and + (@siProp_exist) siProp_sep siProp_persistently. +Proof. + split. + - solve_proper. - (* (P ⊢ Q) → <pers> P ⊢ <pers> Q *) done. - (* <pers> P ⊢ <pers> <pers> P *) @@ -119,7 +126,9 @@ Qed. Canonical Structure siPropI : bi := {| bi_ofe_mixin := ofe_mixin_of siProp; - bi_bi_mixin := siProp_bi_mixin; bi_bi_later_mixin := siProp_bi_later_mixin |}. + bi_bi_mixin := siProp_bi_mixin; + bi_bi_persistently_mixin := siProp_bi_persistently_mixin; + bi_bi_later_mixin := siProp_bi_later_mixin |}. Global Instance siProp_pure_forall : BiPureForall siPropI. Proof. exact: @pure_forall_2. Qed. diff --git a/tests/heapprop.v b/tests/heapprop.v index 31b5edbee..69634a3c0 100644 --- a/tests/heapprop.v +++ b/tests/heapprop.v @@ -94,7 +94,7 @@ Local Definition heapProp_wand_unseal: @heapProp_wand = @heapProp_wand_def := seal_eq heapProp_wand_aux. Local Definition heapProp_persistently_def (P : heapProp) : heapProp := - {| heapProp_holds σ := P ∅ |}. + heapProp_pure (heapProp_entails heapProp_emp P). Local Definition heapProp_persistently_aux : seal (@heapProp_persistently_def). Proof. by eexists. Qed. Definition heapProp_persistently := unseal heapProp_persistently_aux. @@ -120,7 +120,7 @@ Section mixins. BiMixin heapProp_entails heapProp_emp heapProp_pure heapProp_and heapProp_or heapProp_impl (@heapProp_forall) (@heapProp_exist) - heapProp_sep heapProp_wand heapProp_persistently. + heapProp_sep heapProp_wand. Proof. split. - (* [PreOrder heapProp_entails] *) @@ -145,8 +145,6 @@ Section mixins. unseal=> n P1 P2 [HP] Q1 Q2 [HQ]; split; naive_solver. - (* [NonExpansive2 bi_wand] *) unseal=> n P1 P2 [HP] Q1 Q2 [HQ]; split; naive_solver. - - (* [NonExpansive2 bi_persistently] *) - unseal=> n P1 P2 [HP]; split; naive_solver. - (* [φ → P ⊢ ⌜ φ âŒ] *) unseal=> φ P ?; by split. - (* [(φ → True ⊢ P) → ⌜ φ ⌠⊢ P] *) @@ -193,21 +191,18 @@ Section mixins. unseal=> P Q R [HPQR]; split=> σ1 ? σ2 ??. apply HPQR. by exists σ1, σ2. - (* [(P ⊢ Q -∗ R) → P ∗ Q ⊢ R] *) unseal=> P Q R [HPQR]; split; intros ? (σ1&σ2&->&?&?&?). by apply HPQR. - - (* [(P ⊢ Q) → <pers> P ⊢ <pers> Q] *) - unseal=> P Q [HPQ]; split=> σ. apply HPQ. - - (* [<pers> P ⊢ <pers> <pers> P] *) - unseal=> P; split=> σ; done. - - (* [emp ⊢ <pers> emp] *) - unseal; split=> σ; done. - - (* [(∀ a, <pers> (Ψ a)) ⊢ <pers> (∀ a, Ψ a)] *) - unseal=> A Ψ; split=> σ; done. - - (* [<pers> (∃ a, Ψ a) ⊢ ∃ a, <pers> (Ψ a)] *) - unseal=> A Ψ; split=> σ; done. - - (* [<pers> P ∗ Q ⊢ <pers> P] *) - unseal=> P Q; split; intros ? (σ1&σ2&->&?&?&?); done. - - (* [<pers> P ∧ Q ⊢ P ∗ Q] *) - unseal=> P Q; split=> σ [??]. eexists ∅, σ. rewrite left_id_L. - split_and!; done || apply map_disjoint_empty_l. + Qed. + + Lemma heapProp_bi_persistently_mixin : + BiPersistentlyMixin + heapProp_entails heapProp_emp heapProp_and + (@heapProp_exist) heapProp_sep heapProp_persistently. + Proof. + eapply bi_persistently_mixin_discrete, heapProp_bi_mixin; [done|..]. + - (* [(emp ⊢ ∃ x, Φ x) → ∃ x, emp ⊢ Φ x] *) + unseal. intros A Φ [H]. destruct (H ∅) as [x ?]; [done|]. + exists x. by split=> σ ->. + - by rewrite heapProp_persistently_unseal. Qed. Lemma heapProp_bi_later_mixin : @@ -220,7 +215,9 @@ End mixins. Canonical Structure heapPropI : bi := {| bi_ofe_mixin := ofe_mixin_of heapProp; - bi_bi_mixin := heapProp_bi_mixin; bi_bi_later_mixin := heapProp_bi_later_mixin |}. + bi_bi_mixin := heapProp_bi_mixin; + bi_bi_persistently_mixin := heapProp_bi_persistently_mixin; + bi_bi_later_mixin := heapProp_bi_later_mixin |}. Global Instance heapProp_pure_forall : BiPureForall heapPropI. Proof. intros A φ. rewrite /bi_forall /bi_pure /=. unseal. by split. Qed. -- GitLab