Skip to content
Snippets Groups Projects
Commit d800f42e authored by Ralf Jung's avatar Ralf Jung
Browse files

port the testcase to use the heap.v

parent b6b2f174
No related branches found
No related tags found
No related merge requests found
From heap_lang Require Export derived. From heap_lang Require Export derived.
From program_logic Require Export invariants ghost_ownership.
From program_logic Require Import ownership auth. From program_logic Require Import ownership auth.
From heap_lang Require Import notation. From heap_lang Require Import notation.
Import uPred. Import uPred.
......
(** This file is essentially a bunch of testcases. *) (** This file is essentially a bunch of testcases. *)
From program_logic Require Import ownership. From program_logic Require Import ownership.
From heap_lang Require Import substitution tactics notation. From heap_lang Require Import substitution tactics heap notation.
Import uPred. Import uPred.
Module LangTests. Section LangTests.
Definition add := ('21 + '21)%L. Definition add := ('21 + '21)%L.
Goal σ, prim_step add σ ('42) σ None. Goal σ, prim_step add σ ('42) σ None.
Proof. intros; do_step done. Qed. Proof. intros; do_step done. Qed.
...@@ -21,36 +21,30 @@ Module LangTests. ...@@ -21,36 +21,30 @@ Module LangTests.
Qed. Qed.
End LangTests. End LangTests.
Module LiftingTests. Section LiftingTests.
Context {Σ : iFunctor}. Context {Σ : iFunctorG} (HeapI : gid) `{!HeapInG Σ HeapI}.
Implicit Types P : iProp heap_lang Σ. Implicit Types P : iPropG heap_lang Σ.
Implicit Types Q : val iProp heap_lang Σ. Implicit Types Q : val iPropG heap_lang Σ.
Definition e : expr := Definition e : expr :=
let: "x" := ref '1 in "x" <- !"x" + '1;; !"x". let: "x" := ref '1 in "x" <- !"x" + '1;; !"x".
Goal σ E, ownP (Σ:=Σ) σ wp E e (λ v, v = '2). Goal γh N, heap_ctx HeapI γh N wp N e (λ v, v = '2).
Proof. Proof.
move=> σ E. rewrite /e. move=> γh N. rewrite /e.
rewrite -(wp_bindi (LetCtx _ _)) -wp_alloc_pst //=. rewrite -(wp_bindi (LetCtx _ _)). eapply wp_alloc; eauto; [].
apply sep_intro_True_r; first done.
rewrite -later_intro; apply forall_intro=>l; apply wand_intro_l. rewrite -later_intro; apply forall_intro=>l; apply wand_intro_l.
rewrite right_id; apply const_elim_l=> _.
rewrite -wp_let //= -later_intro. rewrite -wp_let //= -later_intro.
rewrite -(wp_bindi (SeqCtx (Load (Loc _)))) /=. rewrite -(wp_bindi (SeqCtx (Load (Loc _)))) /=.
rewrite -(wp_bindi (StoreRCtx (LocV _))) /=. rewrite -(wp_bindi (StoreRCtx (LocV _))) /=.
rewrite -(wp_bindi (BinOpLCtx PlusOp _)) /=. rewrite -(wp_bindi (BinOpLCtx PlusOp _)) /=.
rewrite -wp_load_pst; first (apply sep_intro_True_r; first done); last first. eapply wp_load; eauto with I; []. apply sep_mono; first done.
{ by rewrite lookup_insert. } (* RJ FIXME: figure out why apply and eapply fail. *) rewrite -later_intro; apply wand_intro_l.
rewrite -later_intro; apply wand_intro_l; rewrite right_id.
rewrite -wp_bin_op // -later_intro. rewrite -wp_bin_op // -later_intro.
rewrite -wp_store_pst; first (apply sep_intro_True_r; first done); last first. eapply wp_store; eauto with I; []. apply sep_mono; first done.
{ by rewrite lookup_insert. } rewrite -later_intro. apply wand_intro_l.
{ done. }
rewrite -later_intro. apply wand_intro_l. rewrite right_id.
rewrite -wp_seq -wp_value' -later_intro. rewrite -wp_seq -wp_value' -later_intro.
rewrite -wp_load_pst; first (apply sep_intro_True_r; first done); last first. eapply wp_load; eauto with I; []. apply sep_mono; first done.
{ by rewrite lookup_insert. } rewrite -later_intro. apply wand_intro_l.
rewrite -later_intro. apply wand_intro_l. rewrite right_id.
by apply const_intro. by apply const_intro.
Qed. Qed.
...@@ -106,7 +100,7 @@ Module LiftingTests. ...@@ -106,7 +100,7 @@ Module LiftingTests.
Qed. Qed.
Goal E, Goal E,
True wp (Σ:=Σ) E (let: "x" := Pred '42 in Pred "x") (λ v, v = '40). True wp (Σ:=globalF Σ) E (let: "x" := Pred '42 in Pred "x") (λ v, v = '40).
Proof. Proof.
intros E. intros E.
rewrite -(wp_bindi (LetCtx _ _)) -Pred_spec //= -wp_let' //=. rewrite -(wp_bindi (LetCtx _ _)) -Pred_spec //= -wp_let' //=.
......
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment