diff --git a/tests/proofmode.v b/tests/proofmode.v index 2d395c27ebb6b39b32344675e1722aa8de104b3d..3905f8bb75bd8dc857930ed5bc7d9c27b080d08d 100644 --- a/tests/proofmode.v +++ b/tests/proofmode.v @@ -111,6 +111,10 @@ Lemma test_iSpecialize_auto_frame P Q R : (P -∗ True -∗ True -∗ Q -∗ R) -∗ P -∗ Q -∗ R. Proof. iIntros "H ? HQ". by iApply ("H" with "[$]"). Qed. +Lemma test_iSpecialize_pure (φ : Prop) Q R: + φ → (⌜φ⌠-∗ Q) → Q. +Proof. iIntros (HP HPQ). iDestruct (HPQ $! HP) as "?". done. Qed. + Lemma test_iSpecialize_Coq_entailment P Q R : P → (P -∗ Q) → Q. Proof. iIntros (HP HPQ). iDestruct (HPQ $! HP) as "?". done. Qed. diff --git a/theories/proofmode/class_instances_bi.v b/theories/proofmode/class_instances_bi.v index ac035af9326a8f726b5c8e87006c31fd214d50d1..cf4a2743c2a5f0c1052e511027d8bccc1f1d62d5 100644 --- a/theories/proofmode/class_instances_bi.v +++ b/theories/proofmode/class_instances_bi.v @@ -806,8 +806,13 @@ Proof. rewrite /IntoForall=> HP. by rewrite HP persistently_forall. Qed. Global Instance into_forall_embed `{BiEmbed PROP PROP'} {A} P (Φ : A → PROP) : IntoForall P Φ → IntoForall ⎡P⎤ (λ a, ⎡Φ a⎤%I). Proof. by rewrite /IntoForall -embed_forall => <-. Qed. -Global Instance into_forall_wand P Q : IntoForall (P -∗ Q) (λ _ : bi_emp_valid P, Q). +(* These instances must be used only after [into_forall_wand_pure] in +[class_instances_sbi] failed. *) +Global Instance into_forall_wand P Q : IntoForall (P -∗ Q) (λ _ : bi_emp_valid 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. +Proof. rewrite /IntoForall. apply forall_intro=><-. rewrite -True_emp True_impl //. Qed. (* FromForall *) Global Instance from_forall_forall {A} (Φ : A → PROP) :