From d52296934b85858784c32cc766895192e099feea Mon Sep 17 00:00:00 2001
From: Robbert Krebbers <mail@robbertkrebbers.nl>
Date: Tue, 19 Jul 2016 18:38:44 +0200
Subject: [PATCH] Use SOME/NONE in tests/one_shot.

---
 tests/one_shot.v | 28 ++++++++++++++--------------
 1 file changed, 14 insertions(+), 14 deletions(-)

diff --git a/tests/one_shot.v b/tests/one_shot.v
index 6b7a0a198..108f0b637 100644
--- a/tests/one_shot.v
+++ b/tests/one_shot.v
@@ -5,17 +5,17 @@ From iris.proofmode Require Import invariants ghost_ownership.
 Import uPred.
 
 Definition one_shot_example : val := λ: <>,
-  let: "x" := ref (InjL #0) in (
+  let: "x" := ref NONE in (
   (* tryset *) (λ: "n",
-    CAS "x" (InjL #0) (InjR "n")),
+    CAS "x" NONE (SOME "n")),
   (* check  *) (λ: <>,
     let: "y" := !"x" in λ: <>,
     match: "y" with
-      InjL <> => #()
-    | InjR "n" =>
+      NONE => #()
+    | SOME "n" =>
        match: !"x" with
-         InjL <> => assert: #false
-       | InjR "m" => assert: "n" = "m"
+         NONE => assert: #false
+       | SOME "m" => assert: "n" = "m"
        end
     end)).
 Global Opaque one_shot_example.
@@ -35,8 +35,8 @@ Context (heapN N : namespace) (HN : heapN ⊥ N).
 Local Notation iProp := (iPropG heap_lang Σ).
 
 Definition one_shot_inv (γ : gname) (l : loc) : iProp :=
-  (l ↦ InjLV #0 ★ own γ Pending ∨
-  ∃ n : Z, l ↦ InjRV #n ★ own γ (Cinr (DecAgree n)))%I.
+  (l ↦ NONEV ★ own γ Pending ∨
+  ∃ n : Z, l ↦ SOMEV #n ★ own γ (Cinr (DecAgree n)))%I.
 
 Lemma wp_one_shot (Φ : val → iProp) :
   heap_ctx heapN ★ (∀ f1 f2 : val,
@@ -58,14 +58,14 @@ Proof.
       iPvsIntro; iRight; iExists n; by iSplitL "Hl".
     + wp_cas_fail. rewrite /one_shot_inv; eauto 10.
   - iIntros "!". wp_seq. wp_focus (! _)%E. iInv> N as "Hγ".
-    iAssert (∃ v, l ↦ v ★ ((v = InjLV #0 ★ own γ Pending) ∨
-       ∃ n : Z, v = InjRV #n ★ own γ (Cinr (DecAgree n))))%I with "[-]" as "Hv".
+    iAssert (∃ v, l ↦ v ★ ((v = NONEV ★ own γ Pending) ∨
+       ∃ n : Z, v = SOMEV #n ★ own γ (Cinr (DecAgree n))))%I with "[-]" as "Hv".
     { iDestruct "Hγ" as "[[Hl Hγ]|Hl]"; last iDestruct "Hl" as (m) "[Hl Hγ]".
-      + iExists (InjLV #0). iFrame. eauto.
-      + iExists (InjRV #m). iFrame. eauto. }
+      + iExists NONEV. iFrame. eauto.
+      + iExists (SOMEV #m). iFrame. eauto. }
     iDestruct "Hv" as (v) "[Hl Hv]". wp_load; iPvsIntro.
-    iAssert (one_shot_inv γ l ★ (v = InjLV #0 ∨ ∃ n : Z,
-      v = InjRV #n ★ own γ (Cinr (DecAgree n))))%I with "[-]" as "[$ #Hv]".
+    iAssert (one_shot_inv γ l ★ (v = NONEV ∨ ∃ n : Z,
+      v = SOMEV #n ★ own γ (Cinr (DecAgree n))))%I with "[-]" as "[$ #Hv]".
     { iDestruct "Hv" as "[[% ?]|Hv]"; last iDestruct "Hv" as (m) "[% ?]"; subst.
       + iSplit. iLeft; by iSplitL "Hl". eauto.
       + iSplit. iRight; iExists m; by iSplitL "Hl". eauto. }
-- 
GitLab