1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| A.Asperti, C.Sacerdoti Coen, *)
8 (* ||A|| E.Tassi, S.Zacchiroli *)
10 (* \ / Matita is distributed under the terms of the *)
11 (* v GNU Lesser General Public License Version 2.1 *)
13 (**************************************************************************)
15 set "baseuri" "cic:/matita/nat/minimization".
17 include "nat/minus.ma".
25 | (S j) \Rightarrow max j f ]].
27 theorem max_O_f : \forall f: nat \to bool. max O f = O.
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.
42 theorem le_max_n : \forall f: nat \to bool. \forall n:nat.
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.
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.
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)).
58 apply le_S.apply le_max_n.
59 elim H3.rewrite > H5.apply le_n.
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).
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.
75 theorem max_f_g: \forall f,g,n. (\forall i. i \le n \to f i = g i) \to
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
103 elim (f O);apply le_O_n
105 apply (bool_elim ? (f (S n1)));intro
106 [rewrite > (H1 (S n1) ? H2)
118 [apply le_S.assumption
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.
130 [simplify.rewrite > H
134 |simplify.rewrite > H1
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.
148 elim n.elim H.elim H1.generalize in match H3.
149 apply (le_n_O_elim a H2).intro.simplify.rewrite > H4.
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).
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.
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).
173 [apply (bool_elim ? (f O));intro
174 [left.apply (ex_intro ? ? O).
175 split[apply le_n|assumption]
177 apply (le_n_O_elim ? H1).
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]
188 elim (le_to_or_lt_eq ? ? H3)
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).
204 elim (exists_forall_le f n)
207 |apply f_max_true.assumption
211 |apply max_O.assumption
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.
219 elim (le_to_or_lt_eq ? ? H2)
221 |elim (exists_max_forall_false f m)
224 apply not_eq_true_false.
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.
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).
244 absurd (m \le S n1).assumption.
245 apply lt_to_not_le.rewrite < H6.assumption.
247 apply (le_n_Sm_elim m n1 H2).
249 apply H.rewrite < H6.assumption.
250 apply le_S_S_to_le.assumption.
251 intro.rewrite > H7.assumption.
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)
258 apply not_lt_to_le.intro.
259 apply not_eq_true_false.
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)).
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.
273 [elim H.elim H1.apply (le_n_O_elim ? H3).
276 elim (max_S_max f n1)
282 apply not_eq_true_false.
285 [assumption|apply le_n]
294 [elim (le_to_or_lt_eq ? ? H7)
295 [apply le_S_S_to_le.assumption
297 apply not_eq_true_false.
306 [assumption|apply le_S.assumption]
312 let rec min_aux off n f \def
318 | (S p) \Rightarrow min_aux p (S n) f]].
320 definition min : nat \to (nat \to bool) \to nat \def
321 \lambda n.\lambda f. min_aux n O f.
323 theorem min_aux_O_f: \forall f:nat \to bool. \forall i :nat.
326 elim (f i).reflexivity.
327 simplify.reflexivity.
330 theorem min_O_f : \forall f:nat \to bool.
332 intro.apply (min_aux_O_f f O).
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.
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.
347 elim off.elim H.elim H1.elim H2.
349 rewrite > (min_aux_O_f f).rewrite < Hcut.assumption.
350 apply (antisym_le a m).assumption.assumption.
352 apply (bool_ind (\lambda b:bool.
353 (f m = b) \to (f (match b in bool with
355 | false \Rightarrow (min_aux n (S m) f)])) = true)).
358 elim H1.elim H3.elim H4.
359 elim (le_to_or_lt_eq ? a H6).
360 apply (ex_intro nat ? a).
363 rewrite < plus_n_Sm; assumption.
365 absurd (f a = false).rewrite < H8.assumption.
367 apply not_eq_true_false.
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.
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);
381 | elim (min_aux_S f n n1);
383 elim (not_le_Sn_n n1);
385 apply (lt_to_le (S n1) n1 ?).
386 apply (le_to_lt_to_lt (S n1) m n1 ? ?);
395 | rewrite < H3 in H2 ⊢ %.
396 elim (min_aux_S f n n1);
400 elim (not_le_Sn_n ? H2)
408 lemma le_min_aux : \forall f:nat \to bool.
409 \forall n,off:nat. n \leq (min_aux off n f).
411 generalize in match n. clear n.
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));
423 theorem le_min_aux_r : \forall f:nat \to bool.
424 \forall n,off:nat. (min_aux off n f) \le n+off.
426 generalize in match n. clear n.
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).
432 apply (le_plus_n_r (S n) n1).
433 simplify.rewrite < plus_n_Sm.