Skip to content
GitLab
Projects
Groups
Snippets
/
Help
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
Menu
Open sidebar
Iris
Iris
Commits
228dff90
Commit
228dff90
authored
Feb 02, 2016
by
Robbert Krebbers
Browse files
Define heap_lang.fill using fold_right.
parent
93765638
Changes
1
Hide whitespace changes
Inline
Side-by-side
barrier/heap_lang.v
View file @
228dff90
...
...
@@ -111,7 +111,7 @@ Notation ectx := (list ectx_item).
Implicit
Types
Ki
:
ectx_item
.
Implicit
Types
K
:
ectx
.
Definition
ectx
_item
_fill
(
Ki
:
ectx_item
)
(
e
:
expr
)
:
expr
:
=
Definition
fill
_item
(
Ki
:
ectx_item
)
(
e
:
expr
)
:
expr
:
=
match
Ki
with
|
AppLCtx
e2
=>
App
e
e2
|
AppRCtx
v1
=>
App
(
of_val
v1
)
e
...
...
@@ -134,12 +134,7 @@ Definition ectx_item_fill (Ki : ectx_item) (e : expr) : expr :=
|
CasMCtx
v0
e2
=>
Cas
(
of_val
v0
)
e
e2
|
CasRCtx
v0
v1
=>
Cas
(
of_val
v0
)
(
of_val
v1
)
e
end
.
Fixpoint
fill
K
e
:
=
(* FIXME RJ: This really is fold_left, but if I use that all automation breaks:
fold_left (fun e Ki => ectx_item_fill Ki e).
Or maybe we even have a combinator somewhere to swap the arguments? *)
match
K
with
[]
=>
e
|
Ki
::
K
=>
ectx_item_fill
Ki
(
fill
K
e
)
end
.
Definition
fill
(
K
:
ectx
)
(
e
:
expr
)
:
expr
:
=
fold_right
fill_item
e
K
.
(** The stepping relation *)
Inductive
head_step
:
expr
->
state
->
expr
->
state
->
option
expr
->
Prop
:
=
...
...
@@ -216,7 +211,7 @@ Qed.
Instance
:
Injective
(=)
(=)
of_val
.
Proof
.
by
intros
??
Hv
;
apply
(
injective
Some
)
;
rewrite
-!
to_of_val
Hv
.
Qed
.
Instance
ectx
_item_
fill_
inj
Ki
:
Injective
(=)
(=)
(
ectx
_item
_fill
Ki
).
Instance
fill
_item_inj
Ki
:
Injective
(=)
(=)
(
fill
_item
Ki
).
Proof
.
destruct
Ki
;
intros
???
;
simplify_equality'
;
auto
with
f_equal
.
Qed
.
Instance
ectx_fill_inj
K
:
Injective
(=)
(=)
(
fill
K
).
...
...
@@ -263,12 +258,12 @@ Proof.
Qed
.
Lemma
head_ctx_step_val
Ki
e
σ
1 e2
σ
2
ef
:
head_step
(
ectx
_item
_fill
Ki
e
)
σ
1 e2
σ
2
ef
→
is_Some
(
to_val
e
).
head_step
(
fill
_item
Ki
e
)
σ
1 e2
σ
2
ef
→
is_Some
(
to_val
e
).
Proof
.
destruct
Ki
;
inversion_clear
1
;
simplify_option_equality
;
eauto
.
Qed
.
Lemma
fill_item_inj
Ki1
Ki2
e1
e2
:
Lemma
fill_item_
no_val_
inj
Ki1
Ki2
e1
e2
:
to_val
e1
=
None
→
to_val
e2
=
None
→
ectx
_item
_fill
Ki1
e1
=
ectx
_item
_fill
Ki2
e2
→
Ki1
=
Ki2
.
fill
_item
Ki1
e1
=
fill
_item
Ki2
e2
→
Ki1
=
Ki2
.
Proof
.
destruct
Ki1
,
Ki2
;
intros
;
try
discriminate
;
simplify_equality'
;
repeat
match
goal
with
...
...
@@ -289,7 +284,7 @@ Proof.
{
destruct
(
proj1
(
eq_None_not_Some
(
to_val
(
fill
K
e1
))))
;
eauto
using
fill_not_val
,
head_ctx_step_val
.
}
cut
(
Ki
=
Ki'
)
;
[
naive_solver
eauto
using
prefix_of_cons
|].
eauto
using
fill_item_inj
,
values_head_stuck
,
fill_not_val
.
eauto
using
fill_item_
no_val_
inj
,
values_head_stuck
,
fill_not_val
.
Qed
.
Lemma
alloc_fresh
e
v
σ
:
...
...
Write
Preview
Supports
Markdown
0%
Try again
or
attach a new 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