From 09663be3e518f3cd8d949c4b5c3225c5d1d7e66e Mon Sep 17 00:00:00 2001
From: Jacques-Henri Jourdan <jacques-henri.jourdan@normalesup.org>
Date: Thu, 1 Mar 2018 19:17:14 +0100
Subject: [PATCH] Make iFrame able to accumulate assertions in an evar.

This requires changing the Hint Mode of the [Frame] type class because it should not fail if its parameter is an evar, but instantiate it instead. In order to prevent all the other instances of [Frame] to intantiate this evar themselves, we create a new type class [KnwonFrame], which corresponds to the old behavior.
---
 theories/base_logic/lib/boxes.v               |   2 +-
 .../base_logic/lib/cancelable_invariants.v    |   2 +-
 theories/base_logic/lib/invariants.v          |   2 +-
 theories/base_logic/lib/na_invariants.v       |   4 +-
 theories/program_logic/total_weakestpre.v     |   5 +-
 theories/program_logic/weakestpre.v           |   5 +-
 theories/proofmode/class_instances.v          | 139 ++++++++++--------
 theories/proofmode/classes.v                  |  16 +-
 theories/proofmode/monpred.v                  |   8 +-
 theories/tests/proofmode.v                    |  15 +-
 10 files changed, 117 insertions(+), 81 deletions(-)

