Skip to content
GitLab
Projects
Groups
Snippets
Help
Loading...
Help
What's new
7
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
Open sidebar
Iris
Iris
Commits
074339d3
Commit
074339d3
authored
Nov 29, 2017
by
Robbert Krebbers
Browse files
Options
Browse Files
Download
Email Patches
Plain Diff
Let `wp_bind` only `simpl` locally.
parent
53c8199e
Changes
1
Hide whitespace changes
Inline
Side-by-side
Showing
1 changed file
with
5 additions
and
4 deletions
+5
-4
theories/heap_lang/proofmode.v
theories/heap_lang/proofmode.v
+5
-4
No files found.
theories/heap_lang/proofmode.v
View file @
074339d3
...
...
@@ -61,15 +61,16 @@ Tactic Notation "wp_proj" := wp_pure (Fst _) || wp_pure (Snd _).
Tactic
Notation
"wp_case"
:
=
wp_pure
(
Case
_
_
_
).
Tactic
Notation
"wp_match"
:
=
wp_case
;
wp_let
.
Lemma
tac_wp_bind
`
{
heapG
Σ
}
K
Δ
E
Φ
e
:
envs_entails
Δ
(
WP
e
@
E
{{
v
,
WP
fill
K
(
of_val
v
)
@
E
{{
Φ
}}
}})%
I
→
Lemma
tac_wp_bind
`
{
heapG
Σ
}
K
Δ
E
Φ
e
f
:
f
=
(
λ
e
,
fill
K
e
)
→
(* as an eta expanded hypothesis so that we can `simpl` it *)
envs_entails
Δ
(
WP
e
@
E
{{
v
,
WP
f
(
of_val
v
)
@
E
{{
Φ
}}
}})%
I
→
envs_entails
Δ
(
WP
fill
K
e
@
E
{{
Φ
}}).
Proof
.
rewrite
/
envs_entails
=>
->.
by
apply
:
wp_bind
.
Qed
.
Proof
.
rewrite
/
envs_entails
=>
->
->.
by
apply
:
wp_bind
.
Qed
.
Ltac
wp_bind_core
K
:
=
lazymatch
eval
hnf
in
K
with
|
[]
=>
idtac
|
_
=>
apply
(
tac_wp_bind
K
)
;
simpl
|
_
=>
e
apply
(
tac_wp_bind
K
)
;
[
simpl
;
reflexivity
|
lazy
beta
]
end
.
Tactic
Notation
"wp_bind"
open_constr
(
efoc
)
:
=
...
...
Write
Preview
Markdown
is supported
0%
Try again
or
attach a new file
.
Attach a file
Cancel
You are about to add
0
people
to the discussion. Proceed with caution.
Finish editing this message first!
Cancel
Please
register
or
sign in
to comment