Require Import List Permutation EquivDec Program.
Require Import RelationClasses Morphisms.
Require Import Lia Lra Rbase.
Require Import Relation_Definitions Sorted.
Require FinFun.
Require Import LibUtils BasicUtils.
Import ListNotations.
Section Map.
Lemma removelast_map {
A B :
Type} (
f:
A->
B) (
l :
list A) :
removelast (
map f l) =
map f (
removelast l).
Proof.
induction l; simpl; trivial.
rewrite IHl.
destruct l; simpl; trivial.
Qed.
Lemma tl_map {
A B :
Type} (
f:
A->
B) (
l :
list A) :
tl (
map f l) =
map f (
tl l).
Proof.
destruct l; simpl; trivial.
Qed.
Lemma concat_map_swap_perm {
A B C} (
f:
A->
B->
C) (
l1:
list A) (
l2:
list B) :
Permutation (
concat (
map (
fun x =>
map (
fun y =>
f x y)
l2)
l1))
(
concat (
map (
fun x =>
map (
fun y =>
f y x)
l1)
l2)) .
Proof.
revert l2.
induction l1;
simpl;
intros l2.
-
induction l2;
simpl;
trivial.
-
rewrite IHl1.
clear IHl1.
induction l2;
simpl;
trivial.
apply perm_skip.
rewrite <-
IHl2.
repeat rewrite <-
app_ass.
apply Permutation_app;
trivial.
apply Permutation_app_swap.
Qed.
Lemma concat_map_concat {
A} (
l:
list (
list (
list A))) :
concat (
map (@
concat _)
l) =
concat (
concat l).
Proof.
induction l;
simpl;
trivial.
rewrite IHl.
now rewrite concat_app.
Qed.
Lemma concat_map_concat' {
A B} (
f:
A->
list (
list B)) (
l:
list A) :
concat (
map (
fun x =>
concat (
f x))
l) =
concat (
concat (
map f l)).
Proof.
induction l;
simpl;
trivial.
rewrite IHl.
now rewrite concat_app.
Qed.
Lemma map_app_interleave_perm {
A B} (
l:
list A) (
fl:
list (
A->
B)) :
Permutation (
concat (
map (
fun f =>
map (
fun x =>
f x)
l)
fl)) (
concat (
map (
fun x =>
map (
fun f =>
f x)
fl)
l)).
Proof.
Lemma map2_app_interleave_perm {
A B} (
l:
list A) (
f g:
A->
B) :
Permutation (
map f l ++
map g l) (
concat (
map (
fun x =>
f x::
g x::
nil)
l)).
Proof.
End Map.
Section Fold.
Context {
A B C:
Type}.
Lemma fold_right_map
(
f:
C ->
A ->
A) (
g:
B->
C) (
l:
list B) (
init:
A) :
fold_right f init (
map g l) =
fold_right (
fun a b =>
f (
g a)
b)
init l.
Proof.
revert init.
induction l; simpl; trivial; intros.
rewrite IHl; trivial.
Qed.
End Fold.
Lemma fold_right_assoc_abs {
A} (
f:
A->
A->
A) (
init:
A) (
l :
list (
list A))
(
assoc:
forall x y z :
A,
f x (
f y z) =
f (
f x y)
z)
(
abs:
forall x,
f init x =
x) :
fold_right f init (
concat l) =
fold_right f init (
map (
fold_right f init)
l).
Proof.
Lemma fold_right_plus_concat (
l :
list (
list R)) :
fold_right Rplus R0 (
concat l) =
fold_right Rplus R0 (
map (
fold_right Rplus R0)
l).
Proof.
Lemma fold_right_plus_acc f acc l :
fold_right (
fun (
a :
nat) (
b :
R) =>
f a +
b)%
R acc l =
(
fold_right (
fun (
a :
nat) (
b :
R) =>
f a +
b)%
R R0 l +
acc)%
R.
Proof.
induction l; simpl.
- lra.
- rewrite IHl; lra.
Qed.
Section Seq.
Lemma seq_shiftn {
A:
Type} (
l:
list A) (
n:
nat) :
seq n (
length l) =
map (
fun x =>
x +
n)%
nat (
seq 0 (
length l)).
Proof.
Lemma list_as_nthseq_start {
A:
Type} (
l:
list A) (
d:
A) (
c:
nat) :
l =
map (
fun n =>
nth (
n-
c)
l d) (
seq c%
nat (
length l)).
Proof.
induction l;
simpl;
trivial.
rewrite <-
seq_shift.
rewrite map_map.
simpl.
replace (
c-
c)%
nat with 0%
nat by lia.
rewrite IHl.
f_equal.
rewrite map_length.
rewrite seq_length.
apply map_ext_in;
intros x inn.
apply in_seq in inn.
rewrite <-
IHl.
destruct c.
-
f_equal;
lia.
-
assert (
x-
c > 0)%
nat by lia.
replace (
x -
S c)%
nat with ((
x -
c) - 1)%
nat by lia.
destruct (
x-
c)%
nat.
+
lia.
+
f_equal;
lia.
Qed.
Lemma list_as_nthseq {
A:
Type} (
l:
list A) (
d:
A) :
l =
map (
fun n =>
nth n l d) (
seq 0%
nat (
length l)).
Proof.
Lemma seq_Sn s n :
seq s (
S n) =
seq s n ++ [(
s+
n)]%
nat.
Proof.
replace (
S n)
with (
n + 1)%
nat by lia.
rewrite seq_plus.
simpl;
trivial.
Qed.
Lemma tl_seq n :
tl (
seq 0
n) =
seq 1 (
n-1).
Proof.
destruct n;
simpl;
trivial.
rewrite Nat.sub_0_r.
trivial.
Qed.
Lemma removelast_seq (
n:
nat) :
removelast (
seq 0
n) =
seq 0 (
n-1).
Proof.
Lemma seq_shiftn_map start len :
seq start len =
map (
plus start) (
seq 0
len).
Proof.
induction start;
simpl.
-
rewrite map_id;
trivial.
-
rewrite <-
seq_shift.
rewrite IHstart.
rewrite map_map.
trivial.
Qed.
Lemma seq_not_nil :
forall m n, (0 <
n)%
nat -> [] <>
seq m n.
Proof.
induction n;
simpl;
intuition.
generalize (
nil_cons H0);
trivial.
Qed.
Lemma sublist_seq_le :
forall n k, (
n <=
k)%
nat ->
sublist (
seq 0
n) (
seq 0
k).
Proof.
Lemma map_shiftn_seq (
f :
nat ->
R) (
n m :
nat) :
map f (
seq n m) =
map (
fun n0 =>
f (
n +
n0)%
nat) (
seq 0
m).
Proof.
Lemma map_shift1_seq (
f :
nat ->
R) (
n m :
nat) :
map f (
seq (
S n)
m) =
map (
fun n0 =>
f (
S n0)%
nat) (
seq n m).
Proof.
End Seq.
Section fp.
Instance ForallOrdPairs_sub1 {
A:
Type} :
Proper (
subrelation ==>
eq ==>
impl) (@
ForallOrdPairs A).
Proof.
intros R1 R2 rsub l1 l2 lsub FO.
subst.
induction FO;
constructor;
trivial.
eapply Forall_impl;
try eapply H.
apply rsub.
Qed.
Instance ForallOrdPairs_sub2 {
A:
Type} :
Proper (
eq ==>
sublist -->
impl) (@
ForallOrdPairs A).
Proof.
unfold flip.
intros R1 R2 rsub l1 l2 lsub FO;
subst.
induction lsub.
-
constructor.
-
invcs FO.
constructor;
auto 3.
now rewrite lsub.
-
invcs FO.
auto.
Qed.
Global Instance ForallOrdPairs_sub {
A:
Type} :
Proper (
subrelation ==>
sublist -->
impl) (@
ForallOrdPairs A).
Proof.
unfold flip.
intros R1 R2 rsub l1 l2 lsub FO.
now rewrite <-
rsub,
lsub.
Qed.
Lemma ForallOrdPairs_impl {
A:
Type} (
R:
A->
A->
Prop) (
l:
list A) (
f:
A->
A) :
ForallOrdPairs R l ->
ForallOrdPairs (
fun x y =>
R x y ->
R (
f x) (
f y))
l ->
ForallOrdPairs R (
map f l).
Proof.
induction l;
intros FP;
inversion FP;
clear FP;
subst;
intros FP;
simpl.
-
constructor.
-
inversion FP;
clear FP;
subst.
constructor.
+
rewrite Forall_forall in *.
intros x inn.
apply in_map_iff in inn.
destruct inn as [
xx [
eqxx inxx]];
subst.
auto.
+
intuition.
Qed.
Lemma ForallPairs_all {
A:
Type} (
R:
A->
A->
Prop) (
l:
list A) :
(
forall x1 x2,
R x1 x2) ->
ForallPairs R l.
Proof.
firstorder.
Qed.
Global Instance ForallOrdPairs_perm {
A}
R {
sym:
Symmetric R} :
Proper (@
Permutation A ==>
iff) (
ForallOrdPairs R).
Proof.
cut (
forall l l',
Permutation l l' -> (
fun l l' =>
ForallOrdPairs R l ->
ForallOrdPairs R l')
l l').
{
unfold Proper,
respectful;
simpl;
split;
intros.
-
eapply H;
try eapply H1;
eauto.
-
eapply H;
try eapply H1.
symmetry;
eauto.
}
apply Permutation_ind_bis;
simpl;
intros.
-
trivial.
-
invcs H1.
constructor;
auto 2.
now rewrite <-
H.
-
invcs H1.
invcs H5.
invcs H4.
constructor.
+
rewrite <-
H.
constructor;
trivial.
now symmetry.
+
constructor.
*
now rewrite <-
H.
*
eauto.
-
eauto.
Qed.
Instance ForallOrdPairs_Forall2_prop {
A:
Type} (
R1 R2:
A->
A->
Prop) (
Rprop:
Proper (
R1 ==>
R1 ==>
impl)
R2) :
Proper (
Forall2 R1 ==>
impl) (
ForallOrdPairs R2).
Proof.
intros x y xyeq;
unfold impl in *.
induction xyeq.
-
trivial.
-
intros FO.
invcs FO.
constructor.
+
rewrite Forall_forall in *.
intros ?
inn.
unfold Proper,
respectful in Rprop.
destruct (
Forall2_In_r xyeq inn)
as [?[??]].
eauto.
+
eauto.
Qed.
Lemma ForallOrdPairs_In_nth_symm {
A} {
R} {
sym:
Symmetric R} (
l:
list A)
d1 d2 n1 n2
(
FP:
ForallOrdPairs R l)
(
n1bound:(
n1 <
length l)%
nat)
(
n2bound:(
n2 <
length l)%
nat) :
n1 =
n2 \/
R (
nth n1 l d1) (
nth n2 l d2).
Proof.
revert n1 n2 n1bound n2bound.
induction FP;
simpl;
intros.
-
eelim Nat.nlt_0_r;
eauto.
-
simpl in *.
destruct n1;
destruct n2;
simpl in *.
+
auto.
+
right.
rewrite Forall_forall in H.
apply H.
apply nth_In.
now apply lt_S_n.
+
right.
rewrite Forall_forall in H.
symmetry.
apply H.
apply nth_In.
now apply lt_S_n.
+
apply lt_S_n in n1bound.
apply lt_S_n in n2bound.
destruct (
IHFP _ _ n1bound n2bound)
as [?|?];
auto.
Qed.
Lemma ForallOrdPairs_In_nth {
A} {
R} (
l:
list A)
d1 d2 n1 n2
(
FP:
ForallOrdPairs R l)
(
n1bound:(
n1 <
length l)%
nat)
(
n2bound:(
n2 <
length l)%
nat) :
n1 =
n2 \/
R (
nth n1 l d1) (
nth n2 l d2) \/
R (
nth n2 l d2) (
nth n1 l d1).
Proof.
End fp.
Lemma nth_tl {
A}
idx (
l:
list A)
d :
nth idx (
tl l)
d =
nth (
S idx)
l d.
Proof.
destruct l; simpl; trivial.
destruct idx; trivial.
Qed.
Lemma nth_removelast_in {
A}
idx (
l:
list A)
d :
idx <
pred (
length l) ->
nth idx (
removelast l)
d =
nth idx l d.
Proof.
revert idx.
induction l; simpl; trivial; intros idx inn.
destruct l.
- destruct idx; simpl in *; lia.
- simpl in *.
destruct idx; trivial.
rewrite IHl by lia.
trivial.
Qed.
Lemma nth_last {
A} (
l:
list A)
d:
nth (
pred (
length l))
l d =
last l d.
Proof.
induction l; simpl; trivial.
destruct l; simpl in *; trivial.
Qed.
Lemma nth_hd {
A} (
l:
list A)
d:
nth 0
l d =
hd d l.
Proof.
destruct l; simpl in *; trivial.
Qed.
Lemma hd_app {
A} (
l1 l2:
list A)
d :
l1 <>
nil ->
hd d (
l1 ++
l2) =
hd d l1.
Proof.
induction l1; simpl; congruence.
Qed.
Lemma last_rev {
A} (
l:
list A)
d :
last l d =
hd d (
rev l).
Proof.
induction l;
trivial.
simpl rev.
destruct l;
trivial.
rewrite hd_app;
trivial.
intros eqq.
apply (
f_equal (@
length A))
in eqq.
simpl in eqq.
rewrite app_length in eqq.
simpl in eqq.
lia.
Qed.
Lemma last_app {
A} (
l1 l2:
list A)
d :
l2 <>
nil ->
last (
l1 ++
l2)
d =
last l2 d.
Proof.
Lemma last_cons {
A} (
x:
A)
l y :
last (
x::
l)
y =
last l x.
Proof.
revert y.
induction l; simpl; trivial; intros.
destruct l; trivial.
simpl in *.
apply IHl.
Qed.
Lemma seq_last s n d :
(
n > 0)%
nat ->
last (
seq s n)
d = (
s+
n-1)%
nat.
Proof.
intros.
destruct n.
-
simpl;
lia.
-
rewrite seq_Sn.
rewrite last_app by congruence.
simpl.
lia.
Qed.
Lemma last_map {
A B} (
f:
A->
B) (
l:
list A)
d :
last (
map f l) (
f d) =
f (
last l d).
Proof.
induction l; simpl; trivial.
destruct l; simpl; trivial.
Qed.
Lemma map_nth_in {
A B} (
f:
A->
B)
l d1 n d2 :
(
n <
length l)%
nat ->
nth n (
map f l)
d1 =
f (
nth n l d2).
Proof.
revert n.
induction l; simpl.
- destruct n; lia.
- destruct n; trivial.
intros; eauto.
rewrite IHl; trivial; lia.
Qed.
Lemma map_nth_in_exists {
A B} (
f:
A->
B)
l d1 n :
(
n <
length l)%
nat ->
exists d2,
nth n (
map f l)
d1 =
f (
nth n l d2).
Proof.
revert n.
induction l; simpl.
- destruct n; lia.
- destruct n; trivial.
+ intros; eauto.
+ intros.
destruct (IHl n).
* lia.
* rewrite H0.
eauto.
Qed.
Lemma nth_in_default {
A} (
l:
list A)
d1 d2 n :
(
n <
length l)%
nat ->
nth n l d1 =
nth n l d2.
Proof.
revert n.
induction l; simpl.
- destruct n; lia.
- destruct n; trivial.
+ intros; eauto.
rewrite (IHl n); trivial.
lia.
Qed.
Lemma Forall_app_iff {
A} {
P:
A->
Prop} {
l1 l2} :
Forall P (
l1 ++
l2) <->
Forall P l1 /\
Forall P l2.
Proof.
Lemma StronglySorted_app_inv {
A} {
R:
relation A} {
l1 l2} :
StronglySorted R (
l1 ++
l2) ->
StronglySorted R l1 /\
StronglySorted R l2.
Proof.
Hint Constructors StronglySorted :
list.
revert l2.
induction l1;
intros l2 ss;
simpl in *.
-
simpl in *;
split;
trivial with list.
-
inversion ss;
subst;
clear ss.
destruct (
IHl1 _ H1).
split;
trivial.
constructor;
trivial.
apply Forall_app_iff in H2.
tauto.
Qed.
Lemma StronglySorted_sub {
A} (
R1 R2:
relation A) :
subrelation R1 R2 ->
forall l,
StronglySorted R1 l ->
StronglySorted R2 l.
Proof.
Hint Constructors StronglySorted :
list.
intros sub.
induction l;
simpl;
intros ssl;
trivial with list.
inversion ssl;
clear ssl;
subst.
simpl in *.
constructor.
-
apply IHl;
intuition.
-
rewrite Forall_forall in *.
eauto.
Qed.
Lemma StronglySorted_map_in {
A B} (
R1:
relation A) (
R2:
relation B) (
f:
A->
B)
l :
(
forall x y,
In x l /\
In y l ->
R1 x y ->
R2 (
f x) (
f y)) ->
StronglySorted R1 l ->
StronglySorted R2 (
map f l).
Proof.
Hint Constructors StronglySorted :
list.
intros prop.
induction l;
simpl;
intros ssl;
trivial with list.
inversion ssl;
clear ssl;
subst.
simpl in *.
constructor.
-
apply IHl;
intuition.
-
rewrite Forall_forall in *.
intros x inn.
apply in_map_iff in inn.
destruct inn as [
a' [
eqq inn]].
subst;
auto.
Qed.
Lemma StronglySorted_map {
A B} (
R1:
relation A) (
R2:
relation B) (
f:
A->
B) :
Proper (
R1 ==>
R2)
f ->
forall l,
StronglySorted R1 l ->
StronglySorted R2 (
map f l).
Proof.
Lemma StronglySorted_compose {
A B}
R (
f:
A->
B) (
l:
list A) :
StronglySorted R (
map f l) <->
StronglySorted (
fun x y =>
R (
f x) (
f y))
l.
Proof.
induction l;
simpl.
-
intuition.
-
split;
inversion 1;
subst;
constructor;
intuition.
+
now rewrite Forall_map in H3.
+
now rewrite Forall_map.
Qed.
Lemma StronglySorted_break {
A}
R (
l:
list A)
x :
StronglySorted R l ->
In x l ->
exists b c,
l =
b++
x::
c /\
Forall (
fun y =>
R y x)
b /\
Forall (
R x)
c.
Proof.
induction l;
simpl;
intros ss inn; [
tauto | ].
invcs ss.
destruct inn.
-
subst.
exists nil,
l.
simpl.
intuition.
-
destruct IHl as [
b [
c [
p1 [
p2 p3]]]];
trivial.
subst.
exists (
a::
b),
c.
simpl;
intuition.
constructor;
trivial.
rewrite Forall_forall in H2.
specialize (
H2 x).
rewrite in_app_iff in H2;
simpl in H2.
eauto.
Qed.
Lemma StronglySorted_nth_lt {
A}
R (
l:
list A)
idx1 idx2 d1 d2 :
StronglySorted R l ->
(
idx2 <
length l)%
nat ->
(
idx1 <
idx2)%
nat ->
R (
nth idx1 l d1) (
nth idx2 l d2).
Proof.
intros.
destruct (@
nth_split _ idx1 l d1)
as [
l1 [
l2 [
leqq l1len]]]
; [
lia | ].
rewrite leqq in H.
apply StronglySorted_app_inv in H.
destruct H as [
_ ssl2].
inversion ssl2;
clear ssl2;
subst.
rewrite leqq in *.
rewrite app_length in H0.
simpl in H0.
revert H4.
generalize (
nth (
length l1)
l d1).
clear l leqq.
intros a Fa.
rewrite Forall_forall in Fa.
apply Fa.
rewrite app_nth2 by lia.
simpl.
case_eq (
idx2 -
length l1);
try lia.
intros.
apply nth_In.
lia.
Qed.
Lemma StronglySorted_nth_le {
A}
R (
l:
list A)
idx1 idx2 d1 d2 :
reflexive _ R ->
StronglySorted R l ->
(
idx2 <
length l)%
nat ->
(
idx1 <=
idx2)%
nat ->
R (
nth idx1 l d1) (
nth idx2 l d2).
Proof.
Section bucket.
Context {
A:
Type} {
R:
relation A} (
R_dec :
forall x y, {
R x y} + {~
R x y}).
Fixpoint find_bucket (
needle:
A) (
haystack:
list A)
:=
match haystack with
|
x::((
y::
_)
as more) =>
if R_dec x needle
then if R_dec needle y
then Some (
x,
y)
else find_bucket needle more
else None
|
_ =>
None
end.
Lemma find_bucket_break {
needle l a1 a2}:
find_bucket needle l =
Some (
a1,
a2) ->
exists l1 l2,
l =
l1 ++ [
a1;
a2] ++
l2.
Proof.
induction l;
simpl;
try discriminate.
destruct l;
try discriminate.
destruct (
R_dec a needle);
try discriminate.
destruct (
R_dec needle a0).
-
inversion 1;
subst.
exists nil,
l.
reflexivity.
-
intros HH.
destruct (
IHl HH)
as [
l1 [
l2 eqq]].
rewrite eqq.
exists (
a::
l1),
l2.
reflexivity.
Qed.
Lemma middle_find_bucket needle l1 l2 a1 a2:
transitive _ R ->
antisymmetric _ R ->
StronglySorted R (
l1++[
a1]) ->
R a1 needle ->
R needle a2 ->
~
R needle a1 ->
find_bucket needle (
l1 ++
a1::
a2::
l2) =
Some (
a1,
a2).
Proof.
intros trans antisymm.
intros sorted r1 r2 nr1.
revert sorted.
induction l1;
intros sorted.
-
simpl.
destruct (
R_dec a1 needle); [ |
tauto].
destruct (
R_dec needle a2); [ |
tauto].
trivial.
-
simpl in *.
inversion sorted;
clear sorted;
subst.
specialize (
IHl1 H1).
rewrite IHl1;
trivial.
destruct (
R_dec a needle).
+
destruct l1;
simpl.
*
destruct (
R_dec needle a1);
tauto.
*
destruct (
R_dec needle a0);
trivial.
elim nr1.
apply (
trans _ a0);
trivial.
inversion H1;
clear H1;
subst.
rewrite Forall_forall in H4.
apply H4.
rewrite in_app_iff.
simpl;
tauto.
+
rewrite Forall_forall in H2.
elim n.
apply (
trans _ a1);
trivial.
apply H2.
rewrite in_app_iff.
simpl;
tauto.
Qed.
Lemma find_bucket_nth_finds needle l idx d1 d2:
transitive _ R ->
antisymmetric _ R ->
StronglySorted R l ->
S idx <
length l ->
R (
nth idx l d1)
needle ->
R needle (
nth (
S idx)
l d2) ->
~
R needle (
nth idx l d1) ->
find_bucket needle l =
Some (
nth idx l d1,
nth (
S idx)
l d2).
Proof.
intros trans antisymm ss idx_bound.
assert (
idx_bound':
idx <
length l)
by lia.
destruct (
nth_split l d1 idx_bound')
as [
l1 [
l2 [
eqq leneq]]].
revert eqq.
generalize (
nth idx l d1);
intros a1.
intros eqq r1 r2 nr1.
subst.
rewrite app_nth2 in *
by lia.
replace ((
S (
length l1) -
length l1))
with 1
in *
by lia.
rewrite app_length in idx_bound.
simpl in *.
destruct l2;
simpl in *; [
lia | ].
apply middle_find_bucket;
trivial.
replace (
l1 ++
a1 ::
a ::
l2)
with ((
l1 ++
a1::
nil) ++ (
a ::
l2))
in ss.
-
apply StronglySorted_app_inv in ss.
tauto.
-
rewrite app_ass;
simpl;
trivial.
Qed.
Lemma find_bucket_needle_in needle l a1 a2:
find_bucket needle l =
Some (
a1,
a2) ->
R a1 needle /\
R needle a2.
Proof.
induction l;
simpl;
try discriminate.
destruct l;
try discriminate.
destruct (
R_dec a needle);
try discriminate.
destruct (
R_dec needle a0).
-
inversion 1;
subst.
tauto.
-
intuition.
Qed.
Lemma find_bucket_bucket_in needle l a1 a2 d1 d2:
reflexive _ R ->
StronglySorted R l ->
find_bucket needle l =
Some (
a1,
a2) ->
R (
hd d1 l)
a1 /\
R a2 (
last l d2).
Proof.
intros refl ssl eqq1.
destruct (
find_bucket_break eqq1)
as [
l1 [
l2 eqq2]].
replace (
l1 ++ [
a1;
a2] ++
l2)
with ((
l1 ++
a1::
nil) ++ (
a2::
l2))
in eqq2.
-
subst.
apply StronglySorted_app_inv in ssl.
destruct ssl as [
ssl1 ssl2].
split.
+
destruct l1;
simpl.
*
apply refl.
*
inversion ssl1;
subst.
rewrite Forall_forall in H2.
apply H2.
rewrite in_app_iff;
simpl;
tauto.
+
inversion ssl2;
clear ssl2;
subst.
rewrite last_app by congruence.
rewrite Forall_forall in H2.
simpl.
destruct l2.
*
apply refl.
*
apply H2.
clear.
{
revert a.
induction l2.
-
simpl;
tauto.
-
simpl in *.
eauto.
}
-
rewrite app_ass;
simpl;
reflexivity.
Qed.
Lemma find_bucket_bounded_le_exists a b l (
needle:
A) :
(
forall x y,
R x y \/
R y x) ->
R a needle ->
R needle b ->
exists lower upper,
find_bucket needle (
a::
l++[
b]) =
Some (
lower,
upper).
Proof.
intros total r1 r2.
simpl.
destruct (
R_dec a needle); [ |
tauto].
induction l;
simpl.
-
destruct (
R_dec needle b); [ |
tauto].
eauto.
-
destruct (
IHl)
as [
lower [
upper eqq]].
destruct (
R_dec needle a0).
+
eauto.
+
clear IHl r.
{
destruct (
R_dec a0 needle).
-
destruct l;
simpl in *.
+
destruct (
R_dec needle b);
simpl;
eauto.
+
destruct (
R_dec needle a1);
eauto.
-
destruct (
total needle a0);
tauto.
}
Qed.
End bucket.
Lemma StronglySorted_seq s n :
StronglySorted lt (
seq s n).
Proof.
revert s.
induction n;
intros s;
simpl.
-
constructor.
-
constructor;
trivial.
rewrite Forall_forall;
intros.
apply in_seq in H.
lia.
Qed.
Lemma length_S_tl {
A :
Type} (
l :
list A) :
l <>
nil ->
length l =
S (
length (
tl l)).
Proof.
intros.
destruct l; simpl; congruence.
Qed.
Lemma tl_length {
A :
Type} (
l :
list A) :
length (
tl l) =
pred (
length l).
Proof.
intros.
destruct l; simpl; congruence.
Qed.
Lemma removelast_length {
A :
Type} (
l :
list A) :
length (
removelast l) =
pred (
length l).
Proof.
induction l; trivial.
destruct l; trivial.
simpl in *.
rewrite IHl; trivial.
Qed.
Section combining.
Lemma combine_nth_in {
A B :
Type} (
l :
list A) (
l' :
list B) (
n :
nat) (
x :
A) (
y :
B) :
n <
min (
length l) (
length l') ->
nth n (
combine l l') (
x,
y) = (
nth n l x,
nth n l'
y).
Proof.
revert l' n.
induction l; simpl; intros l' n nlt.
- lia.
- destruct l'; simpl in *.
+ lia.
+ destruct n; simpl; trivial.
apply IHl.
lia.
Qed.
Lemma combine_map {
A B C D:
Type} (
f:
A->
C) (
g:
B->
D) (
l1:
list A) (
l2:
list B) :
combine (
map f l1) (
map g l2) =
map (
fun '(
x,
y) => (
f x,
g y)) (
combine l1 l2).
Proof.
revert l2.
induction l1; intros l2; simpl; trivial.
destruct l2; simpl; trivial.
f_equal.
auto.
Qed.
Lemma combine_self {
A:
Type} (
l:
list A) :
combine l l =
map (
fun x => (
x,
x))
l.
Proof.
induction l; simpl; trivial.
f_equal; trivial.
Qed.
Lemma combine_nil_l {
A B} (
l:
list B) :
combine (@
nil A)
l =
nil.
Proof.
reflexivity.
Qed.
Lemma combine_nil_r {
A B} (
l:
list A) :
combine l (@
nil B) =
nil.
Proof.
destruct l; trivial.
Qed.
Lemma combine_swap {
A B} (
l1:
list A) (
l2:
list B) :
combine l1 l2 =
map (
fun xy => (
snd xy,
fst xy)) (
combine l2 l1).
Proof.
revert l2; induction l1; simpl; intros
; destruct l2; simpl; trivial.
rewrite IHl1; trivial.
Qed.
Lemma combine_domain_eq {
A B} (
x:
list A) (
y:
list B) :
length x =
length y ->
domain (
combine x y) =
x.
Proof.
revert y.
induction x; destruct y; simpl in *; intros
; try easy.
inversion H.
rewrite IHx; trivial.
Qed.
Lemma list_prod_map {
A B C D:
Type} (
f:
A->
C) (
g:
B->
D) (
l1:
list A) (
l2:
list B) :
list_prod (
map f l1) (
map g l2) =
map (
fun '(
x,
y) => (
f x,
g y)) (
list_prod l1 l2).
Proof.
revert l2.
induction l1;
intros l2;
simpl;
trivial.
rewrite map_app.
repeat rewrite map_map.
rewrite IHl1.
trivial.
Qed.
Lemma list_prod_nil_l {
A B} (
l:
list B) :
list_prod (@
nil A)
l =
nil.
Proof.
reflexivity.
Qed.
Lemma list_prod_nil_r {
A B} (
l:
list A) :
list_prod l (@
nil B) =
nil.
Proof.
induction l; trivial.
Qed.
Lemma list_prod_cons2_pred {
A B} (
l1:
list A)
a (
l2:
list B) :
Permutation (
list_prod l1 (
a ::
l2)) (
map (
fun x :
A => (
x,
a))
l1 ++
list_prod l1 l2).
Proof.
Lemma list_prod_swap {
A B} (
l1:
list A) (
l2:
list B) :
Permutation (
list_prod l1 l2) (
map (
fun xy => (
snd xy,
fst xy)) (
list_prod l2 l1)).
Proof.
Instance list_prod_perm_proper1 {
A B} :
Proper ((@
Permutation A) ==>
eq ==> (@
Permutation (
A*
B))) (@
list_prod A B).
Proof.
intros l1 l1'
perm1 l2'
l2 eqq;
subst.
revert l1 l1'
perm1.
apply Permutation_ind_bis;
intros.
-
reflexivity.
-
simpl.
rewrite H0;
reflexivity.
-
simpl.
rewrite H0.
repeat rewrite <-
app_ass.
apply Permutation_app; [ |
reflexivity ].
rewrite Permutation_app_swap.
reflexivity.
-
etransitivity;
eauto.
Qed.
Global Instance list_prod_perm_proper {
A B} :
Proper ((@
Permutation A) ==> (@
Permutation B) ==> (@
Permutation (
A*
B))) (@
list_prod A B).
Proof.
Lemma combine_incl_list_prod {
A B} (
l1:
list A) (
l2:
list B) :
incl (
combine l1 l2) (
list_prod l1 l2).
Proof.
End combining.
Definition adjacent_pairs {
A:
Type} (
l:
list A) := (
combine l (
tl l)).
Definition adjacent_pairs_alt {
A:
Type} (
l:
list A) := (
combine (
removelast l) (
tl l)).
Lemma adjacent_pairs_alt_eq {
A:
Type} (
l:
list A) :
adjacent_pairs l =
adjacent_pairs_alt l.
Proof.
Lemma adjacent_pairs_length {
A:
Type} (
l:
list A) :
length (
adjacent_pairs l) =
pred (
length l).
Proof.
Lemma adjacent_pairs_nth_in {
A:
Type}
n (
l:
list A)
d1 d2 :
S n <
length l ->
nth n (
adjacent_pairs l) (
d1,
d2) = (
nth n l d1,
nth (
S n)
l d2).
Proof.
Lemma adjacent_pairs_map {
A B:
Type} (
f:
A->
B) (
l:
list A) :
adjacent_pairs (
map f l) =
map (
fun '(
x,
y) => (
f x,
f y)) (
adjacent_pairs l).
Proof.
Lemma adjacent_pairs_seq s n :
adjacent_pairs (
seq s n) =
map (
fun i => (
i,
S i)) (
seq s (
pred n)).
Proof.
unfold adjacent_pairs.
revert s.
induction n;
simpl;
intros s;
trivial.
destruct n;
simpl;
trivial.
f_equal.
apply IHn.
Qed.
Section folds_with_history.
Definition fold_left_with_history_aux {
A B :
Type} (
f:
A ->
B ->
A) (
l:
list B) (
init:
A) (
acc:
list A) :
list A
:=
rev (
let '(
x,
y) := (
fold_left (
fun '(
a,
hs)
b => (
f a b,
a::
hs))
l (
init,
acc))
in
x::
y).
Definition fold_left_with_history {
A B :
Type} (
f:
A ->
B ->
A) (
l:
list B) (
init:
A) :
list A
:=
fold_left_with_history_aux f l init nil.
Lemma fold_left_with_history_aux_len {
A B :
Type} (
f:
A ->
B ->
A) (
l:
list B) (
init:
A) (
acc:
list A) :
(
List.length (
fold_left_with_history_aux f l init acc) =
S (
length l) +
length acc)%
nat.
Proof.
Lemma fold_left_with_history_len {
A B :
Type} (
f:
A ->
B ->
A) (
l:
list B) (
init:
A) :
List.length (
fold_left_with_history f l init) =
S (
length l).
Proof.
Fixpoint fold_left_partial_with_history_aux {
A B :
Type} (
f :
A ->
B ->
option A) (
l:
list B) (
acc:
A*
list A) :
list A +
list A
:=
match l with
| [] =>
inl (
rev (
fst acc ::
snd acc))
|
b ::
t =>
match f (
fst acc)
b with
|
None =>
inr (
rev (
fst acc ::
snd acc))
|
Some a' =>
fold_left_partial_with_history_aux f t (
a',
fst acc::
snd acc)
end
end.
Definition fold_left_partial_with_history {
A B :
Type} (
f :
A ->
B ->
option A) (
l:
list B) (
init:
A) :
list A +
list A
:=
fold_left_partial_with_history_aux f l (
init,
nil).
Lemma fold_left_partial_with_history_aux_ok {
A B :
Type} (
f :
A ->
B ->
option A) (
l:
list B) (
init:
A) (
acc:
list A)
l' :
fold_left_partial_with_history_aux f l (
init,
acc) =
inl l' ->
(
List.length l' =
S (
length l) +
length acc)%
nat.
Proof.
revert init acc l'.
induction l;
simpl;
intros init acc l'
eqq.
-
invcs eqq.
rewrite app_length,
rev_length;
simpl.
lia.
-
match_destr_in eqq.
rewrite (
IHl _ _ _ eqq);
simpl.
lia.
Qed.
Lemma fold_left_partial_with_history_ok {
A B :
Type} (
f :
A ->
B ->
option A) (
l:
list B) (
init:
A)
l' :
fold_left_partial_with_history f l init =
inl l' ->
List.length l' =
S (
length l).
Proof.
Lemma fold_left_partial_with_history_aux_err {
A B :
Type} (
f :
A ->
B ->
option A) (
l:
list B) (
init:
A) (
acc:
list A)
l' :
fold_left_partial_with_history_aux f l (
init,
acc) =
inr l' ->
(
List.length l' <
S (
length l) +
length acc)%
nat.
Proof.
revert init acc l'.
induction l;
simpl;
intros init acc l'
eqq.
-
invcs eqq.
-
match_destr_in eqq.
+
specialize (
IHl _ _ _ eqq);
simpl in IHl.
lia.
+
invcs eqq.
rewrite app_length,
rev_length;
simpl.
lia.
Qed.
Lemma fold_left_partial_with_history_aux_total {
A B :
Type}
(
f :
A ->
B ->
A) (
l:
list B) (
init:
A) (
acc:
list A) :
fold_left_partial_with_history_aux (
fun a b =>
Some (
f a b))
l (
init,
acc) =
inl (
fold_left_with_history_aux f l init acc).
Proof.
Lemma fold_left_partial_with_history_total {
A B :
Type} (
f :
A ->
B ->
A) (
l:
list B) (
init:
A) :
fold_left_partial_with_history (
fun a b =>
Some (
f a b))
l init =
inl (
fold_left_with_history f l init).
Proof.
End folds_with_history.
Lemma nodup_not_nil (
T :
Type) (
l :
list T) (
dec:
forall (
x y:
T), {
x =
y} + {
x <>
y}) :
l <>
nil <->
nodup dec l <>
nil.
Proof.
intros.
destruct l;
intros;
simpl.
-
intuition.
-
intuition try congruence.
match_destr_in H0.
apply (
nodup_In dec)
in i.
rewrite H0 in i.
simpl in i;
trivial.
Qed.
Lemma map_in_inj_strong {
A B} (
f:
A->
B)
a (
l:
list A) :
(
forall a b,
In (
f a) (
map f l) ->
In (
f b) (
map f l) ->
f a =
f b ->
a =
b) ->
In (
f a) (
map f l) ->
In a l.
Proof.
intros inj HH.
apply in_map_iff in HH.
destruct HH as [
x [
eqqx inx]].
rewrite (
inj a x);
trivial.
-
rewrite <-
eqqx.
now apply in_map.
-
now apply in_map.
-
congruence.
Qed.
Lemma nodup_map_inj {
A B}
decA decB (
f:
A->
B) (
l:
list A) :
(
forall a b,
In (
f a) (
map f l) ->
In (
f b) (
map f l) ->
f a =
f b ->
a =
b) ->
nodup decB (
map f l) =
map f (
nodup decA l).
Proof.
intros inj.
induction l;
simpl;
trivial.
assert (
forall a b :
A,
In (
f a) (
map f l) ->
In (
f b) (
map f l) ->
f a =
f b ->
a =
b).
{
simpl in inj.
intuition.
}
rewrite IHl by trivial.
match_destr;
match_destr.
-
apply map_in_inj_strong in i;
trivial.
congruence.
-
elim n.
now apply in_map.
Qed.
Lemma list_prod_concat {
A B} (
l1:
list A) (
l2:
list B) :
list_prod l1 l2 =
concat (
map (
fun x =>
map (
fun y => (
x,
y))
l2)
l1).
Proof.
induction l1; simpl; trivial.
now rewrite IHl1.
Qed.
Lemma nodup_const_map {
A} (
c r:
A)
dec (
l :
list A) :
[
c] =
nodup dec (
map (
fun _ =>
c) (
r ::
l)).
Proof.
induction l; simpl; trivial.
rewrite IHl.
match_destr.
simpl.
intuition.
Qed.
Lemma concat_NoDup {
A} (
l:
list (
list A)) :
NoDup (
concat l) ->
Forall (@
NoDup A)
l.
Proof.
induction l;
simpl;
intros nd.
-
constructor.
-
constructor.
+
eapply NoDup_app_inv;
eauto.
+
apply IHl.
eapply NoDup_app_inv2;
eauto.
Qed.
Lemma nodup_app2_incl {
A}
decA (
l1 l2:
list A) :
incl l1 l2 ->
nodup decA (
l1 ++
l2) =
nodup decA l2.
Proof.
unfold incl;
intros inn.
induction l1;
simpl;
trivial;
simpl in *.
match_destr.
-
eauto.
-
elim n.
apply in_app_iff.
eauto.
Qed.
Lemma nodup_app_distr {
A}
decA (
l1 l2:
list A) :
disjoint l1 l2 ->
nodup decA (
l1 ++
l2) =
nodup decA l1 ++
nodup decA l2.
Proof.
unfold disjoint.
intros disj.
induction l1;
simpl;
trivial.
rewrite IHl1 by firstorder.
destruct (
in_dec decA a l1).
-
match_destr.
elim n.
apply in_app_iff;
auto.
-
match_destr.
apply in_app_iff in i.
elim (
disj a);
simpl;
intuition.
Qed.
Lemma list_prod_nodup {
A B}
decA decB decAB (
l1:
list A) (
l2:
list B):
nodup decAB (
list_prod l1 l2) =
list_prod (
nodup decA l1) (
nodup decB l2).
Proof.
repeat rewrite list_prod_concat.
revert l2.
induction l1;
simpl;
trivial.
intros l2.
match_destr.
-
rewrite <-
IHl1.
apply nodup_app2_incl.
intros x inn.
apply concat_In.
eexists.
split;
try eassumption.
apply in_map_iff.
eauto.
-
simpl.
rewrite <-
IHl1.
rewrite nodup_app_distr.
+
f_equal.
induction l2;
simpl;
trivial.
rewrite IHl2.
match_destr.
*
apply in_map_iff in i.
destruct i as [
x [
eqq xin]].
invcs eqq.
match_destr.
congruence.
*
match_destr.
elim n0.
apply in_map_iff.
eauto.
+
unfold disjoint.
intros [
x y]
inn HH.
apply concat_In in HH.
destruct HH as [
xx [
xxin xinn]].
apply in_map_iff in xxin.
destruct xxin as [
xxx [?
xxxin]];
subst.
apply in_map_iff in inn.
destruct inn as [? [
eqq ?]].
invcs eqq;
subst.
apply in_map_iff in xinn.
destruct xinn as [? [
eqq ?]].
invcs eqq.
congruence.
Qed.
Lemma nodup_map_nodup {
A B}
decA decB (
f:
A->
B) (
l:
list A) :
nodup decB (
map f (
nodup decA l)) =
nodup decB (
map f l).
Proof.
induction l;
simpl;
trivial.
match_destr;
match_destr.
+
elim n.
apply in_map_iff;
eauto.
+
simpl.
match_destr.
elim n0.
eapply in_map_iff.
eapply in_map_iff in i.
destruct i as [? [?
inn]].
eapply nodup_In in inn.
eauto.
+
simpl.
rewrite IHl.
match_destr.
elim n0.
eapply in_map_iff in i.
destruct i as [? [?
inn]].
apply nodup_In in inn.
apply in_map_iff.
eauto.
Qed.
Lemma nodup_equiv {
A}
dec (
l:
list A) :
equivlist (
nodup dec l)
l.
Proof.
induction l;
simpl.
-
reflexivity.
-
match_destr.
+
rewrite IHl.
unfold equivlist;
simpl;
intuition congruence.
+
now rewrite IHl.
Qed.
Lemma incl_nil_r {
A} (
l:
list A) :
incl l nil ->
l =
nil.
Proof.
unfold incl.
destruct l;
simpl;
trivial.
intros HH.
elim (
HH a);
auto.
Qed.
Lemma remove_one_nin {
A} {
dec:
EqDec A eq}
a (
l:
list A) :
~
In a l ->
remove_one a l =
l.
Proof.
induction l; simpl; trivial.
match_destr.
- intuition.
- intros; f_equal; apply IHl.
eauto.
Qed.
Lemma remove_one_app_nin {
A} {
dec:
EqDec A eq}
a (
l1 l2:
list A) :
~
In a l1 ->
remove_one a (
l1 ++
l2) =
l1 ++
remove_one a l2.
Proof.
induction l1; simpl; trivial.
intros ninn.
match_destr.
- red in e.
intuition.
- rewrite IHl1 by intuition.
trivial.
Qed.
Lemma remove_one_in_perm {
A} {
dec :
EqDec A eq} (
a:
A)
l :
In a l ->
Permutation l (
a::
remove_one a l).
Proof.
induction l;
simpl;
intros inn.
-
tauto.
-
match_destr.
+
red in e;
subst.
reflexivity.
+
rewrite perm_swap.
apply perm_skip.
intuition.
Qed.
Lemma remove_other_in {
A} {
dec :
EqDec A eq} (
a1 a2:
A)
l :
a1 <>
a2 ->
In a1 l <->
In a1 (
remove_one a2 l).
Proof.
intros.
induction l; simpl.
- intuition.
- match_destr.
+ red in e; subst.
intuition.
+ simpl.
intuition.
Qed.
Lemma bminus_in_nin {
A} {
decA:
EqDec A eq}
a (
l1 l2 :
list A) :
In a l1 -> ~
In a l2 ->
In a (
bminus l2 l1).
Proof.
revert l1.
induction l2;
simpl in *.
-
intuition.
-
intros.
apply IHl2.
+
apply remove_other_in;
eauto.
+
eauto.
Qed.
Lemma incl_front_perm {
A} {
decA:
EqDec A eq} (
l1 l2 :
list A) :
incl l2 l1 ->
NoDup l2 ->
{
l3:
list A |
Permutation l1 (
l2 ++
l3)}.
Proof.
Instance equivlist_incl_part {
A} :
PartialOrder equivlist (@
incl A).
Proof.
Lemma NoDup_app_disj {
A} (
a b :
list A) :
NoDup (
a ++
b) ->
disjoint a b.
Proof.
unfold disjoint.
induction a;
simpl.
-
intuition.
-
intros.
invcs H.
destruct H0.
+
subst.
apply H4.
apply in_app_iff;
tauto.
+
eauto.
Qed.
Lemma NoDup_perm_disj {
A} (
l1 l2 l3 :
list A) :
Permutation l1 (
l2 ++
l3) ->
NoDup l1 ->
disjoint l2 l3.
Proof.
Lemma incl_front_perm_nodup {
A} (
decA:
EqDec A eq) (
l1 l2 :
list A) :
incl l2 l1 ->
{
l3:
list A |
Permutation (
nodup decA l1) (
nodup decA l2 ++
l3)}.
Proof.
Lemma Forallt_in {
A} (
decA:
forall x y:
A, {
x=
y} + {
x <>
y}) {
X:
A->
Type} {
l:
list A} (
ft:
Forallt X l) {
a} (
pf:
In a l) :
X a.
Proof.
induction l; simpl in *.
- elim pf.
- inversion ft.
destruct (decA a a0).
+ congruence.
+ apply IHl; trivial.
intuition congruence.
Defined.
Fixpoint Forallt_map {
A B:
Type} {
X:
A->
Type} {
l:
list A} (
f:
forall a,
X a ->
B) (
ft:
Forallt X l) :
list B
:=
match ft with
|
Forallt_nil _ =>
nil
|
Forallt_cons _ x l px pl =>
f x px ::
Forallt_map f pl
end.
Lemma Forallt_map_length {
A B} {
X:
A->
Type} (
f:
forall a :
A,
X a ->
B) (
l:
list A) (
ft:
Forallt X l) :
length (
Forallt_map f ft) =
length l.
Proof.
induction ft; simpl; trivial.
now rewrite IHft.
Qed.
Lemma Forallt_map_irrel {
A B:
Type} {
X:
A->
Type} {
l:
list A} (
f:
forall a,
X a ->
B) (
ft1 ft2:
Forallt X l) :
(
forall a,
In a l ->
forall x y,
f a x =
f a y) ->
Forallt_map f ft1 =
Forallt_map f ft2.
Proof.
intros.
revert ft1 ft2.
induction l; intros
; dependent destruction ft1
; dependent destruction ft2
; simpl in *
; trivial.
f_equal; eauto.
Qed.
Lemma map_nil' {
A B} (
f:
A->
B)
l :
List.map f l =
nil <->
l =
nil.
Proof.
split; intros.
- induction l; try reflexivity; simpl in *.
congruence.
- rewrite H; reflexivity.
Qed.
Lemma map_nil {
A B} (
f :
A ->
B) (
l :
list A) :
List.map f l = (@
nil B) <->
l = (@
nil A).
Proof.
split; intros.
- induction l; try reflexivity; simpl in *.
congruence.
- rewrite H; reflexivity.
Qed.
Lemma map_not_nil {
A B} (
l :
list A) (
f :
A ->
B):
[] <>
List.map f l <-> [] <>
l.
Proof.
rewrite ne_symm ;
rewrite (
ne_symm _ l).
split ;
intros.
*
intro Hl.
rewrite <-(
map_nil f)
in Hl ;
firstorder.
*
intro Hl.
rewrite (
map_nil f)
in Hl ;
firstorder.
Qed.
Lemma not_nil_exists {
A} (
l :
list A) :
[] <>
l <->
exists a,
In a l.
Proof.
split.
*
intros Hl.
induction l.
-
firstorder.
-
destruct l.
--
exists a.
simpl;
now left.
--
set (
Hnc := @
nil_cons _ a0 l).
specialize (
IHl Hnc).
destruct IHl as [
a1 Ha1].
exists a1.
simpl in * ;
intuition.
*
intros [
a Ha]
not.
rewrite <-
not in Ha ;
firstorder.
Qed.
Lemma list_prod_not_nil {
A B} {
la :
list A} {
lb :
list B}(
Hla : [] <>
la) (
Hlb : [] <>
lb) :
[] <>
list_prod la lb.
Proof.
Fixpoint applyn {
A} (
init :
A) (
g :
A ->
A) (
n :
nat) :
A :=
match n with
| 0 =>
init
|
S k =>
g (
applyn init g k)
end.
Lemma concat_map_map {
A} (
l :
list(
list A)) (
f :
A ->
R) :
concat (
map (
map f)
l) =
map f (
concat l).
Proof.
induction l.
simpl ;
reflexivity.
simpl.
rewrite map_app.
rewrite IHl.
reflexivity.
Qed.
Lemma ForallPairs_filter {
A}
R (
l :
list A) (
p :
A ->
bool) :
ForallPairs R l ->
ForallPairs R (
filter p l).
Proof.
Lemma ForallOrdPairs_filter {
A}
R (
l :
list A) (
p :
A ->
bool) :
ForallOrdPairs R l ->
ForallOrdPairs R (
filter p l).
Proof.
intros H.
induction l.
-
simpl ;
constructor.
-
simpl.
case_eq (
p a).
+
invcs H.
specialize (
IHl H3).
intro Hpa.
constructor ;
trivial.
apply Forall_filter ;
trivial.
+
intro Hpa.
invcs H.
specialize (
IHl H3).
assumption.
Qed.
Lemma ForallOrdPairs_app {
A :
Type} {
R :
A ->
A ->
Prop} {
l1 l2 :
list A} :
ForallOrdPairs R l1 ->
ForallOrdPairs R l2 ->
(
forall x y,
In x l1 ->
In y l2 ->
R x y) ->
ForallOrdPairs R (
l1 ++
l2).
Proof.
revert l2.
induction l1;
simpl;
trivial;
intros.
invcs H.
constructor.
-
apply Forall_app;
trivial.
rewrite Forall_forall.
intros.
apply H1;
trivial.
eauto.
-
apply IHl1;
trivial.
eauto.
Qed.
Lemma ForallOrdPairs_app_in {
A R} {
l1 l2:
list A} :
ForallOrdPairs R (
l1 ++
l2) ->
forall x y,
In x l1 ->
In y l2 ->
R x y.
Proof.
revert l2.
induction l1;
simpl;
intros.
-
intuition.
-
invcs H.
destruct H0.
+
subst.
eapply (
Forall_forall _ _)
in H4;
try eassumption.
rewrite in_app_iff ;
eauto.
+
eapply IHl1;
eauto.
Qed.
Lemma filter_true {
A} :
forall p:
list A,
filter (
fun _ =>
true)
p =
p.
Proof.
induction p.
- simpl; reflexivity.
- simpl. rewrite IHp. reflexivity.
Qed.
Lemma filter_false {
A} :
forall p:
list A,
filter (
fun _ =>
false)
p = [].
Proof.
induction p.
- simpl; reflexivity.
- simpl. rewrite IHp. reflexivity.
Qed.
Lemma Forall2_refl_in {
A}
R (
l:
list A) :
Forall (
fun x =>
R x x)
l ->
Forall2 R l l.
Proof.
induction l; simpl; trivial.
intros HH; invcs HH.
constructor; auto.
Qed.
Lemma Forall2_perm {
A}
R (
l1 l1'
l2:
list A) :
Forall2 R l1 l2 ->
Permutation l1 l1' ->
exists l2',
Permutation l2 l2' /\
Forall2 R l1'
l2'.
Proof.
revert l2.
cut (
forall (
l1 l1':
list A),
Permutation l1 l1' ->
(
fun l1 l1' =>
forall l2,
Forall2 R l1 l2 ->
exists l2',
Permutation l2 l2' /\
Forall2 R l1'
l2')
l1 l1')
; [
eauto | ].
apply Permutation_ind_bis;
simpl;
intros.
-
invcs H.
exists nil;
eauto.
-
invcs H1.
destruct (
H0 _ H6)
as [? [??]].
exists (
y::
x0).
split.
+
now rewrite H1.
+
now constructor.
-
invcs H1.
invcs H6.
destruct (
H0 _ H7)
as [? [??]].
exists (
y1 ::
y0 ::
x0).
split.
+
rewrite H1.
apply perm_swap.
+
repeat constructor;
trivial.
-
destruct (
H0 _ H3)
as [? [??]].
destruct (
H2 _ H5)
as [? [??]].
exists x0.
split;
trivial.
etransitivity;
eauto.
Qed.
Global Instance ForallPairs_sublist {
A} {
R} :
Proper (
sublist -->
Basics.impl) (@
ForallPairs A R).
Proof.
Lemma Forall2_map_rel {
A B} (
R:
A->
A->
Prop)
l1 l2 (
f:
A->
B) :
(
forall x1 x2,
In x1 l1 ->
In x2 l2 ->
R x1 x2 ->
f x1 =
f x2) ->
Forall2 R l1 l2 ->
map f l1 =
map f l2.
Proof.
intros H H0.
induction H0.
- simpl ; reflexivity.
- simpl in *. rewrite IHForall2.
specialize (H x y).
rewrite H ; trivial.
+ left ; trivial.
+ left ; trivial.
+ intros x1 x2 H2 H3 H4.
apply H.
-- now right.
-- now right.
-- assumption.
Qed.
Section equivlist.
Global Instance app_equivlist_proper {
A} :
Proper (
equivlist ==>
equivlist ==>
equivlist) (@
app A).
Proof.
split;
intros inn
;
apply in_app_iff
;
apply in_app_iff in inn.
-
destruct inn.
+
rewrite H in H1;
tauto.
+
rewrite H0 in H1;
tauto.
-
destruct inn.
+
rewrite <-
H in H1;
tauto.
+
rewrite <-
H0 in H1;
tauto.
Qed.
Lemma equivlist_const {
A B} (
b:
B) (
l:
list A) :
l <>
nil ->
equivlist (
map (
fun _ =>
b)
l) (
b::
nil).
Proof.
induction l; simpl; intros.
- congruence.
- destruct l; simpl.
+ reflexivity.
+ simpl in IHl.
rewrite IHl by congruence.
red; simpl; tauto.
Qed.
Lemma list_prod_fst_equiv {
A B} (
a:
list A) (
b:
list B) :
b <>
nil ->
equivlist (
map fst (
list_prod a b))
a.
Proof.
intros.
induction a;
simpl.
-
reflexivity.
-
intros.
simpl.
rewrite map_app.
rewrite map_map;
simpl.
rewrite IHa.
rewrite equivlist_const by trivial.
reflexivity.
Qed.
Lemma list_prod_snd_equiv {
A B} (
a:
list A) (
b:
list B) :
a <>
nil ->
equivlist (
map snd (
list_prod a b))
b.
Proof.
Global Instance list_prod_equivlist {
A B} :
Proper (
equivlist ==>
equivlist ==>
equivlist) (@
list_prod A B).
Proof.
Global Instance map_equivlist_proper {
A B}:
Proper (
pointwise_relation _ eq ==>
equivlist ==>
equivlist) (@
map A B).
Proof.
End equivlist.
Global Instance nodup_perm {
A}
dec :
Proper (@
Permutation A ==> @
Permutation A) (
nodup dec).
Proof.
repeat red;
intros.
revert x y H.
apply Permutation_ind_bis;
simpl;
intros.
-
trivial.
-
repeat match_destr.
+
rewrite H in i;
congruence.
+
rewrite <-
H in i;
congruence.
+
apply perm_skip;
trivial.
-
destruct (
dec x y)
;
destruct (
dec y x)
;
try congruence.
+
subst.
destruct (
in_dec dec y l)
;
destruct (
in_dec dec y l')
;
try congruence.
*
rewrite H in i;
congruence.
*
rewrite <-
H in i;
congruence.
*
apply perm_skip;
congruence.
+
destruct (
in_dec dec y l)
;
destruct (
in_dec dec x l)
;
destruct (
in_dec dec x l')
;
destruct (
in_dec dec y l')
;
try congruence
;
try solve [
rewrite H in i;
congruence
|
rewrite H in i0;
congruence
|
rewrite H in i1;
congruence
|
rewrite <-
H in i;
congruence
|
rewrite <-
H in i0;
congruence
|
rewrite <-
H in i1;
congruence
|
apply perm_skip;
congruence
] .
rewrite H0.
apply perm_swap.
-
now rewrite H0.
Qed.
Lemma Forall_nodup {
A}
dec P (
l:
list A) :
Forall P l <->
Forall P (
nodup dec l).
Proof.
Lemma fold_right_ext {
A B:
Type} (
f1 f2 :
B ->
A ->
A) (
init :
A) (
l:
list B) :
(
forall a b,
In b l ->
f1 b a =
f2 b a) ->
fold_right f1 init l =
fold_right f2 init l.
Proof.
induction l; simpl; trivial; intros.
rewrite IHl; auto.
Qed.
Lemma listo_to_olist_length {
A:
Type} (
l:
list (
option A)) (
r:
list A)
:
listo_to_olist l =
Some r ->
length l =
length r.
Proof.
revert r.
induction l; simpl; intros.
- now invcs H; simpl.
- destruct a; try discriminate.
match_option_in H; try discriminate.
invcs H.
simpl.
now rewrite (IHl _ eqq).
Qed.
Section cross_product.
Definition list_cross_product {
T} (
l:
list (
list T)) : (
list (
list T))
:=
match l with
|
nil =>
nil
|
x::
l' =>
fold_left (
fun (
acc:
list (
list T)) (
b:
list T) =>
map (
fun '(
a,
b) =>
b ++ [
a]) (
list_prod b acc))
l' (
map singleton x)
end.
Example list_cross_product_example :
list_cross_product [[1;2];[3];[4;5;6]] =
[[1; 3; 4]; [2; 3; 4]; [1; 3; 5]; [2; 3; 5]; [1; 3; 6]; [2; 3; 6]]
:=
eq_refl _.
Lemma list_cross_product_length {
T} (
l:
list (
list T)) (
lnnil:
l <>
nil)
:
length (
list_cross_product l) =
fold_left Peano.mult (
List.map (@
length T)
l) 1%
nat.
Proof.
Lemma list_cross_product_inner_length {
T} (
l:
list (
list T))
:
Forall (
fun x =>
length x =
length l) (
list_cross_product l).
Proof.
Lemma in_list_cross_product {
T} (
l:
list (
list T)) (
nnil:
l<>
nil) (
x:
list T) :
In x (
list_cross_product l) ->
Forall2 (@
In T)
x l.
Proof.
destruct l;
simpl; [
congruence |
clear nnil].
revert x.
induction l0 using rev_ind;
simpl.
-
intros.
apply in_map_iff in H.
destruct H as [? [?
inn]];
subst.
unfold singleton.
constructor;
trivial.
-
intros ?
HH.
rewrite fold_left_app in HH;
simpl in HH.
apply in_map_iff in HH.
destruct HH as [[??][??]];
subst.
apply in_prod_iff in H0.
destruct H0 as [
inn1 inn2].
replace (
l::
l0 ++ [
x])
with ((
l::
l0)++[
x])
by reflexivity.
apply Forall2_app;
auto.
Qed.
Lemma list_cross_product_in {
T} (
l:
list (
list T)) (
nnil:
l<>
nil) (
x:
list T) :
Forall2 (@
In T)
x l ->
In x (
list_cross_product l).
Proof.
destruct l;
simpl; [
congruence |
clear nnil].
revert x.
induction l0 using rev_ind;
simpl;
intros xx HH.
-
invcs HH.
apply in_map_iff.
unfold singleton.
invcs H3.
eauto.
-
invcs HH.
apply Forall2_app_inv_r in H3.
destruct H3 as [?[?[
F2 [
F2'
eqq]]]];
subst.
invcs F2'.
invcs H4.
specialize (
IHl0 (
x0::
x1)).
rewrite fold_left_app;
simpl.
apply in_map_iff.
exists (
x3,
x0::
x1).
simpl.
split;
trivial.
apply in_prod_iff.
split;
auto.
Qed.
Lemma list_cross_product_in_iff {
T} (
l:
list (
list T)) (
nnil:
l<>
nil) (
x:
list T) :
In x (
list_cross_product l) <->
Forall2 (@
In T)
x l.
Proof.
End cross_product.
Section list_dep.
Program Fixpoint list_dep_zip {
T} {
P:
T->
Prop} (
l:
list T)
:
Forall P l ->
list (
sig P)
:=
match l as l'
return Forall P l' ->
list (
sig P)
with
|
nil =>
fun _ =>
nil
|
x::
l' =>
fun fp =>
exist _ x _ ::
list_dep_zip l'
_
end.
Next Obligation.
now invcs fp.
Qed.
Next Obligation.
now invcs fp.
Qed.
Lemma list_dep_zip_map1 {
T} {
P:
T->
Prop} (
l:
list T)
Fp
:
map (@
proj1_sig _ P) (
list_dep_zip l Fp) =
l.
Proof.
revert Fp.
induction l; simpl; trivial; intros.
now rewrite IHl.
Qed.
Lemma list_dep_zip_length {
T} {
P:
T->
Prop} (
l:
list T) (
Fp:
Forall P l) :
length (
list_dep_zip l Fp) =
length l.
Proof.
Lemma list_dep_zip_ext2 {
T} (
P :
T ->
Prop) (
l :
list T) (
F1 F2:
Forall P l) :
Forall2 (
fun x y =>
proj1_sig x =
proj1_sig y) (
list_dep_zip l F1) (
list_dep_zip l F2).
Proof.
induction l; simpl; trivial.
constructor; trivial.
Qed.
Lemma list_dep_zip_ext_map {
T B} (
P :
T ->
Prop) (
l :
list T) (
F1 F2:
Forall P l) (
f1 f2:
sig P->
B) :
(
forall x pf1 pf2,
In x l ->
f1 (
exist _ x pf1) =
f2 (
exist _ x pf2)) ->
map f1 (
list_dep_zip l F1) =
map f2 (
list_dep_zip l F2).
Proof.
intros.
induction l; simpl in *; trivial.
f_equal; firstorder.
Qed.
End list_dep.
Program Fixpoint map_onto {
A B} (
l:
list A) (
f:
forall a,
In a l ->
B) :
list B
:=
match l with
| [] => []
|
x::
l' =>
f x _ ::
map_onto l' (
fun a pf =>
f a _)
end.
Next Obligation.
simpl; auto.
Qed.
Next Obligation.
simpl; auto.
Qed.
Lemma map_onto_length {
A B} (
l:
list A) (
f:
forall a,
In a l ->
B) :
length (
map_onto l f) =
length l.
Proof.
induction l; simpl; congruence.
Qed.
Lemma filter_map_swap {
A B} (
P:
B->
bool) (
f:
A->
B) (
l:
list A) :
filter P (
map f l) =
map f (
filter (
fun x =>
P (
f x))
l).
Proof.
induction l; simpl; trivial.
rewrite IHl.
match_destr.
Qed.
Lemma ForallOrdPairs_map {
A B} (
P:
B->
B->
Prop) (
f:
A->
B) (
l:
list A) :
ForallOrdPairs P (
map f l) <->
ForallOrdPairs (
fun x y =>
P (
f x) (
f y))
l.
Proof.
induction l;
simpl;
split;
intros HH
;
invcs HH
;
constructor
;
trivial.
-
now rewrite Forall_map in H1.
-
now apply IHl.
-
now rewrite Forall_map.
-
now apply IHl.
Qed.
Lemma list_max_in l :
l <>
nil ->
In (
list_max l)
l.
Proof.
induction l;
simpl; [
eauto |];
intros _.
destruct (
Max.max_dec a (
list_max l))
;
rewrite e in *
;
eauto.
destruct l.
-
simpl in e.
rewrite Max.max_0_r in e;
simpl
;
eauto.
-
right;
apply IHl;
congruence.
Qed.
Lemma list_max_upper l :
List.Forall (
fun k :
nat => (
k <=
list_max l)%
nat)
l.
Proof.
Lemma NoDup_list_max_count l :
NoDup l ->
(
length l <=
S (
list_max l))%
nat.
Proof.
Lemma NoDup_prod {
A B} {
decA:
EqDec A eq} {
decB:
EqDec B eq} (
l1 :
list A) (
l2 :
list B) :
NoDup l1 ->
NoDup l2 ->
NoDup (
list_prod l1 l2).
Proof.
Lemma Forall2_concat {
A}
R (
l1 l2:
list (
list A)):
Forall2 (
Forall2 R)
l1 l2 ->
Forall2 R (
concat l1) (
concat l2).
Proof.
induction 1;
simpl.
-
constructor.
-
apply Forall2_app;
auto.
Qed.
Lemma map_const_repeat {
A B} (
c:
A) (
l:
list B) :
map (
fun _ =>
c)
l =
repeat c (
length l).
Proof.
induction l; simpl; congruence.
Qed.
Lemma filter_le_length {
A:
Type} (
f1 f2:
A->
bool) (
l:
list A) :
(
forall x,
In x l ->
f1 x =
true ->
f2 x =
true) ->
(
length (
filter f1 l) <=
length (
filter f2 l))%
nat.
Proof.
induction l; simpl; trivial; intros.
cut_to IHl; [| firstorder].
specialize (H a).
destruct (f1 a).
- rewrite H by tauto.
simpl; lia.
- destruct (f2 a); simpl; lia.
Qed.
Lemma incl_seq (
s1 n1 s2 n2:
nat) :
incl (
seq s1 (
S n1)) (
seq s2 n2) <-> ((
s2 <=
s1)%
nat /\ (
s1 +
S n1 <=
s2 +
n2)%
nat).
Proof.
transitivity (
forall a, (
s1 <=
a <
s1 +
S n1)%
nat -> (
s2 <=
a <
s2 +
n2)%
nat).
-
split.
+
intros.
apply in_seq.
apply H.
now apply in_seq.
+
intros ???.
apply in_seq.
apply H.
now apply in_seq.
-
split.
+
intros.
split.
*
specialize (
H s1);
lia.
*
specialize (
H (
s1 +
n1))%
nat;
lia.
+
lia.
Qed.