From 357f1cc229dcaa030b5e66aef4687e591ba6f8da Mon Sep 17 00:00:00 2001
From: Ralf Jung <jung@mpi-sws.org>
Date: Sat, 9 Jun 2018 09:01:29 +0200
Subject: [PATCH] allow specializing a wand with a Coq-level proof of the
 premise

---
 tests/proofmode.v                       | 4 ++++
 theories/proofmode/class_instances_bi.v | 2 ++
 2 files changed, 6 insertions(+)

diff --git a/tests/proofmode.v b/tests/proofmode.v
index b0168e1f8..2d395c27e 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_Coq_entailment P Q R :
+  P → (P -∗ Q) → Q.
+Proof. iIntros (HP HPQ). iDestruct (HPQ $! HP) as "?". done. Qed.
+
 Lemma test_iEmp_intro P Q R `{!Affine P, !Persistent Q, !Affine R} :
   P -∗ Q → R -∗ emp.
 Proof. iIntros "HP #HQ HR". iEmpIntro. Qed.
diff --git a/theories/proofmode/class_instances_bi.v b/theories/proofmode/class_instances_bi.v
index 8e7d610ff..ac035af93 100644
--- a/theories/proofmode/class_instances_bi.v
+++ b/theories/proofmode/class_instances_bi.v
@@ -806,6 +806,8 @@ 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).
+Proof. rewrite /IntoForall. apply forall_intro=><-. rewrite emp_wand //. Qed.
 
 (* FromForall *)
 Global Instance from_forall_forall {A} (Φ : A → PROP) :
-- 
GitLab