From aa9b2a0b736000e5d4b18a5a6f82d3c956d7b444 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jkas@itu.dk> Date: Wed, 8 May 2019 13:14:42 +0200 Subject: [PATCH] Moved split to list library --- theories/encodings/list.v | 21 +++++++++++++++++++++ theories/examples/list_sort.v | 25 +------------------------ 2 files changed, 22 insertions(+), 24 deletions(-) diff --git a/theories/encodings/list.v b/theories/encodings/list.v index c0ccf3e..aca68f3 100644 --- a/theories/encodings/list.v +++ b/theories/encodings/list.v @@ -66,6 +66,13 @@ Definition lsnoc : val := | NONE => lcons "x" NONE end. +Definition lsplit : val := + λ: "xs", + let: "hd" := lhead "xs" in + let: "ys" := lcons "hd" NONEV in + let: "zs" := ltail "xs" in + ("ys", "zs"). + Section lists. Context `{heapG Σ}. Implicit Types i : nat. @@ -175,4 +182,18 @@ Proof. iIntros (_). by wp_apply lcons_spec=> //. Qed. +Lemma lsplit_spec xs x : + {{{ True }}} lsplit (encode (x::xs)) {{{ RET (encode [x],encode xs); True }}}. +Proof. + iIntros (Φ _) "HΦ". + wp_lam. wp_apply (lhead_spec)=> //. + iIntros (_). + wp_apply (lcons_spec _ [])=> //. + iIntros (_). + wp_apply (ltail_spec)=> //. + iIntros (_). + wp_pures. + by iApply "HΦ". +Qed. + End lists. \ No newline at end of file diff --git a/theories/examples/list_sort.v b/theories/examples/list_sort.v index 9706add..555e874 100644 --- a/theories/examples/list_sort.v +++ b/theories/examples/list_sort.v @@ -11,29 +11,6 @@ Require Import Omega. Section ListSortExample. Context `{!heapG Σ} `{logrelG val Σ} (N : namespace). - Definition lsplit : val := - λ: "xs", - let: "hd" := lhead "xs" in - let: "ys" := lcons "hd" NONEV in - let: "zs" := ltail "xs" in - ("ys", "zs"). - - Lemma lsplit_spec (x : Z) (xs : list Z) : - {{{ True }}} - lsplit (encode (x::xs)) - {{{ RET (encode [x],encode xs); True }}}. - Proof. - iIntros (Φ _) "HΦ". - wp_lam. wp_apply (lhead_spec (T := Z))=> //. - iIntros (_). - wp_apply (lcons_spec (T := Z) _ [])=> //. - iIntros (_). - wp_apply (ltail_spec (T := Z))=> //. - iIntros (_). - wp_pures. - by iApply "HΦ". - Qed. - Definition lsplit_ref : val := λ: "xs", let: "ys_zs" := lsplit !"xs" in (ref (Fst "ys_zs"), ref (Snd ("ys_zs"))). @@ -48,7 +25,7 @@ Section ListSortExample. lzs ↦ encode zs }}}. Proof. iIntros (Φ) "Hhd HΦ". - wp_lam. wp_load. wp_apply (lsplit_spec)=> //; iIntros (_). + wp_lam. wp_load. wp_apply (lsplit_spec (T:=Z))=> //; iIntros (_). wp_alloc lzs as "Hlzs". wp_alloc lys as "Hlys". wp_pures. -- GitLab