diff --git a/CHANGELOG.md b/CHANGELOG.md index 518d58d5d46c76f3f6e8517934b74c3fb3f0c648..10fbc83ae2255664ffb5986d96982df336e34622 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -21,6 +21,8 @@ lemma. the conclusion. The old behavior can be emulated with`iExFalso. iExact "H".` * `iInduction` now supports induction schemes that involve `Forall` and `Forall2` (for example, for trees with finite branching). +* Change `iRevert` of a pure hypothesis to generate a magic wand instead of an + implication. **Changes in `base_logic`:** diff --git a/iris/proofmode/coq_tactics.v b/iris/proofmode/coq_tactics.v index a2708b041d45ae268b9e521448b152e5f32a03b0..060f818cd72079f5a9826e21ab022840dcffb0cf 100644 --- a/iris/proofmode/coq_tactics.v +++ b/iris/proofmode/coq_tactics.v @@ -170,8 +170,14 @@ Proof. rewrite -persistent_and_affinely_sep_l. apply pure_elim_l=> ?. by rewrite HQ. Qed. -Lemma tac_pure_revert Δ φ Q : envs_entails Δ (⌜φ⌝ → Q) → (φ → envs_entails Δ Q). -Proof. rewrite envs_entails_eq. intros HΔ ?. by rewrite HΔ pure_True // left_id. Qed. +Lemma tac_pure_revert Δ φ P Q : + FromAffinely P ⌜ φ ⌝ → + envs_entails Δ (P -∗ Q) → + (φ → envs_entails Δ Q). +Proof. + rewrite /FromAffinely envs_entails_eq. intros <- -> ?. + by rewrite pure_True // affinely_True_emp affinely_emp left_id. +Qed. (** * Intuitionistic *) Lemma tac_intuitionistic Δ i j p P P' Q : diff --git a/iris/proofmode/ltac_tactics.v b/iris/proofmode/ltac_tactics.v index f4da0ed786ded955e152a2d499e932b067b69484..b651b3fbf069fb0c040d11c87a71dad0c90415da 100644 --- a/iris/proofmode/ltac_tactics.v +++ b/iris/proofmode/ltac_tactics.v @@ -651,7 +651,10 @@ Local Tactic Notation "iForallRevert" ident(x) := first [let A := type of x in idtac|fail 1 "iRevert:" x "not in scope"]; let A := type of x in lazymatch type of A with - | Prop => revert x; first [apply tac_pure_revert|err x] + | Prop => + revert x; first + [eapply tac_pure_revert; [iSolveTC (* [FromAffinely], should never fail *)|] + |err x] | _ => revert x; first [apply tac_forall_revert|err x] end. diff --git a/tests/proofmode.ref b/tests/proofmode.ref index ce21d1c9d1c2c3b7629c0e461fa8a4ff09977bb1..58739cd69d4e2b275403abd12ee31e493d190595 100644 --- a/tests/proofmode.ref +++ b/tests/proofmode.ref @@ -396,6 +396,29 @@ No applicable tactic. "HQ" : Q --------------------------------------∗ ▷ P ∗ Q +"test_iRevert_pure" + : string +1 goal + + PROP : bi + φ : Prop + P : PROP + ============================ + "H" : ⌜φ⌝ -∗ P + --------------------------------------∗ + ⌜φ⌝ -∗ P +"test_iRevert_pure_affine" + : string +1 goal + + PROP : bi + BiAffine0 : BiAffine PROP + φ : Prop + P : PROP + ============================ + "H" : ⌜φ⌝ -∗ P + --------------------------------------∗ + ⌜φ⌝ -∗ P "elim_mod_accessor" : string 1 goal diff --git a/tests/proofmode.v b/tests/proofmode.v index 8ad3f6da4b95f11a1f86e15748228e9556d8b504..e83bd50bfcfc45ea2a06fae6bd001eaa466f6dc0 100644 --- a/tests/proofmode.v +++ b/tests/proofmode.v @@ -1146,6 +1146,22 @@ Proof. iIntros (?) "HP HQ". iModIntro. Show. by iFrame. Qed. +Check "test_iRevert_pure". +Lemma test_iRevert_pure (φ : Prop) P : + φ → ( ⌜ φ ⌝ -∗ P) -∗ P. +Proof. + (* Check that iRevert creates a wand instead of an implication *) + iIntros (Hφ) "H". iRevert (Hφ). Show. done. +Qed. + +Check "test_iRevert_pure_affine". +Lemma test_iRevert_pure_affine `{!BiAffine PROP} (φ : Prop) P : + φ → (⌜ φ ⌝ -∗ P) -∗ P. +Proof. + (* If the BI is affine, no affine modality should be added *) + iIntros (Hφ) "H". iRevert (Hφ). Show. done. +Qed. + End tests. Section parsing_tests.