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 include "nat/minus.ma".
23 | (S j) \Rightarrow max j f ]].
25 theorem max_O_f : \forall f: nat \to bool. max O f = O.
32 theorem max_S_max : \forall f: nat \to bool. \forall n:nat.
33 (f (S n) = true \land max (S n) f = (S n)) \lor
34 (f (S n) = false \land max (S n) f = max n f).
35 intros.simplify.elim (f (S n)).
36 simplify.left.split.reflexivity.reflexivity.
37 simplify.right.split.reflexivity.reflexivity.
40 theorem le_max_n : \forall f: nat \to bool. \forall n:nat.
42 intros.elim n.rewrite > max_O_f.apply le_n.
43 simplify.elim (f (S n1)).simplify.apply le_n.
44 simplify.apply le_S.assumption.
47 theorem le_to_le_max : \forall f: nat \to bool. \forall n,m:nat.
48 n\le m \to max n f \le max m f.
51 apply (trans_le ? (max n1 f)).apply H2.
52 cut ((f (S n1) = true \land max (S n1) f = (S n1)) \lor
53 (f (S n1) = false \land max (S n1) f = max n1 f)).
56 apply le_S.apply le_max_n.
57 elim H3.rewrite > H5.apply le_n.
61 theorem f_m_to_le_max: \forall f: nat \to bool. \forall n,m:nat.
62 m\le n \to f m = true \to m \le max n f.
63 intros 3.elim n.apply (le_n_O_elim m H).
65 apply (le_n_Sm_elim m n1 H1).
66 intro.apply (trans_le ? (max n1 f)).
67 apply H.apply le_S_S_to_le.assumption.assumption.
68 apply le_to_le_max.apply le_n_Sn.
69 intro.simplify.rewrite < H3.
70 rewrite > H2.simplify.apply le_n.
73 theorem max_f_g: \forall f,g,n. (\forall i. i \le n \to f i = g i) \to
96 theorem le_max_f_max_g: \forall f,g,n. (\forall i. i \le n \to f i = true \to g i =true) \to
101 elim (f O);apply le_O_n
103 apply (bool_elim ? (f (S n1)));intro
104 [rewrite > (H1 (S n1) ? H2)
116 [apply le_S.assumption
125 theorem max_O : \forall f:nat \to bool. \forall n:nat.
126 (\forall i:nat. le i n \to f i = false) \to max n f = O.
128 [simplify.rewrite > H
132 |simplify.rewrite > H1
143 theorem f_max_true : \forall f:nat \to bool. \forall n:nat.
144 (\exists i:nat. le i n \land f i = true) \to f (max n f) = true.
146 elim n.elim H.elim H1.generalize in match H3.
147 apply (le_n_O_elim a H2).intro.simplify.rewrite > H4.
150 apply (bool_ind (\lambda b:bool.
151 (f (S n1) = b) \to (f (match b in bool with
152 [ true \Rightarrow (S n1)
153 | false \Rightarrow (max n1 f)])) = true)).
154 simplify.intro.assumption.
155 simplify.intro.apply H.
156 elim H1.elim H3.generalize in match H5.
157 apply (le_n_Sm_elim a n1 H4).
159 apply (ex_intro nat ? a).
160 split.apply le_S_S_to_le.assumption.assumption.
161 intros.apply False_ind.apply not_eq_true_false.
162 rewrite < H2.rewrite < H7.rewrite > H6. reflexivity.
166 theorem exists_forall_le:\forall f,n.
167 (\exists i. i \le n \land f i = true) \lor
168 (\forall i. i \le n \to f i = false).
171 [apply (bool_elim ? (f O));intro
172 [left.apply (ex_intro ? ? O).
173 split[apply le_n|assumption]
175 apply (le_n_O_elim ? H1).
180 left.apply (ex_intro ? ? a).
181 split[apply le_S.assumption|assumption]
182 |apply (bool_elim ? (f (S n1)));intro
183 [left.apply (ex_intro ? ? (S n1)).
184 split[apply le_n|assumption]
186 elim (le_to_or_lt_eq ? ? H3)
198 theorem exists_max_forall_false:\forall f,n.
199 ((\exists i. i \le n \land f i = true) \land (f (max n f) = true))\lor
200 ((\forall i. i \le n \to f i = false) \land (max n f) = O).
202 elim (exists_forall_le f n)
205 |apply f_max_true.assumption
209 |apply max_O.assumption
214 theorem false_to_lt_max: \forall f,n,m.O < n \to
215 f n = false \to max m f \le n \to max m f < n.
217 elim (le_to_or_lt_eq ? ? H2)
219 |elim (exists_max_forall_false f m)
222 apply not_eq_true_false.
233 theorem lt_max_to_false : \forall f:nat \to bool.
234 \forall n,m:nat. (max n f) < m \to m \leq n \to f m = false.
236 elim n.absurd (le m O).assumption.
237 cut (O < m).apply (lt_O_n_elim m Hcut).exact not_le_Sn_O.
238 rewrite < (max_O_f f).assumption.
239 elim (max_S_max f n1) in H1 ⊢ %.
241 absurd (m \le S n1).assumption.
242 apply lt_to_not_le.rewrite < H5.assumption.
244 apply (le_n_Sm_elim m n1 H2).
246 apply H.rewrite < H5.assumption.
247 apply le_S_S_to_le.assumption.
248 intro.rewrite > H6.assumption.
251 theorem f_false_to_le_max: \forall f,n,p. (∃i:nat.i≤n∧f i=true) \to
252 (\forall m. p < m \to f m = false)
255 apply not_lt_to_le.intro.
256 apply not_eq_true_false.
263 definition max_spec \def \lambda f:nat \to bool.\lambda n,m: nat.
264 m \le n \land (f m)=true \land (\forall i. m < i \to i \le n \to (f i = false)).
266 theorem max_spec_to_max: \forall f:nat \to bool.\forall n,m:nat.
267 max_spec f n m \to max n f = m.
270 [elim H.elim H1.apply (le_n_O_elim ? H3).
273 elim (max_S_max f n1)
279 apply not_eq_true_false.
282 [assumption|apply le_n]
291 [elim (le_to_or_lt_eq ? ? H7)
292 [apply le_S_S_to_le.assumption
294 apply not_eq_true_false.
303 [assumption|apply le_S.assumption]
309 let rec min_aux off n f \def
315 | (S p) \Rightarrow min_aux p (S n) f]].
317 definition min : nat \to (nat \to bool) \to nat \def
318 \lambda n.\lambda f. min_aux n O f.
320 theorem min_aux_O_f: \forall f:nat \to bool. \forall i :nat.
323 elim (f i).reflexivity.
324 simplify.reflexivity.
327 theorem min_O_f : \forall f:nat \to bool.
329 intro.apply (min_aux_O_f f O).
332 theorem min_aux_S : \forall f: nat \to bool. \forall i,n:nat.
333 ((f n) = true \land min_aux (S i) n f = n) \lor
334 ((f n) = false \land min_aux (S i) n f = min_aux i (S n) f).
335 intros.simplify.elim (f n).
336 simplify.left.split.reflexivity.reflexivity.
337 simplify.right.split.reflexivity.reflexivity.
340 theorem f_min_aux_true: \forall f:nat \to bool. \forall off,m:nat.
341 (\exists i. le m i \land le i (off + m) \land f i = true) \to
342 f (min_aux off m f) = true.
344 elim off.elim H.elim H1.elim H2.
346 rewrite > (min_aux_O_f f).rewrite < Hcut.assumption.
347 apply (antisym_le a m).assumption.assumption.
349 apply (bool_ind (\lambda b:bool.
350 (f m = b) \to (f (match b in bool with
352 | false \Rightarrow (min_aux n (S m) f)])) = true)).
355 elim H1.elim H3.elim H4.
356 elim (le_to_or_lt_eq ? a H6).
357 apply (ex_intro nat ? a).
360 rewrite < plus_n_Sm; assumption.
362 absurd (f a = false).rewrite < H8.assumption.
364 apply not_eq_true_false.
368 theorem f_min_true: \forall f:nat \to bool. \forall m:nat.
369 (\exists i. le i m \land f i = true) \to
372 apply f_min_aux_true.
373 elim H.clear H.elim H1.clear H1.
374 apply (ex_intro ? ? a).
378 |rewrite < plus_n_O.assumption
384 theorem lt_min_aux_to_false : \forall f:nat \to bool.
385 \forall n,off,m:nat. n \leq m \to m < (min_aux off n f) \to f m = false.
387 generalize in match n; clear n;
388 elim off.absurd (le n1 m).assumption.
389 apply lt_to_not_le.rewrite < (min_aux_O_f f n1).assumption.
390 elim (le_to_or_lt_eq ? ? H1);
394 | elim (min_aux_S f n n1);
396 elim (not_le_Sn_n n1);
398 apply (lt_to_le (S n1) n1 ?).
399 apply (le_to_lt_to_lt (S n1) m n1 ? ?);
408 | rewrite < H3 in H2 ⊢ %.
409 elim (min_aux_S f n n1);
413 elim (not_le_Sn_n ? H2)
421 lemma le_min_aux : \forall f:nat \to bool.
422 \forall n,off:nat. n \leq (min_aux off n f).
425 rewrite > (min_aux_O_f f n1).apply le_n.
426 elim (min_aux_S f n n1).
427 elim H1.rewrite > H3.apply le_n.
428 elim H1.rewrite > H3.
429 apply (transitive_le ? (S n1));
435 theorem le_min_aux_r : \forall f:nat \to bool.
436 \forall n,off:nat. (min_aux off n f) \le n+off.
438 elim off in n ⊢ %.simplify.
439 elim (f n1).simplify.rewrite < plus_n_O.apply le_n.
440 simplify.rewrite < plus_n_O.apply le_n.
441 simplify.elim (f n1).
443 apply (le_plus_n_r (S n) n1).
444 simplify.rewrite < plus_n_Sm.