]> matita.cs.unibo.it Git - helm.git/blob - matita/library/nat/minimization.ma
Towards chebyshev.
[helm.git] / matita / library / nat / minimization.ma
1 (**************************************************************************)
2 (*       ___                                                                *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||       A.Asperti, C.Sacerdoti Coen,                          *)
8 (*      ||A||       E.Tassi, S.Zacchiroli                                 *)
9 (*      \   /                                                             *)
10 (*       \ /        Matita is distributed under the terms of the          *)
11 (*        v         GNU Lesser General Public License Version 2.1         *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 set "baseuri" "cic:/matita/nat/minimization".
16
17 include "nat/minus.ma".
18
19 let rec max i f \def
20   match (f i) with 
21   [ true \Rightarrow i
22   | false \Rightarrow 
23       match i with 
24       [ O \Rightarrow O
25       | (S j) \Rightarrow max j f ]].
26
27 theorem max_O_f : \forall f: nat \to bool. max O f = O.
28 intro. simplify.
29 elim (f O).
30 simplify.reflexivity.
31 simplify.reflexivity.
32 qed. 
33
34 theorem max_S_max : \forall f: nat \to bool. \forall n:nat.
35 (f (S n) = true \land max (S n) f = (S n)) \lor 
36 (f (S n) = false \land max (S n) f = max n f).
37 intros.simplify.elim (f (S n)).
38 simplify.left.split.reflexivity.reflexivity.
39 simplify.right.split.reflexivity.reflexivity.
40 qed.
41
42 theorem le_max_n : \forall f: nat \to bool. \forall n:nat.
43 max n f \le n.
44 intros.elim n.rewrite > max_O_f.apply le_n.
45 simplify.elim (f (S n1)).simplify.apply le_n.
46 simplify.apply le_S.assumption.
47 qed.
48
49 theorem le_to_le_max : \forall f: nat \to bool. \forall n,m:nat.
50 n\le m  \to max n f \le max m f.
51 intros.elim H.
52 apply le_n.
53 apply (trans_le ? (max n1 f)).apply H2.
54 cut ((f (S n1) = true \land max (S n1) f = (S n1)) \lor 
55 (f (S n1) = false \land max (S n1) f = max n1 f)).
56 elim Hcut.elim H3.
57 rewrite > H5.
58 apply le_S.apply le_max_n.
59 elim H3.rewrite > H5.apply le_n.
60 apply max_S_max.
61 qed.
62
63 theorem f_m_to_le_max: \forall f: nat \to bool. \forall n,m:nat.
64 m\le n \to f m = true \to m \le max n f.
65 intros 3.elim n.apply (le_n_O_elim m H).
66 apply le_O_n.
67 apply (le_n_Sm_elim m n1 H1).
68 intro.apply (trans_le ? (max n1 f)).
69 apply H.apply le_S_S_to_le.assumption.assumption.
70 apply le_to_le_max.apply le_n_Sn.
71 intro.simplify.rewrite < H3. 
72 rewrite > H2.simplify.apply le_n.
73 qed.
74
75 theorem max_f_g: \forall f,g,n. (\forall i. i \le n \to f i = g i) \to
76 max n f = max n g.
77 intros 3.
78 elim n
79   [simplify.
80    rewrite > (H O)
81     [reflexivity
82     |apply le_n
83     ]
84   |simplify.
85    rewrite > H
86     [rewrite > H1
87       [reflexivity
88       |apply le_n
89       ]
90     |intros.
91      apply H1.
92      apply le_S.
93      assumption
94     ]
95   ]
96 qed.
97
98 theorem le_max_f_max_g: \forall f,g,n. (\forall i. i \le n \to f i = true \to g i =true) \to
99 max n f \le max n g.
100 intros 3.
101 elim n
102   [simplify.
103    elim (f O);apply le_O_n
104   |simplify.
105    apply (bool_elim ? (f (S n1)));intro
106     [rewrite > (H1 (S n1) ? H2)
107       [apply le_n
108       |apply le_n
109       ]
110     |cases (g(S n1))
111       [simplify.
112        apply le_S.
113        apply le_max_n
114       |simplify.
115        apply H.
116        intros.
117        apply H1
118         [apply le_S.assumption
119         |assumption
120         ]
121       ]
122     ]
123   ]
124 qed.
125
126
127 theorem max_O : \forall f:nat \to bool. \forall n:nat.
128 (\forall i:nat. le i n \to f i = false) \to max n f = O.
129 intros 2.elim n
130   [simplify.rewrite > H
131     [reflexivity
132     |apply le_O_n
133     ]
134   |simplify.rewrite > H1
135     [simplify.apply H.
136      intros.
137      apply H1.
138      apply le_S.
139      assumption
140     |apply le_n
141     ]
142   ]
143 qed.
144
145 theorem f_max_true : \forall f:nat \to bool. \forall n:nat.
146 (\exists i:nat. le i n \land f i = true) \to f (max n f) = true. 
147 intros 2.
148 elim n.elim H.elim H1.generalize in match H3.
149 apply (le_n_O_elim a H2).intro.simplify.rewrite > H4.
150 simplify.assumption.
151 simplify.
152 apply (bool_ind (\lambda b:bool.
153 (f (S n1) = b) \to (f (match b in bool with
154 [ true \Rightarrow (S n1)
155 | false  \Rightarrow (max n1 f)])) = true)).
156 simplify.intro.assumption.
157 simplify.intro.apply H.
158 elim H1.elim H3.generalize in match H5.
159 apply (le_n_Sm_elim a n1 H4).
160 intros.
161 apply (ex_intro nat ? a).
162 split.apply le_S_S_to_le.assumption.assumption.
163 intros.apply False_ind.apply not_eq_true_false.
164 rewrite < H2.rewrite < H7.rewrite > H6. reflexivity.
165 reflexivity.
166 qed.
167
168 theorem exists_forall_le:\forall f,n. 
169 (\exists i. i \le n \land f i = true) \lor
170 (\forall i. i \le n \to f i = false).
171 intros.
172 elim n
173   [apply (bool_elim ? (f O));intro
174     [left.apply (ex_intro ? ? O).
175      split[apply le_n|assumption]
176     |right.intros.
177      apply (le_n_O_elim ? H1).
178      assumption
179     ]
180   |elim H
181     [elim H1.elim H2.
182      left.apply (ex_intro ? ? a).
183      split[apply le_S.assumption|assumption]
184     |apply (bool_elim ? (f (S n1)));intro
185       [left.apply (ex_intro ? ? (S n1)).
186        split[apply le_n|assumption]
187       |right.intros.
188        elim (le_to_or_lt_eq ? ? H3)
189         [apply H1.
190          apply le_S_S_to_le.
191          apply H4
192         |rewrite > H4.
193          assumption
194         ]
195       ]
196     ]
197   ]
198 qed.
199      
200 theorem exists_max_forall_false:\forall f,n. 
201 ((\exists i. i \le n \land f i = true) \land (f (max n f) = true))\lor
202 ((\forall i. i \le n \to f i = false) \land (max n f) = O).
203 intros.
204 elim (exists_forall_le f n)
205   [left.split
206     [assumption
207     |apply f_max_true.assumption
208     ]
209   |right.split
210     [assumption
211     |apply max_O.assumption
212     ]
213   ]
214 qed.
215
216 theorem false_to_lt_max: \forall f,n,m.O < n \to
217 f n = false \to max m f \le n \to max m f < n.
218 intros.
219 elim (le_to_or_lt_eq ? ? H2)
220   [assumption
221   |elim (exists_max_forall_false f m)
222     [elim H4.
223      apply False_ind.
224      apply not_eq_true_false.
225      rewrite < H6.
226      rewrite > H3.
227      assumption
228     |elim H4.
229      rewrite > H6.
230      assumption
231     ]
232   ]
233 qed.
234
235 theorem lt_max_to_false : \forall f:nat \to bool. 
236 \forall n,m:nat. (max n f) < m \to m \leq n \to f m = false.
237 intros 2.
238 elim n.absurd (le m O).assumption.
239 cut (O < m).apply (lt_O_n_elim m Hcut).exact not_le_Sn_O.
240 rewrite < (max_O_f f).assumption.
241 generalize in match H1.
242 elim (max_S_max f n1).
243 elim H3.
244 absurd (m \le S n1).assumption.
245 apply lt_to_not_le.rewrite < H6.assumption.
246 elim H3.
247 apply (le_n_Sm_elim m n1 H2).
248 intro.
249 apply H.rewrite < H6.assumption.
250 apply le_S_S_to_le.assumption.
251 intro.rewrite > H7.assumption.
252 qed.
253
254 theorem f_false_to_le_max: \forall f,n,p. (∃i:nat.i≤n∧f i=true) \to
255 (\forall m. p < m \to f m = false)
256 \to max n f \le p.
257 intros.
258 apply not_lt_to_le.intro.
259 apply not_eq_true_false.
260 rewrite < (H1 ? H2).
261 apply sym_eq.
262 apply f_max_true.
263 assumption.
264 qed.
265
266 definition max_spec \def \lambda f:nat \to bool.\lambda n,m: nat.
267 m \le n \land (f m)=true \land (\forall i. m < i \to i \le n \to (f i = false)).
268
269 theorem max_spec_to_max: \forall f:nat \to bool.\forall n,m:nat.
270 max_spec f n m \to max n f = m.
271 intros 2.
272 elim n
273   [elim H.elim H1.apply (le_n_O_elim ? H3).
274    apply max_O_f
275   |elim H1.
276    elim (max_S_max f n1)
277     [elim H4.
278      rewrite > H6.
279      apply le_to_le_to_eq
280       [apply not_lt_to_le.
281        unfold Not.intro.
282        apply not_eq_true_false.
283        rewrite < H5.
284        apply H3
285         [assumption|apply le_n]
286       |elim H2.assumption
287       ]
288     |elim H4.
289      rewrite > H6.
290      apply H.
291      elim H2.
292      split
293       [split
294         [elim (le_to_or_lt_eq ? ? H7)
295           [apply le_S_S_to_le.assumption
296           |apply False_ind.
297            apply not_eq_true_false.
298            rewrite < H8.
299            rewrite > H9.
300            assumption
301           ]
302         |assumption
303         ]
304       |intros.
305        apply H3
306         [assumption|apply le_S.assumption]
307       ]
308     ]
309   ]
310 qed.
311  
312 let rec min_aux off n f \def
313   match f n with 
314   [ true \Rightarrow n
315   | false \Rightarrow 
316       match off with
317       [ O \Rightarrow n
318       | (S p) \Rightarrow min_aux p (S n) f]].
319
320 definition min : nat \to (nat \to bool) \to nat \def
321 \lambda n.\lambda f. min_aux n O f.
322
323 theorem min_aux_O_f: \forall f:nat \to bool. \forall i :nat.
324 min_aux O i f = i.
325 intros.simplify.
326 elim (f i).reflexivity.
327 simplify.reflexivity.
328 qed.
329
330 theorem min_O_f : \forall f:nat \to bool.
331 min O f = O.
332 intro.apply (min_aux_O_f f O).
333 qed.
334
335 theorem min_aux_S : \forall f: nat \to bool. \forall i,n:nat.
336 ((f n) = true \land min_aux (S i) n f = n) \lor 
337 ((f n) = false \land min_aux (S i) n f = min_aux i (S n) f).
338 intros.simplify.elim (f n).
339 simplify.left.split.reflexivity.reflexivity.
340 simplify.right.split.reflexivity.reflexivity.
341 qed.
342
343 theorem f_min_aux_true: \forall f:nat \to bool. \forall off,m:nat.
344 (\exists i. le m i \land le i (off + m) \land f i = true) \to
345 f (min_aux off m f) = true. 
346 intros 2.
347 elim off.elim H.elim H1.elim H2.
348 cut (a = m).
349 rewrite > (min_aux_O_f f).rewrite < Hcut.assumption.
350 apply (antisym_le a m).assumption.assumption.
351 simplify.
352 apply (bool_ind (\lambda b:bool.
353 (f m = b) \to (f (match b in bool with
354 [ true \Rightarrow m
355 | false  \Rightarrow (min_aux n (S m) f)])) = true)).
356 intro; assumption.
357 intro. apply H.
358 elim H1.elim H3.elim H4.
359 elim (le_to_or_lt_eq ? a H6). 
360 apply (ex_intro nat ? a).
361 split.split.
362 assumption.
363 rewrite < plus_n_Sm; assumption.
364 assumption.
365 absurd (f a = false).rewrite < H8.assumption.
366 rewrite > H5.
367 apply not_eq_true_false.
368 reflexivity.
369 qed.
370
371 theorem lt_min_aux_to_false : \forall f:nat \to bool. 
372 \forall n,off,m:nat. n \leq m \to m < (min_aux off n f) \to f m = false.
373 intros 3.
374 generalize in match n; clear n.
375 elim off.absurd (le n1 m).assumption.
376 apply lt_to_not_le.rewrite < (min_aux_O_f f n1).assumption.
377 elim (le_to_or_lt_eq ? ? H1);
378  [ unfold lt in H3;
379    apply (H (S n1));
380    [ assumption
381    | elim (min_aux_S f n n1);
382      [ elim H4;
383        elim (not_le_Sn_n n1);
384        rewrite > H6 in H2;
385        apply (lt_to_le (S n1) n1 ?).
386        apply (le_to_lt_to_lt (S n1) m n1 ? ?);
387         [apply (H3).
388         |apply (H2).
389         ]
390      | elim H4;
391        rewrite < H6;
392        assumption
393      ]
394    ]
395  | rewrite < H3 in H2 ⊢ %.
396    elim (min_aux_S f n n1);
397     [ elim H4;
398       rewrite > H6 in H2;
399       unfold lt in H2;
400       elim (not_le_Sn_n ? H2)
401     | elim H4;
402       assumption
403     ]
404  ]
405 qed.
406
407
408 lemma le_min_aux : \forall f:nat \to bool. 
409 \forall n,off:nat. n \leq (min_aux off n f).
410 intros 3.
411 generalize in match n. clear n.
412 elim off.
413 rewrite > (min_aux_O_f f n1).apply le_n.
414 elim (min_aux_S f n n1).
415 elim H1.rewrite > H3.apply le_n.
416 elim H1.rewrite > H3.
417 apply (transitive_le ? (S n1));
418  [ apply le_n_Sn
419  | apply (H (S n1))
420  ]
421 qed.
422
423 theorem le_min_aux_r : \forall f:nat \to bool. 
424 \forall n,off:nat. (min_aux off n f) \le n+off.
425 intros.
426 generalize in match n. clear n.
427 elim off.simplify.
428 elim (f n1).simplify.rewrite < plus_n_O.apply le_n.
429 simplify.rewrite < plus_n_O.apply le_n.
430 simplify.elim (f n1).
431 simplify.
432 apply (le_plus_n_r (S n) n1).
433 simplify.rewrite < plus_n_Sm.
434 apply (H (S n1)).
435 qed.