diff --git a/tests/proofmode.v b/tests/proofmode.v index bb095de55daec40862f2388fdd534a5761c7e10f..9628ea7e45c40f602167edd5fdbf94a1fc504314 100644 --- a/tests/proofmode.v +++ b/tests/proofmode.v @@ -457,11 +457,21 @@ Proof. iIntros (?) "HP #HQ HR". iPureIntro; eauto. Qed. (* Ensure that [% ...] works as a pattern when the left-hand side of and/sep is pure. *) -Lemma test_pure_and_sep_destruct `{!BiAffine PROP} (φ : Prop) (P : PROP) : +Lemma test_pure_and_sep_destruct_affine `{!BiAffine PROP} (φ : Prop) (P : PROP) : ⌜φ⌠∧ (⌜φ⌠∗ P) -∗ P. Proof. iIntros "[% [% $]]". Qed. +Lemma test_pure_and_sep_destruct_1 (φ : Prop) (P : PROP) : + ⌜φ⌠∧ (<affine> ⌜φ⌠∗ P) -∗ P. +Proof. + iIntros "[% [% $]]". +Qed. +Lemma test_pure_and_sep_destruct_2 (φ : Prop) (P : PROP) : + ⌜φ⌠∧ (⌜φ⌠∗ <absorb> P) -∗ <absorb> P. +Proof. + iIntros "[% [% $]]". +Qed. (* Ensure that [% %] also works when the conjunction is *inside* ⌜...⌠*) Lemma test_pure_inner_and_destruct : ⌜False ∧ True⌠⊢@{PROP} False.