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
Rodolphe Lepigre
Iris
Commits
3d553078
Commit
3d553078
authored
Feb 13, 2016
by
Ralf Jung
Browse files
prove non-step-indexed fin_map validity lemmas
parent
dd6f44a4
Changes
1
Hide whitespace changes
Inline
Side-by-side
algebra/fin_maps.v
View file @
3d553078
...
...
@@ -192,6 +192,13 @@ Proof.
by
move
=>/(
_
i
)
;
simplify_map_equality
.
Qed
.
Lemma
map_lookup_valid
m
i
x
:
✓
m
→
m
!!
i
≡
Some
x
→
✓
x
.
Proof
.
move
=>
Hm
Hi
n
.
move
:
(
Hm
n
i
).
by
rewrite
Hi
.
Qed
.
Lemma
map_insert_valid
m
i
x
:
✓
x
→
✓
m
→
✓
<[
i
:
=
x
]>
m
.
Proof
.
intros
??
n
j
;
apply
map_insert_validN
;
auto
.
Qed
.
Lemma
map_singleton_valid
i
x
:
✓
({[
i
↦
x
]}
:
gmap
K
A
)
↔
✓
x
.
Proof
.
split
;
intros
?
n
;
eapply
map_singleton_validN
;
eauto
.
Qed
.
Lemma
map_insert_op_None
m1
m2
i
x
:
m2
!!
i
=
None
→
<[
i
:
=
x
]>(
m1
⋅
m2
)
=
<[
i
:
=
x
]>
m1
⋅
m2
.
Proof
.
by
intros
Hi
;
apply
(
insert_merge_l
_
m1
m2
)
;
rewrite
Hi
.
Qed
.
...
...
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