Skip to content
Snippets Groups Projects
spawn.v 2.96 KiB
Newer Older
From iris.program_logic Require Export weakestpre.
From iris.base_logic.lib Require Export invariants.
From iris.heap_lang Require Export lang.
From iris.proofmode Require Import tactics.
Robbert Krebbers's avatar
Robbert Krebbers committed
From iris.heap_lang Require Import proofmode notation.
From iris.algebra Require Import excl.
Set Default Proof Using "Type".
Ralf Jung's avatar
Ralf Jung committed

Definition spawn : val :=
Robbert Krebbers's avatar
Robbert Krebbers committed
    let: "c" := ref NONE in
    Fork ("c" <- SOME ("f" #())) ;; "c".
Ralf Jung's avatar
Ralf Jung committed
Definition join : val :=
Robbert Krebbers's avatar
Robbert Krebbers committed
      SOME "x" => "x"
    | NONE => "join" "c"
(** The CMRA & functor we need. *)
Ralf Jung's avatar
Ralf Jung committed
(* Not bundling heapG, as it may be shared with other users. *)
Class spawnG Σ := SpawnG { spawn_tokG :> inG Σ (exclR unitC) }.
Definition spawnΣ : gFunctors := #[GFunctor (exclR unitC)].
Instance subG_spawnΣ {Σ} : subG spawnΣ Σ  spawnG Σ.
Ralf Jung's avatar
Ralf Jung committed

(** Now we come to the Iris part of the proof. *)
Section proof.
Context `{!heapG Σ, !spawnG Σ} (N : namespace).
Definition spawn_inv (γ : gname) (l : loc) (Ψ : val  iProp Σ) : iProp Σ :=
  ( lv, l  lv  (lv = NONEV 
                    v, lv = SOMEV v  (Ψ v  own γ (Excl ()))))%I.
Definition join_handle (l : loc) (Ψ : val  iProp Σ) : iProp Σ :=
  ( γ, own γ (Excl ())  inv N (spawn_inv γ l Ψ))%I.
Ralf Jung's avatar
Ralf Jung committed

Global Instance spawn_inv_ne n γ l :
  Proper (pointwise_relation val (dist n) ==> dist n) (spawn_inv γ l).
Proof. solve_proper. Qed.
Global Instance join_handle_ne n l :
  Proper (pointwise_relation val (dist n) ==> dist n) (join_handle l).
Proof. solve_proper. Qed.

(** The main proofs. *)
Lemma spawn_spec (Ψ : val  iProp Σ) e (f : val) `{Hef : !IntoVal e f} :
  {{{ WP f #() {{ Ψ }} }}} spawn e {{{ l, RET #l; join_handle l Ψ }}}.
Ralf Jung's avatar
Ralf Jung committed
Proof.
  apply of_to_val in Hef as <-. iIntros (Φ) "Hf HΦ". rewrite /spawn /=.
  wp_let. wp_alloc l as "Hl". wp_let.
  iMod (own_alloc (Excl ())) as (γ) "Hγ"; first done.
  iMod (inv_alloc N _ (spawn_inv γ l Ψ) with "[Hl]") as "#?".
Robbert Krebbers's avatar
Robbert Krebbers committed
  { iNext. iExists NONEV. iFrame; eauto. }
  wp_apply wp_fork; simpl. iSplitR "Hf".
  - wp_seq. iApply "HΦ". rewrite /join_handle. eauto.
  - wp_bind (f _). iApply (wp_wand with "Hf"); iIntros (v) "Hv".
    iInv N as (v') "[Hl _]" "Hclose".
    wp_store. iApply "Hclose". iNext. iExists (SOMEV v). iFrame. eauto.
Ralf Jung's avatar
Ralf Jung committed
Qed.

Lemma join_spec (Ψ : val  iProp Σ) l :
  {{{ join_handle l Ψ }}} join #l {{{ v, RET v; Ψ v }}}.
Ralf Jung's avatar
Ralf Jung committed
Proof.
  iIntros (Φ) "H HΦ". iDestruct "H" as (γ) "[Hγ #?]".
  iLöb as "IH". wp_rec. wp_bind (! _)%E. iInv N as (v) "[Hl Hinv]" "Hclose".
Robbert Krebbers's avatar
Robbert Krebbers committed
  wp_load. iDestruct "Hinv" as "[%|Hinv]"; subst.
  - iMod ("Hclose" with "[Hl]"); [iNext; iExists _; iFrame; eauto|].
Ralf Jung's avatar
Ralf Jung committed
    iModIntro. wp_match. iApply ("IH" with "Hγ [HΦ]"). auto.
  - iDestruct "Hinv" as (v') "[% [HΨ|Hγ']]"; simplify_eq/=.
    + iMod ("Hclose" with "[Hl Hγ]"); [iNext; iExists _; iFrame; eauto|].
Ralf Jung's avatar
Ralf Jung committed
      iModIntro. wp_match. by iApply "HΦ".
    + iDestruct (own_valid_2 with "Hγ Hγ'") as %[].
Ralf Jung's avatar
Ralf Jung committed
Qed.
Ralf Jung's avatar
Ralf Jung committed
End proof.
Robbert Krebbers's avatar
Robbert Krebbers committed

Typeclasses Opaque join_handle.