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

make `wp_alloc l as "?"` behave as expected

parent aa6ecfbb
No related branches found
No related tags found
No related merge requests found
...@@ -18,6 +18,18 @@ ...@@ -18,6 +18,18 @@
--------------------------------------∗ --------------------------------------∗
WP #l <- #1 + #1;; ! #l @ E {{ v, ⌜v = #2⌝ }} WP #l <- #1 + #1;; ! #l @ E {{ v, ⌜v = #2⌝ }}
1 subgoal
Σ : gFunctors
H : heapG Σ
E : coPset
l : loc
============================
"Hl" : l ↦ #1
--------------------------------------∗
WP let: "x" := #l in let: "y" := ref #1 in "x" <- ! "x" + #1;; ! "x"
@ E [{ v, ⌜v = #2⌝ }]
1 subgoal 1 subgoal
Σ : gFunctors Σ : gFunctors
......
...@@ -27,7 +27,7 @@ Section tests. ...@@ -27,7 +27,7 @@ Section tests.
Lemma heap_e_spec E : WP heap_e @ E {{ v, v = #2 }}%I. Lemma heap_e_spec E : WP heap_e @ E {{ v, v = #2 }}%I.
Proof. Proof.
iIntros "". rewrite /heap_e. Show. iIntros "". rewrite /heap_e. Show.
wp_alloc l. wp_let. wp_load. Show. wp_alloc l as "?". wp_let. wp_load. Show.
wp_op. wp_store. by wp_load. wp_op. wp_store. by wp_load.
Qed. Qed.
...@@ -39,7 +39,7 @@ Section tests. ...@@ -39,7 +39,7 @@ Section tests.
Lemma heap_e2_spec E : WP heap_e2 @ E [{ v, v = #2 }]%I. Lemma heap_e2_spec E : WP heap_e2 @ E [{ v, v = #2 }]%I.
Proof. Proof.
iIntros "". rewrite /heap_e2. iIntros "". rewrite /heap_e2.
wp_alloc l. wp_let. wp_alloc l'. wp_let. wp_alloc l as "Hl". Show. wp_let. wp_alloc l'. wp_let.
wp_load. wp_op. wp_store. wp_load. done. wp_load. wp_op. wp_store. wp_load. done.
Qed. Qed.
......
...@@ -360,31 +360,32 @@ Tactic Notation "wp_apply" open_constr(lem) := ...@@ -360,31 +360,32 @@ Tactic Notation "wp_apply" open_constr(lem) :=
end). end).
Tactic Notation "wp_alloc" ident(l) "as" constr(H) := Tactic Notation "wp_alloc" ident(l) "as" constr(H) :=
let Htmp := iFresh in
let finish _ := let finish _ :=
first [intros l | fail 1 "wp_alloc:" l "not fresh"]; first [intros l | fail 1 "wp_alloc:" l "not fresh"];
eexists; split; eexists; split;
[pm_reflexivity || fail "wp_alloc:" H "not fresh" [pm_reflexivity || fail "wp_alloc:" H "not fresh"
|wp_expr_simpl; try wp_value_head] in |iDestruct Htmp as H; wp_expr_simpl; try wp_value_head] in
iStartProof; iStartProof;
lazymatch goal with lazymatch goal with
| |- envs_entails _ (wp ?s ?E ?e ?Q) => | |- envs_entails _ (wp ?s ?E ?e ?Q) =>
first first
[reshape_expr e ltac:(fun K e' => [reshape_expr e ltac:(fun K e' =>
eapply (tac_wp_alloc _ _ _ _ H K); [iSolveTC|..]) eapply (tac_wp_alloc _ _ _ _ Htmp K); [iSolveTC|..])
|fail 1 "wp_alloc: cannot find 'Alloc' in" e]; |fail 1 "wp_alloc: cannot find 'Alloc' in" e];
[iSolveTC [iSolveTC
|finish ()] |finish ()]
| |- envs_entails _ (twp ?s ?E ?e ?Q) => | |- envs_entails _ (twp ?s ?E ?e ?Q) =>
first first
[reshape_expr e ltac:(fun K e' => [reshape_expr e ltac:(fun K e' =>
eapply (tac_twp_alloc _ _ _ H K); [iSolveTC|..]) eapply (tac_twp_alloc _ _ _ Htmp K); [iSolveTC|..])
|fail 1 "wp_alloc: cannot find 'Alloc' in" e]; |fail 1 "wp_alloc: cannot find 'Alloc' in" e];
finish () finish ()
| _ => fail "wp_alloc: not a 'wp'" | _ => fail "wp_alloc: not a 'wp'"
end. end.
Tactic Notation "wp_alloc" ident(l) := Tactic Notation "wp_alloc" ident(l) :=
let H := iFresh in wp_alloc l as H. wp_alloc l as "?".
Tactic Notation "wp_load" := Tactic Notation "wp_load" :=
let solve_mapsto _ := let solve_mapsto _ :=
......
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