diff --git a/theories/base_logic/lib/boxes.v b/theories/base_logic/lib/boxes.v
index 3bda9e53e..561c1a24a 100644
--- a/theories/base_logic/lib/boxes.v
+++ b/theories/base_logic/lib/boxes.v
@@ -255,7 +255,7 @@ Proof.
     iAlways. by iSplit; iIntros "[? $]"; iApply "HQQ'".
   - iMod (slice_delete_empty with "Hs Hb") as (P') "(Heq & Hb)"; try done.
     iMod (slice_insert_empty with "Hb") as (γ' ?) "[#Hs' Hb]"; try done.
-    iExists γ', _. iIntros "{$∗ $# $%} !>". do 2 iNext. iRewrite "Heq".
+    iExists γ', (Q' ∗ P')%I. iIntros "{$∗ $# $%} !>".  do 2 iNext. iRewrite "Heq".
     iAlways. by iSplit; iIntros "[? $]"; iApply "HQQ'".
 Qed.
 
diff --git a/theories/base_logic/lib/cancelable_invariants.v b/theories/base_logic/lib/cancelable_invariants.v
index 294019fae..42744be61 100644
--- a/theories/base_logic/lib/cancelable_invariants.v
+++ b/theories/base_logic/lib/cancelable_invariants.v
@@ -98,7 +98,7 @@ Section proofs.
       (▷ P ∗ cinv_own γ p) (▷ P ={E∖↑N,E}=∗ True) Q Q'.
   Proof.
     rewrite /ElimInv /ElimModal. iIntros (Helim ?) "(#H1&Hown&H2)".
-    iApply (Helim with "[- $H2]"); first done.
+    iApply Helim; [done|]. iSplitR "H2"; [|done].
     iMod (cinv_open E N γ p P with "[#] [Hown]") as "(HP&Hown&Hclose)"; auto. 
     by iFrame.
   Qed.
diff --git a/theories/base_logic/lib/invariants.v b/theories/base_logic/lib/invariants.v
index b18cb809f..12654be80 100644
--- a/theories/base_logic/lib/invariants.v
+++ b/theories/base_logic/lib/invariants.v
@@ -101,7 +101,7 @@ Global Instance elim_inv_inv E N P Q Q' :
   ElimInv (↑N ⊆ E) (inv N P) True (▷ P) (▷ P ={E∖↑N,E}=∗ True) Q Q'.
 Proof.
   rewrite /ElimInv /ElimModal. iIntros (Helim ?) "(#H1&_&H2)".
-  iApply (Helim with "[$H2]"); first done.
+  iApply (Helim with "[H2]"); [done|]. iSplitR "H2"; [|done].
   iMod (inv_open _ N with "[#]") as "(HP&Hclose)"; auto with iFrame.
 Qed.
 
diff --git a/theories/base_logic/lib/na_invariants.v b/theories/base_logic/lib/na_invariants.v
index 6b5296060..2cd1a8a1a 100644
--- a/theories/base_logic/lib/na_invariants.v
+++ b/theories/base_logic/lib/na_invariants.v
@@ -118,8 +118,8 @@ Section proofs.
       (▷ P ∗ na_own p (F∖↑N)) (▷ P ∗ na_own p (F∖↑N) ={E}=∗ na_own p F) Q Q'.
   Proof.
     rewrite /ElimInv /ElimModal. iIntros (Helim (?&?)) "(#H1&Hown&H2)".
-    iApply (Helim with "[- $H2]"); first done.
-    iMod (na_inv_open p E F N P with "[#] [Hown]") as "(HP&Hown&Hclose)"; auto. 
+    iApply Helim; [done|]. iSplitR "H2"; [|done].
+    iMod (na_inv_open p E F N P with "[#] [Hown]") as "(HP&Hown&Hclose)"; auto.
     by iFrame.
   Qed.
 End proofs.
diff --git a/theories/program_logic/total_weakestpre.v b/theories/program_logic/total_weakestpre.v
index dd0d0628b..b188a5b34 100644
--- a/theories/program_logic/total_weakestpre.v
+++ b/theories/program_logic/total_weakestpre.v
@@ -377,8 +377,9 @@ Section proofmode_classes.
   Implicit Types Φ : val Λ → iProp Σ.
 
   Global Instance frame_twp p s E e R Φ Ψ :
-    (∀ v, Frame p R (Φ v) (Ψ v)) → Frame p R (WP e @ s; E [{ Φ }]) (WP e @ s; E [{ Ψ }]).
-  Proof. rewrite /Frame=> HR. rewrite twp_frame_l. apply twp_mono, HR. Qed.
+    (∀ v, Frame p R (Φ v) (Ψ v)) →
+    KnownFrame p R (WP e @ s; E [{ Φ }]) (WP e @ s; E [{ Ψ }]).
+  Proof. rewrite /KnownFrame /Frame=> HR. rewrite twp_frame_l. apply twp_mono, HR. Qed.
 
   Global Instance is_except_0_wp s E e Φ : IsExcept0 (WP e @ s; E [{ Φ }]).
   Proof. by rewrite /IsExcept0 -{2}fupd_twp -except_0_fupd -fupd_intro. Qed.
diff --git a/theories/program_logic/weakestpre.v b/theories/program_logic/weakestpre.v
index 05416490d..a9cb74349 100644
--- a/theories/program_logic/weakestpre.v
+++ b/theories/program_logic/weakestpre.v
@@ -371,8 +371,9 @@ Section proofmode_classes.
   Implicit Types Φ : val Λ → iProp Σ.
 
   Global Instance frame_wp p s E e R Φ Ψ :
-    (∀ v, Frame p R (Φ v) (Ψ v)) → Frame p R (WP e @ s; E {{ Φ }}) (WP e @ s; E {{ Ψ }}).
-  Proof. rewrite /Frame=> HR. rewrite wp_frame_l. apply wp_mono, HR. Qed.
+    (∀ v, Frame p R (Φ v) (Ψ v)) →
+    KnownFrame p R (WP e @ s; E {{ Φ }}) (WP e @ s; E {{ Ψ }}).
+  Proof. rewrite /KnownFrame /Frame=> HR. rewrite wp_frame_l. apply wp_mono, HR. Qed.
 
   Global Instance is_except_0_wp s E e Φ : IsExcept0 (WP e @ s; E {{ Φ }}).
   Proof. by rewrite /IsExcept0 -{2}fupd_wp -except_0_fupd -fupd_intro. Qed.
diff --git a/theories/proofmode/class_instances.v b/theories/proofmode/class_instances.v
index 730dea755..40a67b10d 100644
--- a/theories/proofmode/class_instances.v
+++ b/theories/proofmode/class_instances.v
@@ -821,23 +821,26 @@ Proof.
 Qed.
 
 (* Frame *)
-Global Instance frame_here_absorbing p R : Absorbing R → Frame p R R True | 0.
-Proof. intros. by rewrite /Frame affinely_persistently_if_elim sep_elim_l. Qed.
-Global Instance frame_here p R : Frame p R R emp | 1.
-Proof. intros. by rewrite /Frame affinely_persistently_if_elim sep_elim_l. Qed.
+Global Instance frame_here_absorbing p R : Absorbing R → KnownFrame p R R True | 0.
+Proof. intros. by rewrite /KnownFrame /Frame affinely_persistently_if_elim sep_elim_l. Qed.
+Global Instance frame_here p R : KnownFrame p R R emp | 1.
+Proof. intros. by rewrite /KnownFrame /Frame affinely_persistently_if_elim sep_elim_l. Qed.
 Global Instance frame_affinely_here_absorbing p R :
-  Absorbing R → Frame p (bi_affinely R) R True | 0.
+  Absorbing R → KnownFrame p (bi_affinely R) R True | 0.
 Proof.
-  intros. by rewrite /Frame affinely_persistently_if_elim affinely_elim sep_elim_l.
+  intros. rewrite /KnownFrame /Frame affinely_persistently_if_elim affinely_elim.
+  apply sep_elim_l, _.
 Qed.
-Global Instance frame_affinely_here p R : Frame p (bi_affinely R) R emp | 1.
+Global Instance frame_affinely_here p R : KnownFrame p (bi_affinely R) R emp | 1.
 Proof.
-  intros. by rewrite /Frame affinely_persistently_if_elim affinely_elim sep_elim_l.
+  intros. rewrite /KnownFrame /Frame affinely_persistently_if_elim affinely_elim.
+  apply sep_elim_l, _.
 Qed.
 
-Global Instance frame_here_pure p φ Q : FromPure false Q φ → Frame p ⌜φ⌝ Q True.
+Global Instance frame_here_pure p φ Q : FromPure false Q φ → KnownFrame p ⌜φ⌝ Q True.
 Proof.
-  rewrite /FromPure /Frame=> <-. by rewrite affinely_persistently_if_elim sep_elim_l.
+  rewrite /FromPure /KnownFrame /Frame=> <-.
+  by rewrite affinely_persistently_if_elim sep_elim_l.
 Qed.
 
 Global Instance make_embed_pure `{BiEmbedding PROP PROP'} φ :
@@ -851,14 +854,16 @@ Global Instance make_embed_default `{BiEmbedding PROP PROP'} P :
 Proof. by rewrite /MakeEmbed. Qed.
 
 Global Instance frame_embed `{BiEmbedding PROP PROP'} p P Q (Q' : PROP') R :
-  Frame p R P Q → MakeEmbed Q Q' → Frame p ⎡R⎤ ⎡P⎤ Q'.
+  Frame p R P Q → MakeEmbed Q Q' → KnownFrame p ⎡R⎤ ⎡P⎤ Q'.
 Proof.
-  rewrite /Frame /MakeEmbed => <- <-.
+  rewrite /KnownFrame /Frame /MakeEmbed => <- <-.
   rewrite bi_embed_sep bi_embed_affinely_if bi_embed_persistently_if => //.
 Qed.
 Global Instance frame_pure_embed `{BiEmbedding PROP PROP'} p P Q (Q' : PROP') φ :
-  Frame p ⌜φ⌝ P Q → MakeEmbed Q Q' → Frame p ⌜φ⌝ ⎡P⎤ Q'.
-Proof. rewrite /Frame /MakeEmbed -bi_embed_pure. apply (frame_embed p P Q). Qed.
+  Frame p ⌜φ⌝ P Q → MakeEmbed Q Q' → KnownFrame p ⌜φ⌝ ⎡P⎤ Q'.
+Proof.
+  rewrite /KnownFrame /Frame /MakeEmbed -bi_embed_pure. apply (frame_embed p P Q).
+Qed.
 
 Global Instance make_sep_emp_l P : KnownLMakeSep emp P P.
 Proof. apply left_id, _. Qed.
@@ -877,29 +882,31 @@ Proof. by rewrite /MakeSep. Qed.
 
 Global Instance frame_sep_persistent_l progress R P1 P2 Q1 Q2 Q' :
   Frame true R P1 Q1 → MaybeFrame true R P2 Q2 progress → MakeSep Q1 Q2 Q' →
-  Frame true R (P1 ∗ P2) Q' | 9.
+  KnownFrame true R (P1 ∗ P2) Q' | 9.
 Proof.
-  rewrite /Frame /MaybeFrame /MakeSep /= => <- <- <-.
+  rewrite /KnownFrame /Frame /MaybeFrame /MakeSep /= => <- <- <-.
   rewrite {1}(affinely_persistently_sep_dup R). solve_sep_entails.
 Qed.
 Global Instance frame_sep_l R P1 P2 Q Q' :
-  Frame false R P1 Q → MakeSep Q P2 Q' → Frame false R (P1 ∗ P2) Q' | 9.
-Proof. rewrite /Frame /MakeSep => <- <-. by rewrite assoc. Qed.
+  Frame false R P1 Q → MakeSep Q P2 Q' → KnownFrame false R (P1 ∗ P2) Q' | 9.
+Proof. rewrite /KnownFrame /Frame /MakeSep => <- <-. by rewrite assoc. Qed.
 Global Instance frame_sep_r p R P1 P2 Q Q' :
-  Frame p R P2 Q → MakeSep P1 Q Q' → Frame p R (P1 ∗ P2) Q' | 10.
-Proof. rewrite /Frame /MakeSep => <- <-. by rewrite assoc -(comm _ P1) assoc. Qed.
+  Frame p R P2 Q → MakeSep P1 Q Q' → KnownFrame p R (P1 ∗ P2) Q' | 10.
+Proof.
+  rewrite /KnownFrame /Frame /MakeSep => <- <-. by rewrite assoc -(comm _ P1) assoc.
+Qed.
 
 Global Instance frame_big_sepL_cons {A} p (Φ : nat → A → PROP) R Q l x l' :
   IsCons l x l' →
   Frame p R (Φ 0 x ∗ [∗ list] k ↦ y ∈ l', Φ (S k) y) Q →
-  Frame p R ([∗ list] k ↦ y ∈ l, Φ k y) Q.
-Proof. rewrite /IsCons=>->. by rewrite /Frame big_sepL_cons. Qed.
+  KnownFrame p R ([∗ list] k ↦ y ∈ l, Φ k y) Q.
+Proof. rewrite /IsCons=>->. by rewrite /KnownFrame /Frame big_sepL_cons. Qed.
 Global Instance frame_big_sepL_app {A} p (Φ : nat → A → PROP) R Q l l1 l2 :
   IsApp l l1 l2 →
   Frame p R (([∗ list] k ↦ y ∈ l1, Φ k y) ∗
            [∗ list] k ↦ y ∈ l2, Φ (length l1 + k) y) Q →
-  Frame p R ([∗ list] k ↦ y ∈ l, Φ k y) Q.
-Proof. rewrite /IsApp=>->. by rewrite /Frame big_opL_app. Qed.
+  KnownFrame p R ([∗ list] k ↦ y ∈ l, Φ k y) Q.
+Proof. rewrite /IsApp=>->. by rewrite /KnownFrame /Frame big_opL_app. Qed.
 
 Global Instance make_and_true_l P : KnownLMakeAnd True P P.
 Proof. apply left_id, _. Qed.
@@ -921,9 +928,9 @@ Global Instance frame_and p progress1 progress2 R P1 P2 Q1 Q2 Q' :
   MaybeFrame p R P2 Q2 progress2 →
   TCEq (progress1 || progress2) true →
   MakeAnd Q1 Q2 Q' →
-  Frame p R (P1 ∧ P2) Q' | 9.
+  KnownFrame p R (P1 ∧ P2) Q' | 9.
 Proof.
-  rewrite /MaybeFrame /Frame /MakeAnd => <- <- _ <-. apply and_intro;
+  rewrite /MaybeFrame /KnownFrame /Frame /MakeAnd => <- <- _ <-. apply and_intro;
   [rewrite and_elim_l|rewrite and_elim_r]; done.
 Qed.
 
@@ -941,9 +948,9 @@ Proof. by rewrite /MakeOr. Qed.
 (* We could in principle write the instance [frame_or_spatial] by a bunch of
 instances, i.e. (omitting the parameter [p = false]):
 
-  Frame R P1 Q1 → Frame R P2 Q2 → Frame R (P1 ∨ P2) (Q1 ∨ Q2)
-  Frame R P1 True → Frame R (P1 ∨ P2) P2
-  Frame R P2 True → Frame R (P1 ∨ P2) P1
+  Frame R P1 Q1 → Frame R P2 Q2 → KnownFrame R (P1 ∨ P2) (Q1 ∨ Q2)
+  Frame R P1 True → KnownFrame R (P1 ∨ P2) P2
+  Frame R P2 True → KnownFrame R (P1 ∨ P2) P1
 
 The problem here is that Coq will try to infer [Frame R P1 ?] and [Frame R P2 ?]
 multiple times, whereas the current solution makes sure that said inference
@@ -957,19 +964,19 @@ Global Instance frame_or_spatial progress1 progress2 R P1 P2 Q1 Q2 Q :
     (TCAnd (TCEq progress1 true) (TCEq Q1 True%I))
     (TCAnd (TCEq progress2 true) (TCEq Q2 True%I))) →
   MakeOr Q1 Q2 Q →
-  Frame false R (P1 ∨ P2) Q | 9.
-Proof. rewrite /Frame /MakeOr => <- <- _ <-. by rewrite -sep_or_l. Qed.
+  KnownFrame false R (P1 ∨ P2) Q | 9.
+Proof. rewrite /KnownFrame /Frame /MakeOr => <- <- _ <-. by rewrite -sep_or_l. Qed.
 
 Global Instance frame_or_persistent progress1 progress2 R P1 P2 Q1 Q2 Q :
   MaybeFrame true R P1 Q1 progress1 → MaybeFrame true R P2 Q2 progress2 →
   TCEq (progress1 || progress2) true →
-  MakeOr Q1 Q2 Q → Frame true R (P1 ∨ P2) Q | 9.
-Proof. rewrite /Frame /MakeOr => <- <- _ <-. by rewrite -sep_or_l. Qed.
+  MakeOr Q1 Q2 Q → KnownFrame true R (P1 ∨ P2) Q | 9.
+Proof. rewrite /KnownFrame /Frame /MakeOr => <- <- _ <-. by rewrite -sep_or_l. Qed.
 
 Global Instance frame_wand p R P1 P2 Q2 :
-  Frame p R P2 Q2 → Frame p R (P1 -∗ P2) (P1 -∗ Q2).
+  Frame p R P2 Q2 → KnownFrame p R (P1 -∗ P2) (P1 -∗ Q2).
 Proof.
-  rewrite /Frame=> ?. apply wand_intro_l.
+  rewrite /KnownFrame /Frame=> ?. apply wand_intro_l.
   by rewrite assoc (comm _ P1) -assoc wand_elim_r.
 Qed.
 
@@ -981,9 +988,9 @@ Global Instance make_affinely_default P : MakeAffinely P (bi_affinely P) | 100.
 Proof. by rewrite /MakeAffinely. Qed.
 
 Global Instance frame_affinely R P Q Q' :
-  Frame true R P Q → MakeAffinely Q Q' → Frame true R (bi_affinely P) Q'.
+  Frame true R P Q → MakeAffinely Q Q' → KnownFrame true R (bi_affinely P) Q'.
 Proof.
-  rewrite /Frame /MakeAffinely=> <- <- /=.
+  rewrite /KnownFrame /Frame /MakeAffinely=> <- <- /=.
   by rewrite -{1}affinely_idemp affinely_sep_2.
 Qed.
 
@@ -999,8 +1006,10 @@ Global Instance make_absorbingly_default P : MakeAbsorbingly P (bi_absorbingly P
 Proof. by rewrite /MakeAbsorbingly. Qed.
 
 Global Instance frame_absorbingly p R P Q Q' :
-  Frame p R P Q → MakeAbsorbingly Q Q' → Frame p R (bi_absorbingly P) Q'.
-Proof. rewrite /Frame /MakeAbsorbingly=> <- <- /=. by rewrite absorbingly_sep_r. Qed.
+  Frame p R P Q → MakeAbsorbingly Q Q' → KnownFrame p R (bi_absorbingly P) Q'.
+Proof.
+  rewrite /KnownFrame /Frame /MakeAbsorbingly=> <- <- /=. by rewrite absorbingly_sep_r.
+Qed.
 
 Global Instance make_persistently_true : @KnownMakePersistently PROP True True.
 Proof. by rewrite /KnownMakePersistently /MakePersistently persistently_pure. Qed.
@@ -1014,38 +1023,44 @@ Global Instance make_persistently_default P :
 Proof. by rewrite /MakePersistently. Qed.
 
 Global Instance frame_persistently R P Q Q' :
-  Frame true R P Q → MakePersistently Q Q' → Frame true R (bi_persistently P) Q'.
+  Frame true R P Q → MakePersistently Q Q' → KnownFrame true R (bi_persistently P) Q'.
 Proof.
-  rewrite /Frame /MakePersistently=> <- <- /=. rewrite -persistently_and_affinely_sep_l.
+  rewrite /KnownFrame /Frame /MakePersistently=> <- <- /=.
+  rewrite -persistently_and_affinely_sep_l.
   by rewrite -persistently_sep_2 -persistently_and_sep_l_1 persistently_affinely
               persistently_idemp.
 Qed.
 
 Global Instance frame_exist {A} p R (Φ Ψ : A → PROP) :
-  (∀ a, Frame p R (Φ a) (Ψ a)) → Frame p R (∃ x, Φ x) (∃ x, Ψ x).
-Proof. rewrite /Frame=> ?. by rewrite sep_exist_l; apply exist_mono. Qed.
+  (∀ a, Frame p R (Φ a) (Ψ a)) → KnownFrame p R (∃ x, Φ x) (∃ x, Ψ x).
+Proof. rewrite /KnownFrame /Frame=> ?. by rewrite sep_exist_l; apply exist_mono. Qed.
 Global Instance frame_forall {A} p R (Φ Ψ : A → PROP) :
-  (∀ a, Frame p R (Φ a) (Ψ a)) → Frame p R (∀ x, Φ x) (∀ x, Ψ x).
-Proof. rewrite /Frame=> ?. by rewrite sep_forall_l; apply forall_mono. Qed.
+  (∀ a, Frame p R (Φ a) (Ψ a)) → KnownFrame p R (∀ x, Φ x) (∀ x, Ψ x).
+Proof. rewrite /KnownFrame /Frame=> ?. by rewrite sep_forall_l; apply forall_mono. Qed.
 
 Global Instance frame_impl_persistent R P1 P2 Q2 :
-  Frame true R P2 Q2 → Frame true R (P1 → P2) (P1 → Q2).
+  Frame true R P2 Q2 → KnownFrame true R (P1 → P2) (P1 → Q2).
 Proof.
-  rewrite /Frame /= => ?. apply impl_intro_l.
+  rewrite /KnownFrame /Frame /= => ?. apply impl_intro_l.
   by rewrite -persistently_and_affinely_sep_l assoc (comm _ P1) -assoc impl_elim_r
              persistently_and_affinely_sep_l.
 Qed.
 Global Instance frame_impl R P1 P2 Q2 :
   Persistent P1 → Absorbing P1 →
-  Frame false R P2 Q2 → Frame false R (P1 → P2) (P1 → Q2).
+  Frame false R P2 Q2 → KnownFrame false R (P1 → P2) (P1 → Q2).
 Proof.
-  rewrite /Frame /==> ???. apply impl_intro_l.
+  rewrite /KnownFrame /Frame /==> ???. apply impl_intro_l.
   rewrite {1}(persistent P1) persistently_and_affinely_sep_l assoc.
   rewrite (comm _ (â–¡ P1)%I) -assoc -persistently_and_affinely_sep_l.
   rewrite persistently_elim impl_elim_r //.
 Qed.
 
-(* ElimModal *)
+(* In the case nothing else suceeded, then we try to instantiate the
+   evar [P] with [R ∗ Q]. This only does so in non-persistent mode:
+   The idea is that persistent assumptions will still be available
+   when [P] will be used, so there is no need to keep it in [P]. *)
+Global Instance frame_sep_instantiate R Q : Frame false R (R ∗ Q) Q | 100.
+Proof. by rewrite /Frame. Qed.
 
 (* IntoEmbed *)
 Global Instance into_embed_embed {PROP' : bi} `{BiEmbed PROP PROP'} P :
@@ -1504,8 +1519,8 @@ Proof. by rewrite /AddModal fupd_frame_r wand_elim_r fupd_trans. Qed.
 (* Frame *)
 Global Instance frame_eq_embed `{SbiEmbedding PROP PROP'} p P Q (Q' : PROP')
        {A : ofeT} (a b : A) :
-  Frame p (a ≡ b) P Q → MakeEmbed Q Q' → Frame p (a ≡ b) ⎡P⎤ Q'.
-Proof. rewrite /Frame /MakeEmbed -sbi_embed_internal_eq. apply (frame_embed p P Q). Qed.
+  Frame p (a ≡ b) P Q → MakeEmbed Q Q' → KnownFrame p (a ≡ b) ⎡P⎤ Q'.
+Proof. rewrite /KnownFrame /Frame /MakeEmbed -sbi_embed_internal_eq. apply (frame_embed p P Q). Qed.
 
 Global Instance make_laterN_true n : @KnownMakeLaterN PROP n True True | 0.
 Proof. by rewrite /KnownMakeLaterN /MakeLaterN laterN_True. Qed.
@@ -1518,25 +1533,25 @@ Proof. by rewrite /MakeLaterN. Qed.
 
 Global Instance frame_later p R R' P Q Q' :
   NoBackTrack (MaybeIntoLaterN true 1 R' R) →
-  Frame p R P Q → MakeLaterN 1 Q Q' → Frame p R' (▷ P) Q'.
+  Frame p R P Q → MakeLaterN 1 Q Q' → KnownFrame p R' (▷ P) Q'.
 Proof.
-  rewrite /Frame /MakeLaterN /MaybeIntoLaterN=>-[->] <- <-.
+  rewrite /KnownFrame /Frame /MakeLaterN /MaybeIntoLaterN=>-[->] <- <-.
   by rewrite later_affinely_persistently_if_2 later_sep.
 Qed.
 Global Instance frame_laterN p n R R' P Q Q' :
   NoBackTrack (MaybeIntoLaterN true n R' R) →
-  Frame p R P Q → MakeLaterN n Q Q' → Frame p R' (▷^n P) Q'.
+  Frame p R P Q → MakeLaterN n Q Q' → KnownFrame p R' (▷^n P) Q'.
 Proof.
-  rewrite /Frame /MakeLaterN /MaybeIntoLaterN=>-[->] <- <-.
+  rewrite /KnownFrame /Frame /MakeLaterN /MaybeIntoLaterN=>-[->] <- <-.
   by rewrite laterN_affinely_persistently_if_2 laterN_sep.
 Qed.
 
 Global Instance frame_bupd `{BUpdFacts PROP} p R P Q :
-  Frame p R P Q → Frame p R (|==> P) (|==> Q).
-Proof. rewrite /Frame=><-. by rewrite bupd_frame_l. Qed.
+  Frame p R P Q → KnownFrame p R (|==> P) (|==> Q).
+Proof. rewrite /KnownFrame /Frame=><-. by rewrite bupd_frame_l. Qed.
 Global Instance frame_fupd `{FUpdFacts PROP} p E1 E2 R P Q :
-  Frame p R P Q → Frame p R (|={E1,E2}=> P) (|={E1,E2}=> Q).
-Proof. rewrite /Frame=><-. by rewrite fupd_frame_l. Qed.
+  Frame p R P Q → KnownFrame p R (|={E1,E2}=> P) (|={E1,E2}=> Q).
+Proof. rewrite /KnownFrame /Frame=><-. by rewrite fupd_frame_l. Qed.
 
 Global Instance make_except_0_True : @KnownMakeExcept0 PROP True True.
 Proof. by rewrite /KnownMakeExcept0 /MakeExcept0 except_0_True. Qed.
@@ -1544,9 +1559,9 @@ Global Instance make_except_0_default P : MakeExcept0 P (â—‡ P) | 100.
 Proof. by rewrite /MakeExcept0. Qed.
 
 Global Instance frame_except_0 p R P Q Q' :
-  Frame p R P Q → MakeExcept0 Q Q' → Frame p R (◇ P) Q'.
+  Frame p R P Q → MakeExcept0 Q Q' → KnownFrame p R (◇ P) Q'.
 Proof.
-  rewrite /Frame /MakeExcept0=><- <-.
+  rewrite /KnownFrame /Frame /MakeExcept0=><- <-.
   by rewrite except_0_sep -(except_0_intro (â–¡?p R)%I).
 Qed.
 
diff --git a/theories/proofmode/classes.v b/theories/proofmode/classes.v
index 07c35206a..04ae2c701 100644
--- a/theories/proofmode/classes.v
+++ b/theories/proofmode/classes.v
@@ -253,7 +253,19 @@ Proof. done. Qed.
 Class Frame {PROP : bi} (p : bool) (R P Q : PROP) := frame : □?p R ∗ Q ⊢ P.
 Arguments Frame {_} _ _%I _%I _%I.
 Arguments frame {_ _} _%I _%I _%I {_}.
-Hint Mode Frame + + ! ! - : typeclass_instances.
+(* [P] has - hint mode even though it is morally an input. In the case
+   P is an evar, it will be instantiated with [R * ?P'], where [?P']
+   is a new evar. This mechanism should be the only instance able to
+   instanciate [P] when it is an evar. *)
+Hint Mode Frame + + ! - - : typeclass_instances.
+
+(* [KnownFrame] is like [Frame], but with an [Hint Mode] that make
+   sure we do not instantiate [P]. *)
+Class KnownFrame {PROP : bi} (p : bool) (R P Q : PROP) :=
+  known_frame :> Frame p R P Q.
+Arguments KnownFrame {_} _ _%I _%I _%I.
+Arguments known_frame {_ _} _%I _%I _%I {_}.
+Hint Mode KnownFrame + + ! ! - : typeclass_instances.
 
 (* The boolean [progress] indicates whether actual framing has been performed.
 If it is [false], then the default instance [maybe_frame_default] below has been
@@ -262,7 +274,7 @@ Class MaybeFrame {PROP : bi} (p : bool) (R P Q : PROP) (progress : bool) :=
   maybe_frame : □?p R ∗ Q ⊢ P.
 Arguments MaybeFrame {_} _ _%I _%I _%I _.
 Arguments maybe_frame {_} _ _%I _%I _%I _ {_}.
-Hint Mode MaybeFrame + + ! ! - - : typeclass_instances.
+Hint Mode MaybeFrame + + ! - - - : typeclass_instances.
 
 Instance maybe_frame_frame {PROP : bi} p (R P Q : PROP) :
   Frame p R P Q → MaybeFrame p R P Q true.
diff --git a/theories/proofmode/monpred.v b/theories/proofmode/monpred.v
index b5a92ce01..d89bfb20c 100644
--- a/theories/proofmode/monpred.v
+++ b/theories/proofmode/monpred.v
@@ -342,16 +342,16 @@ Qed.
    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) 𝓠.
+  KnownFrame p (R i) (P j) 𝓠.
 Proof.
-  rewrite /Frame /MakeMonPredAt /bi_affinely_if /bi_persistently_if
+  rewrite /KnownFrame /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) 𝓠.
+  Frame p ⎡𝓡⎤ P Q → MakeMonPredAt i Q 𝓠 → KnownFrame p 𝓡 (P i) 𝓠.
 Proof.
-  rewrite /Frame /MakeMonPredAt /bi_affinely_if /bi_persistently_if=> <- <-.
+  rewrite /KnownFrame /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.
diff --git a/theories/tests/proofmode.v b/theories/tests/proofmode.v
index 6b9e3804b..d97463a4b 100644
--- a/theories/tests/proofmode.v
+++ b/theories/tests/proofmode.v
@@ -371,13 +371,20 @@ Proof.
 Qed.
 
 Lemma iFrame_with_evar_r P Q :
-  P -∗ Q -∗ ∃ R, P ∗ R.
+  ∃ R, (P -∗ Q -∗ P ∗ R) ∧ R = (Q ∗ emp)%I.
 Proof.
-  iIntros "HP HQ". iExists _. iFrame. iApply "HQ".
+  eexists. split. iIntros "HP HQ".  iFrame. iEmpIntro. done.
 Qed.
 Lemma iFrame_with_evar_l P Q :
-  P -∗ Q -∗ ∃ R, R ∗ P.
+  ∃ R, (P -∗ Q -∗ R ∗ P) ∧ R = (Q ∗ emp)%I.
 Proof.
-  iIntros "HP HQ". iExists _. iFrame. iApply "HQ".
+  eexists. split. iIntros "HP HQ". iFrame "HQ".
+  iSplitR. iEmpIntro. done. done.
 Qed.
+Lemma iFrame_with_evar_persistent P Q :
+  ∃ R, (P -∗ □ Q -∗ P ∗ R ∗ Q) ∧ R = emp%I.
+Proof.
+  eexists. split. iIntros "HP #HQ". iFrame "HQ HP". iEmpIntro. done.
+Qed.
+
 End tests.
-- 
GitLab