From f6a4683d0710d7aa22ee9c0181ebe2edef704bd4 Mon Sep 17 00:00:00 2001
From: Robbert Krebbers <mail@robbertkrebbers.nl>
Date: Tue, 9 Jul 2019 13:02:52 +0200
Subject: [PATCH] More consistent naming & more consistent use of in-place.

---
 theories/examples/map.v            | 85 +++++++++++++++---------------
 theories/examples/map_reduce.v     | 55 +++++++++----------
 theories/examples/sort_fg_client.v | 35 ++++++------
 3 files changed, 88 insertions(+), 87 deletions(-)

diff --git a/theories/examples/map.v b/theories/examples/map.v
index 610eb64..cdef133 100644
--- a/theories/examples/map.v
+++ b/theories/examples/map.v
@@ -3,7 +3,7 @@ From iris.heap_lang Require Import proofmode notation lib.spin_lock.
 From actris.utils Require Import llist contribution.
 From iris.algebra Require Import gmultiset.
 
-Definition map_worker : val :=
+Definition par_map_worker : val :=
   rec: "go" "map" "l" "c" :=
     acquire "l";;
       send "c" #true;;
@@ -17,18 +17,17 @@ Definition map_worker : val :=
     release "l";;
     "go" "map" "l" "c".
 
-Definition start_map_workers : val :=
+Definition par_map_workers : val :=
   rec: "go" "n" "map" "l" "c" :=
     if: "n" = #0 then #() else
