From bf42c06565bb7e999e12b014b8e3193cc2905616 Mon Sep 17 00:00:00 2001
From: Robbert Krebbers <mail@robbertkrebbers.nl>
Date: Tue, 19 Jul 2016 21:58:41 +0200
Subject: [PATCH] Repair tree_sum test case and add to _CoqProject.

---
 _CoqProject      |  1 +
 tests/tree_sum.v | 15 +++++++--------
 2 files changed, 8 insertions(+), 8 deletions(-)

diff --git a/_CoqProject b/_CoqProject
index 92fe4e429..7f98bf17f 100644
--- a/_CoqProject
+++ b/_CoqProject
@@ -106,6 +106,7 @@ tests/joining_existentials.v
 tests/proofmode.v
 tests/barrier_client.v
 tests/list_reverse.v
+tests/tree_sum.v
 proofmode/coq_tactics.v
 proofmode/pviewshifts.v
 proofmode/environments.v
diff --git a/tests/tree_sum.v b/tests/tree_sum.v
index b96686ebd..beb650663 100644
--- a/tests/tree_sum.v
+++ b/tests/tree_sum.v
@@ -21,16 +21,15 @@ Fixpoint sum (t : tree) : Z :=
 
 Definition sum_loop : val :=
   rec: "sum_loop" "t" "l" :=
-    match: '"t" with
-      InjL "n" => '"l" <- '"n" + !'"l"
-    | InjR "tt" =>
-       '"sum_loop" !(Fst '"tt") '"l" ;; '"sum_loop" !(Snd '"tt") '"l"
+    match: "t" with
+      InjL "n" => "l" <- "n" + !"l"
+    | InjR "tt" => "sum_loop" !(Fst "tt") "l" ;; "sum_loop" !(Snd "tt") "l"
     end.
 
 Definition sum' : val := λ: "t",
   let: "l" := ref #0 in
-  ^sum_loop '"t" '"l";;
-  !'"l".
+  sum_loop "t" "l";;
+  !"l".
 
 Global Opaque sum_loop sum'.
 
@@ -40,12 +39,12 @@ Lemma sum_loop_wp `{!heapG Σ} heapN v t l (n : Z) (Φ : val → iPropG heap_lan
   ⊢ WP sum_loop v #l {{ Φ }}.
 Proof.
   iIntros "(#Hh & Hl & Ht & HΦ)".
-  iLöb {v t l n Φ} as "IH". wp_rec; wp_let.
+  iLöb (v t l n Φ) as "IH". wp_rec. wp_let.
   destruct t as [n'|tl tr]; simpl in *.
   - iDestruct "Ht" as "%"; subst.
     wp_match. wp_load. wp_op. wp_store.
     by iApply ("HΦ" with "Hl").
-  - iDestruct "Ht" as {ll lr vl vr} "(% & Hll & Htl & Hlr & Htr)"; subst.
+  - iDestruct "Ht" as (ll lr vl vr) "(% & Hll & Htl & Hlr & Htr)"; subst.
     wp_match. wp_proj. wp_load.
     wp_apply ("IH" with "Hl Htl"). iIntros "Hl Htl".
     wp_seq. wp_proj. wp_load.
-- 
GitLab