diff --git a/tests/proofmode_monpred.ref b/tests/proofmode_monpred.ref index 8426e3b26ad338a12b75f2ac4db7d96e4a862575..39b3fb6e4c3606c2a5488613dc999da08d61637a 100644 --- a/tests/proofmode_monpred.ref +++ b/tests/proofmode_monpred.ref @@ -46,3 +46,14 @@ --------------------------------------∗ (Q -∗ emp) i +1 subgoal + + I : biIndex + PROP : sbi + FU0 : BiFUpd PROP * () + P : monpred.monPred I PROP + i : I + ============================ + --------------------------------------∗ + ∀ _ : (), ∃ _ : (), emp + diff --git a/tests/proofmode_monpred.v b/tests/proofmode_monpred.v index a61e964708ea4714bc051e3057b97575ab4c743b..c683a11a74f73442789fca7307f5b408cd74442f 100644 --- a/tests/proofmode_monpred.v +++ b/tests/proofmode_monpred.v @@ -120,12 +120,50 @@ Section tests. iIntros "[$ #HP]". iFrame "HP". Qed. + Lemma test_iNext_Bi P : + â–· P ⊢@{monPredI} â–· P. + Proof. iIntros "H". by iNext. Qed. + + (** Test monPred_at framing *) Lemma test_iFrame_monPred_at_wand (P Q : monPred) i : P i -∗ (Q -∗ P) i. Proof. iIntros "$". Show. Abort. - Lemma test_iNext_Bi P : - @bi_entails monPredI (â–· P) (â–· P). - Proof. iIntros "H". by iNext. Qed. + Program Definition monPred_id (R : monPred) : monPred := + MonPred (λ V, R V) _. + Next Obligation. intros ? ???. eauto. Qed. + + Lemma test_iFrame_monPred_id (Q R : monPred) i : + Q i ∗ R i -∗ (Q ∗ monPred_id R) i. + Proof. + iIntros "(HQ & HR)". iFrame "HR". iAssumption. + Qed. + + Lemma test_iFrame_rel P i j ij : + IsBiIndexRel i ij → IsBiIndexRel j ij → + P i -∗ P j -∗ P ij ∗ P ij. + Proof. iIntros (??) "HPi HPj". iFrame. Qed. + + Lemma test_iFrame_later_rel `{!BiAffine PROP} P i j : + IsBiIndexRel i j → + â–· (P i) -∗ (â–· P) j. + Proof. iIntros (?) "?". iFrame. Qed. + + Lemma test_iFrame_laterN n P i : + â–·^n (P i) -∗ (â–·^n P) i. + Proof. iIntros "?". iFrame. Qed. + + Lemma test_iFrame_quantifiers P i : + P i -∗ (∀ _:(), ∃ _:(), P) i. + Proof. iIntros "?". iFrame. Show. iIntros ([]). iExists (). iEmpIntro. Qed. + + Lemma test_iFrame_embed (P : PROP) i : + P -∗ (embed (B:=monPredI) P) i. + Proof. iIntros "$". Qed. + + (* Make sure search doesn't diverge on an evar *) + Lemma test_iFrame_monPred_at_evar (P : monPred) i j : + P i -∗ ∃ Q, (Q j). + Proof. iIntros "HP". iExists _. Fail iFrame "HP". Abort. End tests. diff --git a/theories/bi/monpred.v b/theories/bi/monpred.v index 0a01f4e0d7a5b014ab134772c2630efc1c2aef01..2a25581675dd91c329f44d53811c2096408b0a5b 100644 --- a/theories/bi/monpred.v +++ b/theories/bi/monpred.v @@ -825,6 +825,8 @@ Lemma monPred_at_internal_eq {A : ofeT} i (a b : A) : Proof. rewrite monPred_internal_eq_unfold. by apply monPred_at_embed. Qed. Lemma monPred_at_later i P : (â–· P) i ⊣⊢ â–· P i. Proof. by unseal. Qed. +Lemma monPred_at_laterN n i P : (â–·^n P) i ⊣⊢ â–·^n P i. +Proof. induction n; first done. rewrite /= monPred_at_later IHn //. Qed. Lemma monPred_at_except_0 i P : (â—‡ P) i ⊣⊢ â—‡ P i. Proof. by unseal. Qed. diff --git a/theories/proofmode/monpred.v b/theories/proofmode/monpred.v index ddf4fda5c673b169c7800bb052b871ff24b4ec45..68920fb9a8678799eed5a46c2ab4edb7ae9534b7 100644 --- a/theories/proofmode/monpred.v +++ b/theories/proofmode/monpred.v @@ -15,6 +15,14 @@ Proof. by rewrite /IsBiIndexRel. Qed. Hint Extern 1 (IsBiIndexRel _ _) => unfold IsBiIndexRel; assumption : typeclass_instances. +(** Frame [ð“¡] into the goal [monPred_at P i] and determine the remainder [ð“ ]. + Used when framing encounters a monPred_at in the goal. *) +Class FrameMonPredAt {I : biIndex} {PROP : bi} (p : bool) (i : I) + (ð“¡ : PROP) (P : monPred I PROP) (ð“ : PROP) := + frame_monPred_at : â–¡?p 𓡠∗ ð“ -∗ P i. +Arguments FrameMonPredAt {_ _} _ _ _%I _%I _%I. +Hint Mode FrameMonPredAt + + + + ! ! - : typeclass_instances. + Section modalities. Context {I : biIndex} {PROP : bi}. @@ -340,24 +348,77 @@ Proof. by rewrite monPred_at_forall. Qed. -(* FIXME : there are two good ways to frame under a call to - monPred_at. This may cause some backtracking in the typeclass - search, and hence performance issues. *) -Global Instance frame_monPred_at p P Q ð“ R i j : - IsBiIndexRel i j → Frame p R P Q → MakeMonPredAt j Q ð“ → - Frame p (R i) (P j) ð“ . -Proof. - rewrite /Frame /MakeMonPredAt /bi_affinely_if /bi_persistently_if - /IsBiIndexRel=> Hij <- <-. - destruct p; by rewrite Hij monPred_at_sep ?monPred_at_affinely ?monPred_at_persistently. -Qed. -Global Instance frame_monPred_at_embed i p P Q ð“ ð“¡ : - Frame p ⎡ð“¡âŽ¤ P Q → MakeMonPredAt i Q ð“ → Frame p ð“¡ (P i) ð“ . -Proof. - rewrite /Frame /MakeMonPredAt /bi_affinely_if /bi_persistently_if=> <- <-. - destruct p; by rewrite monPred_at_sep ?monPred_at_affinely - ?monPred_at_persistently monPred_at_embed. -Qed. +(* Framing. *) +Global Instance frame_monPred_at_enter p i ð“¡ P ð“ : + FrameMonPredAt p i ð“¡ P ð“ → Frame p ð“¡ (P i) ð“ . +Proof. intros. done. Qed. +Global Instance frame_monPred_at_here p P i j : + IsBiIndexRel i j → FrameMonPredAt p j (P i) P emp | 0. +Proof. + rewrite /FrameMonPredAt /IsBiIndexRel right_id bi.intuitionistically_if_elim=> -> //. +Qed. + +Global Instance frame_monPred_at_embed p ð“¡ ð“ ð“Ÿ i : + Frame p ð“¡ ð“Ÿ ð“ → FrameMonPredAt p i ð“¡ (embed (B:=monPredI) ð“Ÿ) ð“ . +Proof. rewrite /Frame /FrameMonPredAt=> ->. by rewrite monPred_at_embed. Qed. +Global Instance frame_monPred_at_sep p P Q ð“¡ ð“ i : + Frame p ð“¡ (P i ∗ Q i) ð“ → FrameMonPredAt p i ð“¡ (P ∗ Q) ð“ . +Proof. rewrite /Frame /FrameMonPredAt=> ->. by rewrite monPred_at_sep. Qed. +Global Instance frame_monPred_at_and p P Q ð“¡ ð“ i : + Frame p ð“¡ (P i ∧ Q i) ð“ → FrameMonPredAt p i ð“¡ (P ∧ Q) ð“ . +Proof. rewrite /Frame /FrameMonPredAt=> ->. by rewrite monPred_at_and. Qed. +Global Instance frame_monPred_at_or p P Q ð“¡ ð“ i : + Frame p ð“¡ (P i ∨ Q i) ð“ → FrameMonPredAt p i ð“¡ (P ∨ Q) ð“ . +Proof. rewrite /Frame /FrameMonPredAt=> ->. by rewrite monPred_at_or. Qed. +Global Instance frame_monPred_at_wand p P R Q1 Q2 i j : + IsBiIndexRel i j → + Frame p R Q1 Q2 → + FrameMonPredAt p j (R i) (P -∗ Q1) ((P -∗ Q2) i). +Proof. + rewrite /Frame /FrameMonPredAt=>-> Hframe. + rewrite -monPred_at_intuitionistically_if -monPred_at_sep. apply monPred_in_entails. + change ((â–¡?p R ∗ (P -∗ Q2)) -∗ P -∗ Q1). apply bi.wand_intro_r. + rewrite -assoc bi.wand_elim_l. done. +Qed. +Global Instance frame_monPred_at_impl P R Q1 Q2 i j : + IsBiIndexRel i j → + Frame true R Q1 Q2 → + FrameMonPredAt true j (R i) (P → Q1) ((P → Q2) i). +Proof. + rewrite /Frame /FrameMonPredAt=>-> Hframe. + rewrite -monPred_at_intuitionistically_if -monPred_at_sep. apply monPred_in_entails. + change ((â–¡ R ∗ (P → Q2)) -∗ P → Q1). + rewrite -bi.persistently_and_intuitionistically_sep_l. apply bi.impl_intro_r. + rewrite -assoc bi.impl_elim_l bi.persistently_and_intuitionistically_sep_l. done. +Qed. +Global Instance frame_monPred_at_forall {X : Type} p (Ψ : X → monPred) ð“¡ ð“ i : + Frame p ð“¡ (∀ x, Ψ x i) ð“ → FrameMonPredAt p i ð“¡ (∀ x, Ψ x) ð“ . +Proof. rewrite /Frame /FrameMonPredAt=> ->. by rewrite monPred_at_forall. Qed. +Global Instance frame_monPred_at_exist {X : Type} p (Ψ : X → monPred) ð“¡ ð“ i : + Frame p ð“¡ (∃ x, Ψ x i) ð“ → FrameMonPredAt p i ð“¡ (∃ x, Ψ x) ð“ . +Proof. rewrite /Frame /FrameMonPredAt=> ->. by rewrite monPred_at_exist. Qed. + +Global Instance frame_monPred_at_absorbingly p P ð“¡ ð“ i : + Frame p ð“¡ (<absorb> P i) ð“ → FrameMonPredAt p i ð“¡ (<absorb> P) ð“ . +Proof. rewrite /Frame /FrameMonPredAt=> ->. by rewrite monPred_at_absorbingly. Qed. +Global Instance frame_monPred_at_affinely p P ð“¡ ð“ i : + Frame p ð“¡ (<affine> P i) ð“ → FrameMonPredAt p i ð“¡ (<affine> P) ð“ . +Proof. rewrite /Frame /FrameMonPredAt=> ->. by rewrite monPred_at_affinely. Qed. +Global Instance frame_monPred_at_persistently p P ð“¡ ð“ i : + Frame p ð“¡ (<pers> P i) ð“ → FrameMonPredAt p i ð“¡ (<pers> P) ð“ . +Proof. rewrite /Frame /FrameMonPredAt=> ->. by rewrite monPred_at_persistently. Qed. +Global Instance frame_monPred_at_intuitionistically p P ð“¡ ð“ i : + Frame p ð“¡ (â–¡ P i) ð“ → FrameMonPredAt p i ð“¡ (â–¡ P) ð“ . +Proof. rewrite /Frame /FrameMonPredAt=> ->. by rewrite monPred_at_intuitionistically. Qed. +Global Instance frame_monPred_at_objectively p P ð“¡ ð“ i : + Frame p ð“¡ (∀ i, P i) ð“ → FrameMonPredAt p i ð“¡ (<obj> P) ð“ . +Proof. rewrite /Frame /FrameMonPredAt=> ->. by rewrite monPred_at_objectively. Qed. +Global Instance frame_monPred_at_subjectively p P ð“¡ ð“ i : + Frame p ð“¡ (∃ i, P i) ð“ → FrameMonPredAt p i ð“¡ (<subj> P) ð“ . +Proof. rewrite /Frame /FrameMonPredAt=> ->. by rewrite monPred_at_subjectively. Qed. +Global Instance frame_monPred_at_bupd `{BiBUpd PROP} p P ð“¡ ð“ i : + Frame p ð“¡ (|==> P i) ð“ → FrameMonPredAt p i ð“¡ (|==> P) ð“ . +Proof. rewrite /Frame /FrameMonPredAt=> ->. by rewrite monPred_at_bupd. Qed. Global Instance into_embed_objective P : Objective P → IntoEmbed P (∀ i, P i). @@ -472,6 +533,16 @@ Proof. by rewrite monPred_at_later. Qed. +Global Instance frame_monPred_at_later p P ð“¡ ð“ i : + Frame p ð“¡ (â–· P i) ð“ → FrameMonPredAt p i ð“¡ (â–· P) ð“ . +Proof. rewrite /Frame /FrameMonPredAt=> ->. by rewrite monPred_at_later. Qed. +Global Instance frame_monPred_at_laterN p n P ð“¡ ð“ i : + Frame p ð“¡ (â–·^n P i) ð“ → FrameMonPredAt p i ð“¡ (â–·^n P) ð“ . +Proof. rewrite /Frame /FrameMonPredAt=> ->. by rewrite monPred_at_laterN. Qed. +Global Instance frame_monPred_at_fupd `{BiFUpd PROP} E1 E2 p P ð“¡ ð“ i : + Frame p ð“¡ (|={E1,E2}=> P i) ð“ → FrameMonPredAt p i ð“¡ (|={E1,E2}=> P) ð“ . +Proof. rewrite /Frame /FrameMonPredAt=> ->. by rewrite monPred_at_fupd. Qed. + Global Instance elim_modal_at_fupd_goal `{BiFUpd PROP} φ p p' E1 E2 E3 ð“Ÿ ð“Ÿ' Q Q' i : ElimModal φ p p' ð“Ÿ ð“Ÿ' (|={E1,E3}=> Q i) (|={E2,E3}=> Q' i) → ElimModal φ p p' ð“Ÿ ð“Ÿ' ((|={E1,E3}=> Q) i) ((|={E2,E3}=> Q') i).