From abd4de4583f7297c83c99ea13db4dc03029a0f53 Mon Sep 17 00:00:00 2001 From: Ralf Jung <jung@mpi-sws.org> Date: Mon, 19 Feb 2018 18:02:50 +0100 Subject: [PATCH] define logically atomic heap and instantiate it with physical heap --- _CoqProject | 1 + theories/heap_lang/lib/atomic_heap.v | 81 ++++++++++++++++++++++++++++ theories/program_logic/atomic.v | 4 +- 3 files changed, 84 insertions(+), 2 deletions(-) create mode 100644 theories/heap_lang/lib/atomic_heap.v diff --git a/_CoqProject b/_CoqProject index 799b53c49..8b36d717b 100644 --- a/_CoqProject +++ b/_CoqProject @@ -85,6 +85,7 @@ theories/heap_lang/lib/lock.v theories/heap_lang/lib/spin_lock.v theories/heap_lang/lib/ticket_lock.v theories/heap_lang/lib/counter.v +theories/heap_lang/lib/atomic_heap.v theories/heap_lang/proofmode.v theories/heap_lang/adequacy.v theories/heap_lang/total_adequacy.v diff --git a/theories/heap_lang/lib/atomic_heap.v b/theories/heap_lang/lib/atomic_heap.v new file mode 100644 index 000000000..fe99f8ead --- /dev/null +++ b/theories/heap_lang/lib/atomic_heap.v @@ -0,0 +1,81 @@ +From iris.heap_lang Require Export lifting notation. +From iris.base_logic.lib Require Export invariants. +From iris.program_logic Require Export atomic. +From iris.proofmode Require Import tactics. +From iris.heap_lang Require Import proofmode notation. +Set Default Proof Using "Type". + +(** A general logically atomic interface for a heap. *) +Structure atomic_heap Σ `{!heapG Σ} := AtomicHeap { + (* -- operations -- *) + alloc : val; + load : val; + store : val; + (* -- predicates -- *) + (* name is used to associate locked with is_lock *) + mapsto (l : loc) (q: Qp) (v : val) : iProp Σ; + (* -- general properties -- *) + mapsto_timeless l q v : Timeless (mapsto l q v); + (* -- operation specs -- *) + alloc_spec v : + {{{ True }}} alloc v {{{ l, RET #l; mapsto l 1 v }}}; + load_spec (l : loc) : + atomic_wp (load #l)%E + (λ '(v, q), mapsto l q v) + (λ '(v, q) (_:()), mapsto l q v) + ⊤ ∅ + (λ '(v, q) _, v); + store_spec (l : loc) (w : val) : + atomic_wp (store (#l, w))%E + (λ v, mapsto l 1 v) + (λ v (_:()), mapsto l 1 w) + ⊤ ∅ + (λ _ _, #()%V); +}. + +(** Proof that the primitive physical operations of heap_lang satisfy said interface. *) +Definition primitive_alloc : val := + λ: "v", ref "v". +Definition primitive_load : val := + λ: "l", !"l". +Definition primitive_store : val := + λ: "p", (Fst "p") <- (Snd "p"). + +Section proof. + Context `{!heapG Σ}. + + Lemma primitive_alloc_spec v : + {{{ True }}} primitive_alloc v {{{ l, RET #l; l ↦ v }}}. + Proof. + iIntros (Φ) "_ HΦ". wp_let. wp_alloc l. iApply "HΦ". done. + Qed. + + Lemma primitive_load_spec (l : loc) : + atomic_wp (primitive_load #l)%E + (λ '(v, q), l ↦{q} v)%I + (λ '(v, q) (_:()), l ↦{q} v)%I + ⊤ ∅ + (λ '(v, q) _, v). + Proof. + iIntros (Φ) "Aupd". wp_let. + iMod (aupd_acc with "Aupd") as ((v, q)) "[H↦ [_ Hclose]]"; first solve_ndisj. + wp_load. iMod ("Hclose" $! () with "H↦"). done. + Qed. + + Lemma primitive_store_spec (l : loc) (w : val) : + atomic_wp (primitive_store (#l, w))%E + (λ v, l ↦ v)%I + (λ v (_:()), l ↦ w)%I + ⊤ ∅ + (λ _ _, #()%V). + Proof. + iIntros (Φ) "Aupd". wp_let. wp_proj. wp_proj. + iMod (aupd_acc with "Aupd") as (v) "[H↦ [_ Hclose]]"; first solve_ndisj. + wp_store. iMod ("Hclose" $! () with "H↦"). done. + Qed. + + Definition primitive_atomic_heap : atomic_heap Σ := + {| alloc_spec := primitive_alloc_spec; + load_spec := primitive_load_spec; + store_spec := primitive_store_spec |}. +End proof. diff --git a/theories/program_logic/atomic.v b/theories/program_logic/atomic.v index 74e3311e7..7852c21ad 100644 --- a/theories/program_logic/atomic.v +++ b/theories/program_logic/atomic.v @@ -1,6 +1,6 @@ From iris.program_logic Require Export weakestpre. From iris.proofmode Require Import tactics classes. -From iris.bi.lib Require Import atomic. +From iris.bi.lib Require Export atomic. Set Default Proof Using "Type". Definition atomic_wp `{irisG Λ Σ} {A B : Type} @@ -13,4 +13,4 @@ Definition atomic_wp `{irisG Λ Σ} {A B : Type} (∀ Φ, atomic_update α β Eo Em (λ x y, Φ (f x y)) -∗ WP e {{ Φ }})%I. (* Note: To add a private postcondition, use - atomic_shift α β Eo Em (λ x y, POST x y -∗ Φ (f x y)) *) + atomic_update α β Eo Em (λ x y, POST x y -∗ Φ (f x y)) *) -- GitLab