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
Jonas Kastberg
iris
Commits
ac0d54da
Commit
ac0d54da
authored
Nov 04, 2017
by
Ralf Jung
Browse files
Options
Browse Files
Download
Email Patches
Plain Diff
rename sub_values -> sub_redexes_are_values
parent
0582920d
Changes
3
Hide whitespace changes
Inline
Side-by-side
Showing
3 changed files
with
10 additions
and
7 deletions
+10
-7
theories/heap_lang/tactics.v
theories/heap_lang/tactics.v
+1
-1
theories/program_logic/ectx_language.v
theories/program_logic/ectx_language.v
+6
-4
theories/program_logic/ectxi_language.v
theories/program_logic/ectxi_language.v
+3
-2
No files found.
theories/heap_lang/tactics.v
View file @
ac0d54da
...
@@ -195,7 +195,7 @@ Proof.
...
@@ -195,7 +195,7 @@ Proof.
destruct
e
=>
//=
;
repeat
(
simplify_eq
/=
;
case_match
=>//)
;
destruct
e
=>
//=
;
repeat
(
simplify_eq
/=
;
case_match
=>//)
;
inversion
1
;
simplify_eq
/=
;
rewrite
?to_of_val
;
eauto
.
inversion
1
;
simplify_eq
/=
;
rewrite
?to_of_val
;
eauto
.
unfold
subst'
;
repeat
(
simplify_eq
/=
;
case_match
=>//)
;
eauto
.
unfold
subst'
;
repeat
(
simplify_eq
/=
;
case_match
=>//)
;
eauto
.
-
apply
ectxi_language_sub_values
=>
/=
Ki
e'
Hfill
.
-
apply
ectxi_language_sub_
redexes_are_
values
=>
/=
Ki
e'
Hfill
.
destruct
e
=>
//
;
destruct
Ki
;
repeat
(
simplify_eq
/=
;
case_match
=>//)
;
destruct
e
=>
//
;
destruct
Ki
;
repeat
(
simplify_eq
/=
;
case_match
=>//)
;
naive_solver
eauto
using
to_val_is_Some
.
naive_solver
eauto
using
to_val_is_Some
.
Qed
.
Qed
.
...
...
theories/program_logic/ectx_language.v
View file @
ac0d54da
...
@@ -62,7 +62,9 @@ Section ectx_language.
...
@@ -62,7 +62,9 @@ Section ectx_language.
Definition
head_irreducible
(
e
:
expr
)
(
σ
:
state
)
:
=
Definition
head_irreducible
(
e
:
expr
)
(
σ
:
state
)
:
=
∀
e'
σ
'
efs
,
¬
head_step
e
σ
e'
σ
'
efs
.
∀
e'
σ
'
efs
,
¬
head_step
e
σ
e'
σ
'
efs
.
Definition
sub_values
(
e
:
expr
)
:
=
(* All non-value redexes are at the root. In other words, all sub-redexes are
values. *)
Definition
sub_redexes_are_values
(
e
:
expr
)
:
=
∀
K
e'
,
e
=
fill
K
e'
→
to_val
e'
=
None
→
K
=
empty_ectx
.
∀
K
e'
,
e
=
fill
K
e'
→
to_val
e'
=
None
→
K
=
empty_ectx
.
Inductive
prim_step
(
e1
:
expr
)
(
σ
1
:
state
)
Inductive
prim_step
(
e1
:
expr
)
(
σ
1
:
state
)
...
@@ -103,21 +105,21 @@ Section ectx_language.
...
@@ -103,21 +105,21 @@ Section ectx_language.
Qed
.
Qed
.
Lemma
prim_head_reducible
e
σ
:
Lemma
prim_head_reducible
e
σ
:
reducible
e
σ
→
sub_values
e
→
head_reducible
e
σ
.
reducible
e
σ
→
sub_
redexes_are_
values
e
→
head_reducible
e
σ
.
Proof
.
Proof
.
intros
(
e'
&
σ
'
&
efs
&[
K
e1'
e2'
->
->
Hstep
])
?.
intros
(
e'
&
σ
'
&
efs
&[
K
e1'
e2'
->
->
Hstep
])
?.
assert
(
K
=
empty_ectx
)
as
->
by
eauto
10
using
val_stuck
.
assert
(
K
=
empty_ectx
)
as
->
by
eauto
10
using
val_stuck
.
rewrite
fill_empty
/
head_reducible
;
eauto
.
rewrite
fill_empty
/
head_reducible
;
eauto
.
Qed
.
Qed
.
Lemma
prim_head_irreducible
e
σ
:
Lemma
prim_head_irreducible
e
σ
:
head_irreducible
e
σ
→
sub_values
e
→
irreducible
e
σ
.
head_irreducible
e
σ
→
sub_
redexes_are_
values
e
→
irreducible
e
σ
.
Proof
.
Proof
.
rewrite
-
not_reducible
-
not_head_reducible
.
eauto
using
prim_head_reducible
.
rewrite
-
not_reducible
-
not_head_reducible
.
eauto
using
prim_head_reducible
.
Qed
.
Qed
.
Lemma
ectx_language_atomic
e
:
Lemma
ectx_language_atomic
e
:
(
∀
σ
e'
σ
'
efs
,
head_step
e
σ
e'
σ
'
efs
→
irreducible
e'
σ
'
)
→
(
∀
σ
e'
σ
'
efs
,
head_step
e
σ
e'
σ
'
efs
→
irreducible
e'
σ
'
)
→
sub_values
e
→
sub_
redexes_are_
values
e
→
Atomic
e
.
Atomic
e
.
Proof
.
Proof
.
intros
Hatomic_step
Hatomic_fill
σ
e'
σ
'
efs
[
K
e1'
e2'
->
->
Hstep
].
intros
Hatomic_step
Hatomic_fill
σ
e'
σ
'
efs
[
K
e1'
e2'
->
->
Hstep
].
...
...
theories/program_logic/ectxi_language.v
View file @
ac0d54da
...
@@ -90,8 +90,9 @@ Section ectxi_language.
...
@@ -90,8 +90,9 @@ Section ectxi_language.
fill_not_val
,
fill_app
,
step_by_val
,
foldl_app
.
fill_not_val
,
fill_app
,
step_by_val
,
foldl_app
.
Next
Obligation
.
intros
K1
K2
?%
app_eq_nil
;
tauto
.
Qed
.
Next
Obligation
.
intros
K1
K2
?%
app_eq_nil
;
tauto
.
Qed
.
Lemma
ectxi_language_sub_values
e
:
Lemma
ectxi_language_sub_redexes_are_values
e
:
(
∀
Ki
e'
,
e
=
fill_item
Ki
e'
→
is_Some
(
to_val
e'
))
→
sub_values
e
.
(
∀
Ki
e'
,
e
=
fill_item
Ki
e'
→
is_Some
(
to_val
e'
))
→
sub_redexes_are_values
e
.
Proof
.
Proof
.
intros
Hsub
K
e'
->.
destruct
K
as
[|
Ki
K
_
]
using
@
rev_ind
=>
//=.
intros
Hsub
K
e'
->.
destruct
K
as
[|
Ki
K
_
]
using
@
rev_ind
=>
//=.
intros
[]%
eq_None_not_Some
.
eapply
fill_val
,
Hsub
.
by
rewrite
/=
fill_app
.
intros
[]%
eq_None_not_Some
.
eapply
fill_val
,
Hsub
.
by
rewrite
/=
fill_app
.
...
...
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