diff --git a/iris/base_logic/lib/gen_heap.v b/iris/base_logic/lib/gen_heap.v
index a172b30eaebbad236bcc269cf6a37ccce6e2e872..005a01456d8c28c0cab063bbdb8c1a6b0eaa9464 100644
--- a/iris/base_logic/lib/gen_heap.v
+++ b/iris/base_logic/lib/gen_heap.v
@@ -196,7 +196,7 @@ Section gen_heap.
     meta_token l E1 -∗ meta_token l E2 -∗ meta_token l (E1 ∪ E2).
   Proof.
     rewrite meta_token_eq /meta_token_def.
-    iDestruct 1 as (γm1) "[#Hγm1 Hm1]". iDestruct 1 as (γm2) "[#Hγm2 Hm2]".
+    iIntros "(%γm1 & #Hγm1 & Hm1) (%γm2 & #Hγm2 & Hm2)".
     iDestruct (ghost_map_elem_valid_2 with "Hγm1 Hγm2") as %[_ ->].
     iDestruct (own_valid_2 with "Hm1 Hm2") as %?%reservation_map_token_valid_op.
     iExists γm2. iFrame "Hγm2". rewrite reservation_map_token_union //. by iSplitL "Hm1".
@@ -219,7 +219,7 @@ Section gen_heap.
     meta l i x1 -∗ meta l i x2 -∗ ⌜x1 = x2⌝.
   Proof.
     rewrite meta_eq /meta_def.
-    iDestruct 1 as (γm1) "[Hγm1 Hm1]"; iDestruct 1 as (γm2) "[Hγm2 Hm2]".
+    iIntros "(%γm1 & Hγm1 & Hm1) (%γm2 & Hγm2 & Hm2)".
     iDestruct (ghost_map_elem_valid_2 with "Hγm1 Hγm2") as %[_ ->].
     iDestruct (own_valid_2 with "Hm1 Hm2") as %Hγ; iPureIntro.
     move: Hγ. rewrite -reservation_map_data_op reservation_map_data_valid.
diff --git a/iris/base_logic/lib/na_invariants.v b/iris/base_logic/lib/na_invariants.v
index 4c1faf91d2e49cafcae46d9991c325d094c35425..71131cf0386d88d1da1d3a47068203bb7796994b 100644
--- a/iris/base_logic/lib/na_invariants.v
+++ b/iris/base_logic/lib/na_invariants.v
@@ -45,7 +45,7 @@ Section proofs.
 
   Lemma na_inv_iff p N P Q : na_inv p N P -∗ ▷ □ (P ↔ Q) -∗ na_inv p N Q.
   Proof.
-    iIntros "HI #HPQ". rewrite /na_inv. iDestruct "HI" as (i ?) "HI".
+    rewrite /na_inv. iIntros "(%i & % & HI) #HPQ".
     iExists i. iSplit; first done. iApply (inv_iff with "HI").
     iIntros "!> !>".
     iSplit; iIntros "[[? Ho]|$]"; iLeft; iFrame "Ho"; by iApply "HPQ".
@@ -97,8 +97,7 @@ Section proofs.
     na_inv p N P -∗ na_own p F ={E}=∗ ▷ P ∗ na_own p (F∖↑N) ∗
                        (▷ P ∗ na_own p (F∖↑N) ={E}=∗ na_own p F).
   Proof.
-    rewrite /na_inv. iIntros (??) "#Hnainv Htoks".
-    iDestruct "Hnainv" as (i) "[% Hinv]".
+    rewrite /na_inv. iIntros (??) "#(%i & % & Hinv) Htoks".
     rewrite [F as X in na_own p X](union_difference_L (↑N) F) //.
     rewrite [X in (X ∪ _)](union_difference_L {[i]} (↑N)) ?na_own_union; [|set_solver..].
     iDestruct "Htoks" as "[[Htoki $] $]".
diff --git a/iris_heap_lang/lib/array.v b/iris_heap_lang/lib/array.v
index 4a844ee4e3b8a9af5ea0b9c03482f9450cf04fc7..86598a764c749236289e61dc46eb0341aa83c35e 100644
--- a/iris_heap_lang/lib/array.v
+++ b/iris_heap_lang/lib/array.v
@@ -234,7 +234,7 @@ Section proof.
     Proof.
       iIntros "Hvs". iInduction vs as [|v vs] "IH" forall (Q); simpl.
       { iExists []. by auto. }
-      iDestruct "Hvs" as "[Hv Hvs]"; iDestruct "Hv" as (x ->) "Hv".
+      iDestruct "Hvs" as "[(%x & -> & Hv) Hvs]".
       iDestruct ("IH" with "Hvs") as (xs ->) "Hxs".
       iExists (x :: xs). by iFrame.
     Qed.
diff --git a/iris_heap_lang/lib/ticket_lock.v b/iris_heap_lang/lib/ticket_lock.v
index 05cae40e0ef87778df210a9b573fbb55758b9f5c..a36221a3a58b05d003fa2e563871a48d0bb73256 100644
--- a/iris_heap_lang/lib/ticket_lock.v
+++ b/iris_heap_lang/lib/ticket_lock.v
@@ -65,7 +65,7 @@ Section proof.
 
   Lemma locked_exclusive (γ : gname) : locked γ -∗ locked γ -∗ False.
   Proof.
-    iDestruct 1 as (o1) "H1". iDestruct 1 as (o2) "H2".
+    iIntros "[%σ1 H1] [%σ2 H2]".
     iDestruct (own_valid_2 with "H1 H2") as %[[] _]%auth_frag_op_valid_1.
   Qed.
 
@@ -74,7 +74,7 @@ Section proof.
   Proof.
     iDestruct 1 as (lo ln ->) "#Hinv"; iIntros "#HR".
     iExists lo, ln; iSplit; [done|]. iApply (inv_iff with "Hinv").
-    iIntros "!> !>"; iSplit; iDestruct 1 as (o n) "(Ho & Hn & H● & H)";
+    iIntros "!> !>"; iSplit; iIntros "(%o & %n & Ho & Hn & H● & H)";
       iExists o, n; iFrame "Ho Hn H●";
       (iDestruct "H" as "[[Hâ—¯ H]|Hâ—¯]"; [iLeft; iFrame "Hâ—¯"; by iApply "HR"|by iRight]).
   Qed.