-    Fork (map_worker "map" "l" "c");;
+    Fork (par_map_worker "map" "l" "c");;
     "go" ("n" - #1) "map" "l" "c".
 
-Definition start_map_service : val := λ: "n" "map",
-  start_chan (λ: "c",
-    let: "l" := newlock #() in
-    start_map_workers "n" "map" "l" "c").
+Definition par_map_service : val := λ: "n" "map" "c",
+  let: "l" := newlock #() in
+  par_map_workers "n" "map" "l" "c".
 
-Definition par_map_server : val :=
+Definition par_map_client_loop : val :=
   rec: "go" "n" "c" "xs" "ys" :=
     if: "n" = #0 then #() else
     if: recv "c" then (* send item to map_worker *)
@@ -44,11 +43,11 @@ Definition par_map_server : val :=
       lprep "ys" "zs";;
       "go" "n" "c" "xs" "ys".
 
-Definition par_map : val := λ: "n" "map" "xs",
-  let: "c" := start_map_service "n" "map" in
+Definition par_map_client : val := λ: "n" "map" "xs",
+  let: "c" := start_chan (λ: "c", par_map_service "n" "map" "c") in
   let: "ys" := lnil #() in
-  par_map_server "n" "c" "xs" "ys";;
-  "ys".
+  par_map_client_loop "n" "c" "xs" "ys";;
+  lapp "xs" "ys".
 
 Class mapG Σ A `{Countable A} := {
   map_contributionG :> contributionG Σ (gmultisetUR A);
@@ -65,7 +64,7 @@ Section map.
   Definition map_spec (vmap : val) : iProp Σ := (∀ x v,
     {{{ IA x v }}} vmap v {{{ l, RET #l; llist IB l (map x) }}})%I.
 
-  Definition map_worker_protocol_aux (rec : nat -d> gmultiset A -d> iProto Σ) :
+  Definition par_map_protocol_aux (rec : nat -d> gmultiset A -d> iProto Σ) :
       nat -d> gmultiset A -d> iProto Σ := λ i X,
     let rec : nat → gmultiset A → iProto Σ := rec in
     (if i is 0 then END else
@@ -75,20 +74,20 @@ Section map.
      ) <{⌜ i ≠ 1 ∨ X = ∅ ⌝}&>
          <?> x (l : loc), MSG #l {{ ⌜ x ∈ X ⌝ ∗ llist IB l (map x) }};
          rec i (X ∖ {[ x ]}))%proto.
-  Instance map_worker_protocol_aux_contractive : Contractive map_worker_protocol_aux.
+  Instance par_map_protocol_aux_contractive : Contractive par_map_protocol_aux.
   Proof. solve_proper_prepare. f_equiv. solve_proto_contractive. Qed.
-  Definition map_worker_protocol := fixpoint map_worker_protocol_aux.
-  Global Instance map_worker_protocol_unfold n X :
-    ProtoUnfold (map_worker_protocol n X) (map_worker_protocol_aux map_worker_protocol n X).
-  Proof. apply proto_unfold_eq, (fixpoint_unfold map_worker_protocol_aux). Qed.
+  Definition par_map_protocol := fixpoint par_map_protocol_aux.
+  Global Instance par_map_protocol_unfold n X :
+    ProtoUnfold (par_map_protocol n X) (par_map_protocol_aux par_map_protocol n X).
+  Proof. apply proto_unfold_eq, (fixpoint_unfold par_map_protocol_aux). Qed.
 
   Definition map_worker_lock_inv (γ : gname) (c : val) : iProp Σ :=
-    (∃ i X, server γ i X ∗ c ↣ iProto_dual (map_worker_protocol i X) @ N)%I.
+    (∃ i X, server γ i X ∗ c ↣ iProto_dual (par_map_protocol i X) @ N)%I.
 
-  Lemma map_worker_spec γl γ vmap lk c :
+  Lemma par_map_worker_spec γl γ vmap lk c :
     map_spec vmap -∗
     {{{ is_lock N γl lk (map_worker_lock_inv γ c) ∗ client γ (∅ : gmultiset A) }}}
-      map_worker vmap lk c
+      par_map_worker vmap lk c
     {{{ RET #(); True }}}.
   Proof.
     iIntros "#Hmap !>" (Φ) "[#Hlk Hγ] HΦ". iLöb as "IH".
@@ -125,10 +124,10 @@ Section map.
     iIntros "Hu". by wp_apply ("IH" with "[$] [$]").
   Qed.
 
-  Lemma start_map_workers_spec γl γ n vmap lk c :
+  Lemma par_map_workers_spec γl γ n vmap lk c :
     map_spec vmap -∗
     {{{ is_lock N γl lk (map_worker_lock_inv γ c) ∗ client γ (∅:gmultiset A) ^ n }}}
-      start_map_workers #n vmap lk c
+      par_map_workers #n vmap lk c
     {{{ RET #(); True }}}.
   Proof.
     iIntros "#Hmap !>" (Φ) "[#Hlk Hγs] HΦ".
@@ -136,32 +135,32 @@ Section map.
     { by iApply "HΦ". }
     iDestruct "Hγs" as "[Hγ Hγs]".
     wp_apply (wp_fork with "[Hγ]").
-    { iNext. wp_apply (map_worker_spec with "Hmap [$]"); auto. }
+    { iNext. wp_apply (par_map_worker_spec with "Hmap [$]"); auto. }
     wp_pures. rewrite Nat2Z.inj_succ Z.sub_1_r Z.pred_succ.
     wp_apply ("IH" with "[$] [$]").
   Qed.
 
-  Lemma start_map_service_spec n vmap :
+  Lemma par_map_service_spec n vmap c :
     map_spec vmap -∗
-    {{{ True }}}
-      start_map_service #n vmap
-    {{{ c, RET c; c ↣ map_worker_protocol n ∅ @ N }}}.
+    {{{ c ↣ iProto_dual (par_map_protocol n ∅) @ N }}}
+      par_map_service #n vmap c
+    {{{ RET #(); True }}}.
   Proof.
-    iIntros "#Hf !>"; iIntros (Φ _) "HΦ". wp_lam; wp_pures.
-    wp_apply (start_chan_proto_spec N (map_worker_protocol n ∅)); iIntros (c) "// Hc".
-    wp_lam.
+    iIntros "#Hf !>"; iIntros (Φ) "Hc HΦ". wp_lam; wp_pures.
     iMod (contribution_init_pow (A:=gmultisetUR A) n) as (γ) "[Hs Hγs]".
     wp_apply (newlock_spec N (map_worker_lock_inv γ c) with "[Hc Hs]").
     { iExists n, ∅. iFrame. }
     iIntros (lk γl) "#Hlk".
-    wp_apply (start_map_workers_spec with "Hf [$Hlk $Hγs]"); auto.
+    wp_apply (par_map_workers_spec with "Hf [$Hlk $Hγs]"); auto.
   Qed.
 
-  Lemma par_map_server_spec n c l k xs X ys :
+  Lemma par_map_client_loop_spec n c l k xs X ys :
     (n = 0 → X = ∅ ∧ xs = []) →
-    {{{ llist IA l xs ∗ llist IB k ys ∗ c ↣ map_worker_protocol n X @ N }}}
-      par_map_server #n c #l #k
-    {{{ ys', RET #(); ⌜ys' ≡ₚ (xs ++ elements X) ≫= map⌝ ∗ llist IB k (ys' ++ ys) }}}.
+    {{{ llist IA l xs ∗ llist IB k ys ∗ c ↣ par_map_protocol n X @ N }}}
+      par_map_client_loop #n c #l #k
+    {{{ ys', RET #();
+      ⌜ys' ≡ₚ (xs ++ elements X) ≫= map⌝ ∗ llist IA l [] ∗ llist IB k (ys' ++ ys)
+    }}}.
   Proof.
     iIntros (Hn Φ) "(Hl & Hk & Hc) HΦ".
     iLöb as "IH" forall (n l xs X ys Hn Φ); wp_rec; wp_pures; simpl.
@@ -190,18 +189,20 @@ Section map.
       + by rewrite -assoc_L.
   Qed.
 
-  Lemma par_map_spec n vmap l xs :
+  Lemma par_map_client_spec n vmap l xs :
     0 < n →
     map_spec vmap -∗
     {{{ llist IA l xs }}}
-      par_map #n vmap #l
-    {{{ k ys, RET #k; ⌜ys ≡ₚ xs ≫= map⌝ ∗ llist IB k ys }}}.
+      par_map_client #n vmap #l
+    {{{ ys, RET #(); ⌜ys ≡ₚ xs ≫= map⌝ ∗ llist IB l ys }}}.
   Proof.
     iIntros (?) "#Hmap !>"; iIntros (Φ) "Hl HΦ". wp_lam; wp_pures.
-    wp_apply (start_map_service_spec with "Hmap [//]"); iIntros (c) "Hc".
+    wp_apply (start_chan_proto_spec N (par_map_protocol n ∅)); iIntros (c) "// Hc".
+    { wp_apply (par_map_service_spec with "Hmap Hc"); auto. }
     wp_pures. wp_apply (lnil_spec with "[//]"); iIntros (k) "Hk".
-    wp_apply (par_map_server_spec with "[$Hl $Hk $Hc //]"); first lia.
-    iIntros (ys) "[??]". rewrite /= gmultiset_elements_empty !right_id_L .
+    wp_apply (par_map_client_loop_spec with "[$Hl $Hk $Hc //]"); first lia.
+    iIntros (ys) "(?&Hl&Hk)". rewrite /= gmultiset_elements_empty !right_id_L.
+    wp_apply (lapp_spec IB _ _ [] with "[$Hl $Hk]"); iIntros "[Hk _] /=".
     wp_pures. iApply "HΦ"; auto.
   Qed.
 End map.
diff --git a/theories/examples/map_reduce.v b/theories/examples/map_reduce.v
index f4e4c8a..9adc22e 100644
--- a/theories/examples/map_reduce.v
+++ b/theories/examples/map_reduce.v
@@ -31,7 +31,7 @@ Instance: Params (@group) 7.
 
 
 (** Distributed version *)
-Definition par_map_reduce_map_server : val :=
+Definition par_map_reduce_map : val :=
   rec: "go" "n" "cmap" "csort" "xs" :=
     if: "n" = #0 then #() else
     if: recv "cmap" then (* send item to mapper *)
@@ -55,7 +55,7 @@ Definition par_map_reduce_collect : val :=
     if: "i" = "j" then lcons "y" "ys";; "go" "csort" "j" "ys" else
     SOME ("j", "y").
 
-Definition par_map_reduce_reduce_server : val :=
+Definition par_map_reduce_reduce : val :=
   rec: "go" "n" "csort" "cred" "acc" "zs" :=
     if: "n" = #0 then #() else
     if: recv "cred" then (* Send item to mapper *)
@@ -79,16 +79,15 @@ Definition par_map_reduce_reduce_server : val :=
 Definition cmpZfst : val := λ: "x" "y", Fst "x" ≤ Fst "y".
 
 Definition par_map_reduce : val := λ: "n" "map" "red" "xs",
-  let: "cmap" := start_map_service "n" "map" in
+  let: "cmap" := start_chan (λ: "c", par_map_service "n" "map" "c") in
   let: "csort" := start_chan (λ: "c", sort_service_fg cmpZfst "c") in
-  par_map_reduce_map_server "n" "cmap" "csort" "xs";;
+  par_map_reduce_map "n" "cmap" "csort" "xs";;
   send "csort" #stop;;
-  let: "cred" := start_map_service "n" "red" in
+  let: "cred" := start_chan (λ: "c", par_map_service "n" "red" "c") in
   (* We need the first sorted element in the loop to compare subsequent elements *)
-  if: ~recv "csort" then lnil #() else (* Handle the empty case *)
+  if: ~recv "csort" then #() else (* Handle the empty case *)
   let: "jy" := recv "csort" in
-  let: "zs" := lnil #() in
-  par_map_reduce_reduce_server "n" "csort" "cred" (SOME "jy") "zs";; "zs".
+  par_map_reduce_reduce "n" "csort" "cred" (SOME "jy") "xs".
 
 
 (** Properties about the functional version *)
@@ -240,24 +239,24 @@ Section mapper.
     repeat case_bool_decide=> //; unfold RZB, prod_relation in *; naive_solver.
   Qed.
 
-  Lemma par_map_reduce_map_server_spec n cmap csort l xs X ys :
+  Lemma par_map_reduce_map_spec n cmap csort l xs X ys :
     (n = 0 → X = ∅ ∧ xs = []) →
     {{{
       llist IA l xs ∗
-      cmap ↣ map_worker_protocol IA IZB map n (X : gmultiset A) @ N ∗
+      cmap ↣ par_map_protocol IA IZB map n (X : gmultiset A) @ N ∗
       csort ↣ sort_fg_head_protocol IZB RZB ys @ N
     }}}
-      par_map_reduce_map_server #n cmap csort #l
+      par_map_reduce_map #n cmap csort #l
     {{{ ys', RET #();
       ⌜ys' ≡ₚ (xs ++ elements X) ≫= map⌝ ∗
-      csort ↣ sort_fg_head_protocol IZB RZB (ys ++ ys') @ N
+      llist IA l [] ∗ csort ↣ sort_fg_head_protocol IZB RZB (ys ++ ys') @ N
     }}}.
   Proof.
     iIntros (Hn Φ) "(Hl & Hcmap & Hcsort) HΦ".
     iLöb as "IH" forall (n xs X ys Hn Φ); wp_rec; wp_pures; simpl.
     case_bool_decide; wp_pures; simplify_eq/=.
     { destruct Hn as [-> ->]; first lia.
-      iApply ("HΦ" $! []). rewrite right_id_L. auto. }
+      iApply ("HΦ" $! []). rewrite right_id_L. auto with iFrame. }
     destruct n as [|n]=> //=. wp_branch as %?|%_; wp_pures.
     - wp_apply (lisnil_spec with "Hl"); iIntros "Hl".
       destruct xs as [|x xs]; csimpl; wp_pures.
@@ -270,7 +269,7 @@ Section mapper.
         rewrite assoc_L -(comm _ [x]). iApply "HΦ".
     - wp_recv (x k) as (Hx) "Hk".
       rewrite -(right_id END%proto _ (sort_fg_head_protocol _ _ _)).
-      wp_apply (send_all_spec with "[$Hk $Hcsort]"); iIntros "Hcsort".
+      wp_apply (send_all_spec with "[$Hk $Hcsort]"); iIntros "[_ Hcsort]".
       rewrite right_id.
       wp_apply ("IH" with "[] Hl Hcmap Hcsort"); first done.
       iIntros (ys'). iDestruct 1 as (Hys) "Hcsort"; simplify_eq/=.
@@ -330,7 +329,7 @@ Section mapper.
       eapply elem_of_StronglySorted_app; set_solver.
   Qed.
 
-  Lemma par_map_reduce_reduce_server_spec n iys iys_sorted miy zs l Y csort cred :
+  Lemma par_map_reduce_reduce_spec n iys iys_sorted miy zs l Y csort cred :
     let acc := from_option (λ '(i,y,w), [(i,y)]) [] in
     let accv := from_option (λ '(i,y,w), SOMEV (#(i:Z),w)) NONEV in
     (n = 0 → miy = None ∧ Y = ∅) →
@@ -340,10 +339,10 @@ Section mapper.
       llist IC l zs ∗
       csort ↣ from_option (λ _, sort_fg_tail_protocol IZB RZB iys
         (iys_sorted ++ acc miy)) END%proto miy @ N ∗
-      cred ↣ map_worker_protocol IZBs IC (curry red) n (Y : gmultiset (Z * list B)) @ N ∗
+      cred ↣ par_map_protocol IZBs IC (curry red) n (Y : gmultiset (Z * list B)) @ N ∗
       from_option (λ '(i,y,w), IB i y w) True miy
     }}}
-      par_map_reduce_reduce_server #n csort cred (accv miy) #l
+      par_map_reduce_reduce #n csort cred (accv miy) #l
     {{{ zs', RET #();
        ⌜ (group iys_sorted ≫= curry red) ++ zs' ≡ₚ (group iys ++ elements Y) ≫= curry red ⌝ ∗
        llist IC l (zs' ++ zs)
@@ -394,27 +393,29 @@ Section mapper.
     map_spec IZBs IC (curry red) vred -∗
     {{{ llist IA l xs }}}
       par_map_reduce #n vmap vred #l
-    {{{ k zs, RET #k; ⌜zs ≡ₚ map_reduce map red xs⌝ ∗ llist IC k zs }}}.
+    {{{ zs, RET #(); ⌜zs ≡ₚ map_reduce map red xs⌝ ∗ llist IC l zs }}}.
   Proof.
     iIntros (?) "#Hmap #Hred !>"; iIntros (Φ) "Hl HΦ". wp_lam; wp_pures.
-    wp_apply (start_map_service_spec with "Hmap [//]"); iIntros (cmap) "Hcmap".
+    wp_apply (start_chan_proto_spec N (par_map_protocol IA IZB map n ∅));
+      iIntros (cmap) "// Hcmap".
+    { wp_pures. wp_apply (par_map_service_spec with "Hmap Hcmap"); auto. }
     wp_apply (start_chan_proto_spec N (sort_fg_protocol IZB RZB <++> END)%proto);
       iIntros (csort) "Hcsort".
     { wp_apply (sort_service_fg_spec N with "[] Hcsort"); last by auto.
       iApply RZB_cmp_spec. }
     rewrite right_id.
-    wp_apply (par_map_reduce_map_server_spec with "[$Hl $Hcmap $Hcsort]"); first lia.
+    wp_apply (par_map_reduce_map_spec with "[$Hl $Hcmap $Hcsort]"); first lia.
     iIntros (iys). rewrite gmultiset_elements_empty right_id_L.
-    iDestruct 1 as (Hiys) "Hcsort /=". wp_select; wp_pures; simpl.
-    wp_apply (start_map_service_spec with "Hred [//]"); iIntros (cred) "Hcred".
+    iDestruct 1 as (Hiys) "[Hl Hcsort] /=". wp_select; wp_pures; simpl.
+    wp_apply (start_chan_proto_spec N (par_map_protocol IZBs IC (curry red) n ∅));
+      iIntros (cred) "// Hcred".
+    { wp_pures. wp_apply (par_map_service_spec with "Hred Hcred"); auto. }
     wp_branch as %_|%Hnil; last first.
-    { wp_pures. wp_apply (lnil_spec with "[//]"); iIntros (k) "Hk".
-      iApply ("HΦ" $! k [] with "[$Hk]"); simpl.
+    { wp_pures. iApply ("HΦ" $! [] with "[$Hl]"); simpl.
       by rewrite /map_reduce /= -Hiys -Hnil. }
     wp_recv ([i y] ?) as (_ w ->) "HIB /="; wp_pures.
-    wp_apply (lnil_spec with "[//]"); iIntros (k) "Hk".
-    wp_apply (par_map_reduce_reduce_server_spec _ _ [] (Some (i, y, w))
-      with "[$Hk $Hcsort $Hcred $HIB]"); simpl; auto; [lia|set_solver|].
+    wp_apply (par_map_reduce_reduce_spec _ _ [] (Some (i, y, w)) []
+      with "[$Hl $Hcsort $Hcred $HIB]"); simpl; auto; [lia|set_solver|].
     iIntros (zs). rewrite /= gmultiset_elements_empty !right_id.
     iDestruct 1 as (Hzs) "Hk". wp_pures.
     iApply ("HΦ" with "[$Hk]"). by rewrite Hzs Hiys.
diff --git a/theories/examples/sort_fg_client.v b/theories/examples/sort_fg_client.v
index a8c8306..267b7f7 100644
--- a/theories/examples/sort_fg_client.v
+++ b/theories/examples/sort_fg_client.v
@@ -11,16 +11,16 @@ Definition send_all : val :=
     send "c" #cont;; send "c" (lpop "xs");; "go" "c" "xs".
 
 Definition recv_all : val :=
-  rec: "go" "c" :=
-    if: ~recv "c" then lnil #() else
+  rec: "go" "c" "ys" :=
+    if: ~recv "c" then #() else
     let: "x" := recv "c" in
-    let: "l" := "go" "c" in lcons "x" "l";; "l".
+    "go" "c" "ys";; lcons "x" "ys".
 
 Definition sort_client_fg : val := λ: "cmp" "xs",
   let: "c" := start_chan (λ: "c", sort_service_fg "cmp" "c") in
   send_all "c" "xs";;
   send "c" #stop;;
-  recv_all "c".
+  recv_all "c" "xs".
 
 Section sort_fg_client.
   Context `{!heapG Σ, !proto_chanG Σ} (N : namespace).
@@ -29,7 +29,7 @@ Section sort_fg_client.
   Lemma send_all_spec c p xs' l xs :
     {{{ llist I l xs ∗ c ↣ sort_fg_head_protocol I R xs' <++> p @ N }}}
       send_all c #l
-    {{{ RET #(); c ↣ sort_fg_head_protocol I R (xs' ++ xs) <++> p @ N }}}.
+    {{{ RET #(); llist I l [] ∗ c ↣ sort_fg_head_protocol I R (xs' ++ xs) <++> p @ N }}}.
   Proof.
     iIntros (Φ) "[Hl Hc] HΦ".
     iInduction xs as [|x xs] "IH" forall (xs').
@@ -40,40 +40,39 @@ Section sort_fg_client.
     wp_send with "[$HIx]". wp_apply ("IH" with "Hl Hc"). by rewrite -assoc_L.
   Qed.
 
-  Lemma recv_all_spec c p xs ys' :
+  Lemma recv_all_spec c p l xs ys' :
     Sorted R ys' →
-    {{{ c ↣ sort_fg_tail_protocol I R xs ys' <++> p @ N }}}
-      recv_all c
-    {{{ l ys, RET #l;
+    {{{ llist I l [] ∗ c ↣ sort_fg_tail_protocol I R xs ys' <++> p @ N }}}
+      recv_all c #l
+    {{{ ys, RET #();
         ⌜Sorted R (ys' ++ ys)⌝ ∗ ⌜ys' ++ ys ≡ₚ xs⌝ ∗ llist I l ys ∗ c ↣ p @ N
     }}}.
   Proof.
-    iIntros (Hsort Φ) "Hc HΦ".
+    iIntros (Hsort Φ) "[Hl Hc] HΦ".
     iLöb as "IH" forall (xs ys' Φ Hsort).
     wp_lam. wp_branch as %_ | %Hperm; wp_pures.
     - wp_recv (y v) as (Htl) "HIx".
-      wp_apply ("IH" with "[] Hc"); first by auto using Sorted_snoc.
-      iIntros (l ys). rewrite -!assoc_L. iDestruct 1 as (??) "[Hl Hc]".
+      wp_apply ("IH" with "[] Hl Hc"); first by auto using Sorted_snoc.
+      iIntros (ys). rewrite -!assoc_L. iDestruct 1 as (??) "[Hl Hc]".
       wp_apply (lcons_spec with "[$Hl $HIx]"); iIntros "Hl"; wp_pures.
       iApply ("HΦ" with "[$Hl $Hc]"); simpl; eauto.
-    - wp_apply (lnil_spec with "[//]"); iIntros (l) "Hl".
-      iApply ("HΦ" $! _ []); rewrite /= right_id_L; by iFrame.
+    - iApply ("HΦ" $! []); rewrite /= right_id_L; by iFrame.
   Qed.
 
   Lemma sort_client_fg_spec cmp l xs :
     cmp_spec I R cmp -∗
     {{{ llist I l xs }}}
       sort_client_fg cmp #l
-    {{{ k ys, RET #k; ⌜Sorted R ys⌝ ∗ ⌜ys ≡ₚ xs⌝ ∗ llist I k ys }}}.
+    {{{ ys, RET #(); ⌜Sorted R ys⌝ ∗ ⌜ys ≡ₚ xs⌝ ∗ llist I l ys }}}.
   Proof.
     iIntros "#Hcmp !>" (Φ) "Hl HΦ". wp_lam.
     wp_apply (start_chan_proto_spec N (sort_fg_protocol I R <++> END)%proto);
       iIntros (c) "Hc".
     { wp_apply (sort_service_fg_spec N with "Hcmp Hc"); auto. }
-    wp_apply (send_all_spec with "[$Hl $Hc]"); iIntros "Hc".
+    wp_apply (send_all_spec with "[$Hl $Hc]"); iIntros "[Hl Hc]".
     wp_select.
-    wp_apply (recv_all_spec with "Hc"); auto.
-    iIntros (k ys) "/=". iDestruct 1 as (??) "[Hk _]".
+    wp_apply (recv_all_spec with "[$Hl $Hc]"); auto.
+    iIntros (ys) "/=". iDestruct 1 as (??) "[Hk _]".
     iApply "HΦ"; auto with iFrame.
   Qed.
 End sort_fg_client.
-- 
GitLab