diff --git a/program_logic/hoare.v b/program_logic/hoare.v
index 4fa4299a1a3d0c278af18fe6e8c540f760124ad2..38178d2197b4229ce458adf308718d9ce39313d0 100644
--- a/program_logic/hoare.v
+++ b/program_logic/hoare.v
@@ -114,8 +114,7 @@ Proof.
   iSplitL "HR"; [|by iApply "Hwp"].
   iPvs "Hvs1" "HR"; first by set_solver.
   iPvsIntro. iNext.
-  iPvs "Hvs2" "Hvs1"; first by set_solver.
-  by iPvsIntro.
+  by iPvs "Hvs2" "Hvs1"; first by set_solver.
 Qed.
 
 Lemma ht_frame_step_r E E1 E2 P R1 R2 R3 e Φ :
@@ -125,7 +124,7 @@ Lemma ht_frame_step_r E E1 E2 P R1 R2 R3 e Φ :
 Proof.
   iIntros {???} "[#Hvs1 [#Hvs2 #Hwp]]".
   setoid_rewrite (comm _ _ R3).
-  iApply ht_frame_step_l; try eassumption. repeat iSplit; by iIntros "!".
+  iApply (ht_frame_step_l _ _ E2); by repeat iSplit.
 Qed.
 
 Lemma ht_frame_step_l' E P R e Φ :
diff --git a/program_logic/viewshifts.v b/program_logic/viewshifts.v
index 5bf652bac7b7c199e149c0ec6330ebceff6fd5ac..f513c6cdf3524f81cc4c7f05b684283054f81517 100644
--- a/program_logic/viewshifts.v
+++ b/program_logic/viewshifts.v
@@ -71,10 +71,10 @@ Qed.
 Lemma vs_transitive' E P Q R : ((P ={E}=> Q) ∧ (Q ={E}=> R)) ⊢ (P ={E}=> R).
 Proof. apply vs_transitive; set_solver. Qed.
 Lemma vs_reflexive E P : P ={E}=> P.
-Proof. iIntros "! HP"; by iPvsIntro. Qed.
+Proof. by iIntros "! HP". Qed.
 
 Lemma vs_impl E P Q : □ (P → Q) ⊢ (P ={E}=> Q).
-Proof. iIntros "#HPQ ! HP". iPvsIntro. by iApply "HPQ". Qed.
+Proof. iIntros "#HPQ ! HP". by iApply "HPQ". Qed.
 
 Lemma vs_frame_l E1 E2 P Q R : (P ={E1,E2}=> Q) ⊢ (R ★ P ={E1,E2}=> R ★ Q).
 Proof. iIntros "#Hvs ! [HR HP]". iFrame "HR". by iApply "Hvs". Qed.
diff --git a/proofmode/coq_tactics.v b/proofmode/coq_tactics.v
index 3502c54ba39ca0ef4b7f5a39536b577b84f16f47..67a1e6c1fd322cf3056813e04c110fbdbe312e83 100644
--- a/proofmode/coq_tactics.v
+++ b/proofmode/coq_tactics.v
@@ -260,8 +260,17 @@ Proof.
 Qed.
 
 (** * Basic rules *)
-Lemma tac_exact Δ i p P : envs_lookup i Δ = Some (p,P) → Δ ⊢ P.
-Proof. intros. by rewrite envs_lookup_sound' // sep_elim_l. Qed.
+Class ToAssumption (p : bool) (P Q : uPred M) :=
+  to_assumption : (if p then □ P else P) ⊢ Q.
+Global Instance to_assumption_exact p P : ToAssumption p P P.
+Proof. destruct p; by rewrite /ToAssumption ?always_elim. Qed.
+Global Instance to_assumption_always P Q :
+  ToAssumption true P Q → ToAssumption true P (□ Q).
+Proof. rewrite /ToAssumption=><-. by rewrite always_always. Qed.
+
+Lemma tac_assumption Δ i p P Q :
+  envs_lookup i Δ = Some (p,P) → ToAssumption p P Q → Δ ⊢ Q.
+Proof. intros. by rewrite envs_lookup_sound // sep_elim_l. Qed.
 
 Lemma tac_rename Δ Δ' i j p P Q :
   envs_lookup i Δ = Some (p,P) →
diff --git a/proofmode/pviewshifts.v b/proofmode/pviewshifts.v
index 12ed0cd38be92d55aa53e1e27715fa02efafdb56..8f77deafe28c52097d0808a47e3234e8325d3e85 100644
--- a/proofmode/pviewshifts.v
+++ b/proofmode/pviewshifts.v
@@ -7,6 +7,9 @@ Section pvs.
 Context {Λ : language} {Σ : iFunctor}.
 Implicit Types P Q : iProp Λ Σ.
 
+Global Instance to_assumption_pvs E p P Q :
+  ToAssumption p P Q → ToAssumption p P (|={E}=> Q)%I.
+Proof. rewrite /ToAssumption=>->. apply pvs_intro. Qed.
 Global Instance sep_split_pvs E P Q1 Q2 :
   SepSplit P Q1 Q2 → SepSplit (|={E}=> P) (|={E}=> Q1) (|={E}=> Q2).
 Proof. rewrite /SepSplit=><-. apply pvs_sep. Qed.
diff --git a/proofmode/tactics.v b/proofmode/tactics.v
index 6a005141679edad918809f2646cf2686d89a293a..8b65fdc2372319fe22b98e40a8336ba92084ace4 100644
--- a/proofmode/tactics.v
+++ b/proofmode/tactics.v
@@ -63,13 +63,10 @@ Tactic Notation "iClear" "★" :=
 
 (** * Assumptions *)
 Tactic Notation "iExact" constr(H) :=
-  eapply tac_exact with H _; (* (i:=H) *)
-    env_cbv;
-    lazymatch goal with
-    | |- None = Some _ => fail "iExact:" H "not found"
-    | |- Some (_, ?P) = Some _ =>
-       reflexivity || fail "iExact:" H ":" P "does not match goal"
-    end.
+  eapply tac_assumption with H _ _; (* (i:=H) *)
+    [env_cbv; reflexivity || fail "iExact:" H "not found"
+    |let P := match goal with |- ToAssumption _ ?P _ => P end in
+     apply _ || fail "iExact:" H ":" P "does not match goal"].
 
 Tactic Notation "iAssumptionCore" :=
   let rec find Γ i P :=
@@ -82,9 +79,21 @@ Tactic Notation "iAssumptionCore" :=
   | |- envs_lookup ?i (Envs ?Γp ?Γs) = Some (_, ?P) =>
      is_evar i; first [find Γp i P | find Γs i P]; env_cbv; reflexivity
   end.
+
 Tactic Notation "iAssumption" :=
-  eapply tac_exact; iAssumptionCore;
-  match goal with |- _ = Some (_, ?P) => fail "iAssumption:" P "not found" end.
+  let Hass := fresh in
+  let rec find p Γ Q :=
+    match Γ with
+    | Esnoc ?Γ ?j ?P => first
+       [pose proof (_ : ToAssumption p P Q) as Hass;
+        apply (tac_assumption _ j p P); [env_cbv; reflexivity|apply Hass]
+       |find p Γ Q]
+    end in
+  match goal with
+  | |- of_envs (Envs ?Γp ?Γs) ⊢ ?Q =>
+     first [find true Γp Q | find false Γs Q
+           |fail "iAssumption:" Q "not found"]
+  end.
 
 (** * False *)
 Tactic Notation "iExFalso" := apply tac_ex_falso.