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
Rodolphe Lepigre
Iris
Commits
bdfa43b9
Commit
bdfa43b9
authored
Jun 01, 2019
by
Ralf Jung
Browse files
show that there are infinitely many locations
parent
5c07c3be
Changes
1
Show whitespace changes
Inline
Side-by-side
theories/heap_lang/locations.v
View file @
bdfa43b9
...
...
@@ -11,6 +11,20 @@ Instance loc_inhabited : Inhabited loc := populate {|loc_car := 0 |}.
Instance
loc_countable
:
Countable
loc
.
Proof
.
by
apply
(
inj_countable'
loc_car
(
λ
i
,
{|
loc_car
:
=
i
|}))
;
intros
[].
Qed
.
Program
Definition
loc_infinite
:
Infinite
loc
:
=
{|
infinite_fresh
l
:
=
{|
loc_car
:
=
infinite_fresh
(
loc_car
<$>
l
)
|}
|}.
Next
Obligation
.
intros
xs
Hfresh
.
unfold
fresh
in
Hfresh
.
eapply
(
infinite_is_fresh
(
loc_car
<$>
xs
)).
eapply
elem_of_list_fmap
.
eexists
.
split
;
last
done
.
reflexivity
.
Qed
.
Next
Obligation
.
intros
xs
ys
Hperm
.
unfold
fresh
.
f_equal
.
eapply
infinite_fresh_Permutation
.
rewrite
Hperm
//.
Qed
.
Definition
loc_add
(
l
:
loc
)
(
off
:
Z
)
:
loc
:
=
{|
loc_car
:
=
loc_car
l
+
off
|}.
...
...
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