From 270596d48395d9e09f6afeb3a4fe5985d6ae8c34 Mon Sep 17 00:00:00 2001 From: Robbert Krebbers <mail@robbertkrebbers.nl> Date: Wed, 1 Apr 2020 02:36:50 +0200 Subject: [PATCH] Some examples. --- theories/examples/basics.v | 48 ++++++++++++++++++++++++++++++++++++++ 1 file changed, 48 insertions(+) diff --git a/theories/examples/basics.v b/theories/examples/basics.v index 533cb99..08b0cf0 100644 --- a/theories/examples/basics.v +++ b/theories/examples/basics.v @@ -86,6 +86,24 @@ Definition prog_lock : val := λ: <>, acquire "l";; send "c'" #21;; release "l") in recv "c" + recv "c". +(** Swapping of sends *) +Definition prog_swap : val := λ: <>, + let: "c" := start_chan (λ: "c'", + send "c'" #20;; + let: "y" := recv "c'" in + send "c'" "y") in + send "c" #22;; + recv "c" + recv "c". + +Definition prog_swap_twice : val := λ: <>, + let: "c" := start_chan (λ: "c'", + send "c'" #20;; + let: "y1" := recv "c'" in + let: "y2" := recv "c'" in + send "c'" ("y1" + "y2")) in + send "c" #11;; send "c" #11;; + recv "c" + recv "c". + Section proofs. Context `{heapG Σ, chanG Σ}. @@ -142,6 +160,17 @@ Fixpoint prot_lock (n : nat) : iProto Σ := | S n' => <?> MSG #21; prot_lock n' end%proto. +Definition prot_swap : iProto Σ := + (<!> x : Z, MSG #x; + <?> MSG #20; + <?> MSG #x; END)%proto. + +Definition prot_swap_twice : iProto Σ := + (<!> x : Z, MSG #x; + <!> y : Z, MSG #y; + <?> MSG #20; + <?> MSG #(x+y); END)%proto. + (** Specs and proofs of the respective programs *) Lemma prog_spec : {{{ True }}} prog #() {{{ RET #42; True }}}. Proof. @@ -270,4 +299,23 @@ Proof. by wp_apply "Hhelp". - wp_recv as "_". wp_recv as "_". wp_pures. by iApply "HΦ". Qed. + +Lemma prog_swap_spec : {{{ True }}} prog_swap #() {{{ RET #42; True }}}. +Proof. + iIntros (Φ) "_ HΦ". wp_lam. + wp_apply (start_chan_spec prot_swap); iIntros (c) "Hc". + - wp_send with "[//]". wp_recv (x) as "_". by wp_send with "[//]". + - wp_send with "[//]". wp_recv as "_". wp_recv as "_". + wp_pures. by iApply "HΦ". +Qed. + +Lemma prog_swap_twice_spec : {{{ True }}} prog_swap_twice #() {{{ RET #42; True }}}. +Proof. + iIntros (Φ) "_ HΦ". wp_lam. + wp_apply (start_chan_spec prot_swap_twice); iIntros (c) "Hc". + - wp_send with "[//]". wp_recv (x1) as "_". wp_recv (x2) as "_". + by wp_send with "[//]". + - wp_send with "[//]". wp_send with "[//]". wp_recv as "_". wp_recv as "_". + wp_pures. by iApply "HΦ". +Qed. End proofs. -- GitLab