diff --git a/tests/list_reverse.ref b/tests/list_reverse.ref
index e69de29bb2d1d6434b8b29ae775ad8c2e48c5391..baa2732685cdddb4b868065a8973d881f6dddc52 100644
--- a/tests/list_reverse.ref
+++ b/tests/list_reverse.ref
@@ -0,0 +1,33 @@
+1 subgoal
+  
+  Σ : gFunctors
+  heapG0 : heapG Σ
+  hd, acc : val
+  xs, ys : list val
+  Φ : val → iPropI Σ
+  ============================
+  "Hxs" : is_list hd xs
+  "Hys" : is_list acc ys
+  "HΦ" : ∀ w : val, is_list w (reverse xs ++ ys) -∗ Φ w
+  --------------------------------------∗
+  WP (rev hd) acc [{ v, Φ v }]
+  
+1 subgoal
+  
+  Σ : gFunctors
+  heapG0 : heapG Σ
+  acc : val
+  ys : list val
+  Φ : val → iPropI Σ
+  ============================
+  "Hys" : is_list acc ys
+  "HΦ" : ∀ w : val, is_list w ys -∗ Φ w
+  --------------------------------------∗
+  WP match: InjL #() with
+       InjL <> => acc
+     | InjR "l" =>
+       let: "tmp1" := Fst ! "l" in
+       let: "tmp2" := Snd ! "l" in
+       "l" <- ("tmp1", acc);; (rev "tmp2") (InjL #())
+     end [{ v, Φ v }]
+  
diff --git a/tests/list_reverse.v b/tests/list_reverse.v
index d5388c0e156e86b880197c0b57ad24addb2f9c75..89b5963bba51711f4faf3223472f10b1c35e5f5f 100644
--- a/tests/list_reverse.v
+++ b/tests/list_reverse.v
@@ -31,10 +31,10 @@ Lemma rev_acc_wp hd acc xs ys :
     rev hd acc
   [[{ w, RET w; is_list w (reverse xs ++ ys) }]].
 Proof.
-  iIntros (Φ) "[Hxs Hys] HΦ".
+  iIntros (Φ) "[Hxs Hys] HΦ". Show.
   iInduction xs as [|x xs] "IH" forall (hd acc ys Φ);
     iSimplifyEq; wp_rec; wp_let.
-  - wp_match. by iApply "HΦ".
+  - Show. wp_match. by iApply "HΦ".
   - iDestruct "Hxs" as (l hd' ->) "[Hx Hxs]".
     wp_match. wp_load. wp_proj. wp_let. wp_load. wp_proj. wp_let. wp_store.
     iApply ("IH" $! hd' (SOMEV #l) (x :: ys) with "Hxs [Hx Hys]"); simpl.
diff --git a/tests/one_shot.ref b/tests/one_shot.ref
index e69de29bb2d1d6434b8b29ae775ad8c2e48c5391..a64d6fb6cb8b67a93f76fc9b6cd42a6ee8490614 100644
--- a/tests/one_shot.ref
+++ b/tests/one_shot.ref
@@ -0,0 +1,42 @@
+1 subgoal
+  
+  Σ : gFunctors
+  heapG0 : heapG Σ
+  one_shotG0 : one_shotG Σ
+  Φ : val → iProp Σ
+  N : namespace
+  l : loc
+  γ : gname
+  ============================
+  "HN" : inv N (one_shot_inv γ l)
+  --------------------------------------â–¡
+  "Hl" : l ↦ InjLV #()
+  _ : own γ Pending
+  --------------------------------------∗
+  one_shot_inv γ l
+  ∗ (⌜InjLV #() = InjLV #()⌝
+     ∨ (∃ n : Z, ⌜InjLV #() = InjRV #n⌝ ∗ own γ (Shot n)))
+  
+1 subgoal
+  
+  Σ : gFunctors
+  heapG0 : heapG Σ
+  one_shotG0 : one_shotG Σ
+  Φ : val → iProp Σ
+  N : namespace
+  l : loc
+  γ : gname
+  m, m' : Z
+  ============================
+  "HN" : inv N (one_shot_inv γ l)
+  "Hγ'" : own γ (Shot m)
+  --------------------------------------â–¡
+  "Hl" : l ↦ InjRV #m'
+  "Hγ" : own γ (Shot m')
+  --------------------------------------∗
+  |={⊤ ∖ ↑N}=> ▷ one_shot_inv γ l
+               ∗ WP match: InjR #m' with
+                      InjL <> => assert: #false
+                    | InjR "m" => assert: #m = "m"
+                    end {{ _, True }}
+  
diff --git a/tests/one_shot.v b/tests/one_shot.v
index 53218278ef905cf98a380cd6a3e8ad0b5e7ac048..79faa4fd39c5d4c55f2f0b1a82d9743133437877 100644
--- a/tests/one_shot.v
+++ b/tests/one_shot.v
@@ -67,16 +67,16 @@ Proof.
     iAssert (one_shot_inv γ l ∗ (⌜v = NONEV⌝ ∨ ∃ n : Z,
       ⌜v = SOMEV #n⌝ ∗ own γ (Shot n)))%I with "[Hl Hv]" as "[Hinv #Hv]".
     { iDestruct "Hv" as "[[% ?]|Hv]"; last iDestruct "Hv" as (m) "[% ?]"; subst.
-      + iSplit. iLeft; by iSplitL "Hl". eauto.
+      + Show. iSplit. iLeft; by iSplitL "Hl". eauto.
       + iSplit. iRight; iExists m; by iSplitL "Hl". eauto. }
     iSplitL "Hinv"; first by eauto.
     iModIntro. wp_let. iIntros "!#". wp_seq.
     iDestruct "Hv" as "[%|Hv]"; last iDestruct "Hv" as (m) "[% Hγ']"; subst.
     { by wp_match. }
     wp_match. wp_bind (! _)%E.
-    iInv N as ">[[Hl Hγ]|H]"; last iDestruct "H" as (m') "[Hl Hγ]".
+    iInv N as "[[Hl >Hγ]|H]"; last iDestruct "H" as (m') "[Hl Hγ]".
     { by iDestruct (own_valid_2 with "Hγ Hγ'") as %?. }
-    wp_load.
+    wp_load. Show.
     iDestruct (own_valid_2 with "Hγ Hγ'") as %?%agree_op_invL'; subst.
     iModIntro. iSplitL "Hl".
     { iNext; iRight; by eauto. }