diff --git a/theories/heap_lang/proofmode.v b/theories/heap_lang/proofmode.v
index 4f97b4544cc81ad8b7eceb51f9be11b45dc3be1b..8263ff881c0f6b20696397596bcbb48828d02423 100644
--- a/theories/heap_lang/proofmode.v
+++ b/theories/heap_lang/proofmode.v
@@ -5,15 +5,15 @@ From iris.heap_lang Require Export tactics lifting.
 Set Default Proof Using "Type".
 Import uPred.
 
-Lemma tac_wp_pure `{heapG Σ} K Δ Δ' E e1 e2 φ Φ :
+Lemma tac_wp_pure `{heapG Σ} Δ Δ' E e1 e2 φ Φ :
   PureExec φ e1 e2 →
   φ →
   IntoLaterNEnvs 1 Δ Δ' →
-  envs_entails Δ' (WP fill K e2 @ E {{ Φ }}) →
-  envs_entails Δ (WP fill K e1 @ E {{ Φ }}).
+  envs_entails Δ' (WP e2 @ E {{ Φ }}) →
+  envs_entails Δ (WP e1 @ E {{ Φ }}).
 Proof.
   rewrite /envs_entails=> ??? HΔ'. rewrite into_laterN_env_sound /=.
-  rewrite -wp_bind HΔ' -wp_pure_step_later //. by rewrite -wp_bind_inv.
+  rewrite HΔ' -wp_pure_step_later //.
 Qed.
 
 Lemma tac_wp_value `{heapG Σ} Δ E Φ e v :
@@ -26,13 +26,16 @@ Ltac wp_value_head := eapply tac_wp_value; [apply _|lazy beta].
 Tactic Notation "wp_pure" open_constr(efoc) :=
   iStartProof;
   lazymatch goal with
-  | |- envs_entails _ (wp ?E ?e ?Q) => reshape_expr e ltac:(fun K e' =>
-    unify e' efoc;
-    eapply (tac_wp_pure K);
-    [simpl; apply _                 (* PureExec *)
-    |try fast_done                  (* The pure condition for PureExec *)
-    |apply _                        (* IntoLaters *)
-    |simpl_subst; try wp_value_head (* new goal *)])
+  | |- envs_entails _ (wp ?E ?e ?Q) =>
+    let e := eval simpl in e in
+    reshape_expr e ltac:(fun K e' =>
+      unify e' efoc;
+      eapply (tac_wp_pure _ _ _ (fill K e'));
+      [apply _                        (* PureExec *)
+      |try fast_done                  (* The pure condition for PureExec *)
+      |apply _                        (* IntoLaters *)
+      |simpl_subst; try wp_value_head (* new goal *)
+      ])
    || fail "wp_pure: cannot find" efoc "in" e "or" efoc "is not a reduct"
   | _ => fail "wp_pure: not a 'wp'"
   end.
diff --git a/theories/program_logic/ectx_language.v b/theories/program_logic/ectx_language.v
index 34a366329c31312b5899a34103ca6515d157b02c..3c20309d309e3874a64b1d9ed8932fb65f55fa28 100644
--- a/theories/program_logic/ectx_language.v
+++ b/theories/program_logic/ectx_language.v
@@ -205,6 +205,12 @@ Section ectx_language.
       eexists e2', σ2, efs. by apply head_prim_step.
     - intros σ1 e2' σ2 efs ? ?%head_reducible_prim_step; eauto.
   Qed.
+
+  Global Instance pure_exec_fill K e1 e2 φ :
+    PureExec φ e1 e2 →
+    PureExec φ (fill K e1) (fill K e2).
+  Proof. apply: pure_exec_ctx. Qed.
+
 End ectx_language.
 
 Arguments ectx_lang : clear implicits.
diff --git a/theories/program_logic/language.v b/theories/program_logic/language.v
index 660965699c6e6b5d91f7a630b067b4129809a79d..a3704c7029b09f3873f4bd844ad2cfcede647c92 100644
--- a/theories/program_logic/language.v
+++ b/theories/program_logic/language.v
@@ -140,6 +140,19 @@ Section language.
     PureExec P e1 e2.
   Proof. intros HPE. split; intros; eapply HPE; eauto. Qed.
 
+  (* We do not make this an instance because it is awfully general. *)
+  Lemma pure_exec_ctx K `{LanguageCtx Λ K} e1 e2 φ :
+    PureExec φ e1 e2 →
+    PureExec φ (K e1) (K e2).
+  Proof.
+    intros [Hred Hstep]. split.
+    - unfold reducible in *. naive_solver eauto using fill_step.
+    - intros σ1 e2' σ2 efs ? Hpstep.
+      destruct (fill_step_inv e1 σ1 e2' σ2 efs) as (e2'' & -> & ?); [|exact Hpstep|].
+      + destruct (Hred σ1) as (? & ? & ? & ?); eauto using val_stuck.
+      + edestruct (Hstep σ1 e2'' σ2 efs) as (-> & -> & ->); auto.
+  Qed.
+
   (* This is a family of frequent assumptions for PureExec *)
   Class IntoVal (e : expr Λ) (v : val Λ) :=
     into_val : to_val e = Some v.