From e354bede03ca7e7be232c1bc556ddf72aacec8cd Mon Sep 17 00:00:00 2001
From: Ralf Jung <jung@mpi-sws.org>
Date: Tue, 1 Nov 2016 16:25:13 +0100
Subject: [PATCH] New notation for Texan triple postconditions: Use 'RET' as
 keyword to indicate the return value

This should make theme asier to parse, "{{{ v, v; l |-> v }}}" looks rather funny.
---
 heap_lang/heap.v              | 10 +++++-----
 heap_lang/lib/barrier/proof.v |  6 +++---
 heap_lang/lib/counter.v       | 14 +++++++-------
 heap_lang/lib/lock.v          |  6 +++---
 heap_lang/lib/par.v           |  4 ++++
 heap_lang/lib/spawn.v         |  4 ++--
 heap_lang/lib/spin_lock.v     |  8 ++++----
 heap_lang/lib/ticket_lock.v   |  8 ++++----
 heap_lang/lifting.v           | 10 +++++-----
 program_logic/ectx_lifting.v  |  2 +-
 program_logic/weakestpre.v    | 32 ++++++++++++++++----------------
 11 files changed, 54 insertions(+), 50 deletions(-)

diff --git a/heap_lang/heap.v b/heap_lang/heap.v
index 996b6bd2e..43d374169 100644
--- a/heap_lang/heap.v
+++ b/heap_lang/heap.v
@@ -122,7 +122,7 @@ Section heap.
   (** Weakest precondition *)
   Lemma wp_alloc E e v :
     to_val e = Some v → nclose heapN ⊆ E →
-    {{{ heap_ctx }}} Alloc e @ E {{{ l; LitV (LitLoc l), l ↦ v }}}.
+    {{{ heap_ctx }}} Alloc e @ E {{{ l, RET LitV (LitLoc l); l ↦ v }}}.
   Proof.
     iIntros (<-%of_to_val ? Φ) "#Hinv HΦ". rewrite /heap_ctx.
     iMod (auth_empty heap_name) as "Ha".
@@ -137,7 +137,7 @@ Section heap.
   Lemma wp_load E l q v :
     nclose heapN ⊆ E →
     {{{ heap_ctx ★ ▷ l ↦{q} v }}} Load (Lit (LitLoc l)) @ E
-    {{{; v, l ↦{q} v }}}.
+    {{{ RET v; l ↦{q} v }}}.
   Proof.
     iIntros (? Φ) "[#Hinv >Hl] HΦ".
     rewrite /heap_ctx heap_mapsto_eq /heap_mapsto_def.
