Skip to content
GitLab
Menu
Projects
Groups
Snippets
Help
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
Menu
Open sidebar
Rice Wine
Iris
Commits
393f02ed
Commit
393f02ed
authored
Jun 21, 2018
by
Ralf Jung
Browse files
get rid of to_val_is_Some'
parent
e04c9826
Changes
2
Hide whitespace changes
Inline
Side-by-side
theories/heap_lang/tactics.v
View file @
393f02ed
...
...
@@ -138,16 +138,13 @@ Fixpoint to_val (e : expr) : option val :=
|
_
=>
None
end
.
Lemma
to_val_Some
e
v
:
to_val
e
=
Some
v
→
of_val
v
=
W
.
to_expr
e
.
to_val
e
=
Some
v
→
heap_lang
.
of_val
v
=
W
.
to_expr
e
.
Proof
.
revert
v
.
induction
e
;
intros
;
simplify_option_eq
;
try
f_equal
;
auto
using
of_to_val
.
Qed
.
Lemma
to_val_is_Some
e
:
is_Some
(
to_val
e
)
→
∃
v
,
of_val
v
=
to_expr
e
.
is_Some
(
to_val
e
)
→
∃
v
,
heap_lang
.
of_val
v
=
to_expr
e
.
Proof
.
intros
[
v
?]
;
exists
v
;
eauto
using
to_val_Some
.
Qed
.
Lemma
to_val_is_Some'
e
:
is_Some
(
to_val
e
)
→
is_Some
(
heap_lang
.
to_val
(
to_expr
e
)).
Proof
.
intros
[
v
?]%
to_val_is_Some
.
exists
v
.
rewrite
-
to_of_val
.
by
f_equal
.
Qed
.
Fixpoint
subst
(
x
:
string
)
(
es
:
expr
)
(
e
:
expr
)
:
expr
:
=
match
e
with
...
...
@@ -202,8 +199,8 @@ Proof.
inversion
1
;
simplify_eq
/=
;
rewrite
?to_of_val
;
eauto
.
unfold
subst'
;
repeat
(
simplify_eq
/=
;
case_match
=>//)
;
eauto
.
-
apply
ectxi_language_sub_redexes_are_values
=>
/=
Ki
e'
Hfill
.
destruct
e
=>
//
;
destruct
Ki
;
repeat
(
simplify_eq
/=
;
case_match
=>//)
;
naive_solver
eauto
using
to_val_is_Some
'
.
destruct
e
=>
//
;
destruct
Ki
;
repeat
(
simplify_eq
/=
;
case_match
=>//)
;
try
naive_solver
eauto
using
as_val_is_Some
,
to_val_is_Some
.
Qed
.
End
W
.
...
...
theories/program_logic/language.v
View file @
393f02ed
...
...
@@ -170,4 +170,7 @@ Section language.
apply
TCForall_Forall
,
Forall_fmap
,
Forall_true
=>
v
.
rewrite
/
AsVal
/=
;
eauto
.
Qed
.
Lemma
as_val_is_Some
e
:
(
∃
v
,
of_val
v
=
e
)
→
is_Some
(
to_val
e
).
Proof
.
intros
[
v
<-].
rewrite
to_of_val
.
eauto
.
Qed
.
End
language
.
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