Newer
Older
(** This file is essentially a bunch of testcases. *)
From iris.program_logic Require Import ownership hoare auth.
Import uPred.
Definition add : expr := (#21 + #21)%E.
Goal ∀ σ, head_step add σ (#42) σ None.
Definition rec_app : expr := ((rec: "f" "x" := "f" "x") #0)%E.
Goal ∀ σ, head_step rec_app σ rec_app σ None.
Proof. intros. rewrite /rec_app. do_head_step done. Qed.
Definition lam : expr := (λ: "x", "x" + #21)%E.
Goal ∀ σ, head_step (lam #21)%E σ add σ None.
Proof. intros. rewrite /lam. do_head_step done. Qed.
Local Notation iProp := (iPropG heap_lang Σ).
Implicit Types P Q : iPropG heap_lang Σ.
Implicit Types Φ : val → iPropG heap_lang Σ.
Definition heap_e : expr :=
let: "x" := ref #1 in "x" <- !"x" + #1 ;; !"x".
Lemma heap_e_spec E :
nclose heapN ⊆ E → heap_ctx ⊢ WP heap_e @ E {{ v, v = #2 }}.
iIntros (HN) "#?". rewrite /heap_e.
wp_alloc l. wp_let. wp_load. wp_op. wp_store. by wp_load.
Definition heap_e2 : expr :=
let: "x" := ref #1 in
let: "y" := ref #1 in
"x" <- !"x" + #1 ;; !"x".
Lemma heap_e2_spec E :
nclose heapN ⊆ E → heap_ctx ⊢ WP heap_e2 @ E {{ v, v = #2 }}.
iIntros (HN) "#?". rewrite /heap_e2.
wp_alloc l. wp_let. wp_alloc l'. wp_let.
wp_load. wp_op. wp_store. wp_load. done.
let: "yp" := "y" + #1 in
if: "yp" < "x" then "pred" "x" "yp" else "y".
if: "x" ≤ #0 then -FindPred (-"x" + #2) #0 else FindPred "x" #0.
Lemma FindPred_spec n1 n2 E Φ :
Φ #(n2 - 1) ⊢ WP FindPred #n2 #n1 @ E {{ Φ }}.
iIntros (Hn) "HΦ". iLöb (n1 Hn) as "IH".
- iApply ("IH" with "[%] HΦ"). omega.
- iApply pvs_intro. by assert (n1 = n2 - 1) as -> by omega.
Lemma Pred_spec n E Φ : ▷ Φ #(n - 1) ⊢ WP Pred #n @ E {{ Φ }}.
wp_op. by replace (n - 1) with (- (-n + 2 - 1)) by omega.
(True : iProp) ⊢ WP let: "x" := Pred #42 in Pred "x" @ E {{ v, v = #40 }}.
Proof. iIntros "". wp_apply Pred_spec. wp_let. by wp_apply Pred_spec. Qed.
Definition Σ : gFunctors := #[ heapGF ].
Notation iProp := (iPropG heap_lang Σ).
Lemma heap_e_closed σ : {{ ownP σ : iProp }} heap_e {{ v, v = #2 }}.
iPvs (heap_alloc with "Hσ") as (h) "[? _]"; first solve_ndisj.
by iApply heap_e_spec; first solve_ndisj.