-
Notifications
You must be signed in to change notification settings - Fork 0
/
ListFacts.v
135 lines (108 loc) · 3.61 KB
/
ListFacts.v
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
Load Common.
Require Import UserTactics.
Lemma in_cons_iff : forall {A : Type} {a b : A} {l : list A}, In b (a :: l) <-> (a = b \/ In b l).
Proof.
auto with *.
Qed.
Lemma Forall_tl (T : Type) (P : T -> Prop) : forall (ss : list T), Forall P ss -> Forall P (tl ss).
Proof.
case => //.
intros; gimme Forall; inversion; assumption.
Qed.
Lemma Forall_app (T : Type) (P : T -> Prop) : forall (A B : list T), Forall P (A ++ B) <-> Forall P A /\ Forall P B.
Proof.
elim; first by firstorder done.
intros * => IH.
intros.
constructor.
inversion.
firstorder (by constructor).
move => [HP1 HP2].
inversion_clear HP1.
rewrite <- app_comm_cons.
firstorder (by constructor).
Qed.
Lemma Forall_cons : forall (T : Type) (P : T -> Prop) (A : list T) (a : T),
Forall P (a :: A) -> P a /\ Forall P A.
Proof.
intros; gimme Forall; inversion; auto.
Qed.
Lemma Forall_In : forall (T : Type) (P : T -> Prop) (A : list T) (a : T), In a A -> Forall P A -> P a.
Proof.
move => T P.
elim => [a | b A IH a]; first case.
move /in_cons_iff.
case => [-> | ?]; move /Forall_cons => [? ?]; auto.
Qed.
Lemma Forall_flat_map (T U: Type) (P : T -> Prop) : forall (ds : list U) (f : U -> list T) (d : U),
Forall P (flat_map f ds) -> In d ds -> Forall P (f d).
Proof.
elim; first done.
cbn.
intros.
gimme Forall; move /Forall_app => [? ?].
firstorder (subst; done).
Qed.
(*destructs all assumptions of the shape Forall P l where l matches cons, nil or app*)
Ltac decompose_Forall :=
do ? (
match goal with
| [H : Forall _ (_ :: _) |- _] => inversion_clear H
| [H : Forall _ nil |- _] => inversion_clear H
| [H : Forall _ (_ ++ _) |- _] => move /Forall_app : H => [? ?]
end).
(*tactic to decide list membership*)
Ltac list_element :=
(try assumption);
lazymatch goal with
| [|- In _ (_ :: _)] => first [by left | right; list_element]
| [|- In _ (_ ++ _)] => apply in_or_app; first [left; list_element | right; list_element]
| [|- In ?a ?l] => let l' := eval hnf in l in progress(change (In a l')); list_element
end.
(*tactic to decide list inclusion*)
Ltac list_inclusion :=
intros; do ? (match goal with | [H : In ?a _ |- context [?a]] => move: H end);
match goal with
| [ |- In _ _] => list_element
| _ => rewrite ? (in_app_iff, in_cons_iff); intuition (subst; tauto)
end.
Ltac list_inclusion_veryfast :=
let rec list_inclusion_rec :=
lazymatch goal with
| [ |- In _ _] => list_element
| [ |- In ?b (?a :: _) -> _] =>
case /(in_inv (a := a) (b := b)) => [?|]; [subst; list_inclusion_rec | list_inclusion_rec]
| [ |- In ?b (_ ++ _) -> _] =>
case /(in_app_or _ _ b); list_inclusion_rec
| [ |- In _ _ -> _] => move => ?; list_inclusion_rec
end
in
intros;
try (match goal with | [H : In ?a _ |- In ?a _] => move: H end); clear;
list_inclusion_rec.
Lemma in_sub : forall (T : Type) (A B : list T) (a : T), In a A -> (forall (b : T), In b A -> In b B) -> In a B.
Proof.
auto.
Qed.
Lemma Forall_exists_monotone : forall (A : Type) (P : nat -> A -> Prop) (l : list A),
(forall (n m : nat) (a : A), P n a -> n <= m -> P m a) -> Forall (fun (a : A) => exists (n : nat), P n a) l ->
exists (n : nat), Forall (P n) l.
Proof.
move => A P l H. elim : l.
intros; exists 0; auto.
move => a l IH; inversion.
gimme Forall. move /IH.
gimme where P; move => [n1 ?].
move => [n2 ?].
exists (n1+n2); constructor.
apply : H; [eassumption | omega].
apply : Forall_impl; last eassumption.
intros; apply : H; [eassumption | omega].
Qed.
Lemma Forall_cons_iff : forall (T : Type) (P : T -> Prop) (l : list T) (a : T),
Forall P (a :: l) <-> P a /\ Forall P l.
Proof.
intros. split.
inversion. auto.
case. intros. by constructor.
Qed.