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
Tej Chajed
stdpp
Commits
9774ce9c
Commit
9774ce9c
authored
Feb 17, 2016
by
Robbert Krebbers
Browse files
Use scheme - then + then * for bullets.
parent
2e9c3f77
Changes
23
Hide whitespace changes
Inline
Side-by-side
theories/bsets.v
View file @
9774ce9c
...
...
@@ -21,11 +21,11 @@ Instance bset_collection {A} `{∀ x y : A, Decision (x = y)} :
Collection
A
(
bset
A
).
Proof
.
split
;
[
split
|
|].
*
by
intros
x
?.
*
by
intros
x
y
;
rewrite
<-(
bool_decide_spec
(
x
=
y
)).
*
split
.
apply
orb_prop_elim
.
apply
orb_prop_intro
.
*
split
.
apply
andb_prop_elim
.
apply
andb_prop_intro
.
*
intros
X
Y
x
;
unfold
elem_of
,
bset_elem_of
;
simpl
.
-
by
intros
x
?.
-
by
intros
x
y
;
rewrite
<-(
bool_decide_spec
(
x
=
y
)).
-
split
.
apply
orb_prop_elim
.
apply
orb_prop_intro
.
-
split
.
apply
andb_prop_elim
.
apply
andb_prop_intro
.
-
intros
X
Y
x
;
unfold
elem_of
,
bset_elem_of
;
simpl
.
destruct
(
bset_car
X
x
),
(
bset_car
Y
x
)
;
simpl
;
tauto
.
Qed
.
Instance
bset_elem_of_dec
{
A
}
x
(
X
:
bset
A
)
:
Decision
(
x
∈
X
)
:
=
_
.
...
...
theories/co_pset.v
View file @
9774ce9c
...
...
@@ -65,10 +65,10 @@ Lemma coPset_eq t1 t2 :
Proof
.
revert
t2
.
induction
t1
as
[
b1
|
b1
l1
IHl
r1
IHr
]
;
intros
[
b2
|
b2
l2
r2
]
Ht
??
;
simpl
in
*.
*
f_equal
;
apply
(
Ht
1
).
*
by
discriminate
(
coPLeaf_wf
(
coPNode
b2
l2
r2
)
b1
).
*
by
discriminate
(
coPLeaf_wf
(
coPNode
b1
l1
r1
)
b2
).
*
f_equal
;
[
apply
(
Ht
1
)|
|].
-
f_equal
;
apply
(
Ht
1
).
-
by
discriminate
(
coPLeaf_wf
(
coPNode
b2
l2
r2
)
b1
).
-
by
discriminate
(
coPLeaf_wf
(
coPNode
b1
l1
r1
)
b2
).
-
f_equal
;
[
apply
(
Ht
1
)|
|].
+
apply
IHl
;
try
apply
(
λ
x
,
Ht
(
x
~
0
))
;
eauto
.
+
apply
IHr
;
try
apply
(
λ
x
,
Ht
(
x
~
1
))
;
eauto
.
Qed
.
...
...
@@ -163,13 +163,13 @@ Instance coPset_elem_of_dec (p : positive) (X : coPset) : Decision (p ∈ X) :=
Instance
coPset_collection
:
Collection
positive
coPset
.
Proof
.
split
;
[
split
|
|].
*
by
intros
??.
*
intros
p
q
.
apply
elem_to_Pset_singleton
.
*
intros
[
t
]
[
t'
]
p
;
unfold
elem_of
,
coPset_elem_of
,
coPset_union
;
simpl
.
-
by
intros
??.
-
intros
p
q
.
apply
elem_to_Pset_singleton
.
-
intros
[
t
]
[
t'
]
p
;
unfold
elem_of
,
coPset_elem_of
,
coPset_union
;
simpl
.
by
rewrite
elem_to_Pset_union
,
orb_True
.
*
intros
[
t
]
[
t'
]
p
;
unfold
elem_of
,
coPset_elem_of
,
coPset_intersection
;
simpl
.
-
intros
[
t
]
[
t'
]
p
;
unfold
elem_of
,
coPset_elem_of
,
coPset_intersection
;
simpl
.
by
rewrite
elem_to_Pset_intersection
,
andb_True
.
*
intros
[
t
]
[
t'
]
p
;
unfold
elem_of
,
coPset_elem_of
,
coPset_difference
;
simpl
.
-
intros
[
t
]
[
t'
]
p
;
unfold
elem_of
,
coPset_elem_of
,
coPset_difference
;
simpl
.
by
rewrite
elem_to_Pset_intersection
,
elem_to_Pset_opp
,
andb_True
,
negb_True
.
Qed
.
...
...
@@ -192,7 +192,7 @@ Lemma coPset_finite_spec X : set_finite X ↔ coPset_finite (`X).
Proof
.
destruct
X
as
[
t
Ht
].
unfold
set_finite
,
elem_of
at
1
,
coPset_elem_of
;
simpl
;
clear
Ht
;
split
.
*
induction
t
as
[
b
|
b
l
IHl
r
IHr
]
;
simpl
.
-
induction
t
as
[
b
|
b
l
IHl
r
IHr
]
;
simpl
.
{
destruct
b
;
simpl
;
[
intros
[
l
Hl
]|
done
].
by
apply
(
is_fresh
(
of_list
l
:
Pset
)),
elem_of_of_list
,
Hl
.
}
intros
[
ll
Hll
]
;
rewrite
andb_True
;
split
.
...
...
@@ -200,7 +200,7 @@ Proof.
rewrite
elem_of_list_omap
;
intros
;
exists
(
i
~
0
)
;
auto
.
+
apply
IHr
;
exists
(
omap
(
maybe
(~
1
))
ll
)
;
intros
i
.
rewrite
elem_of_list_omap
;
intros
;
exists
(
i
~
1
)
;
auto
.
*
induction
t
as
[
b
|
b
l
IHl
r
IHr
]
;
simpl
;
[
by
exists
[]
;
destruct
b
|].
-
induction
t
as
[
b
|
b
l
IHl
r
IHr
]
;
simpl
;
[
by
exists
[]
;
destruct
b
|].
rewrite
andb_True
;
intros
[??]
;
destruct
IHl
as
[
ll
?],
IHr
as
[
rl
?]
;
auto
.
exists
([
1
]
++
((~
0
)
<$>
ll
)
++
((~
1
)
<$>
rl
))%
list
;
intros
[
i
|
i
|]
;
simpl
;
rewrite
elem_of_cons
,
elem_of_app
,
!
elem_of_list_fmap
;
naive_solver
.
...
...
@@ -237,8 +237,8 @@ Qed.
Lemma
coPpick_elem_of
X
:
¬
set_finite
X
→
coPpick
X
∈
X
.
Proof
.
destruct
X
as
[
t
?]
;
unfold
coPpick
;
destruct
(
coPpick_raw
_
)
as
[
j
|]
eqn
:
?.
*
by
intros
;
apply
coPpick_raw_elem_of
.
*
by
intros
[]
;
apply
coPset_finite_spec
,
coPpick_raw_None
.
-
by
intros
;
apply
coPpick_raw_elem_of
.
-
by
intros
[]
;
apply
coPset_finite_spec
,
coPpick_raw_None
.
Qed
.
(** * Conversion to psets *)
...
...
@@ -270,8 +270,8 @@ Fixpoint of_Pset_raw (t : Pmap_raw ()) : coPset_raw :=
Lemma
of_Pset_wf
t
:
Pmap_wf
t
→
coPset_wf
(
of_Pset_raw
t
).
Proof
.
induction
t
as
[|[]
l
IHl
r
IHr
]
;
simpl
;
rewrite
?andb_True
;
auto
.
*
intros
[??]
;
destruct
l
as
[|[]],
r
as
[|[]]
;
simpl
in
*
;
auto
.
*
destruct
l
as
[|[]],
r
as
[|[]]
;
simpl
in
*
;
rewrite
?andb_true_r
;
-
intros
[??]
;
destruct
l
as
[|[]],
r
as
[|[]]
;
simpl
in
*
;
auto
.
-
destruct
l
as
[|[]],
r
as
[|[]]
;
simpl
in
*
;
rewrite
?andb_true_r
;
rewrite
?andb_True
;
rewrite
?andb_True
in
IHl
,
IHr
;
intuition
.
Qed
.
Lemma
elem_of_of_Pset_raw
i
t
:
e_of
i
(
of_Pset_raw
t
)
↔
t
!!
i
=
Some
().
...
...
@@ -327,9 +327,9 @@ Definition coPset_suffixes (p : positive) : coPset :=
Lemma
elem_coPset_suffixes
p
q
:
p
∈
coPset_suffixes
q
↔
∃
q'
,
p
=
q'
++
q
.
Proof
.
unfold
elem_of
,
coPset_elem_of
;
simpl
;
split
.
*
revert
p
;
induction
q
;
intros
[?|?|]
;
simpl
;
-
revert
p
;
induction
q
;
intros
[?|?|]
;
simpl
;
rewrite
?coPset_elem_of_node
;
naive_solver
.
*
by
intros
[
q'
->]
;
induction
q
;
simpl
;
rewrite
?coPset_elem_of_node
.
-
by
intros
[
q'
->]
;
induction
q
;
simpl
;
rewrite
?coPset_elem_of_node
.
Qed
.
Lemma
coPset_suffixes_infinite
p
:
¬
set_finite
(
coPset_suffixes
p
).
Proof
.
...
...
theories/collections.v
View file @
9774ce9c
...
...
@@ -48,8 +48,8 @@ Section simple_collection.
Lemma
elem_of_subseteq_singleton
x
X
:
x
∈
X
↔
{[
x
]}
⊆
X
.
Proof
.
split
.
*
intros
??.
rewrite
elem_of_singleton
.
by
intros
->.
*
intros
Ex
.
by
apply
(
Ex
x
),
elem_of_singleton
.
-
intros
??.
rewrite
elem_of_singleton
.
by
intros
->.
-
intros
Ex
.
by
apply
(
Ex
x
),
elem_of_singleton
.
Qed
.
Global
Instance
singleton_proper
:
Proper
((=)
==>
(
≡
))
(
singleton
(
B
:
=
C
)).
Proof
.
by
repeat
intro
;
subst
.
Qed
.
...
...
@@ -59,9 +59,9 @@ Section simple_collection.
Lemma
elem_of_union_list
Xs
x
:
x
∈
⋃
Xs
↔
∃
X
,
X
∈
Xs
∧
x
∈
X
.
Proof
.
split
.
*
induction
Xs
;
simpl
;
intros
HXs
;
[
by
apply
elem_of_empty
in
HXs
|].
-
induction
Xs
;
simpl
;
intros
HXs
;
[
by
apply
elem_of_empty
in
HXs
|].
setoid_rewrite
elem_of_cons
.
apply
elem_of_union
in
HXs
.
naive_solver
.
*
intros
[
X
[]].
induction
1
;
simpl
;
[
by
apply
elem_of_union_l
|].
-
intros
[
X
[]].
induction
1
;
simpl
;
[
by
apply
elem_of_union_l
|].
intros
.
apply
elem_of_union_r
;
auto
.
Qed
.
Lemma
non_empty_singleton
x
:
({[
x
]}
:
C
)
≢
∅
.
...
...
@@ -113,9 +113,9 @@ Section of_option_list.
Lemma
elem_of_of_list
(
x
:
A
)
l
:
x
∈
of_list
l
↔
x
∈
l
.
Proof
.
split
.
*
induction
l
;
simpl
;
[
by
rewrite
elem_of_empty
|].
-
induction
l
;
simpl
;
[
by
rewrite
elem_of_empty
|].
rewrite
elem_of_union
,
elem_of_singleton
;
intros
[->|?]
;
constructor
;
auto
.
*
induction
1
;
simpl
;
rewrite
elem_of_union
,
elem_of_singleton
;
auto
.
-
induction
1
;
simpl
;
rewrite
elem_of_union
,
elem_of_singleton
;
auto
.
Qed
.
End
of_option_list
.
...
...
@@ -356,11 +356,11 @@ Section collection_ops.
Forall2
(
∈
)
xs
Xs
∧
y
∈
Y
∧
foldr
(
λ
x
,
(
≫
=
f
x
))
(
Some
y
)
xs
=
Some
x
.
Proof
.
split
.
*
revert
x
.
induction
Xs
;
simpl
;
intros
x
HXs
;
[
eexists
[],
x
;
intuition
|].
-
revert
x
.
induction
Xs
;
simpl
;
intros
x
HXs
;
[
eexists
[],
x
;
intuition
|].
rewrite
elem_of_intersection_with
in
HXs
;
destruct
HXs
as
(
x1
&
x2
&?&?&?).
destruct
(
IHXs
x2
)
as
(
xs
&
y
&
hy
&
?
&
?)
;
trivial
.
eexists
(
x1
::
xs
),
y
.
intuition
(
simplify_option_equality
;
auto
).
*
intros
(
xs
&
y
&
Hxs
&
?
&
Hx
).
revert
x
Hx
.
-
intros
(
xs
&
y
&
Hxs
&
?
&
Hx
).
revert
x
Hx
.
induction
Hxs
;
intros
;
simplify_option_equality
;
[
done
|].
rewrite
elem_of_intersection_with
.
naive_solver
.
Qed
.
...
...
@@ -389,8 +389,8 @@ Section NoDup.
Global
Instance
:
Proper
(
R
==>
(
≡
)
==>
iff
)
elem_of_upto
.
Proof
.
intros
??
E1
??
E2
.
split
;
intros
[
z
[??]]
;
exists
z
.
*
rewrite
<-
E1
,
<-
E2
;
intuition
.
*
rewrite
E1
,
E2
;
intuition
.
-
rewrite
<-
E1
,
<-
E2
;
intuition
.
-
rewrite
E1
,
E2
;
intuition
.
Qed
.
Global
Instance
:
Proper
((
≡
)
==>
iff
)
set_NoDup
.
Proof
.
firstorder
.
Qed
.
...
...
@@ -575,8 +575,8 @@ Section collection_monad.
l
∈
mapM
f
k
↔
Forall2
(
λ
x
y
,
x
∈
f
y
)
l
k
.
Proof
.
split
.
*
revert
l
.
induction
k
;
solve_elem_of
.
*
induction
1
;
solve_elem_of
.
-
revert
l
.
induction
k
;
solve_elem_of
.
-
induction
1
;
solve_elem_of
.
Qed
.
Lemma
collection_mapM_length
{
A
B
}
(
f
:
A
→
M
B
)
l
k
:
l
∈
mapM
f
k
→
length
l
=
length
k
.
...
...
theories/fin_collections.v
View file @
9774ce9c
...
...
@@ -20,9 +20,9 @@ Proof. by exists (elements X); intros; rewrite elem_of_elements. Qed.
Global
Instance
elements_proper
:
Proper
((
≡
)
==>
(
≡
ₚ
))
(
elements
(
C
:
=
C
)).
Proof
.
intros
??
E
.
apply
NoDup_Permutation
.
*
apply
NoDup_elements
.
*
apply
NoDup_elements
.
*
intros
.
by
rewrite
!
elem_of_elements
,
E
.
-
apply
NoDup_elements
.
-
apply
NoDup_elements
.
-
intros
.
by
rewrite
!
elem_of_elements
,
E
.
Qed
.
Global
Instance
collection_size_proper
:
Proper
((
≡
)
==>
(=))
(@
size
C
_
).
Proof
.
intros
??
E
.
apply
Permutation_length
.
by
rewrite
E
.
Qed
.
...
...
@@ -45,9 +45,9 @@ Lemma size_singleton (x : A) : size {[ x ]} = 1.
Proof
.
change
(
length
(
elements
{[
x
]})
=
length
[
x
]).
apply
Permutation_length
,
NoDup_Permutation
.
*
apply
NoDup_elements
.
*
apply
NoDup_singleton
.
*
intros
y
.
-
apply
NoDup_elements
.
-
apply
NoDup_singleton
.
-
intros
y
.
by
rewrite
elem_of_elements
,
elem_of_singleton
,
elem_of_list_singleton
.
Qed
.
Lemma
size_singleton_inv
X
x
y
:
size
X
=
1
→
x
∈
X
→
y
∈
X
→
x
=
y
.
...
...
@@ -59,8 +59,8 @@ Qed.
Lemma
collection_choose_or_empty
X
:
(
∃
x
,
x
∈
X
)
∨
X
≡
∅
.
Proof
.
destruct
(
elements
X
)
as
[|
x
l
]
eqn
:
HX
;
[
right
|
left
].
*
apply
equiv_empty
;
intros
x
.
by
rewrite
<-
elem_of_elements
,
HX
,
elem_of_nil
.
*
exists
x
.
rewrite
<-
elem_of_elements
,
HX
.
by
left
.
-
apply
equiv_empty
;
intros
x
.
by
rewrite
<-
elem_of_elements
,
HX
,
elem_of_nil
.
-
exists
x
.
rewrite
<-
elem_of_elements
,
HX
.
by
left
.
Qed
.
Lemma
collection_choose
X
:
X
≢
∅
→
∃
x
,
x
∈
X
.
Proof
.
intros
.
by
destruct
(
collection_choose_or_empty
X
).
Qed
.
...
...
@@ -75,17 +75,17 @@ Lemma size_1_elem_of X : size X = 1 → ∃ x, X ≡ {[ x ]}.
Proof
.
intros
E
.
destruct
(
size_pos_elem_of
X
)
;
auto
with
lia
.
exists
x
.
apply
elem_of_equiv
.
split
.
*
rewrite
elem_of_singleton
.
eauto
using
size_singleton_inv
.
*
solve_elem_of
.
-
rewrite
elem_of_singleton
.
eauto
using
size_singleton_inv
.
-
solve_elem_of
.
Qed
.
Lemma
size_union
X
Y
:
X
∩
Y
≡
∅
→
size
(
X
∪
Y
)
=
size
X
+
size
Y
.
Proof
.
intros
[
E
_
].
unfold
size
,
collection_size
.
simpl
.
rewrite
<-
app_length
.
apply
Permutation_length
,
NoDup_Permutation
.
*
apply
NoDup_elements
.
*
apply
NoDup_app
;
repeat
split
;
try
apply
NoDup_elements
.
-
apply
NoDup_elements
.
-
apply
NoDup_app
;
repeat
split
;
try
apply
NoDup_elements
.
intros
x
;
rewrite
!
elem_of_elements
;
solve_elem_of
.
*
intros
.
by
rewrite
elem_of_app
,
!
elem_of_elements
,
elem_of_union
.
-
intros
.
by
rewrite
elem_of_app
,
!
elem_of_elements
,
elem_of_union
.
Qed
.
Instance
elem_of_dec_slow
(
x
:
A
)
(
X
:
C
)
:
Decision
(
x
∈
X
)
|
100
.
Proof
.
...
...
@@ -129,9 +129,9 @@ Proof.
intros
?
Hemp
Hadd
.
apply
well_founded_induction
with
(
⊂
).
{
apply
collection_wf
.
}
intros
X
IH
.
destruct
(
collection_choose_or_empty
X
)
as
[[
x
?]|
HX
].
*
rewrite
(
union_difference
{[
x
]}
X
)
by
solve_elem_of
.
-
rewrite
(
union_difference
{[
x
]}
X
)
by
solve_elem_of
.
apply
Hadd
.
solve_elem_of
.
apply
IH
;
solve_elem_of
.
*
by
rewrite
HX
.
-
by
rewrite
HX
.
Qed
.
Lemma
collection_fold_ind
{
B
}
(
P
:
B
→
C
→
Prop
)
(
f
:
A
→
B
→
B
)
(
b
:
B
)
:
Proper
((=)
==>
(
≡
)
==>
iff
)
P
→
...
...
@@ -143,9 +143,9 @@ Proof.
{
intros
help
?.
apply
help
;
[
apply
NoDup_elements
|].
symmetry
.
apply
elem_of_elements
.
}
induction
1
as
[|
x
l
??
IH
]
;
simpl
.
*
intros
X
HX
.
setoid_rewrite
elem_of_nil
in
HX
.
-
intros
X
HX
.
setoid_rewrite
elem_of_nil
in
HX
.
rewrite
equiv_empty
.
done
.
solve_elem_of
.
*
intros
X
HX
.
setoid_rewrite
elem_of_cons
in
HX
.
-
intros
X
HX
.
setoid_rewrite
elem_of_cons
in
HX
.
rewrite
(
union_difference
{[
x
]}
X
)
by
solve_elem_of
.
apply
Hadd
.
solve_elem_of
.
apply
IH
.
solve_elem_of
.
Qed
.
...
...
theories/fin_maps.v
View file @
9774ce9c
...
...
@@ -121,9 +121,9 @@ Section setoid.
Global
Instance
map_equivalence
:
Equivalence
((
≡
)
:
relation
(
M
A
)).
Proof
.
split
.
*
by
intros
m
i
.
*
by
intros
m1
m2
?
i
.
*
by
intros
m1
m2
m3
??
i
;
transitivity
(
m2
!!
i
).
-
by
intros
m
i
.
-
by
intros
m1
m2
?
i
.
-
by
intros
m1
m2
m3
??
i
;
transitivity
(
m2
!!
i
).
Qed
.
Global
Instance
lookup_proper
(
i
:
K
)
:
Proper
((
≡
)
==>
(
≡
))
(
lookup
(
M
:
=
M
A
)
i
).
...
...
@@ -258,9 +258,9 @@ Proof.
{
by
rewrite
lookup_partial_alter_ne
,
!
lookup_partial_alter
,
lookup_partial_alter_ne
.
}
destruct
(
decide
(
jj
=
i
))
as
[->|?].
*
by
rewrite
lookup_partial_alter
,
-
by
rewrite
lookup_partial_alter
,
!
lookup_partial_alter_ne
,
lookup_partial_alter
by
congruence
.
*
by
rewrite
!
lookup_partial_alter_ne
by
congruence
.
-
by
rewrite
!
lookup_partial_alter_ne
by
congruence
.
Qed
.
Lemma
partial_alter_self_alt
{
A
}
(
m
:
M
A
)
i
x
:
x
=
m
!!
i
→
partial_alter
(
λ
_
,
x
)
i
m
=
m
.
...
...
@@ -307,8 +307,8 @@ Lemma lookup_alter_Some {A} (f : A → A) m i j y :
(
i
=
j
∧
∃
x
,
m
!!
j
=
Some
x
∧
y
=
f
x
)
∨
(
i
≠
j
∧
m
!!
j
=
Some
y
).
Proof
.
destruct
(
decide
(
i
=
j
))
as
[->|?].
*
rewrite
lookup_alter
.
naive_solver
(
simplify_option_equality
;
eauto
).
*
rewrite
lookup_alter_ne
by
done
.
naive_solver
.
-
rewrite
lookup_alter
.
naive_solver
(
simplify_option_equality
;
eauto
).
-
rewrite
lookup_alter_ne
by
done
.
naive_solver
.
Qed
.
Lemma
lookup_alter_None
{
A
}
(
f
:
A
→
A
)
m
i
j
:
alter
f
i
m
!!
j
=
None
↔
m
!!
j
=
None
.
...
...
@@ -333,9 +333,9 @@ Lemma lookup_delete_Some {A} (m : M A) i j y :
delete
i
m
!!
j
=
Some
y
↔
i
≠
j
∧
m
!!
j
=
Some
y
.
Proof
.
split
.
*
destruct
(
decide
(
i
=
j
))
as
[->|?]
;
-
destruct
(
decide
(
i
=
j
))
as
[->|?]
;
rewrite
?lookup_delete
,
?lookup_delete_ne
;
intuition
congruence
.
*
intros
[??].
by
rewrite
lookup_delete_ne
.
-
intros
[??].
by
rewrite
lookup_delete_ne
.
Qed
.
Lemma
lookup_delete_is_Some
{
A
}
(
m
:
M
A
)
i
j
:
is_Some
(
delete
i
m
!!
j
)
↔
i
≠
j
∧
is_Some
(
m
!!
j
).
...
...
@@ -412,9 +412,9 @@ Lemma lookup_insert_Some {A} (m : M A) i j x y :
<[
i
:
=
x
]>
m
!!
j
=
Some
y
↔
(
i
=
j
∧
x
=
y
)
∨
(
i
≠
j
∧
m
!!
j
=
Some
y
).
Proof
.
split
.
*
destruct
(
decide
(
i
=
j
))
as
[->|?]
;
-
destruct
(
decide
(
i
=
j
))
as
[->|?]
;
rewrite
?lookup_insert
,
?lookup_insert_ne
;
intuition
congruence
.
*
intros
[[->
->]|[??]]
;
[
apply
lookup_insert
|].
by
rewrite
lookup_insert_ne
.
-
intros
[[->
->]|[??]]
;
[
apply
lookup_insert
|].
by
rewrite
lookup_insert_ne
.
Qed
.
Lemma
lookup_insert_is_Some
{
A
}
(
m
:
M
A
)
i
j
x
:
is_Some
(<[
i
:
=
x
]>
m
!!
j
)
↔
i
=
j
∨
i
≠
j
∧
is_Some
(
m
!!
j
).
...
...
@@ -435,8 +435,8 @@ Lemma insert_included {A} R `{!Reflexive R} (m : M A) i x :
(
∀
y
,
m
!!
i
=
Some
y
→
R
y
x
)
→
map_included
R
m
(<[
i
:
=
x
]>
m
).
Proof
.
intros
?
j
;
destruct
(
decide
(
i
=
j
))
as
[->|].
*
rewrite
lookup_insert
.
destruct
(
m
!!
j
)
;
simpl
;
eauto
.
*
rewrite
lookup_insert_ne
by
done
.
by
destruct
(
m
!!
j
)
;
simpl
.
-
rewrite
lookup_insert
.
destruct
(
m
!!
j
)
;
simpl
;
eauto
.
-
rewrite
lookup_insert_ne
by
done
.
by
destruct
(
m
!!
j
)
;
simpl
.
Qed
.
Lemma
insert_subseteq
{
A
}
(
m
:
M
A
)
i
x
:
m
!!
i
=
None
→
m
⊆
<[
i
:
=
x
]>
m
.
Proof
.
apply
partial_alter_subseteq
.
Qed
.
...
...
@@ -462,8 +462,8 @@ Lemma delete_insert_subseteq {A} (m1 m2 : M A) i x :
Proof
.
rewrite
!
map_subseteq_spec
.
intros
Hix
Hi
j
y
Hj
.
destruct
(
decide
(
i
=
j
))
as
[->|?].
*
rewrite
lookup_insert
.
congruence
.
*
rewrite
lookup_insert_ne
by
done
.
apply
Hi
.
by
rewrite
lookup_delete_ne
.
-
rewrite
lookup_insert
.
congruence
.
-
rewrite
lookup_insert_ne
by
done
.
apply
Hi
.
by
rewrite
lookup_delete_ne
.
Qed
.
Lemma
insert_delete_subset
{
A
}
(
m1
m2
:
M
A
)
i
x
:
m1
!!
i
=
None
→
<[
i
:
=
x
]>
m1
⊂
m2
→
m1
⊂
delete
i
m2
.
...
...
@@ -477,10 +477,10 @@ Lemma insert_subset_inv {A} (m1 m2 : M A) i x :
∃
m2'
,
m2
=
<[
i
:
=
x
]>
m2'
∧
m1
⊂
m2'
∧
m2'
!!
i
=
None
.
Proof
.
intros
Hi
Hm1m2
.
exists
(
delete
i
m2
).
split_ands
.
*
rewrite
insert_delete
.
done
.
eapply
lookup_weaken
,
strict_include
;
eauto
.
-
rewrite
insert_delete
.
done
.
eapply
lookup_weaken
,
strict_include
;
eauto
.
by
rewrite
lookup_insert
.
*
eauto
using
insert_delete_subset
.
*
by
rewrite
lookup_delete
.
-
eauto
using
insert_delete_subset
.
-
by
rewrite
lookup_delete
.
Qed
.
Lemma
insert_empty
{
A
}
i
(
x
:
A
)
:
<[
i
:
=
x
]>
∅
=
{[
i
:
=
x
]}.
Proof
.
done
.
Qed
.
...
...
@@ -510,8 +510,8 @@ Qed.
Lemma
alter_singleton
{
A
}
(
f
:
A
→
A
)
i
x
:
alter
f
i
{[
i
:
=
x
]}
=
{[
i
:
=
f
x
]}.
Proof
.
intros
.
apply
map_eq
.
intros
i'
.
destruct
(
decide
(
i
=
i'
))
as
[->|?].
*
by
rewrite
lookup_alter
,
!
lookup_singleton
.
*
by
rewrite
lookup_alter_ne
,
!
lookup_singleton_ne
.
-
by
rewrite
lookup_alter
,
!
lookup_singleton
.
-
by
rewrite
lookup_alter_ne
,
!
lookup_singleton_ne
.
Qed
.
Lemma
alter_singleton_ne
{
A
}
(
f
:
A
→
A
)
i
j
x
:
i
≠
j
→
alter
f
i
{[
j
:
=
x
]}
=
{[
j
:
=
x
]}.
...
...
@@ -528,15 +528,15 @@ Proof. apply map_empty; intros i. by rewrite lookup_omap, lookup_empty. Qed.
Lemma
fmap_insert
{
A
B
}
(
f
:
A
→
B
)
m
i
x
:
f
<$>
<[
i
:
=
x
]>
m
=
<[
i
:
=
f
x
]>(
f
<$>
m
).
Proof
.
apply
map_eq
;
intros
i'
;
destruct
(
decide
(
i'
=
i
))
as
[->|].
*
by
rewrite
lookup_fmap
,
!
lookup_insert
.
*
by
rewrite
lookup_fmap
,
!
lookup_insert_ne
,
lookup_fmap
by
done
.
-
by
rewrite
lookup_fmap
,
!
lookup_insert
.
-
by
rewrite
lookup_fmap
,
!
lookup_insert_ne
,
lookup_fmap
by
done
.
Qed
.
Lemma
omap_insert
{
A
B
}
(
f
:
A
→
option
B
)
m
i
x
y
:
f
x
=
Some
y
→
omap
f
(<[
i
:
=
x
]>
m
)
=
<[
i
:
=
y
]>(
omap
f
m
).
Proof
.
intros
;
apply
map_eq
;
intros
i'
;
destruct
(
decide
(
i'
=
i
))
as
[->|].
*
by
rewrite
lookup_omap
,
!
lookup_insert
.
*
by
rewrite
lookup_omap
,
!
lookup_insert_ne
,
lookup_omap
by
done
.
-
by
rewrite
lookup_omap
,
!
lookup_insert
.
-
by
rewrite
lookup_omap
,
!
lookup_insert_ne
,
lookup_omap
by
done
.
Qed
.
Lemma
map_fmap_singleton
{
A
B
}
(
f
:
A
→
B
)
i
x
:
f
<$>
{[
i
:
=
x
]}
=
{[
i
:
=
f
x
]}.
Proof
.
...
...
@@ -585,8 +585,8 @@ Proof.
setoid_rewrite
elem_of_cons
.
intros
[?|?]
Hdup
;
simplify_equality
;
[
by
rewrite
lookup_insert
|].
destruct
(
decide
(
i
=
j
))
as
[->|].
*
rewrite
lookup_insert
;
f_equal
;
eauto
.
*
rewrite
lookup_insert_ne
by
done
;
eauto
.
-
rewrite
lookup_insert
;
f_equal
;
eauto
.
-
rewrite
lookup_insert_ne
by
done
;
eauto
.
Qed
.
Lemma
elem_of_map_of_list_1
{
A
}
(
l
:
list
(
K
*
A
))
i
x
:
NoDup
(
l
.*
1
)
→
(
i
,
x
)
∈
l
→
map_of_list
l
!!
i
=
Some
x
.
...
...
@@ -617,8 +617,8 @@ Lemma not_elem_of_map_of_list_2 {A} (l : list (K * A)) i :
Proof
.
induction
l
as
[|[
j
y
]
l
IH
]
;
csimpl
;
[
rewrite
elem_of_nil
;
tauto
|].
rewrite
elem_of_cons
.
destruct
(
decide
(
i
=
j
))
;
simplify_equality
.
*
by
rewrite
lookup_insert
.
*
by
rewrite
lookup_insert_ne
;
intuition
.
-
by
rewrite
lookup_insert
.
-
by
rewrite
lookup_insert_ne
;
intuition
.
Qed
.
Lemma
not_elem_of_map_of_list
{
A
}
(
l
:
list
(
K
*
A
))
i
:
i
∉
l
.*
1
↔
map_of_list
l
!!
i
=
None
.
...
...
@@ -665,11 +665,11 @@ Lemma map_to_list_insert {A} (m : M A) i x :
m
!!
i
=
None
→
map_to_list
(<[
i
:
=
x
]>
m
)
≡
ₚ
(
i
,
x
)
::
map_to_list
m
.
Proof
.
intros
.
apply
map_of_list_inj
;
csimpl
.
*
apply
NoDup_fst_map_to_list
.
*
constructor
;
auto
using
NoDup_fst_map_to_list
.
-
apply
NoDup_fst_map_to_list
.
-
constructor
;
auto
using
NoDup_fst_map_to_list
.
rewrite
elem_of_list_fmap
.
intros
[[??]
[?
Hlookup
]]
;
subst
;
simpl
in
*.
rewrite
elem_of_map_to_list
in
Hlookup
.
congruence
.
*
by
rewrite
!
map_of_to_list
.
-
by
rewrite
!
map_of_to_list
.
Qed
.
Lemma
map_of_list_nil
{
A
}
:
map_of_list
(@
nil
(
K
*
A
))
=
∅
.
Proof
.
done
.
Qed
.
...
...
@@ -702,13 +702,13 @@ Lemma lookup_imap {A B} (f : K → A → option B) m i :
map_imap
f
m
!!
i
=
m
!!
i
≫
=
f
i
.
Proof
.
unfold
map_imap
;
destruct
(
m
!!
i
≫
=
f
i
)
as
[
y
|]
eqn
:
Hi
;
simpl
.
*
destruct
(
m
!!
i
)
as
[
x
|]
eqn
:
?
;
simplify_equality'
.
-
destruct
(
m
!!
i
)
as
[
x
|]
eqn
:
?
;
simplify_equality'
.
apply
elem_of_map_of_list_1_help
.
{
apply
elem_of_list_omap
;
exists
(
i
,
x
)
;
split
;
[
by
apply
elem_of_map_to_list
|
by
simplify_option_equality
].
}
intros
y'
;
rewrite
elem_of_list_omap
;
intros
([
i'
x'
]&
Hi'
&?).
by
rewrite
elem_of_map_to_list
in
Hi'
;
simplify_option_equality
.
*
apply
not_elem_of_map_of_list
;
rewrite
elem_of_list_fmap
.
-
apply
not_elem_of_map_of_list
;
rewrite
elem_of_list_fmap
.
intros
([
i'
x
]&->&
Hi'
)
;
simplify_equality'
.
rewrite
elem_of_list_omap
in
Hi'
;
destruct
Hi'
as
([
j
y
]&
Hj
&?).
rewrite
elem_of_map_to_list
in
Hj
;
simplify_option_equality
.
...
...
@@ -726,8 +726,8 @@ Proof.
by
intros
(?&?&?&?&?)
;
simplify_option_equality
.
}
unfold
map_of_collection
;
rewrite
<-
elem_of_map_of_list
by
done
.
rewrite
elem_of_list_omap
.
setoid_rewrite
elem_of_elements
;
split
.
*
intros
(?&?&?)
;
simplify_option_equality
;
eauto
.
*
intros
[??]
;
exists
i
;
simplify_option_equality
;
eauto
.
-
intros
(?&?&?)
;
simplify_option_equality
;
eauto
.
-
intros
[??]
;
exists
i
;
simplify_option_equality
;
eauto
.
Qed
.
(** ** Induction principles *)
...
...
@@ -741,8 +741,8 @@ Proof.
{
apply
map_to_list_empty_inv_alt
in
Hml
.
by
subst
.
}
inversion_clear
Hnodup
.
apply
map_to_list_insert_inv
in
Hml
;
subst
m
.
apply
Hins
.
*
by
apply
not_elem_of_map_of_list_1
.
*
apply
IH
;
auto
using
map_to_of_list
.
-
by
apply
not_elem_of_map_of_list_1
.
-
apply
IH
;
auto
using
map_to_of_list
.
Qed
.
Lemma
map_to_list_length
{
A
}
(
m1
m2
:
M
A
)
:
m1
⊂
m2
→
length
(
map_to_list
m1
)
<
length
(
map_to_list
m2
).
...
...
@@ -759,8 +759,8 @@ Qed.
Lemma
map_wf
{
A
}
:
wf
(
strict
(@
subseteq
(
M
A
)
_
)).
Proof
.
apply
(
wf_projected
(<)
(
length
∘
map_to_list
)).
*
by
apply
map_to_list_length
.
*
by
apply
lt_wf
.
-
by
apply
map_to_list_length
.
-
by
apply
lt_wf
.
Qed
.
(** ** Properties of the [map_Forall] predicate *)
...
...
@@ -770,8 +770,8 @@ Context {A} (P : K → A → Prop).
Lemma
map_Forall_to_list
m
:
map_Forall
P
m
↔
Forall
(
curry
P
)
(
map_to_list
m
).
Proof
.
rewrite
Forall_forall
.
split
.
*
intros
Hforall
[
i
x
].
rewrite
elem_of_map_to_list
.
by
apply
(
Hforall
i
x
).
*
intros
Hforall
i
x
.
rewrite
<-
elem_of_map_to_list
.
by
apply
(
Hforall
(
i
,
x
)).
-
intros
Hforall
[
i
x
].
rewrite
elem_of_map_to_list
.
by
apply
(
Hforall
i
x
).
-
intros
Hforall
i
x
.
rewrite
<-
elem_of_map_to_list
.
by
apply
(
Hforall
(
i
,
x
)).
Qed
.
Lemma
map_Forall_empty
:
map_Forall
P
∅
.
Proof
.
intros
i
x
.
by
rewrite
lookup_empty
.
Qed
.
...
...
@@ -874,24 +874,24 @@ Lemma partial_alter_merge g g1 g2 m1 m2 i :
merge
f
(
partial_alter
g1
i
m1
)
(
partial_alter
g2
i
m2
).
Proof
.
intro
.
apply
map_eq
.
intros
j
.
destruct
(
decide
(
i
=
j
))
;
subst
.
*
by
rewrite
(
lookup_merge
_
),
!
lookup_partial_alter
,
!(
lookup_merge
_
).
*
by
rewrite
(
lookup_merge
_
),
!
lookup_partial_alter_ne
,
(
lookup_merge
_
).
-
by
rewrite
(
lookup_merge
_
),
!
lookup_partial_alter
,
!(
lookup_merge
_
).
-
by
rewrite
(
lookup_merge
_
),
!
lookup_partial_alter_ne
,
(
lookup_merge
_
).
Qed
.
Lemma
partial_alter_merge_l
g
g1
m1
m2
i
:
g
(
f
(
m1
!!
i
)
(
m2
!!
i
))
=
f
(
g1
(
m1
!!
i
))
(
m2
!!
i
)
→
partial_alter
g
i
(
merge
f
m1
m2
)
=
merge
f
(
partial_alter
g1
i
m1
)
m2
.
Proof
.
intro
.
apply
map_eq
.
intros
j
.
destruct
(
decide
(
i
=
j
))
;
subst
.
*
by
rewrite
(
lookup_merge
_
),
!
lookup_partial_alter
,
!(
lookup_merge
_
).
*
by
rewrite
(
lookup_merge
_
),
!
lookup_partial_alter_ne
,
(
lookup_merge
_
).
-
by
rewrite
(
lookup_merge
_
),
!
lookup_partial_alter
,
!(
lookup_merge
_
).
-
by
rewrite
(
lookup_merge
_
),
!
lookup_partial_alter_ne
,
(
lookup_merge
_
).
Qed
.
Lemma
partial_alter_merge_r
g
g2
m1
m2
i
:
g
(
f
(
m1
!!
i
)
(
m2
!!
i
))
=
f
(
m1
!!
i
)
(
g2
(
m2
!!
i
))
→
partial_alter
g
i
(
merge
f
m1
m2
)
=
merge
f
m1
(
partial_alter
g2
i
m2
).
Proof
.
intro
.
apply
map_eq
.
intros
j
.
destruct
(
decide
(
i
=
j
))
;
subst
.
*
by
rewrite
(
lookup_merge
_
),
!
lookup_partial_alter
,
!(
lookup_merge
_
).
*
by
rewrite
(
lookup_merge
_
),
!
lookup_partial_alter_ne
,
(
lookup_merge
_
).
-
by
rewrite
(
lookup_merge
_
),
!
lookup_partial_alter
,
!(
lookup_merge
_
).
-
by
rewrite
(
lookup_merge
_
),
!
lookup_partial_alter_ne
,
(
lookup_merge
_
).
Qed
.
Lemma
insert_merge
m1
m2
i
x
y
z
:
f
(
Some
y
)
(
Some
z
)
=
Some
x
→
...
...
@@ -928,10 +928,10 @@ Lemma map_relation_alt (m1 : M A) (m2 : M B) :
map_relation
R
P
Q
m1
m2
↔
map_Forall
(
λ
_
,
Is_true
)
(
merge
f
m1
m2
).
Proof
.
split
.
*
intros
Hm
i
P'
;
rewrite
lookup_merge
by
done
;
intros
.
-
intros
Hm
i
P'
;
rewrite
lookup_merge
by
done
;
intros
.
specialize
(
Hm
i
).
destruct
(
m1
!!
i
),
(
m2
!!
i
)
;
simplify_equality'
;
auto
using
bool_decide_pack
.
*
intros
Hm
i
.
specialize
(
Hm
i
).
rewrite
lookup_merge
in
Hm
by
done
.
-
intros
Hm
i
.
specialize
(
Hm
i
).
rewrite
lookup_merge
in
Hm
by
done
.
destruct
(
m1
!!
i
),
(
m2
!!
i
)
;
simplify_equality'
;
auto
;
by
eapply
bool_decide_unpack
,
Hm
.
Qed
.
...
...
@@ -950,10 +950,10 @@ Lemma map_not_Forall2 (m1 : M A) (m2 : M B) :
∨
(
∃
y
,
m1
!!
i
=
None
∧
m2
!!
i
=
Some
y
∧
¬
Q
y
).
Proof
.
split
.
*
rewrite
map_relation_alt
,
(
map_not_Forall
_
).
intros
(
i
&?&
Hm
&?)
;
exists
i
.
-
rewrite
map_relation_alt
,
(
map_not_Forall
_
).
intros
(
i
&?&
Hm
&?)
;
exists
i
.
rewrite
lookup_merge
in
Hm
by
done
.
destruct
(
m1
!!
i
),
(
m2
!!
i
)
;
naive_solver
auto
2
using
bool_decide_pack
.
*
unfold
map_relation
,
option_relation
.
-
unfold
map_relation
,
option_relation
.
by
intros
[
i
[(
x
&
y
&?&?&?)|[(
x
&?&?&?)|(
y
&?&?&?)]]]
Hm
;
specialize
(
Hm
i
)
;
simplify_option_equality
.
Qed
.
...
...
@@ -1003,9 +1003,9 @@ Proof. rewrite (symmetry_iff map_disjoint). apply map_disjoint_Some_l. Qed.
Lemma
map_disjoint_singleton_l
{
A
}
(
m
:
M
A
)
i
x
:
{[
i
:
=
x
]}
⊥
ₘ
m
↔
m
!!
i
=
None
.
<