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 (* \ / This file is distributed under the terms of the *)
11 (* v GNU Lesser General Public License Version 2.1 *)
13 (**************************************************************************)
16 include "higher_order_defs/ordering.ma".
19 inductive le (n:nat) : nat \to Prop \def
21 | le_S : \forall m:nat. le n m \to le n (S m).
23 interpretation "natural 'less or equal to'" 'leq x y = (cic:/matita/nat/orders/le.ind#xpointer(1/1) x y).
25 interpretation "natural 'neither less nor equal to'" 'nleq x y =
26 (cic:/matita/logic/connectives/Not.con
27 (cic:/matita/nat/orders/le.ind#xpointer(1/1) x y)).
29 definition lt: nat \to nat \to Prop \def
30 \lambda n,m:nat.(S n) \leq m.
32 interpretation "natural 'less than'" 'lt x y = (cic:/matita/nat/orders/lt.con x y).
34 interpretation "natural 'not less than'" 'nless x y =
35 (cic:/matita/logic/connectives/Not.con (cic:/matita/nat/orders/lt.con x y)).
37 definition ge: nat \to nat \to Prop \def
38 \lambda n,m:nat.m \leq n.
40 interpretation "natural 'greater or equal to'" 'geq x y = (cic:/matita/nat/orders/ge.con x y).
42 definition gt: nat \to nat \to Prop \def
45 interpretation "natural 'greater than'" 'gt x y = (cic:/matita/nat/orders/gt.con x y).
47 interpretation "natural 'not greater than'" 'ngtr x y =
48 (cic:/matita/logic/connectives/Not.con (cic:/matita/nat/orders/gt.con x y)).
50 theorem transitive_le : transitive nat le.
51 unfold transitive.intros.elim H1.
53 apply le_S.assumption.
56 theorem trans_le: \forall n,m,p:nat. n \leq m \to m \leq p \to n \leq p
59 theorem transitive_lt: transitive nat lt.
60 unfold transitive.unfold lt.intros.elim H1.
61 apply le_S. assumption.
62 apply le_S.assumption.
65 theorem trans_lt: \forall n,m,p:nat. lt n m \to lt m p \to lt n p
68 theorem le_S_S: \forall n,m:nat. n \leq m \to S n \leq S m.
71 apply le_S.assumption.
74 theorem le_O_n : \forall n:nat. O \leq n.
80 theorem le_n_Sn : \forall n:nat. n \leq S n.
81 intros. apply le_S.apply le_n.
84 theorem le_pred_n : \forall n:nat. pred n \leq n.
87 simplify.apply le_n_Sn.
90 theorem le_S_S_to_le : \forall n,m:nat. S n \leq S m \to n \leq m.
91 intros.change with (pred (S n) \leq pred (S m)).
92 elim H.apply le_n.apply (trans_le ? (pred n1)).assumption.
96 theorem lt_S_S_to_lt: \forall n,m.
98 intros. apply le_S_S_to_le. assumption.
101 theorem lt_to_lt_S_S: ∀n,m. n < m → S n < S m.
104 apply (le_S_S ? ? H).
107 theorem leS_to_not_zero : \forall n,m:nat. S n \leq m \to not_zero m.
108 intros.elim H.exact I.exact I.
112 theorem not_le_Sn_O: \forall n:nat. S n \nleq O.
113 intros.unfold Not.simplify.intros.apply (leS_to_not_zero ? ? H).
116 theorem not_le_Sn_n: \forall n:nat. S n \nleq n.
117 intros.elim n.apply not_le_Sn_O.unfold Not.simplify.intros.cut (S n1 \leq n1).
119 apply le_S_S_to_le.assumption.
122 theorem lt_pred: \forall n,m.
123 O < n \to n < m \to pred n < pred m.
125 [intros.apply False_ind.apply (not_le_Sn_O ? H)
126 |intros.apply False_ind.apply (not_le_Sn_O ? H1)
127 |intros.simplify.unfold.apply le_S_S_to_le.assumption
131 theorem S_pred: \forall n:nat.lt O n \to eq nat n (S (pred n)).
132 intro.elim n.apply False_ind.exact (not_le_Sn_O O H).
133 apply eq_f.apply pred_Sn.
136 theorem le_pred_to_le:
137 ∀n,m. O < m → pred n ≤ pred m → n ≤ m.
142 rewrite > (S_pred m);
150 theorem le_to_le_pred:
151 ∀n,m. n ≤ m → pred n ≤ pred m.
157 generalize in match H1;
160 [ elim (not_le_Sn_O ? H1)
169 theorem le_to_or_lt_eq : \forall n,m:nat.
170 n \leq m \to n < m \lor n = m.
173 left.unfold lt.apply le_S_S.assumption.
176 theorem Not_lt_n_n: ∀n. n ≮ n.
181 apply (not_le_Sn_n ? H).
185 theorem lt_to_not_eq : \forall n,m:nat. n<m \to n \neq m.
186 unfold Not.intros.cut ((le (S n) m) \to False).
187 apply Hcut.assumption.rewrite < H1.
192 theorem eq_to_not_lt: \forall a,b:nat.
198 apply (lt_to_not_eq b b)
204 theorem lt_n_m_to_not_lt_m_Sn: ∀n,m. n < m → m ≮ S n.
210 generalize in match (le_S_S ? ? H);
212 generalize in match (transitive_le ? ? ? H2 H1);
214 apply (not_le_Sn_n ? H3).
218 theorem lt_to_le : \forall n,m:nat. n<m \to n \leq m.
219 simplify.intros.unfold lt in H.elim H.
220 apply le_S. apply le_n.
221 apply le_S. assumption.
224 theorem lt_S_to_le : \forall n,m:nat. n < S m \to n \leq m.
226 apply le_S_S_to_le.assumption.
229 theorem not_le_to_lt: \forall n,m:nat. n \nleq m \to m<n.
231 apply (nat_elim2 (\lambda n,m.n \nleq m \to m<n)).
232 intros.apply (absurd (O \leq n1)).apply le_O_n.assumption.
233 unfold Not.unfold lt.intros.apply le_S_S.apply le_O_n.
234 unfold Not.unfold lt.intros.apply le_S_S.apply H.intros.apply H1.apply le_S_S.
238 theorem lt_to_not_le: \forall n,m:nat. n<m \to m \nleq n.
239 unfold Not.unfold lt.intros 3.elim H.
240 apply (not_le_Sn_n n H1).
241 apply H2.apply lt_to_le. apply H3.
244 theorem not_lt_to_le: \forall n,m:nat. Not (lt n m) \to le m n.
247 apply not_le_to_lt.exact H.
250 theorem le_to_not_lt: \forall n,m:nat. le n m \to Not (lt m n).
251 intros.unfold Not.unfold lt.
252 apply lt_to_not_le.unfold lt.
253 apply le_S_S.assumption.
256 theorem not_eq_to_le_to_lt: ∀n,m. n≠m → n≤m → n<m.
258 elim (le_to_or_lt_eq ? ? H1);
265 theorem le_n_O_to_eq : \forall n:nat. n \leq O \to O=n.
266 intro.elim n.reflexivity.
269 [2: apply H1 | skip].
272 theorem le_n_O_elim: \forall n:nat.n \leq O \to \forall P: nat \to Prop.
277 apply (not_le_Sn_O ? H1).
280 theorem le_n_Sm_elim : \forall n,m:nat.n \leq S m \to
281 \forall P:Prop. (S n \leq S m \to P) \to (n=S m \to P) \to P.
283 apply H2.reflexivity.
284 apply H3. apply le_S_S. assumption.
288 lemma le_to_le_to_eq: \forall n,m. n \le m \to m \le n \to n = m.
290 [intros.apply le_n_O_to_eq.assumption
291 |intros.apply sym_eq.apply le_n_O_to_eq.assumption
292 |intros.apply eq_f.apply H
293 [apply le_S_S_to_le.assumption
294 |apply le_S_S_to_le.assumption
299 (* lt and le trans *)
300 theorem lt_O_S : \forall n:nat. O < S n.
301 intro. unfold. apply le_S_S. apply le_O_n.
304 theorem lt_to_le_to_lt: \forall n,m,p:nat. lt n m \to le m p \to lt n p.
306 assumption.unfold lt.apply le_S.assumption.
309 theorem le_to_lt_to_lt: \forall n,m,p:nat. le n m \to lt m p \to lt n p.
311 assumption.apply H2.unfold lt.
312 apply lt_to_le.assumption.
315 theorem lt_S_to_lt: \forall n,m. S n < m \to n < m.
317 apply (trans_lt ? (S n))
318 [apply le_n|assumption]
321 theorem ltn_to_ltO: \forall n,m:nat. lt n m \to lt O m.
322 intros.apply (le_to_lt_to_lt O n).
323 apply le_O_n.assumption.
326 theorem lt_SO_n_to_lt_O_pred_n: \forall n:nat.
327 (S O) \lt n \to O \lt (pred n).
329 apply (ltn_to_ltO (pred (S O)) (pred n) ?).
330 apply (lt_pred (S O) n);
336 theorem lt_O_n_elim: \forall n:nat. lt O n \to
337 \forall P:nat\to Prop. (\forall m:nat.P (S m)) \to P n.
338 intro.elim n.apply False_ind.exact (not_le_Sn_O O H).
342 (* other abstract properties *)
343 theorem antisymmetric_le : antisymmetric nat le.
344 unfold antisymmetric.intros 2.
345 apply (nat_elim2 (\lambda n,m.(n \leq m \to m \leq n \to n=m))).
346 intros.apply le_n_O_to_eq.assumption.
347 intros.apply False_ind.apply (not_le_Sn_O ? H).
348 intros.apply eq_f.apply H.
349 apply le_S_S_to_le.assumption.
350 apply le_S_S_to_le.assumption.
353 theorem antisym_le: \forall n,m:nat. n \leq m \to m \leq n \to n=m
354 \def antisymmetric_le.
356 theorem le_n_m_to_lt_m_Sn_to_eq_n_m: ∀n,m. n ≤ m → m < S n → n=m.
359 generalize in match (le_S_S_to_le ? ? H1);
365 theorem decidable_le: \forall n,m:nat. decidable (n \leq m).
367 apply (nat_elim2 (\lambda n,m.decidable (n \leq m))).
368 intros.unfold decidable.left.apply le_O_n.
369 intros.unfold decidable.right.exact (not_le_Sn_O n1).
370 intros 2.unfold decidable.intro.elim H.
371 left.apply le_S_S.assumption.
372 right.unfold Not.intro.apply H1.apply le_S_S_to_le.assumption.
375 theorem decidable_lt: \forall n,m:nat. decidable (n < m).
376 intros.exact (decidable_le (S n) m).
379 (* well founded induction principles *)
381 theorem nat_elim1 : \forall n:nat.\forall P:nat \to Prop.
382 (\forall m.(\forall p. (p \lt m) \to P p) \to P m) \to P n.
383 intros.cut (\forall q:nat. q \le n \to P q).
384 apply (Hcut n).apply le_n.
385 elim n.apply (le_n_O_elim q H1).
387 intros.apply False_ind.apply (not_le_Sn_O p H2).
388 apply H.intros.apply H1.
390 apply lt_S_to_le.assumption.
391 apply (lt_to_le_to_lt p q (S n1) H3 H2).
394 (* some properties of functions *)
396 definition increasing \def \lambda f:nat \to nat.
397 \forall n:nat. f n < f (S n).
399 theorem increasing_to_monotonic: \forall f:nat \to nat.
400 increasing f \to monotonic nat lt f.
401 unfold monotonic.unfold lt.unfold increasing.unfold lt.intros.elim H1.apply H.
402 apply (trans_le ? (f n1)).
403 assumption.apply (trans_le ? (S (f n1))).
408 theorem le_n_fn: \forall f:nat \to nat. (increasing f)
409 \to \forall n:nat. n \le (f n).
412 apply (trans_le ? (S (f n1))).
413 apply le_S_S.apply H1.
414 simplify in H. unfold increasing in H.unfold lt in H.apply H.
417 theorem increasing_to_le: \forall f:nat \to nat. (increasing f)
418 \to \forall m:nat. \exists i. m \le (f i).
420 apply (ex_intro ? ? O).apply le_O_n.
422 apply (ex_intro ? ? (S a)).
423 apply (trans_le ? (S (f a))).
424 apply le_S_S.assumption.
425 simplify in H.unfold increasing in H.unfold lt in H.
429 theorem increasing_to_le2: \forall f:nat \to nat. (increasing f)
430 \to \forall m:nat. (f O) \le m \to
431 \exists i. (f i) \le m \land m <(f (S i)).
433 apply (ex_intro ? ? O).
434 split.apply le_n.apply H.
436 cut ((S n1) < (f (S a)) \lor (S n1) = (f (S a))).
438 apply (ex_intro ? ? a).
439 split.apply le_S. assumption.assumption.
440 apply (ex_intro ? ? (S a)).
441 split.rewrite < H7.apply le_n.
444 apply le_to_or_lt_eq.apply H6.