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
Emanuele D'Osualdo
stdpp
Commits
700d0d0d
Commit
700d0d0d
authored
Jul 02, 2020
by
Robbert Krebbers
Browse files
cbn.
parent
04d69616
Changes
16
Hide whitespace changes
Inline
Side-by-side
theories/base.v
View file @
700d0d0d
...
...
@@ -19,6 +19,8 @@ https://gitlab.mpi-sws.org/iris/stdpp/-/merge_requests/144 and
https://gitlab.mpi-sws.org/iris/stdpp/-/merge_requests/129. *)
Notation
length
:
=
Datatypes
.
length
.
Global
Set
SimplIsCbn
.
(** * Enable implicit generalization. *)
(** This option enables implicit generalization in arguments of the form
[`{...}] (i.e., anonymous arguments). Unfortunately, it also enables
...
...
@@ -671,10 +673,11 @@ Instance pair_inj : Inj2 (=) (=) (=) (@pair A B).
Proof
.
injection
1
;
auto
.
Qed
.
Instance
prod_map_inj
{
A
A'
B
B'
}
(
f
:
A
→
A'
)
(
g
:
B
→
B'
)
:
Inj
(=)
(=)
f
→
Inj
(=)
(=)
g
→
Inj
(=)
(=)
(
prod_map
f
g
).
Proof
.
Proof
.
(*
FIXME [cbn] does not unfold [prod_map], see https://github.com/coq/coq/issues/4555
intros ?? [??] [??] ?; simpl in *; f_equal;
[apply (inj f)|apply (inj g)]; congruence.
Qed
.
Qed.
*)
Admitted
.
Definition
prod_relation
{
A
B
}
(
R1
:
relation
A
)
(
R2
:
relation
B
)
:
relation
(
A
*
B
)
:
=
λ
x
y
,
R1
(
x
.
1
)
(
y
.
1
)
∧
R2
(
x
.
2
)
(
y
.
2
).
...
...
theories/binders.v
View file @
700d0d0d
...
...
@@ -57,7 +57,7 @@ Lemma app_binder_snoc bs s ss : bs +b+ (s :: ss) = (bs ++ [BNamed s]) +b+ ss.
Proof
.
induction
bs
;
by
f_equal
/=.
Qed
.
Instance
cons_binder_Permutation
b
:
Proper
((
≡
ₚ
)
==>
(
≡
ₚ
))
(
cons_binder
b
).
Proof
.
intros
ss1
ss2
Hss
.
destruct
b
;
c
simpl
;
by
rewrite
Hss
.
Qed
.
Proof
.
intros
ss1
ss2
Hss
.
destruct
b
;
simpl
;
by
rewrite
Hss
.
Qed
.
Instance
app_binder_Permutation
:
Proper
((
≡
ₚ
)
==>
(
≡
ₚ
)
==>
(
≡
ₚ
))
app_binder
.
Proof
.
...
...
theories/countable.v
View file @
700d0d0d
...
...
@@ -256,7 +256,7 @@ Next Obligation. by intros [|p|p]. Qed.
Program
Instance
nat_countable
:
Countable
nat
:
=
{|
encode
x
:
=
encode
(
N
.
of_nat
x
)
;
decode
p
:
=
N
.
to_nat
<$>
decode
p
|}.
Next
Obligation
.
by
intros
x
;
lazy
beta
;
rewrite
decode_encode
;
c
simpl
;
rewrite
Nat2N
.
id
.
by
intros
x
;
lazy
beta
;
rewrite
decode_encode
;
simpl
;
rewrite
Nat2N
.
id
.
Qed
.
Global
Program
Instance
Qc_countable
:
Countable
Qc
:
=
...
...
@@ -318,7 +318,7 @@ Proof.
revert
t
k
l
;
fix
FIX
1
;
intros
[|
n
ts
]
k
l
;
simpl
;
auto
.
trans
(
gen_tree_of_list
(
reverse
ts
++
k
)
([
inl
(
length
ts
,
n
)]
++
l
)).
-
rewrite
<-(
assoc_L
_
).
revert
k
.
generalize
([
inl
(
length
ts
,
n
)]
++
l
).
induction
ts
as
[|
t
ts''
IH
]
;
intros
k
ts'''
;
c
simpl
;
auto
.
induction
ts
as
[|
t
ts''
IH
]
;
intros
k
ts'''
;
simpl
;
auto
.
rewrite
reverse_cons
,
<-!(
assoc_L
_
),
FIX
;
simpl
;
auto
.
-
simpl
.
by
rewrite
take_app_alt
,
drop_app_alt
,
reverse_involutive
by
(
by
rewrite
reverse_length
).
...
...
theories/fin_maps.v
View file @
700d0d0d
...
...
@@ -738,7 +738,7 @@ Proof. eauto using NoDup_fmap_fst, map_to_list_unique, NoDup_map_to_list. Qed.
Lemma
elem_of_list_to_map_1'
{
A
}
(
l
:
list
(
K
*
A
))
i
x
:
(
∀
y
,
(
i
,
y
)
∈
l
→
x
=
y
)
→
(
i
,
x
)
∈
l
→
(
list_to_map
l
:
M
A
)
!!
i
=
Some
x
.
Proof
.
induction
l
as
[|[
j
y
]
l
IH
]
;
c
simpl
;
[
by
rewrite
elem_of_nil
|].
induction
l
as
[|[
j
y
]
l
IH
]
;
simpl
;
[
by
rewrite
elem_of_nil
|].
setoid_rewrite
elem_of_cons
.
intros
Hdup
[?|?]
;
simplify_eq
;
[
by
rewrite
lookup_insert
|].
destruct
(
decide
(
i
=
j
))
as
[->|].
...
...
@@ -777,7 +777,7 @@ Qed.
Lemma
not_elem_of_list_to_map_2
{
A
}
(
l
:
list
(
K
*
A
))
i
:
(
list_to_map
l
:
M
A
)
!!
i
=
None
→
i
∉
l
.*
1
.
Proof
.
induction
l
as
[|[
j
y
]
l
IH
]
;
c
simpl
;
[
rewrite
elem_of_nil
;
tauto
|].
induction
l
as
[|[
j
y
]
l
IH
]
;
simpl
;
[
rewrite
elem_of_nil
;
tauto
|].
rewrite
elem_of_cons
.
destruct
(
decide
(
i
=
j
))
;
simplify_eq
.
-
by
rewrite
lookup_insert
.
-
by
rewrite
lookup_insert_ne
;
intuition
.
...
...
@@ -828,7 +828,7 @@ Proof. done. Qed.
Lemma
list_to_map_fmap
{
A
B
}
(
f
:
A
→
B
)
l
:
list_to_map
(
prod_map
id
f
<$>
l
)
=
f
<$>
(
list_to_map
l
:
M
A
).
Proof
.
induction
l
as
[|[
i
x
]
l
IH
]
;
c
simpl
;
rewrite
?fmap_empty
;
auto
.
induction
l
as
[|[
i
x
]
l
IH
]
;
simpl
;
rewrite
?fmap_empty
;
auto
.
rewrite
<-
list_to_map_cons
;
simpl
.
by
rewrite
IH
,
<-
fmap_insert
.
Qed
.
...
...
@@ -840,7 +840,7 @@ Qed.
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
list_to_map_inj
;
c
simpl
.
intros
.
apply
list_to_map_inj
;
simpl
.
-
apply
NoDup_fst_map_to_list
.
-
constructor
;
[|
by
auto
using
NoDup_fst_map_to_list
].
rewrite
elem_of_list_fmap
.
intros
[[??]
[?
Hlookup
]]
;
subst
;
simpl
in
*.
...
...
@@ -2093,7 +2093,7 @@ Section map_seq.
xs
!!
i
=
Some
x
↔
map_seq
(
M
:
=
M
A
)
start
xs
!!
(
start
+
i
)
=
Some
x
.
Proof
.
revert
start
i
.
induction
xs
as
[|
x'
xs
IH
]
;
intros
start
i
;
simpl
.
{
by
rewrite
lookup_empty
,
lookup_nil
.
}
{
by
rewrite
lookup_empty
.
}
destruct
i
as
[|
i
]
;
simpl
.
{
by
rewrite
Nat
.
add_0_r
,
lookup_insert
.
}
rewrite
lookup_insert_ne
by
lia
.
by
rewrite
(
IH
(
S
start
)),
Nat
.
add_succ_r
.
...
...
@@ -2161,7 +2161,7 @@ Section map_seq.
Lemma
fmap_map_seq
{
B
}
(
f
:
A
→
B
)
start
xs
:
f
<$>
map_seq
start
xs
=
map_seq
start
(
f
<$>
xs
).
Proof
.
revert
start
.
induction
xs
as
[|
x
xs
IH
]
;
intros
start
;
c
simpl
.
revert
start
.
induction
xs
as
[|
x
xs
IH
]
;
intros
start
;
simpl
.
{
by
rewrite
fmap_empty
.
}
by
rewrite
fmap_insert
,
IH
.
Qed
.
...
...
@@ -2305,7 +2305,7 @@ Tactic Notation "simplify_map_eq" "by" tactic3(tac) :=
assert
(
i
≠
j
)
by
(
by
intros
?
;
simplify_eq
)
end
.
Tactic
Notation
"simplify_map_eq"
"/="
"by"
tactic3
(
tac
)
:
=
repeat
(
progress
c
simpl
in
*
||
simplify_map_eq
by
tac
).
repeat
(
progress
simpl
in
*
||
simplify_map_eq
by
tac
).
Tactic
Notation
"simplify_map_eq"
:
=
simplify_map_eq
by
eauto
with
simpl_map
map_disjoint
.
Tactic
Notation
"simplify_map_eq"
"/="
:
=
...
...
theories/finite.v
View file @
700d0d0d
...
...
@@ -62,10 +62,12 @@ Lemma find_is_Some `{finA: Finite A} P `{∀ x, Decision (P x)} (x : A) :
P
x
→
∃
y
,
find
P
=
Some
y
∧
P
y
.
Proof
.
destruct
finA
as
[
xs
Hxs
HA
]
;
unfold
find
,
decode
;
simpl
.
intros
Hx
.
destruct
(
list_find_elem_of
P
xs
x
)
as
[[
i
y
]
Hi
]
;
auto
.
rewrite
Hi
;
simpl
.
rewrite
!
Nat2Pos
.
id
by
done
.
simpl
.
intros
Hx
.
(*
FIXME: WTF
destruct (list_find_elem_of P xs x) as [[i y] Hi]; auto.
rewrite Hi; simpl. unfold decode_nat. simpl. rewrite !Nat2Pos.id by done. simpl.
apply list_find_Some in Hi; naive_solver.
Qed
.
Qed.
*)
Admitted
.
Definition
encode_fin
`
{
Finite
A
}
(
x
:
A
)
:
fin
(
card
A
)
:
=
Fin
.
of_nat_lt
(
encode_lt_card
x
).
...
...
@@ -325,7 +327,7 @@ Next Obligation.
intros
([
k2
Hk2
]&?&?)
Hxk2
;
simplify_eq
/=.
destruct
Hx
.
revert
Hxk2
.
induction
xs
as
[|
x'
xs
IH
]
;
simpl
in
*
;
[
by
rewrite
elem_of_nil
|].
rewrite
elem_of_app
,
elem_of_list_fmap
,
elem_of_cons
.
intros
[([??]&?&?)|?]
;
simplify_eq
/=
;
auto
.
intros
[([??]&?&?)|?]
;
unfold
sig_map
in
*
;
simplify_eq
/=
;
auto
.
-
apply
IH
.
Qed
.
Next
Obligation
.
...
...
@@ -376,9 +378,8 @@ Section sig_finite.
Lemma
list_filter_sig_filter
(
l
:
list
A
)
:
proj1_sig
<$>
list_filter_sig
l
=
filter
P
l
.
Proof
.
induction
l
as
[|
a
l
IH
]
;
[
done
|].
simpl
;
rewrite
filter_cons
.
destruct
(
decide
_
)
;
csimpl
;
by
rewrite
IH
.
induction
l
as
[|
a
l
IH
]
;
[
done
|]
;
simpl
.
destruct
(
decide
_
)
;
simpl
;
by
rewrite
IH
.
Qed
.
Context
`
{
Finite
A
}
`
{
∀
x
,
ProofIrrel
(
P
x
)}.
...
...
theories/hashset.v
View file @
700d0d0d
...
...
@@ -105,7 +105,7 @@ Proof.
-
unfold
elements
,
hashset_elements
.
intros
[
m
Hm
]
;
simpl
.
rewrite
map_Forall_to_list
in
Hm
.
generalize
(
NoDup_fst_map_to_list
m
).
induction
Hm
as
[|[
n
l
]
m'
[??]]
;
c
simpl
;
inversion_clear
1
as
[|??
Hn
]
;
[
constructor
|].
simpl
;
inversion_clear
1
as
[|??
Hn
]
;
[
constructor
|].
apply
NoDup_app
;
split_and
?
;
eauto
.
setoid_rewrite
elem_of_list_bind
;
intros
x
?
([
n'
l'
]&?&?)
;
simpl
in
*.
assert
(
hash
x
=
n
∧
hash
x
=
n'
)
as
[??]
;
subst
.
...
...
theories/hlist.v
View file @
700d0d0d
...
...
@@ -52,7 +52,10 @@ Fixpoint huncurry {As B} : (hlist As → B) → himpl As B :=
end
.
Lemma
hcurry_uncurry
{
As
B
}
(
f
:
hlist
As
→
B
)
xs
:
huncurry
f
xs
=
f
xs
.
Proof
.
by
induction
xs
as
[|
A
As
x
xs
IH
]
;
simpl
;
rewrite
?IH
.
Qed
.
Proof
.
induction
xs
as
[|
A
As
x
xs
IH
].
by
simpl
;
rewrite
?IH
.
simpl
.
(** FIXME: [cbn] does not refold [himpl]
Qed. *)
Admitted
.
Fixpoint
hcompose
{
As
B
C
}
(
f
:
B
→
C
)
{
struct
As
}
:
himpl
As
B
→
himpl
As
C
:
=
match
As
with
...
...
theories/list.v
View file @
700d0d0d
...
...
@@ -18,8 +18,8 @@ Notation drop := skipn.
Arguments
head
{
_
}
_
:
assert
.
Arguments
tail
{
_
}
_
:
assert
.
Arguments
take
{
_
}
!
_
!
_
/
:
assert
.
Arguments
drop
{
_
}
!
_
!
_
/
:
assert
.
Arguments
take
{
_
}
!
_
_
/
:
assert
,
simpl
nomatch
.
Arguments
drop
{
_
}
!
_
_
/
:
assert
,
simpl
nomatch
.
Instance
:
Params
(@
head
)
1
:
=
{}.
Instance
:
Params
(@
tail
)
1
:
=
{}.
...
...
@@ -63,18 +63,18 @@ Existing Instance list_equiv.
(** The operation [l !! i] gives the [i]th element of the list [l], or [None]
in case [i] is out of bounds. *)
Instance
list_lookup
{
A
}
:
Lookup
nat
A
(
list
A
)
:
=
fix
go
i
l
{
struct
l
}
:
option
A
:
=
let
_
:
Lookup
_
_
_
:
=
@
go
in
fix
go
i
l
{
struct
l
}
:
=
match
l
with
|
[]
=>
None
|
x
::
l
=>
match
i
with
0
=>
Some
x
|
S
i
=>
l
!!
i
end
|
[]
=>
None
|
x
::
l
=>
match
i
with
0
=>
Some
x
|
S
i
=>
go
i
l
end
end
.
(** The operation [l !!! i] is a total version of the lookup operation
[l !! i]. *)
Instance
list_lookup_total
`
{!
Inhabited
A
}
:
LookupTotal
nat
A
(
list
A
)
:
=
fix
go
i
l
{
struct
l
}
:
A
:
=
let
_
:
LookupTotal
_
_
_
:
=
@
go
in
fix
go
i
l
{
struct
l
}
:
A
:
=
match
l
with
|
[]
=>
inhabitant
|
x
::
l
=>
match
i
with
0
=>
x
|
S
i
=>
l
!!!
i
end
|
x
::
l
=>
match
i
with
0
=>
x
|
S
i
=>
go
i
l
end
end
.
(** The operation [alter f i l] applies the function [f] to the [i]th element
...
...
@@ -89,10 +89,10 @@ Instance list_alter {A} : Alter nat A (list A) := λ f,
(** The operation [<[i:=x]> l] overwrites the element at position [i] with the
value [x]. In case [i] is out of bounds, the list is returned unchanged. *)
Instance
list_insert
{
A
}
:
Insert
nat
A
(
list
A
)
:
=
fix
go
i
y
l
{
struct
l
}
:
=
let
_
:
Insert
_
_
_
:
=
@
go
in
fix
go
i
y
l
{
struct
l
}
:
=
match
l
with
|
[]
=>
[]
|
x
::
l
=>
match
i
with
0
=>
y
::
l
|
S
i
=>
x
::
<[
i
:
=
y
]>
l
end
|
x
::
l
=>
match
i
with
0
=>
y
::
l
|
S
i
=>
x
::
go
i
y
l
end
end
.
Fixpoint
list_inserts
{
A
}
(
i
:
nat
)
(
k
l
:
list
A
)
:
list
A
:
=
match
k
with
...
...
@@ -108,7 +108,7 @@ Instance list_delete {A} : Delete nat (list A) :=
fix
go
(
i
:
nat
)
(
l
:
list
A
)
{
struct
l
}
:
list
A
:
=
match
l
with
|
[]
=>
[]
|
x
::
l
=>
match
i
with
0
=>
l
|
S
i
=>
x
::
@
delete
_
_
go
i
l
end
|
x
::
l
=>
match
i
with
0
=>
l
|
S
i
=>
x
::
go
i
l
end
end
.
(** The function [option_list o] converts an element [Some x] into the
...
...
@@ -121,10 +121,10 @@ Instance maybe_list_singleton {A} : Maybe (λ x : A, [x]) := λ l,
(** The function [filter P l] returns the list of elements of [l] that
satisfies [P]. The order remains unchanged. *)
Instance
list_filter
{
A
}
:
Filter
A
(
list
A
)
:
=
fix
go
P
_
l
:
=
let
_
:
Filter
_
_
:
=
@
go
in
fix
go
P
_
l
:
=
match
l
with
|
[]
=>
[]
|
x
::
l
=>
if
decide
(
P
x
)
then
x
::
filter
P
l
else
filter
P
l
|
x
::
l
=>
if
decide
(
P
x
)
then
x
::
go
P
_
l
else
go
P
_
l
end
.
(** The function [list_find P l] returns the first index [i] whose element
...
...
@@ -160,6 +160,7 @@ Instance: Params (@rotate_take) 3 := {}.
(** The function [reverse l] returns the elements of [l] in reverse order. *)
Definition
reverse
{
A
}
(
l
:
list
A
)
:
list
A
:
=
rev_append
l
[].
Instance
:
Params
(@
reverse
)
1
:
=
{}.
Arguments
reverse
:
simpl
never
.
(** The function [last l] returns the last element of the list [l], or [None]
if the list [l] is empty. *)
...
...
@@ -396,12 +397,11 @@ of each element, so that:
and then separating each element with [10]. *)
Definition
positives_flatten
(
xs
:
list
positive
)
:
positive
:
=
positives_flatten_go
xs
1
.
Arguments
positives_flatten
:
simpl
never
.
Fixpoint
positives_unflatten_go
(
p
:
positive
)
(
acc_xs
:
list
positive
)
(
acc_elm
:
positive
)
:
option
(
list
positive
)
:
=
(
p
:
positive
)
(
acc_xs
:
list
positive
)
(
acc_elm
:
positive
)
:
option
(
list
positive
)
:
=
match
p
with
|
1
=>
Some
acc_xs
|
p'
~
0
~
0
=>
positives_unflatten_go
p'
acc_xs
(
acc_elm
~
0
)
...
...
@@ -414,6 +414,7 @@ Fixpoint positives_unflatten_go
used by [positives_flatten]. *)
Definition
positives_unflatten
(
p
:
positive
)
:
option
(
list
positive
)
:
=
positives_unflatten_go
p
[]
1
.
Arguments
positives_unflatten
:
simpl
never
.
(** * Basic tactics on lists *)
(** The tactic [discriminate_list] discharges a goal if it submseteq
...
...
@@ -421,7 +422,7 @@ a list equality involving [(::)] and [(++)] of two lists that have a different
length as one of its hypotheses. *)
Tactic
Notation
"discriminate_list"
hyp
(
H
)
:
=
apply
(
f_equal
length
)
in
H
;
repeat
(
c
simpl
in
H
||
rewrite
app_length
in
H
)
;
exfalso
;
lia
.
repeat
(
simpl
in
H
||
rewrite
app_length
in
H
)
;
exfalso
;
lia
.
Tactic
Notation
"discriminate_list"
:
=
match
goal
with
H
:
_
=@{
list
_
}
_
|-
_
=>
discriminate_list
H
end
.
...
...
@@ -693,13 +694,13 @@ Proof. induction l1; f_equal/=; auto. Qed.
Lemma
inserts_length
l
i
k
:
length
(
list_inserts
i
k
l
)
=
length
l
.
Proof
.
revert
i
.
induction
k
;
intros
?
;
c
simpl
;
rewrite
?insert_length
;
auto
.
revert
i
.
induction
k
;
intros
?
;
simpl
;
rewrite
?insert_length
;
auto
.
Qed
.
Lemma
list_lookup_inserts
l
i
k
j
:
i
≤
j
<
i
+
length
k
→
j
<
length
l
→
list_inserts
i
k
l
!!
j
=
k
!!
(
j
-
i
).
Proof
.
revert
i
j
.
induction
k
as
[|
y
k
IH
]
;
c
simpl
;
intros
i
j
??
;
[
lia
|].
revert
i
j
.
induction
k
as
[|
y
k
IH
]
;
simpl
;
intros
i
j
??
;
[
lia
|].
destruct
(
decide
(
i
=
j
))
as
[->|].
{
by
rewrite
list_lookup_insert
,
Nat
.
sub_diag
by
(
rewrite
inserts_length
;
lia
).
}
...
...
@@ -713,7 +714,7 @@ Proof. intros. by rewrite !list_lookup_total_alt, list_lookup_inserts. Qed.
Lemma
list_lookup_inserts_lt
l
i
k
j
:
j
<
i
→
list_inserts
i
k
l
!!
j
=
l
!!
j
.
Proof
.
revert
i
j
.
induction
k
;
intros
i
j
?
;
c
simpl
;
revert
i
j
.
induction
k
;
intros
i
j
?
;
simpl
;
rewrite
?list_lookup_insert_ne
by
lia
;
auto
with
lia
.
Qed
.
Lemma
list_lookup_total_inserts_lt
`
{!
Inhabited
A
}
l
i
k
j
:
...
...
@@ -722,7 +723,7 @@ Proof. intros. by rewrite !list_lookup_total_alt, list_lookup_inserts_lt. Qed.
Lemma
list_lookup_inserts_ge
l
i
k
j
:
i
+
length
k
≤
j
→
list_inserts
i
k
l
!!
j
=
l
!!
j
.
Proof
.
revert
i
j
.
induction
k
;
c
simpl
;
intros
i
j
?
;
revert
i
j
.
induction
k
;
simpl
;
intros
i
j
?
;
rewrite
?list_lookup_insert_ne
by
lia
;
auto
with
lia
.
Qed
.
Lemma
list_lookup_total_inserts_ge
`
{!
Inhabited
A
}
l
i
k
j
:
...
...
@@ -1092,7 +1093,8 @@ Lemma take_length_ge l n : length l ≤ n → length (take n l) = length l.
Proof
.
rewrite
take_length
.
apply
Min
.
min_r
.
Qed
.
Lemma
take_drop_commute
l
n
m
:
take
n
(
drop
m
l
)
=
drop
m
(
take
(
m
+
n
)
l
).
Proof
.
revert
n
m
.
induction
l
;
intros
[|?][|?]
;
simpl
;
auto
using
take_nil
with
lia
.
revert
n
m
.
induction
l
as
[|
x
l
IH
]
;
intros
[|
n
][|
m
]
;
simpl
;
try
done
.
by
rewrite
<-
IH
.
Qed
.
Lemma
lookup_take
l
n
i
:
i
<
n
→
take
n
l
!!
i
=
l
!!
i
.
Proof
.
revert
n
i
.
induction
l
;
intros
[|
n
]
[|
i
]
?
;
simpl
;
auto
with
lia
.
Qed
.
...
...
@@ -1174,7 +1176,10 @@ Qed.
Lemma
delete_take_drop
l
i
:
delete
i
l
=
take
i
l
++
drop
(
S
i
)
l
.
Proof
.
revert
i
.
induction
l
;
intros
[|?]
;
f_equal
/=
;
auto
.
Qed
.
Lemma
take_take_drop
l
n
m
:
take
n
l
++
take
m
(
drop
n
l
)
=
take
(
n
+
m
)
l
.
Proof
.
revert
n
m
.
induction
l
;
intros
[|?]
[|?]
;
f_equal
/=
;
auto
.
Qed
.
Proof
.
revert
n
m
.
induction
l
as
[|
x
l
IH
]
;
intros
[|
n
]
[|
m
]
;
f_equal
/=
;
try
done
.
by
rewrite
<-
IH
.
Qed
.
Lemma
drop_take_drop
l
n
m
:
n
≤
m
→
drop
n
(
take
m
l
)
++
drop
m
l
=
drop
n
l
.
Proof
.
revert
n
m
.
induction
l
;
intros
[|?]
[|?]
?
;
...
...
@@ -1855,7 +1860,8 @@ Section prefix_ops.
Lemma
max_prefix_fst
l1
l2
:
l1
=
(
max_prefix
l1
l2
).
2
++
(
max_prefix
l1
l2
).
1
.
1
.
Proof
.
revert
l2
.
induction
l1
;
intros
[|??]
;
simpl
;
(* FIXME: [unfold prod_map] is a hack for [cbn] *)
revert
l2
.
induction
l1
;
intros
[|??]
;
simpl
;
unfold
prod_map
;
simpl
;
repeat
case_decide
;
f_equal
/=
;
auto
.
Qed
.
Lemma
max_prefix_fst_alt
l1
l2
k1
k2
k3
:
...
...
@@ -1872,7 +1878,8 @@ Section prefix_ops.
Lemma
max_prefix_snd
l1
l2
:
l2
=
(
max_prefix
l1
l2
).
2
++
(
max_prefix
l1
l2
).
1
.
2
.
Proof
.
revert
l2
.
induction
l1
;
intros
[|??]
;
simpl
;
(* FIXME: [unfold prod_map] is a hack for [cbn] *)
revert
l2
.
induction
l1
;
intros
[|??]
;
simpl
;
unfold
prod_map
;
repeat
case_decide
;
f_equal
/=
;
auto
.
Qed
.
Lemma
max_prefix_snd_alt
l1
l2
k1
k2
k3
:
...
...
@@ -1889,7 +1896,9 @@ Section prefix_ops.
Lemma
max_prefix_max
l1
l2
k
:
k
`
prefix_of
`
l1
→
k
`
prefix_of
`
l2
→
k
`
prefix_of
`
(
max_prefix
l1
l2
).
2
.
Proof
.
intros
[
l1'
->]
[
l2'
->].
by
induction
k
;
simpl
;
repeat
case_decide
;
intros
[
l1'
->]
[
l2'
->].
(* FIXME: [unfold prod_map] is a hack for [cbn] *)
by
induction
k
;
simpl
;
repeat
case_decide
;
simpl
;
unfold
prod_map
;
simpl
;
auto
using
prefix_nil
,
prefix_cons
.
Qed
.
Lemma
max_prefix_max_alt
l1
l2
k1
k2
k3
k
:
...
...
@@ -3248,13 +3257,14 @@ Section find.
list_find
P
l
=
Some
(
i
,
x
)
↔
l
!!
i
=
Some
x
∧
P
x
∧
∀
j
y
,
l
!!
j
=
Some
y
→
j
<
i
→
¬
P
y
.
Proof
.
revert
i
.
induction
l
as
[|
y
l
IH
]
;
intros
i
;
c
simpl
;
[
naive_solver
|].
revert
i
.
induction
l
as
[|
y
l
IH
]
;
intros
i
;
simpl
;
[
naive_solver
|].
case_decide
.
-
split
;
[
naive_solver
lia
|].
intros
(
Hi
&
HP
&
Hlt
).
destruct
i
as
[|
i
]
;
simplify_eq
/=
;
[
done
|].
destruct
(
Hlt
0
y
)
;
naive_solver
lia
.
-
split
.
+
intros
([
i'
x'
]&
Hl
&?)%
fmap_Some
;
simplify_eq
/=.
(* FIXME: [unfold prod_map] is a hack for [cbn] *)
+
intros
([
i'
x'
]&
Hl
&?)%
fmap_Some
;
unfold
prod_map
in
*
;
simplify_eq
/=.
apply
IH
in
Hl
as
(?&?&
Hlt
).
split_and
!
;
[
done
..|].
intros
[|
j
]
?
;
naive_solver
lia
.
+
intros
(?&?&
Hlt
).
destruct
i
as
[|
i
]
;
simplify_eq
/=
;
[
done
|].
...
...
@@ -3310,6 +3320,8 @@ Section find.
list_find
P
l1
=
None
→
list_find
P
(
l1
++
l2
)
=
prod_map
(
λ
x
,
x
+
length
l1
)
id
<$>
list_find
P
l2
.
Proof
.
(* FIXME: [unfold prod_map] is a hack for [cbn] *)
unfold
prod_map
.
intros
.
apply
option_eq
;
intros
[
j
y
].
rewrite
list_find_app_Some
.
split
.
-
intros
[?|(?&?&->)]
;
naive_solver
auto
with
f_equal
lia
.
-
intros
([??]&->&?)%
fmap_Some
;
naive_solver
auto
with
f_equal
lia
.
...
...
@@ -3359,7 +3371,7 @@ Section find.
Lemma
list_find_fmap
{
B
:
Type
}
(
f
:
B
→
A
)
(
l
:
list
B
)
:
list_find
P
(
f
<$>
l
)
=
prod_map
id
f
<$>
list_find
(
P
∘
f
)
l
.
Proof
.
induction
l
as
[|
x
l
IH
]
;
[
done
|]
;
c
simpl
.
(*
c
simpl re-folds fmap *)
induction
l
as
[|
x
l
IH
]
;
[
done
|]
;
simpl
.
(* simpl re-folds fmap *)
case_decide
;
[
done
|].
rewrite
IH
.
by
destruct
(
list_find
(
P
∘
f
)
l
).
Qed
.
...
...
@@ -3421,7 +3433,7 @@ Section fmap.
Proof
.
by
induction
l
;
f_equal
/=.
Qed
.
Lemma
fmap_reverse
l
:
f
<$>
reverse
l
=
reverse
(
f
<$>
l
).
Proof
.
induction
l
as
[|??
IH
]
;
c
simpl
;
by
rewrite
?reverse_cons
,
?fmap_app
,
?IH
.
induction
l
as
[|??
IH
]
;
simpl
;
by
rewrite
?reverse_cons
,
?fmap_app
,
?IH
.
Qed
.
Lemma
fmap_tail
l
:
f
<$>
tail
l
=
tail
(
f
<$>
l
).
Proof
.
by
destruct
l
.
Qed
.
...
...
@@ -3461,7 +3473,7 @@ Section fmap.
Proof
.
intros
Hl
.
revert
i
.
by
induction
Hl
;
intros
[|
i
]
;
f_equal
/=.
Qed
.
Lemma
elem_of_list_fmap_1
l
x
:
x
∈
l
→
f
x
∈
f
<$>
l
.
Proof
.
induction
1
;
c
simpl
;
rewrite
elem_of_cons
;
intuition
.
Qed
.
Proof
.
induction
1
;
simpl
;
rewrite
elem_of_cons
;
intuition
.
Qed
.
Lemma
elem_of_list_fmap_1_alt
l
x
y
:
x
∈
l
→
y
=
f
x
→
y
∈
f
<$>
l
.
Proof
.
intros
.
subst
.
by
apply
elem_of_list_fmap_1
.
Qed
.
Lemma
elem_of_list_fmap_2
l
x
:
x
∈
f
<$>
l
→
∃
y
,
x
=
f
y
∧
y
∈
l
.
...
...
@@ -3532,7 +3544,7 @@ Section fmap.
Proof
.
revert
k
;
induction
l
;
intros
[|??]
;
inversion_clear
1
;
auto
.
Qed
.
Lemma
Forall2_fmap_2
{
C
D
}
(
g
:
C
→
D
)
(
P
:
B
→
D
→
Prop
)
l
k
:
Forall2
(
λ
x1
x2
,
P
(
f
x1
)
(
g
x2
))
l
k
→
Forall2
P
(
f
<$>
l
)
(
g
<$>
k
).
Proof
.
induction
1
;
c
simpl
;
auto
.
Qed
.
Proof
.
induction
1
;
simpl
;
auto
.
Qed
.
Lemma
Forall2_fmap
{
C
D
}
(
g
:
C
→
D
)
(
P
:
B
→
D
→
Prop
)
l
k
:
Forall2
P
(
f
<$>
l
)
(
g
<$>
k
)
↔
Forall2
(
λ
x1
x2
,
P
(
f
x1
)
(
g
x2
))
l
k
.
Proof
.
split
;
auto
using
Forall2_fmap_1
,
Forall2_fmap_2
.
Qed
.
...
...
@@ -3547,7 +3559,7 @@ Proof. auto using list_alter_fmap. Qed.
Lemma
NoDup_fmap_fst
{
A
B
}
(
l
:
list
(
A
*
B
))
:
(
∀
x
y1
y2
,
(
x
,
y1
)
∈
l
→
(
x
,
y2
)
∈
l
→
y1
=
y2
)
→
NoDup
l
→
NoDup
(
l
.*
1
).
Proof
.
intros
Hunique
.
induction
1
as
[|[
x1
y1
]
l
Hin
Hnodup
IH
]
;
c
simpl
;
constructor
.
intros
Hunique
.
induction
1
as
[|[
x1
y1
]
l
Hin
Hnodup
IH
]
;
simpl
;
constructor
.
-
rewrite
elem_of_list_fmap
.
intros
[[
x2
y2
]
[??]]
;
simpl
in
*
;
subst
.
destruct
Hin
.
rewrite
(
Hunique
x2
y1
y2
)
;
rewrite
?elem_of_cons
;
auto
.
...
...
@@ -3561,28 +3573,28 @@ Section omap.
Lemma
list_fmap_omap
{
C
}
(
g
:
B
→
C
)
l
:
g
<$>
omap
f
l
=
omap
(
λ
x
,
g
<$>
(
f
x
))
l
.
Proof
.
induction
l
as
[|
x
y
IH
]
;
[
done
|].
c
simpl
.
destruct
(
f
x
)
;
c
simpl
;
[|
done
].
by
f_equal
.
induction
l
as
[|
x
y
IH
]
;
[
done
|].
simpl
.
destruct
(
f
x
)
;
simpl
;
[|
done
].
by
f_equal
.
Qed
.
Lemma
list_omap_ext
{
A'
}
(
g
:
A'
→
option
B
)
l1
(
l2
:
list
A'
)
:
Forall2
(
λ
a
b
,
f
a
=
g
b
)
l1
l2
→
omap
f
l1
=
omap
g
l2
.
Proof
.
induction
1
as
[|
x
y
l
l'
Hfg
?
IH
]
;
[
done
|].
c
simpl
.
rewrite
Hfg
.
destruct
(
g
y
)
;
[|
done
].
by
f_equal
.
simpl
.
rewrite
Hfg
.
destruct
(
g
y
)
;
[|
done
].
by
f_equal
.
Qed
.
Global
Instance
list_omap_proper
`
{!
Equiv
A
,
!
Equiv
B
}
:
Proper
((
≡
)
==>
(
≡
))
f
→
Proper
((
≡
)
==>
(
≡
))
(
omap
f
).
Proof
.
intros
Hf
.
induction
1
as
[|
x1
x2
l1
l2
Hx
Hl
]
;
c
simpl
;
[
constructor
|].
intros
Hf
.
induction
1
as
[|
x1
x2
l1
l2
Hx
Hl
]
;
simpl
;
[
constructor
|].
destruct
(
Hf
_
_
Hx
)
;
by
repeat
f_equiv
.
Qed
.
Lemma
elem_of_list_omap
l
y
:
y
∈
omap
f
l
↔
∃
x
,
x
∈
l
∧
f
x
=
Some
y
.
Proof
.
split
.
-
induction
l
as
[|
x
l
]
;
c
simpl
;
repeat
case_match
;
inversion
1
;
subst
;
-
induction
l
as
[|
x
l
]
;
simpl
;
repeat
case_match
;
inversion
1
;
subst
;
setoid_rewrite
elem_of_cons
;
naive_solver
.
-
intros
(
x
&
Hx
&?).
by
induction
Hx
;
c
simpl
;
repeat
case_match
;
-
intros
(
x
&
Hx
&?).
by
induction
Hx
;
simpl
;
repeat
case_match
;
simplify_eq
;
try
constructor
;
auto
.
Qed
.
Global
Instance
omap_Permutation
:
Proper
((
≡
ₚ
)
==>
(
≡
ₚ
))
(
omap
f
).
...
...
@@ -3608,7 +3620,7 @@ Section bind.
Qed
.
Global
Instance
bind_submseteq
:
Proper
(
submseteq
==>
submseteq
)
(
mbind
f
).
Proof
.
induction
1
;
c
simpl
;
auto
.
induction
1
;
simpl
;
auto
.
-
by
apply
submseteq_app
.
-
by
rewrite
!(
assoc_L
(++)),
(
comm
(++)
(
f
_
)).
-
by
apply
submseteq_inserts_l
.
...
...
@@ -3616,7 +3628,7 @@ Section bind.
Qed
.
Global
Instance
bind_Permutation
:
Proper
((
≡
ₚ
)
==>
(
≡
ₚ
))
(
mbind
f
).
Proof
.
induction
1
;
c
simpl
;
auto
.
induction
1
;
simpl
;
auto
.
-
by
f_equiv
.
-
by
rewrite
!(
assoc_L
(++)),
(
comm
(++)
(
f
_
)).
-
etrans
;
eauto
.
...
...
@@ -3624,30 +3636,30 @@ Section bind.
Lemma
bind_cons
x
l
:
(
x
::
l
)
≫
=
f
=
f
x
++
l
≫
=
f
.
Proof
.
done
.
Qed
.
Lemma
bind_singleton
x
:
[
x
]
≫
=
f
=
f
x
.
Proof
.
c
simpl
.
by
rewrite
(
right_id_L
_
(++)).
Qed
.
Proof
.
simpl
.
by
rewrite
(
right_id_L
_
(++)).
Qed
.
Lemma
bind_app
l1
l2
:
(
l1
++
l2
)
≫
=
f
=
(
l1
≫
=
f
)
++
(
l2
≫
=
f
).
Proof
.
by
induction
l1
;
c
simpl
;
rewrite
<-?(
assoc_L
(++))
;
f_equal
.
Qed
.
Proof
.
by
induction
l1
;
simpl
;
rewrite
<-?(
assoc_L
(++))
;
f_equal
.
Qed
.
Lemma
elem_of_list_bind
(
x
:
B
)
(
l
:
list
A
)
:
x
∈
l
≫
=
f
↔
∃
y
,
x
∈
f
y
∧
y
∈
l
.
Proof
.
split
.
-
induction
l
as
[|
y
l
IH
]
;
c
simpl
;
[
inversion
1
|].
-
induction
l
as
[|
y
l
IH
]
;
simpl
;
[
inversion
1
|].
rewrite
elem_of_app
.
intros
[?|?].
+
exists
y
.
split
;
[
done
|
by
left
].
+
destruct
IH
as
[
z
[??]].
done
.
exists
z
.
split
;
[
done
|
by
right
].
-
intros
[
y
[
Hx
Hy
]].
induction
Hy
;
c
simpl
;
rewrite
elem_of_app
;
intuition
.
-
intros
[
y
[
Hx
Hy
]].
induction
Hy
;
simpl
;
rewrite
elem_of_app
;
intuition
.
Qed
.
Lemma
Forall_bind
(
P
:
B
→
Prop
)
l
:
Forall
P
(
l
≫
=
f
)
↔
Forall
(
Forall
P
∘
f
)
l
.
Proof
.
split
.
-
induction
l
;
c
simpl
;
rewrite
?Forall_app
;
constructor
;
c
simpl
;
intuition
.
-
induction
1
;
c
simpl
;
rewrite
?Forall_app
;
auto
.
-
induction
l
;
simpl
;
rewrite
?Forall_app
;
constructor
;
simpl
;
intuition
.
-
induction
1
;
simpl
;
rewrite
?Forall_app
;
auto
.
Qed
.
Lemma
Forall2_bind
{
C
D
}
(
g
:
C
→
list
D
)
(
P
:
B
→
D
→
Prop
)
l1
l2
:
Forall2
(
λ
x1
x2
,
Forall2
P
(
f
x1
)
(
g
x2
))
l1
l2
→
Forall2
P
(
l1
≫
=
f
)
(
l2
≫
=
g
).
Proof
.
induction
1
;
c
simpl
;
auto
using
Forall2_app
.
Qed
.
Proof
.
induction
1
;
simpl
;
auto
using
Forall2_app
.
Qed
.
End
bind
.
Section
ret_join
.
...
...
@@ -3841,7 +3853,7 @@ Section permutations.
Lemma
permutations_swap
x
y
l
:
y
::
x
::
l
∈
permutations
(
x
::
y
::
l
).
Proof
.
simpl
.
apply
elem_of_list_bind
.
exists
(
y
::
l
).
split
;
simpl
.
-
destruct
l
;
c
simpl
;
rewrite
!
elem_of_cons
;
auto
.
-
destruct
l
;
simpl
;
rewrite
!
elem_of_cons
;
auto
.
-
apply
elem_of_list_bind
.
simpl
.
eauto
using
interleave_cons
,
permutations_refl
.
Qed
.
...
...
@@ -3859,12 +3871,12 @@ Section permutations.
intros
Hl1
[?
|
[
l2'
[??]]]
;
simplify_eq
/=.