diff --git a/theories/base_logic/big_op.v b/theories/base_logic/big_op.v
index 17df7d6928d8d1eeb50ac51ce0a435cb196fb081..a21e335232999815858404e062e9de732a4f42bb 100644
--- a/theories/base_logic/big_op.v
+++ b/theories/base_logic/big_op.v
@@ -312,6 +312,34 @@ Section list.
   Proof. rewrite /big_opL. apply _. Qed.
 End list.
 
+Section list2.
+  Context {A : Type}.
+  Implicit Types l : list A.
+  Implicit Types Φ Ψ : nat → A → uPred M.
+  (* Some lemmas depend on the generalized versions of the above ones. *)
+
+  Lemma big_sepL_zip_with {B C} Φ f (l1 : list B) (l2 : list C) :
+    ([∗ list] k↦x ∈ zip_with f l1 l2, Φ k x) ⊣⊢ ([∗ list] k↦x ∈ l1, ∀ y, ⌜l2 !! k = Some y⌝ → Φ k (f x y)).
+  Proof.
+    revert Φ l2; induction l1; intros Φ l2; first by rewrite /= big_sepL_nil.
+    destruct l2; simpl.
+    { rewrite big_sepL_nil. apply (anti_symm _); last exact: True_intro.
+      (* TODO: Can this be done simpler? *)
+      rewrite -(big_sepL_mono (λ _ _, True%I)).
+      - rewrite big_sepL_forall. apply forall_intro=>k. apply forall_intro=>b.
+        apply impl_intro_r. apply True_intro.
+      - intros k b _. apply forall_intro=>c. apply impl_intro_l. rewrite right_id.
+        apply pure_elim'. done.
+    }
+    rewrite !big_sepL_cons. apply sep_proper; last exact: IHl1.
+    apply (anti_symm _).
+    - apply forall_intro=>c'. simpl. apply impl_intro_r.
+      eapply pure_elim; first exact: and_elim_r. intros [=->].
+      apply and_elim_l.
+    - rewrite (forall_elim c). simpl. eapply impl_elim; first done.
+      apply pure_intro. done.
+  Qed.
+End list2.
 
 (** ** Big ops over finite maps *)
 Section gmap.
diff --git a/theories/heap_lang/lifting.v b/theories/heap_lang/lifting.v
index 807ae26b565fcd4ba3c95c44b626ae1b74a61f45..861e50cefadbf4cc99cb7c6b4b98108399cfbe28 100644
--- a/theories/heap_lang/lifting.v
+++ b/theories/heap_lang/lifting.v
@@ -64,7 +64,7 @@ Lemma wp_bind {E e} K Φ :
   WP e @ E {{ v, WP fill K (of_val v) @ E {{ Φ }} }} ⊢ WP fill K e @ E {{ Φ }}.
 Proof. exact: wp_ectx_bind. Qed.
 
-Lemma wp_bind_ctxi {E e} Ki Φ :
+Lemma wp_bindi {E e} Ki Φ :
   WP e @ E {{ v, WP fill_item Ki (of_val v) @ E {{ Φ }} }} ⊢
      WP fill_item Ki e @ E {{ Φ }}.
 Proof. exact: weakestpre.wp_bind. Qed.