@@ -150,7 +150,7 @@ Section heap.
   Lemma wp_store E l v' e v :
     to_val e = Some v → nclose heapN ⊆ E →
     {{{ heap_ctx ★ ▷ l ↦ v' }}} Store (Lit (LitLoc l)) e @ E
-    {{{; LitV LitUnit, l ↦ v }}}.
+    {{{ RET LitV LitUnit; l ↦ v }}}.
   Proof.
     iIntros (<-%of_to_val ? Φ) "[#Hinv >Hl] HΦ".
     rewrite /heap_ctx heap_mapsto_eq /heap_mapsto_def.
@@ -166,7 +166,7 @@ Section heap.
   Lemma wp_cas_fail E l q v' e1 v1 e2 v2 :
     to_val e1 = Some v1 → to_val e2 = Some v2 → v' ≠ v1 → nclose heapN ⊆ E →
     {{{ heap_ctx ★ ▷ l ↦{q} v' }}} CAS (Lit (LitLoc l)) e1 e2 @ E
-    {{{; LitV (LitBool false), l ↦{q} v' }}}.
+    {{{ RET LitV (LitBool false); l ↦{q} v' }}}.
   Proof.
     iIntros (<-%of_to_val <-%of_to_val ?? Φ) "[#Hinv >Hl] HΦ".
     rewrite /heap_ctx heap_mapsto_eq /heap_mapsto_def.
@@ -179,7 +179,7 @@ Section heap.
   Lemma wp_cas_suc E l e1 v1 e2 v2 :
     to_val e1 = Some v1 → to_val e2 = Some v2 → nclose heapN ⊆ E →
     {{{ heap_ctx ★ ▷ l ↦ v1 }}} CAS (Lit (LitLoc l)) e1 e2 @ E
-    {{{; LitV (LitBool true), l ↦ v2 }}}.
+    {{{ RET LitV (LitBool true); l ↦ v2 }}}.
   Proof.
     iIntros (<-%of_to_val <-%of_to_val ? Φ) "[#Hinv >Hl] HΦ".
     rewrite /heap_ctx heap_mapsto_eq /heap_mapsto_def.
diff --git a/heap_lang/lib/barrier/proof.v b/heap_lang/lib/barrier/proof.v
index d07c09815..764fa2347 100644
--- a/heap_lang/lib/barrier/proof.v
+++ b/heap_lang/lib/barrier/proof.v
@@ -92,7 +92,7 @@ Qed.
 (** Actual proofs *)
 Lemma newbarrier_spec (P : iProp Σ) :
   heapN ⊥ N →
-  {{{ heap_ctx }}} newbarrier #() {{{ l; #l, recv l P ★ send l P }}}.
+  {{{ heap_ctx }}} newbarrier #() {{{ l, RET #l; recv l P ★ send l P }}}.
 Proof.
   iIntros (HN Φ) "#? HΦ".
   rewrite -wp_fupd /newbarrier /=. wp_seq. wp_alloc l as "Hl".
@@ -117,7 +117,7 @@ Proof.
 Qed.
 
 Lemma signal_spec l P :
-  {{{ send l P ★ P }}} signal #l {{{; #(), True }}}.
+  {{{ send l P ★ P }}} signal #l {{{ RET #(); True }}}.
 Proof.
   rewrite /signal /send /barrier_ctx /=.
   iIntros (Φ) "(Hs&HP) HΦ"; iDestruct "Hs" as (γ) "[#(%&Hh&Hsts) Hγ]". wp_let.
@@ -133,7 +133,7 @@ Proof.
 Qed.
 
 Lemma wait_spec l P:
-  {{{ recv l P }}} wait #l {{{ ; #(), P }}}.
+  {{{ recv l P }}} wait #l {{{ RET #(); P }}}.
 Proof.
   rename P into R; rewrite /recv /barrier_ctx.
   iIntros (Φ) "Hr HΦ"; iDestruct "Hr" as (γ P Q i) "(#(%&Hh&Hsts)&Hγ&#HQ&HQR)".
diff --git a/heap_lang/lib/counter.v b/heap_lang/lib/counter.v
index 24b7e4009..16a496730 100644
--- a/heap_lang/lib/counter.v
+++ b/heap_lang/lib/counter.v
@@ -35,7 +35,7 @@ Section mono_proof.
 
   Lemma newcounter_mono_spec (R : iProp Σ) :
     heapN ⊥ N →
-    {{{ heap_ctx }}} newcounter #() {{{ l; #l, mcounter l 0 }}}.
+    {{{ heap_ctx }}} newcounter #() {{{ l, RET #l; mcounter l 0 }}}.
   Proof.
     iIntros (? Φ) "#Hh HΦ". rewrite -wp_fupd /newcounter /=. wp_seq. wp_alloc l as "Hl".
     iMod (own_alloc (● (O:mnat) ⋅ ◯ (O:mnat))) as (γ) "[Hγ Hγ']"; first done.
@@ -45,7 +45,7 @@ Section mono_proof.
   Qed.
 
   Lemma inc_mono_spec l n :
-    {{{ mcounter l n }}} inc #l {{{; #(), mcounter l (S n) }}}.
+    {{{ mcounter l n }}} inc #l {{{ RET #(); mcounter l (S n) }}}.
   Proof.
     iIntros (Φ) "Hl HΦ". iLöb as "IH". wp_rec.
     iDestruct "Hl" as (γ) "(% & #? & #Hinv & Hγf)".
@@ -70,7 +70,7 @@ Section mono_proof.
   Qed.
 
   Lemma read_mono_spec l j :
-    {{{ mcounter l j }}} read #l {{{ i; #i, ■ (j ≤ i)%nat ∧ mcounter l i }}}.
+    {{{ mcounter l j }}} read #l {{{ i, RET #i; ■ (j ≤ i)%nat ∧ mcounter l i }}}.
   Proof.
     iIntros (ϕ) "Hc HΦ". iDestruct "Hc" as (γ) "(% & #? & #Hinv & Hγf)".
     rewrite /read /=. wp_let. iInv N as (c) ">[Hγ Hl]" "Hclose". wp_load.
@@ -112,7 +112,7 @@ Section contrib_spec.
   Lemma newcounter_contrib_spec (R : iProp Σ) :
     heapN ⊥ N →
     {{{ heap_ctx }}} newcounter #()
-    {{{ γ l; #l, ccounter_ctx γ l ★ ccounter γ 1 0 }}}.
+    {{{ γ l, RET #l; ccounter_ctx γ l ★ ccounter γ 1 0 }}}.
   Proof.
     iIntros (? Φ) "#Hh HΦ". rewrite -wp_fupd /newcounter /=. wp_seq. wp_alloc l as "Hl".
     iMod (own_alloc (● (Some (1%Qp, O%nat)) ⋅ ◯ (Some (1%Qp, 0%nat))))
@@ -124,7 +124,7 @@ Section contrib_spec.
 
   Lemma inc_contrib_spec γ l q n :
     {{{ ccounter_ctx γ l ★ ccounter γ q n }}} inc #l
-    {{{; #(), ccounter γ q (S n) }}}.
+    {{{ RET #(); ccounter γ q (S n) }}}.
   Proof.
     iIntros (Φ) "(#(%&?&?) & Hγf) HΦ". iLöb as "IH". wp_rec.
     wp_bind (! _)%E. iInv N as (c) ">[Hγ Hl]" "Hclose".
@@ -145,7 +145,7 @@ Section contrib_spec.
 
   Lemma read_contrib_spec γ l q n :
     {{{ ccounter_ctx γ l ★ ccounter γ q n }}} read #l
-    {{{ c; #c, ■ (n ≤ c)%nat ∧ ccounter γ q n }}}.
+    {{{ c, RET #c; ■ (n ≤ c)%nat ∧ ccounter γ q n }}}.
   Proof.
     iIntros (Φ) "(#(%&?&?) & Hγf) HΦ".
     rewrite /read /=. wp_let. iInv N as (c) ">[Hγ Hl]" "Hclose". wp_load.
@@ -157,7 +157,7 @@ Section contrib_spec.
 
   Lemma read_contrib_spec_1 γ l n :
     {{{ ccounter_ctx γ l ★ ccounter γ 1 n }}} read #l
-    {{{ n; #n, ccounter γ 1 n }}}.
+    {{{ n, RET #n; ccounter γ 1 n }}}.
   Proof.
     iIntros (Φ) "(#(%&?&?) & Hγf) HΦ".
     rewrite /read /=. wp_let. iInv N as (c) ">[Hγ Hl]" "Hclose". wp_load.
diff --git a/heap_lang/lib/lock.v b/heap_lang/lib/lock.v
index 42bc6a572..43f81ecb5 100644
--- a/heap_lang/lib/lock.v
+++ b/heap_lang/lib/lock.v
@@ -18,11 +18,11 @@ Structure lock Σ `{!heapG Σ} := Lock {
   (* -- operation specs -- *)
   newlock_spec N (R : iProp Σ) :
     heapN ⊥ N →
-    {{{ heap_ctx ★ R }}} newlock #() {{{ lk γ; lk, is_lock N γ lk R }}};
+    {{{ heap_ctx ★ R }}} newlock #() {{{ lk γ, RET lk; is_lock N γ lk R }}};
   acquire_spec N γ lk R :
-    {{{ is_lock N γ lk R }}} acquire lk {{{; #(), locked γ ★ R }}};
+    {{{ is_lock N γ lk R }}} acquire lk {{{ RET #(); locked γ ★ R }}};
   release_spec N γ lk R :
-    {{{ is_lock N γ lk R ★ locked γ ★ R }}} release lk {{{; #(), True }}}
+    {{{ is_lock N γ lk R ★ locked γ ★ R }}} release lk {{{ RET #(); True }}}
 }.
 
 Arguments newlock {_ _} _.
diff --git a/heap_lang/lib/par.v b/heap_lang/lib/par.v
index 6d1968aed..5efd170e7 100644
--- a/heap_lang/lib/par.v
+++ b/heap_lang/lib/par.v
@@ -16,6 +16,10 @@ Global Opaque par.
 Section proof.
 Context `{!heapG Σ, !spawnG Σ}.
 
+(* Notice that this allows us to strip a later *after* the two Ψ have been
+   brought together.  That is strictly stronger than first stripping a later
+   and then merging them, as demonstrated by [tests/joining_existentials.v].
+   This is why these are not Texan triples. *)
 Lemma par_spec (Ψ1 Ψ2 : val → iProp Σ) e (f1 f2 : val) (Φ : val → iProp Σ) :
   to_val e = Some (f1,f2)%V →
   (heap_ctx ★ WP f1 #() {{ Ψ1 }} ★ WP f2 #() {{ Ψ2 }} ★
diff --git a/heap_lang/lib/spawn.v b/heap_lang/lib/spawn.v
index ab35c45ec..54053ad1b 100644
--- a/heap_lang/lib/spawn.v
+++ b/heap_lang/lib/spawn.v
@@ -49,7 +49,7 @@ Proof. solve_proper. Qed.
 Lemma spawn_spec (Ψ : val → iProp Σ) e (f : val) :
   to_val e = Some f →
   heapN ⊥ N →
-  {{{ heap_ctx ★ WP f #() {{ Ψ }} }}} spawn e {{{ l; #l, join_handle l Ψ }}}.
+  {{{ heap_ctx ★ WP f #() {{ Ψ }} }}} spawn e {{{ l, RET #l; join_handle l Ψ }}}.
 Proof.
   iIntros (<-%of_to_val ? Φ) "(#Hh & Hf) HΦ". rewrite /spawn /=.
   wp_let. wp_alloc l as "Hl". wp_let.
@@ -64,7 +64,7 @@ Proof.
 Qed.
 
 Lemma join_spec (Ψ : val → iProp Σ) l :
-  {{{ join_handle l Ψ }}} join #l {{{ v; v, Ψ v }}}.
+  {{{ join_handle l Ψ }}} join #l {{{ v, RET v; Ψ v }}}.
 Proof.
   rewrite /join_handle; iIntros (Φ) "[% H] HΦ". iDestruct "H" as (γ) "(#?&Hγ&#?)".
   iLöb as "IH". wp_rec. wp_bind (! _)%E. iInv N as (v) "[Hl Hinv]" "Hclose".
diff --git a/heap_lang/lib/spin_lock.v b/heap_lang/lib/spin_lock.v
index 1bea372cb..44443f4f6 100644
--- a/heap_lang/lib/spin_lock.v
+++ b/heap_lang/lib/spin_lock.v
@@ -47,7 +47,7 @@ Section proof.
 
   Lemma newlock_spec (R : iProp Σ):
     heapN ⊥ N →
-    {{{ heap_ctx ★ R }}} newlock #() {{{ lk γ; lk, is_lock γ lk R }}}.
+    {{{ heap_ctx ★ R }}} newlock #() {{{ lk γ, RET lk; is_lock γ lk R }}}.
   Proof.
     iIntros (? Φ) "[#Hh HR] HΦ". rewrite -wp_fupd /newlock /=.
     wp_seq. wp_alloc l as "Hl".
@@ -59,7 +59,7 @@ Section proof.
 
   Lemma try_acquire_spec γ lk R :
     {{{ is_lock γ lk R }}} try_acquire lk
-    {{{b; #b, if b is true then locked γ ★ R else True }}}.
+    {{{ b, RET #b; if b is true then locked γ ★ R else True }}}.
   Proof.
     iIntros (Φ) "#Hl HΦ". iDestruct "Hl" as (l) "(% & #? & % & #?)". subst.
     wp_rec. iInv N as ([]) "[Hl HR]" "Hclose".
@@ -71,7 +71,7 @@ Section proof.
   Qed.
 
   Lemma acquire_spec γ lk R :
-    {{{ is_lock γ lk R }}} acquire lk {{{; #(), locked γ ★ R }}}.
+    {{{ is_lock γ lk R }}} acquire lk {{{ RET #(); locked γ ★ R }}}.
   Proof.
     iIntros (Φ) "#Hl HΦ". iLöb as "IH". wp_rec.
     wp_apply (try_acquire_spec with "Hl"). iIntros ([]).
@@ -80,7 +80,7 @@ Section proof.
   Qed.
 
   Lemma release_spec γ lk R :
-    {{{ is_lock γ lk R ★ locked γ ★ R }}} release lk {{{; #(), True }}}.
+    {{{ is_lock γ lk R ★ locked γ ★ R }}} release lk {{{ RET #(); True }}}.
   Proof.
     iIntros (Φ) "(Hlock & Hlocked & HR) HΦ".
     iDestruct "Hlock" as (l) "(% & #? & % & #?)". subst.
diff --git a/heap_lang/lib/ticket_lock.v b/heap_lang/lib/ticket_lock.v
index eb3ed6d7c..5300485fc 100644
--- a/heap_lang/lib/ticket_lock.v
+++ b/heap_lang/lib/ticket_lock.v
@@ -76,7 +76,7 @@ Section proof.
 
   Lemma newlock_spec (R : iProp Σ) :
     heapN ⊥ N →
-    {{{ heap_ctx ★ R }}} newlock #() {{{ lk γ; lk, is_lock γ lk R }}}.
+    {{{ heap_ctx ★ R }}} newlock #() {{{ lk γ, RET lk; is_lock γ lk R }}}.
   Proof.
     iIntros (HN Φ) "(#Hh & HR) HΦ". rewrite -wp_fupd /newlock /=.
     wp_seq. wp_alloc lo as "Hlo". wp_alloc ln as "Hln".
@@ -89,7 +89,7 @@ Section proof.
   Qed.
 
   Lemma wait_loop_spec γ lk x R :
-    {{{ issued γ lk x R }}} wait_loop #x lk {{{; #(), locked γ ★ R }}}.
+    {{{ issued γ lk x R }}} wait_loop #x lk {{{ RET #(); locked γ ★ R }}}.
   Proof.
     iIntros (Φ) "Hl HΦ". iDestruct "Hl" as (lo ln) "(% & #? & % & #? & Ht)".
     iLöb as "IH". wp_rec. subst. wp_let. wp_proj. wp_bind (! _)%E.
@@ -110,7 +110,7 @@ Section proof.
   Qed.
 
   Lemma acquire_spec γ lk R :
-    {{{ is_lock γ lk R }}} acquire lk {{{; #(), locked γ ★ R }}}.
+    {{{ is_lock γ lk R }}} acquire lk {{{ RET #(); locked γ ★ R }}}.
   Proof.
     iIntros (ϕ) "Hl HΦ". iDestruct "Hl" as (lo ln) "(% & #? & % & #?)".
     iLöb as "IH". wp_rec. wp_bind (! _)%E. subst. wp_proj.
@@ -141,7 +141,7 @@ Section proof.
   Qed.
 
   Lemma release_spec γ lk R :
-    {{{ is_lock γ lk R ★ locked γ ★ R }}} release lk {{{; #(), True }}}.
+    {{{ is_lock γ lk R ★ locked γ ★ R }}} release lk {{{ RET #(); True }}}.
   Proof.
     iIntros (Φ) "(Hl & Hγ & HR) HΦ".
     iDestruct "Hl" as (lo ln) "(% & #? & % & #?)"; subst.
diff --git a/heap_lang/lifting.v b/heap_lang/lifting.v
index bd65645b2..71ce9d616 100644
--- a/heap_lang/lifting.v
+++ b/heap_lang/lifting.v
@@ -48,7 +48,7 @@ Proof. exact: weakestpre.wp_bind. Qed.
 (** Base axioms for core primitives of the language: Stateful reductions. *)
 Lemma wp_alloc_pst E σ v :
   {{{ ▷ ownP σ }}} Alloc (of_val v) @ E
-  {{{ l; LitV (LitLoc l), σ !! l = None ∧ ownP (<[l:=v]>σ) }}}.
+  {{{ l, RET LitV (LitLoc l); σ !! l = None ∧ ownP (<[l:=v]>σ) }}}.
 Proof.
   iIntros (Φ) "HP HΦ".
   iApply (wp_lift_atomic_head_step (Alloc (of_val v)) σ); eauto.
@@ -59,7 +59,7 @@ Qed.
 
 Lemma wp_load_pst E σ l v :
   σ !! l = Some v →
-  {{{ ▷ ownP σ }}} Load (Lit (LitLoc l)) @ E {{{; v, ownP σ }}}.
+  {{{ ▷ ownP σ }}} Load (Lit (LitLoc l)) @ E {{{ RET v; ownP σ }}}.
 Proof.
   intros ? Φ. apply (wp_lift_atomic_det_head_step' σ v σ); eauto.
   intros; inv_head_step; eauto.
@@ -68,7 +68,7 @@ Qed.
 Lemma wp_store_pst E σ l v v' :
   σ !! l = Some v' →
   {{{ ▷ ownP σ }}} Store (Lit (LitLoc l)) (of_val v) @ E
-  {{{; LitV LitUnit, ownP (<[l:=v]>σ) }}}.
+  {{{ RET LitV LitUnit; ownP (<[l:=v]>σ) }}}.
 Proof.
   intros. apply (wp_lift_atomic_det_head_step' σ (LitV LitUnit) (<[l:=v]>σ)); eauto.
   intros; inv_head_step; eauto.
@@ -77,7 +77,7 @@ Qed.
 Lemma wp_cas_fail_pst E σ l v1 v2 v' :
   σ !! l = Some v' → v' ≠ v1 →
   {{{ ▷ ownP σ }}} CAS (Lit (LitLoc l)) (of_val v1) (of_val v2) @ E
-  {{{; LitV $ LitBool false, ownP σ }}}.
+  {{{ RET LitV $ LitBool false; ownP σ }}}.
 Proof.
   intros. apply (wp_lift_atomic_det_head_step' σ (LitV $ LitBool false) σ); eauto.
   intros; inv_head_step; eauto.
@@ -86,7 +86,7 @@ Qed.
 Lemma wp_cas_suc_pst E σ l v1 v2 :
   σ !! l = Some v1 →
   {{{ ▷ ownP σ }}} CAS (Lit (LitLoc l)) (of_val v1) (of_val v2) @ E
-  {{{; LitV $ LitBool true, ownP (<[l:=v2]>σ) }}}.
+  {{{ RET LitV $ LitBool true; ownP (<[l:=v2]>σ) }}}.
 Proof.
   intros. apply (wp_lift_atomic_det_head_step' σ (LitV $ LitBool true)
     (<[l:=v2]>σ)); eauto.
diff --git a/program_logic/ectx_lifting.v b/program_logic/ectx_lifting.v
index 9522a65e3..28a3e7636 100644
--- a/program_logic/ectx_lifting.v
+++ b/program_logic/ectx_lifting.v
@@ -65,7 +65,7 @@ Lemma wp_lift_atomic_det_head_step' {E e1} σ1 v2 σ2 :
   head_reducible e1 σ1 →
   (∀ e2' σ2' efs', head_step e1 σ1 e2' σ2' efs' →
     σ2 = σ2' ∧ to_val e2' = Some v2 ∧ [] = efs') →
-  {{{ ▷ ownP σ1 }}} e1 @ E {{{; v2, ownP σ2 }}}.
+  {{{ ▷ ownP σ1 }}} e1 @ E {{{ RET v2; ownP σ2 }}}.
 Proof.
   intros. rewrite -(wp_lift_atomic_det_head_step σ1 v2 σ2 []); [|done..].
   rewrite big_sepL_nil right_id. by apply uPred.wand_intro_r.
diff --git a/program_logic/weakestpre.v b/program_logic/weakestpre.v
index 4fde7f70e..d5f94f1c4 100644
--- a/program_logic/weakestpre.v
+++ b/program_logic/weakestpre.v
@@ -68,43 +68,43 @@ Notation "'WP' e {{ v , Q } }" := (wp ⊤ e%E (λ v, Q))
    format "'WP'  e  {{  v ,  Q  } }") : uPred_scope.
 
 (* Texan triples *)
-Notation "'{{{' P } } } e {{{ x .. y ; pat , Q } } }" :=
+Notation "'{{{' P } } } e {{{ x .. y , 'RET' pat ; Q } } }" :=
   (□ ∀ Φ,
       P -★ ▷ (∀ x, .. (∀ y, Q -★ Φ pat%V) .. ) -★ WP e {{ Φ }})%I
     (at level 20, x closed binder, y closed binder,
-     format "{{{  P  } } }  e  {{{ x .. y ;  pat ,  Q } } }") : uPred_scope.
-Notation "'{{{' P } } } e @ E {{{ x .. y ; pat , Q } } }" :=
+     format "{{{  P  } } }  e  {{{  x .. y ,   RET  pat ;  Q } } }") : uPred_scope.
+Notation "'{{{' P } } } e @ E {{{ x .. y , 'RET' pat ; Q } } }" :=
   (□ ∀ Φ,
       P -★ ▷ (∀ x, .. (∀ y, Q -★ Φ pat%V) .. ) -★ WP e @ E {{ Φ }})%I
     (at level 20, x closed binder, y closed binder,
-     format "{{{  P  } } }  e  @  E  {{{ x .. y ;  pat ,  Q } } }") : uPred_scope.
-Notation "'{{{' P } } } e {{{ ; pat , Q } } }" :=
+     format "{{{  P  } } }  e  @  E  {{{  x .. y ,  RET  pat ;  Q } } }") : uPred_scope.
+Notation "'{{{' P } } } e {{{ 'RET' pat ; Q } } }" :=
   (□ ∀ Φ, P -★ ▷ (Q -★ Φ pat%V) -★ WP e {{ Φ }})%I
     (at level 20,
-     format "{{{  P  } } }  e  {{{ ; pat ,  Q } } }") : uPred_scope.
-Notation "'{{{' P } } } e @ E {{{ ; pat , Q } } }" :=
+     format "{{{  P  } } }  e  {{{  RET  pat ;  Q } } }") : uPred_scope.
+Notation "'{{{' P } } } e @ E {{{ 'RET' pat ; Q } } }" :=
   (□ ∀ Φ, P -★ ▷ (Q -★ Φ pat%V) -★ WP e @ E {{ Φ }})%I
     (at level 20,
-     format "{{{  P  } } }  e  @  E  {{{ ; pat ,  Q } } }") : uPred_scope.
+     format "{{{  P  } } }  e  @  E  {{{  RET  pat ;  Q } } }") : uPred_scope.
 
-Notation "'{{{' P } } } e {{{ x .. y ; pat , Q } } }" :=
+Notation "'{{{' P } } } e {{{ x .. y , 'RET' pat ; Q } } }" :=
   (∀ Φ : _ → uPred _,
       P ⊢ ▷ (∀ x, .. (∀ y, Q -★ Φ pat%V) .. ) -★ WP e {{ Φ }})
     (at level 20, x closed binder, y closed binder,
-     format "{{{  P  } } }  e  {{{ x .. y ;  pat ,  Q } } }") : C_scope.
-Notation "'{{{' P } } } e @ E {{{ x .. y ; pat , Q } } }" :=
+     format "{{{  P  } } }  e  {{{  x .. y ,  RET  pat ;  Q } } }") : C_scope.
+Notation "'{{{' P } } } e @ E {{{ x .. y , 'RET' pat ; Q } } }" :=
   (∀ Φ : _ → uPred _,
       P ⊢ ▷ (∀ x, .. (∀ y, Q -★ Φ pat%V) .. ) -★ WP e @ E {{ Φ }})
     (at level 20, x closed binder, y closed binder,
-     format "{{{  P  } } }  e  @  E  {{{ x .. y ;  pat ,  Q } } }") : C_scope.
-Notation "'{{{' P } } } e {{{ ; pat , Q } } }" :=
+     format "{{{  P  } } }  e  @  E  {{{  x .. y ,  RET  pat ;  Q } } }") : C_scope.
+Notation "'{{{' P } } } e {{{ 'RET' pat ; Q } } }" :=
   (∀ Φ : _ → uPred _, P ⊢ ▷ (Q -★ Φ pat%V) -★ WP e {{ Φ }})
     (at level 20,
-     format "{{{  P  } } }  e  {{{ ; pat ,  Q } } }") : C_scope.
-Notation "'{{{' P } } } e @ E {{{ ; pat , Q } } }" :=
+     format "{{{  P  } } }  e  {{{  RET  pat ;  Q } } }") : C_scope.
+Notation "'{{{' P } } } e @ E {{{ 'RET' pat ; Q } } }" :=
   (∀ Φ : _ → uPred _, P ⊢ ▷ (Q -★ Φ pat%V) -★ WP e @ E {{ Φ }})
     (at level 20,
-     format "{{{  P  } } }  e  @  E  {{{ ; pat ,  Q } } }") : C_scope.
+     format "{{{  P  } } }  e  @  E  {{{  RET  pat ;  Q } } }") : C_scope.
 
 Section wp.
 Context `{irisG Λ Σ}.
-- 
GitLab