]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/dama/dama/models/list_support.ma
8dad0c436b9b224bcb73e7527e86bba5fde54e1b
[helm.git] / helm / software / matita / contribs / dama / dama / models / list_support.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 include "nat/minus.ma".
16 include "list/list.ma".
17
18 interpretation "list nth" 'nth = (nth _).
19 interpretation "list nth" 'nth_appl l d i = (nth _ l d i).
20 notation "\nth" with precedence 90 for @{'nth}.
21 notation < "\nth \nbsp term 90 l \nbsp term 90 d \nbsp term 90 i" 
22 with precedence 69 for @{'nth_appl $l $d $i}.
23
24 definition make_list ≝
25   λA:Type.λdef:nat→A.
26     let rec make_list (n:nat) on n ≝
27       match n with [ O ⇒ nil ? | S m ⇒ def m :: make_list m]
28     in make_list.
29
30 interpretation "\mk_list appl" 'mk_list_appl f n = (make_list _ f n).
31 interpretation "\mk_list" 'mk_list = (make_list _).   
32 notation "\mk_list" with precedence 90 for @{'mk_list}.
33 notation < "\mk_list \nbsp term 90 f \nbsp term 90 n" 
34 with precedence 69 for @{'mk_list_appl $f $n}.
35
36 notation "\len" with precedence 90 for @{'len}.
37 interpretation "len" 'len = (length _).
38 notation < "\len \nbsp term 90 l" with precedence 69 for @{'len_appl $l}.
39 interpretation "len appl" 'len_appl l = (length _ l).
40
41 lemma len_mk_list : ∀T:Type.∀f:nat→T.∀n.\len (\mk_list f n) = n.
42 intros; elim n; [reflexivity] simplify; rewrite > H; reflexivity;
43 qed.
44
45 record rel (rel_T : Type) : Type ≝ {
46   rel_op :2> rel_T → rel_T → Prop
47 }.
48
49 record trans_rel : Type ≝ {
50   o_T :> Type;
51   o_rel :> rel o_T;
52   o_tra : ∀x,y,z: o_T.o_rel x y → o_rel y z → o_rel x z 
53 }.
54
55 lemma trans: ∀r:trans_rel.∀x,y,z:r.r x y → r y z → r x z.
56 apply o_tra;
57 qed.
58
59 inductive sorted (lt : trans_rel): list (o_T lt) → Prop ≝ 
60 | sorted_nil : sorted lt []
61 | sorted_one : ∀x. sorted lt [x]
62 | sorted_cons : ∀x,y,tl. lt x y → sorted lt (y::tl) → sorted lt (x::y::tl). 
63
64 lemma nth_nil: ∀T,i.∀def:T. \nth [] def i = def.
65 intros; elim i; simplify; [reflexivity;] assumption; qed.
66
67 lemma len_append: ∀T:Type.∀l1,l2:list T. \len (l1@l2) = \len l1 + \len l2.
68 intros; elim l1; [reflexivity] simplify; rewrite < H; reflexivity;
69 qed.
70
71 inductive non_empty_list (A:Type) : list A → Type := 
72 | show_head: ∀x,l. non_empty_list A (x::l).
73
74 lemma len_gt_non_empty : 
75   ∀T.∀l:list T.O < \len l → non_empty_list T l.
76 intros; cases l in H; [intros; cases (not_le_Sn_O ? H);] intros; constructor 1;
77 qed. 
78
79 lemma sorted_tail: ∀r,x,l.sorted r (x::l) → sorted r l.
80 intros; inversion H; intros; [destruct H1;|destruct H1;constructor 1;]
81 destruct H4; assumption;
82 qed.
83
84 lemma sorted_skip: ∀r,x,y,l. sorted r (x::y::l) → sorted r (x::l).
85 intros (r x y l H1); inversion H1; intros; [1,2: destruct H]
86 destruct H4; inversion H2; intros; [destruct H4]
87 [1: destruct H4; constructor 2;
88 |2: destruct H7; constructor 3; 
89     [ apply (o_tra ? ??? H H4); | apply (sorted_tail ??? H2);]]
90 qed.
91
92 lemma sorted_tail_bigger : ∀r,x,l,d.sorted r (x::l) → ∀i. i < \len l → r x (\nth l d i).
93 intros 4; elim l; [ cases (not_le_Sn_O i H1);]
94 cases i in H2;
95 [2: intros; apply (H ? n);[apply (sorted_skip ???? H1)|apply le_S_S_to_le; apply H2]
96 |1: intros; inversion H1; intros; [1,2: destruct H3]
97     destruct H6; simplify; assumption;]
98 qed. 
99
100 (*
101 lemma all_bases_positive : ∀f:q_f.∀i. OQ < nth_base (bars f) (S i).
102 intro f; generalize in match (bars_begin_OQ f); generalize in match (bars_sorted f);
103 cases (bars_not_nil f); intros;
104 cases (cmp_nat i (len l));
105 [1: lapply (sorted_tail_bigger ?? H ? H2) as K; simplify in H1;
106     rewrite > H1 in K; apply K;
107 |2: rewrite > H2; simplify; elim l; simplify; [apply (q_pos_OQ one)] 
108     assumption;
109 |3: simplify; elim l in i H2;[simplify; rewrite > nth_nil; apply (q_pos_OQ one)]
110     cases n in H3; intros; [cases (not_le_Sn_O ? H3)] apply (H2 n1);
111     apply (le_S_S_to_le ?? H3);]
112 qed.
113 *)
114
115 (* move in nat/ *)
116 lemma lt_n_plus_n_Sm : ∀n,m:nat.n < n + S m.
117 intros; rewrite > sym_plus; apply (le_S_S n (m+n)); alias id "le_plus_n" = "cic:/matita/nat/le_arith/le_plus_n.con".
118 apply (le_plus_n m n); qed.
119
120 lemma nth_append_lt_len:
121   ∀T:Type.∀l1,l2:list T.∀def.∀i.i < \len l1 → \nth (l1@l2) def i = \nth l1 def i.
122 intros 4; elim l1; [cases (not_le_Sn_O ? H)] cases i in H H1; simplify; intros;
123 [reflexivity| rewrite < H;[reflexivity] apply le_S_S_to_le; apply H1]
124 qed. 
125
126 lemma nth_append_ge_len:
127   ∀T:Type.∀l1,l2:list T.∀def.∀i.
128     \len l1 ≤ i → \nth (l1@l2) def i = \nth l2 def (i - \len l1).
129 intros 4; elim l1; [ rewrite < minus_n_O; reflexivity]
130 cases i in H1; simplify; intros; [cases (not_le_Sn_O ? H1)]
131 apply H; apply le_S_S_to_le; apply H1;
132 qed. 
133
134 lemma nth_len:
135   ∀T:Type.∀l1,l2:list T.∀def,x.
136     \nth (l1@x::l2) def (\len l1) = x.
137 intros 2; elim l1;[reflexivity] simplify; apply H; qed.
138
139 lemma sorted_head_smaller:
140   ∀r,l,p,d. sorted r (p::l) → ∀i.i < \len l → r p (\nth l d i).
141 intros 2 (r l); elim l; [cases (not_le_Sn_O ? H1)] cases i in H2; simplify; intros;
142 [1: inversion H1; [1,2: simplify; intros; destruct H3] intros; destruct H6; assumption; 
143 |2: apply (H p ?? n ?); [apply (sorted_skip ???? H1)] apply le_S_S_to_le; apply H2]
144 qed.
145
146 alias symbol "lt" = "natural 'less than'".
147 lemma sorted_pivot:
148   ∀r,l1,l2,p,d. sorted r (l1@p::l2) → 
149     (∀i. i < \len l1 → r (\nth l1 d i) p) ∧
150     (∀i. i < \len l2 → r p (\nth l2 d i)).
151 intros 2 (r l); elim l; 
152 [1: split; [intros; cases (not_le_Sn_O ? H1);] intros;
153     apply sorted_head_smaller; assumption;
154 |2: simplify in H1; cases (H ?? d (sorted_tail ??? H1));
155     lapply depth = 0 (sorted_head_smaller ??? d H1) as Hs;
156     split; simplify; intros;
157     [1: cases i in H4; simplify; intros; 
158         [1: lapply depth = 0 (Hs (\len l1)) as HS; 
159             rewrite > nth_len in HS; apply HS;
160             rewrite > len_append; simplify; apply lt_n_plus_n_Sm;
161         |2: apply (H2 n); apply le_S_S_to_le; apply H4]
162     |2: apply H3; assumption]]
163 qed.
164
165 inductive cases_bool (p:bool) : bool → CProp ≝
166 | case_true : p = true → cases_bool p true
167 | cases_false : p = false → cases_bool p false.
168
169 lemma case_b : ∀A:Type.∀f:A → bool. ∀x.cases_bool (f x) (f x).
170 intros; cases (f x);[left;|right] reflexivity;
171 qed. 
172
173 include "cprop_connectives.ma".
174
175 definition eject_N ≝
176   λP.λp:∃x:nat.P x.match p with [ex_introT p _ ⇒ p].
177 coercion eject_N.
178 definition inject_N ≝ λP.λp:nat.λh:P p. ex_introT ? P p h.
179 coercion inject_N with 0 1 nocomposites.
180
181 inductive find_spec (T:Type) (P:T→bool) (l:list T) (d:T) (res : nat) : nat → CProp ≝
182 | found: 
183     ∀i. i < \len l → P (\nth l d i) = true → res = i →  
184     (∀j. j < i → P (\nth l d j) = false) → find_spec T P l d res i  
185 | not_found: ∀i. i = \len l → res = i →
186     (∀j.j < \len l → P (\nth l d j) = false) → find_spec T P l d res i.
187
188 lemma find_lemma : ∀T.∀P:T→bool.∀l:list T.∀d.∃i.find_spec ? P l d i i.
189 intros;
190 letin find ≝ (
191   let rec aux (acc: nat) (l : list T) on l : nat ≝
192     match l with
193     [ nil ⇒ acc
194     | cons x tl ⇒
195         match P x with
196         [ false ⇒ aux (S acc) tl   
197         | true ⇒ acc]]
198   in aux :
199   ∀acc,l1.∃p:nat.
200     ∀story. story @ l1 = l → acc = \len story →
201     find_spec ? P story d acc acc → 
202     find_spec ? P (story @ l1) d p p);
203 [4: clearbody find; cases (find 0 l);
204     lapply (H [] (refl_eq ??) (refl_eq ??)) as K;
205     [2: apply (not_found ?? [] d); intros; try reflexivity; cases (not_le_Sn_O ? H1);
206     |1: cases K; clear K;
207         [2: exists[apply (\len l)]
208             apply not_found; try reflexivity; apply H3;
209         |1: exists[apply i] apply found;  try reflexivity; assumption;]]
210 |1: intros; cases (aux (S n) l2); simplify; clear aux;
211     lapply depth = 0 (H5 (story@[t])) as K; clear H5;
212     change with (find_spec ? P (story @ ([t] @ l2)) d w w);
213     rewrite < associative_append; apply K; clear K;
214     [1: rewrite > associative_append; apply H2;
215     |2: rewrite > H3; rewrite > len_append; rewrite > sym_plus; reflexivity;
216     |3: cases H4; clear H4; destruct H7;
217         [2: rewrite > H5; rewrite > (?:S (\len story) = \len (story @ [t])); [2:
218               rewrite > len_append; rewrite > sym_plus; reflexivity;]
219             apply not_found; try reflexivity; intros; cases (cmp_nat (S j) (\len story));
220             [1: rewrite > (nth_append_lt_len ????? H8); apply H7; apply H8;
221             |2: rewrite > (nth_append_ge_len ????? (le_S_S_to_le ?? H8));
222                 rewrite > (?: j - \len story = 0);[assumption]
223                 rewrite > (?:j = \len story);[rewrite > minus_n_n; reflexivity]
224                 apply le_to_le_to_eq; [2: apply le_S_S_to_le; assumption;] 
225                 rewrite > len_append in H4;rewrite > sym_plus in H4; 
226                 apply le_S_S_to_le; apply H4;]
227         |1: rewrite < H3 in H5; cases (not_le_Sn_n ? H5);]]
228 |2: intros; cases H4; clear H4; 
229     [1: destruct H7; rewrite > H3 in H5; cases (not_le_Sn_n ? H5); 
230     |2: apply found; try reflexivity;
231         [1: rewrite > len_append; simplify; rewrite > H5; apply lt_n_plus_n_Sm;
232         |2: rewrite > H5; rewrite > nth_append_ge_len; [2: apply le_n]
233             rewrite < minus_n_n; assumption;
234         |3: intros; rewrite > H5 in H4; rewrite > nth_append_lt_len; [2:assumption]
235             apply H7; assumption]]
236 |3: intros; rewrite > append_nil; assumption;]
237 qed.
238
239 lemma find : ∀T:Type.∀P:T→bool.∀l:list T.∀d:T.nat.
240 intros; cases (find_lemma ? f l t); apply w; qed.
241
242 lemma cases_find: ∀T,P,l,d. find_spec T P l d (find T P l d) (find T P l d).
243 intros; unfold find; cases (find_lemma T P l d); simplify; assumption; qed.