]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/nlibrary/arithmetics/nat.ma
fe9ef8267ffc064e1d50f9e3283b6f0c5e0c3660
[helm.git] / helm / software / matita / nlibrary / arithmetics / nat.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 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU Lesser General Public License Version 2.1         *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 (* include "higher_order_defs/functions.ma". *)
16 include "hints_declaration.ma".
17 include "basics/functions.ma".
18 include "basics/eq.ma".
19
20 ntheorem foo: ∀A:Type.∀a,b:A.∀f:A→A.∀g:A→A→A.
21 (∀x,y.f (g x y) = x) → ∀x. g (f a) x = b → f a = f b.
22 //; nqed.
23
24 ninductive nat : Type[0] ≝
25   | O : nat
26   | S : nat → nat.
27   
28 interpretation "Natural numbers" 'N = nat.
29
30 alias num (instance 0) = "nnatural number".
31
32 (*
33 nrecord pos : Type ≝
34  {n:>nat; is_pos: n ≠ 0}.
35
36 ncoercion nat_to_pos: ∀n:nat. n ≠0 →pos ≝ mk_pos on 
37 *)
38
39 (* default "natural numbers" cic:/matita/ng/arithmetics/nat/nat.ind.
40 *)
41
42 ndefinition pred ≝
43  λn. match n with [ O ⇒  O | (S p) ⇒ p].
44
45 ntheorem pred_Sn : ∀n. n = pred (S n).
46 //; nqed.
47
48 ntheorem injective_S : injective nat nat S.
49 //; nqed.
50
51 (*
52 ntheorem inj_S : \forall n,m:nat.(S n)=(S m) \to n=m.
53 //. nqed. *)
54
55 ntheorem not_eq_S: ∀n,m:nat. n ≠ m → S n ≠ S m.
56 /3/; nqed.
57
58 ndefinition not_zero: nat → Prop ≝
59  λn: nat. match n with
60   [ O ⇒ False | (S p) ⇒ True ].
61
62 ntheorem not_eq_O_S : ∀n:nat. O ≠ S n.
63 #n; napply nmk; #eqOS; nchange with (not_zero O); nrewrite > eqOS; //.
64 nqed.
65
66 ntheorem not_eq_n_Sn: ∀n:nat. n ≠ S n.
67 #n; nelim n;/2/; nqed.
68
69 ntheorem nat_case:
70  ∀n:nat.∀P:nat → Prop. 
71   (n=O → P O) → (∀m:nat. (n=(S m) → P (S m))) → P n.
72 #n; #P; nelim n; /2/; nqed.
73
74 ntheorem nat_elim2 :
75  ∀R:nat → nat → Prop.
76   (∀n:nat. R O n) 
77   → (∀n:nat. R (S n) O)
78   → (∀n,m:nat. R n m → R (S n) (S m))
79   → ∀n,m:nat. R n m.
80 #R; #ROn; #RSO; #RSS; #n; nelim n;//;
81 #n0; #Rn0m; #m; ncases m;/2/; nqed.
82
83 ntheorem decidable_eq_nat : ∀n,m:nat.decidable (n=m).
84 napply nat_elim2; #n;
85  ##[ ncases n; /2/;
86  ##| /3/;
87  ##| #m; #Hind; ncases Hind; /3/;
88  ##]
89 nqed. 
90
91 (*************************** plus ******************************)
92
93 nlet rec plus n m ≝ 
94  match n with 
95  [ O ⇒ m
96  | S p ⇒ S (plus p m) ].
97
98 interpretation "natural plus" 'plus x y = (plus x y).
99
100 ntheorem plus_O_n: ∀n:nat. n = 0+n.
101 //; nqed.
102
103 (*
104 ntheorem plus_Sn_m: ∀n,m:nat. S (n + m) = S n + m.
105 //; nqed.
106 *)
107
108 ntheorem plus_n_O: ∀n:nat. n = n+0.
109 #n; nelim n; nnormalize; //; nqed.
110
111 ntheorem plus_n_Sm : ∀n,m:nat. S (n+m) = n + S m.
112 #n; nelim n; nnormalize; //; nqed.
113
114 (*
115 ntheorem plus_Sn_m1: ∀n,m:nat. S m + n = n + S m.
116 #n; nelim n; nnormalize; //; nqed.
117 *)
118
119 (* deleterio?
120 ntheorem plus_n_1 : ∀n:nat. S n = n+1.
121 //; nqed.
122 *)
123
124 ntheorem symmetric_plus: symmetric ? plus.
125 #n; nelim n; nnormalize; //; nqed. 
126
127 ntheorem associative_plus : associative nat plus.
128 #n; nelim n; nnormalize; //; nqed.
129
130 ntheorem assoc_plus1: ∀a,b,c. c + (b + a) = b + c + a.
131 //; nqed.
132
133 ntheorem injective_plus_r: ∀n:nat.injective nat nat (λm.n+m).
134 #n; nelim n; nnormalize; /3/; nqed.
135
136 (* ntheorem inj_plus_r: \forall p,n,m:nat. p+n = p+m \to n=m
137 \def injective_plus_r. 
138
139 ntheorem injective_plus_l: ∀m:nat.injective nat nat (λn.n+m).
140 /2/; nqed. *)
141
142 (* ntheorem inj_plus_l: \forall p,n,m:nat. n+p = m+p \to n=m
143 \def injective_plus_l. *)
144
145 (*************************** times *****************************)
146
147 nlet rec times n m ≝ 
148  match n with 
149  [ O ⇒ O
150  | S p ⇒ m+(times p m) ].
151
152 interpretation "natural times" 'times x y = (times x y).
153
154 ntheorem times_Sn_m: ∀n,m:nat. m+n*m = S n*m.
155 //; nqed.
156
157 ntheorem times_O_n: ∀n:nat. O = O*n.
158 //; nqed.
159
160 ntheorem times_n_O: ∀n:nat. O = n*O.
161 #n; nelim n; //; nqed.
162
163 ntheorem times_n_Sm : ∀n,m:nat. n+(n*m) = n*(S m).
164 #n; nelim n; nnormalize; //; nqed.
165
166 ntheorem symmetric_times : symmetric nat times. 
167 #n; nelim n; nnormalize; //; nqed. 
168
169 (* variant sym_times : \forall n,m:nat. n*m = m*n \def
170 symmetric_times. *)
171
172 ntheorem distributive_times_plus : distributive nat times plus.
173 #n; nelim n; nnormalize; //; nqed.
174
175 ntheorem distributive_times_plus_r :
176   ∀a,b,c:nat. (b+c)*a = b*a + c*a.
177 //; nqed. 
178
179 ntheorem associative_times: associative nat times.
180 #n; nelim n; nnormalize; //; nqed.
181
182 nlemma times_times: ∀x,y,z. x*(y*z) = y*(x*z).
183 //; nqed. 
184
185 (* ci servono questi risultati? 
186 ntheorem times_O_to_O: ∀n,m:nat.n*m=O → n=O ∨ m=O.
187 napply nat_elim2; /2/; 
188 #n; #m; #H; nnormalize; #H1; napply False_ind;napply not_eq_O_S;
189 //; nqed.
190   
191 ntheorem times_n_SO : ∀n:nat. n = n * S O.
192 #n; //; nqed.
193
194 ntheorem times_SSO_n : ∀n:nat. n + n = (S(S O)) * n.
195 nnormalize; //; nqed.
196
197 nlemma times_SSO: \forall n.(S(S O))*(S n) = S(S((S(S O))*n)).
198 //; nqed.
199
200 ntheorem or_eq_eq_S: \forall n.\exists m. 
201 n = (S(S O))*m \lor n = S ((S(S O))*m).
202 #n; nelim n;
203   ##[@; /2/;
204   ##|#a; #H; nelim H; #b;#or;nelim or;#aeq;
205     ##[@ b; @ 2; //;
206     ##|@ (S b); @ 1; /2/;
207     ##]
208 nqed.
209 *)
210
211 (******************** ordering relations ************************)
212
213 ninductive le (n:nat) : nat → Prop ≝
214   | le_n : le n n
215   | le_S : ∀ m:nat. le n m → le n (S m).
216
217 interpretation "natural 'less or equal to'" 'leq x y = (le x y).
218
219 interpretation "natural 'neither less nor equal to'" 'nleq x y = (Not (le x y)).
220
221 ndefinition lt: nat → nat → Prop ≝
222 λn,m:nat. S n ≤ m.
223
224 interpretation "natural 'less than'" 'lt x y = (lt x y).
225
226 interpretation "natural 'not less than'" 'nless x y = (Not (lt x y)).
227
228 (* nlemma eq_lt: ∀n,m. (n < m) = (S n ≤ m).
229 //; nqed. *)
230
231 ndefinition ge: nat → nat → Prop ≝
232 λn,m:nat.m ≤ n.
233
234 interpretation "natural 'greater or equal to'" 'geq x y = (ge x y).
235
236 ndefinition gt: nat → nat → Prop ≝
237 λn,m:nat.m<n.
238
239 interpretation "natural 'greater than'" 'gt x y = (gt x y).
240
241 interpretation "natural 'not greater than'" 'ngtr x y = (Not (gt x y)).
242
243 ntheorem transitive_le : transitive nat le.
244 #a; #b; #c; #leab; #lebc;nelim lebc;/2/;
245 nqed.
246
247 (*
248 ntheorem trans_le: \forall n,m,p:nat. n \leq m \to m \leq p \to n \leq p
249 \def transitive_le. *)
250
251
252 ntheorem transitive_lt: transitive nat lt.
253 #a; #b; #c; #ltab; #ltbc;nelim ltbc;/2/;nqed.
254
255 (*
256 theorem trans_lt: \forall n,m,p:nat. lt n m \to lt m p \to lt n p
257 \def transitive_lt. *)
258
259 ntheorem le_S_S: ∀n,m:nat. n ≤ m → S n ≤ S m.
260 #n; #m; #lenm; nelim lenm; /2/; nqed.
261
262 ntheorem le_O_n : ∀n:nat. O ≤ n.
263 #n; nelim n; /2/; nqed.
264
265 ntheorem le_n_Sn : ∀n:nat. n ≤ S n.
266 /2/; nqed.
267
268 ntheorem le_pred_n : ∀n:nat. pred n ≤ n.
269 #n; nelim n; //; nqed.
270
271 (* XXX global problem 
272 nlemma my_trans_le : ∀x,y,z:nat.x ≤ y → y ≤ z → x ≤ z. 
273 napply transitive_le.
274 nqed. *)
275
276 ntheorem monotonic_pred: monotonic ? le pred.
277 #n; #m; #lenm; nelim lenm; /2/;nqed.
278
279 ntheorem le_S_S_to_le: ∀n,m:nat. S n ≤ S m → n ≤ m.
280 (* XXX *) nletin hint ≝ monotonic. 
281 /2/; nqed.
282
283 (* this are instances of the le versions 
284 ntheorem lt_S_S_to_lt: ∀n,m. S n < S m → n < m.
285 /2/; nqed. 
286
287 ntheorem lt_to_lt_S_S: ∀n,m. n < m → S n < S m.
288 /2/; nqed. *)
289
290 ntheorem lt_to_not_zero : ∀n,m:nat. n < m → not_zero m.
291 #n; #m; #Hlt; nelim Hlt;//; nqed.
292
293 (* lt vs. le *)
294 ntheorem not_le_Sn_O: ∀ n:nat. S n ≰ O.
295 #n; napply nmk; #Hlen0; napply (lt_to_not_zero ?? Hlen0); nqed.
296
297 ntheorem not_le_to_not_le_S_S: ∀ n,m:nat. n ≰ m → S n ≰ S m.
298 /3/; nqed.
299
300 ntheorem not_le_S_S_to_not_le: ∀ n,m:nat. S n ≰ S m → n ≰ m.
301 /3/; nqed.
302
303 ntheorem decidable_le: ∀n,m. decidable (n≤m).
304 napply nat_elim2; #n; /2/;
305 #m; *; /3/; nqed.
306
307 ntheorem decidable_lt: ∀n,m. decidable (n < m).
308 #n; #m; napply decidable_le ; nqed.
309
310 ntheorem not_le_Sn_n: ∀n:nat. S n ≰ n.
311 #n; nelim n; /2/; nqed.
312
313 (* this is le_S_S_to_le
314 ntheorem lt_S_to_le: ∀n,m:nat. n < S m → n ≤ m.
315 /2/; nqed.
316 *)
317
318 ntheorem not_le_to_lt: ∀n,m. n ≰ m → m < n.
319 napply nat_elim2; #n;
320  ##[#abs; napply False_ind;/2/;
321  ##|/2/;
322  ##|#m;#Hind;#HnotleSS; napply le_S_S;/3/;
323  ##]
324 nqed.
325
326 ntheorem lt_to_not_le: ∀n,m. n < m → m ≰ n.
327 #n; #m; #Hltnm; nelim Hltnm;/3/; nqed.
328
329 ntheorem not_lt_to_le: ∀n,m:nat. n ≮ m → m ≤ n.
330 /4/; nqed.
331
332 (*
333 #n; #m; #Hnlt; napply le_S_S_to_le;/2/;
334 (* something strange here: /2/ fails *)
335 napply not_le_to_lt; napply Hnlt; nqed. *)
336
337 ntheorem le_to_not_lt: ∀n,m:nat. n ≤ m → m ≮ n.
338 #n; #m; #H;napply lt_to_not_le; /2/; (* /3/ *) nqed.
339
340 (* lt and le trans *)
341
342 ntheorem lt_to_le_to_lt: ∀n,m,p:nat. n < m → m ≤ p → n < p.
343 #n; #m; #p; #H; #H1; nelim H1; /2/; nqed.
344
345 ntheorem le_to_lt_to_lt: ∀n,m,p:nat. n ≤ m → m < p → n < p.
346 #n; #m; #p; #H; nelim H; /3/; nqed.
347
348 ntheorem lt_S_to_lt: ∀n,m. S n < m → n < m.
349 /2/; nqed.
350
351 ntheorem ltn_to_ltO: ∀n,m:nat. n < m → O < m.
352 /2/; nqed.
353
354 (*
355 theorem lt_SO_n_to_lt_O_pred_n: \forall n:nat.
356 (S O) \lt n \to O \lt (pred n).
357 intros.
358 apply (ltn_to_ltO (pred (S O)) (pred n) ?).
359  apply (lt_pred (S O) n);
360  [ apply (lt_O_S O) 
361  | assumption
362  ]
363 qed. *)
364
365 ntheorem lt_O_n_elim: ∀n:nat. O < n → 
366   ∀P:nat → Prop.(∀m:nat.P (S m)) → P n.
367 #n; nelim n; //; #abs; napply False_ind;/2/;
368 nqed.
369
370 (*
371 theorem lt_pred: \forall n,m. 
372   O < n \to n < m \to pred n < pred m. 
373 apply nat_elim2
374   [intros.apply False_ind.apply (not_le_Sn_O ? H)
375   |intros.apply False_ind.apply (not_le_Sn_O ? H1)
376   |intros.simplify.unfold.apply le_S_S_to_le.assumption
377   ]
378 qed.
379
380 theorem S_pred: \forall n:nat.lt O n \to eq nat n (S (pred n)).
381 intro.elim n.apply False_ind.exact (not_le_Sn_O O H).
382 apply eq_f.apply pred_Sn.
383 qed.
384
385 theorem le_pred_to_le:
386  ∀n,m. O < m → pred n ≤ pred m → n ≤ m.
387 intros 2;
388 elim n;
389 [ apply le_O_n
390 | simplify in H2;
391   rewrite > (S_pred m);
392   [ apply le_S_S;
393     assumption
394   | assumption
395   ]
396 ].
397 qed.
398
399 *)
400
401 (* le to lt or eq *)
402 ntheorem le_to_or_lt_eq: ∀n,m:nat. n ≤ m → n < m ∨ n = m.
403 #n; #m; #lenm; nelim lenm; /3/; nqed.
404
405 (* not eq *)
406 ntheorem lt_to_not_eq : ∀n,m:nat. n < m → n ≠ m.
407 #n; #m; #H; napply not_to_not;/2/; nqed.
408
409 (*not lt 
410 ntheorem eq_to_not_lt: ∀a,b:nat. a = b → a ≮ b.
411 intros.
412 unfold Not.
413 intros.
414 rewrite > H in H1.
415 apply (lt_to_not_eq b b)
416 [ assumption
417 | reflexivity
418 ]
419 qed. 
420
421 theorem lt_n_m_to_not_lt_m_Sn: ∀n,m. n < m → m ≮ S n.
422 intros;
423 unfold Not;
424 intro;
425 unfold lt in H;
426 unfold lt in H1;
427 generalize in match (le_S_S ? ? H);
428 intro;
429 generalize in match (transitive_le ? ? ? H2 H1);
430 intro;
431 apply (not_le_Sn_n ? H3).
432 qed. *)
433
434 ntheorem not_eq_to_le_to_lt: ∀n,m. n≠m → n≤m → n<m.
435 #n; #m; #Hneq; #Hle; ncases (le_to_or_lt_eq ?? Hle); //;
436 #Heq; /3/; nqed.
437 (*
438 nelim (Hneq Heq); nqed. *)
439
440 (* le elimination *)
441 ntheorem le_n_O_to_eq : ∀n:nat. n ≤ O → O=n.
442 #n; ncases n; //; #a ; #abs;
443 napply False_ind; /2/;nqed.
444
445 ntheorem le_n_O_elim: ∀n:nat. n ≤ O → ∀P: nat →Prop. P O → P n.
446 #n; ncases n; //; #a; #abs; 
447 napply False_ind; /2/; nqed. 
448
449 ntheorem le_n_Sm_elim : ∀n,m:nat.n ≤ S m → 
450 ∀P:Prop. (S n ≤ S m → P) → (n=S m → P) → P.
451 #n; #m; #Hle; #P; nelim Hle; /3/; nqed.
452
453 (* le and eq *)
454
455 ntheorem le_to_le_to_eq: ∀n,m. n ≤ m → m ≤ n → n = m.
456 napply nat_elim2; /4/; nqed. 
457
458 ntheorem lt_O_S : ∀n:nat. O < S n.
459 /2/; nqed.
460
461 (*
462 (* other abstract properties *)
463 theorem antisymmetric_le : antisymmetric nat le.
464 unfold antisymmetric.intros 2.
465 apply (nat_elim2 (\lambda n,m.(n \leq m \to m \leq n \to n=m))).
466 intros.apply le_n_O_to_eq.assumption.
467 intros.apply False_ind.apply (not_le_Sn_O ? H).
468 intros.apply eq_f.apply H.
469 apply le_S_S_to_le.assumption.
470 apply le_S_S_to_le.assumption.
471 qed.
472
473 theorem antisym_le: \forall n,m:nat. n \leq m \to m \leq n \to n=m
474 \def antisymmetric_le.
475
476 theorem le_n_m_to_lt_m_Sn_to_eq_n_m: ∀n,m. n ≤ m → m < S n → n=m.
477 intros;
478 unfold lt in H1;
479 generalize in match (le_S_S_to_le ? ? H1);
480 intro;
481 apply antisym_le;
482 assumption.
483 qed.
484 *)
485
486 (* well founded induction principles *)
487
488 ntheorem nat_elim1 : ∀n:nat.∀P:nat → Prop. 
489 (∀m.(∀p. p < m → P p) → P m) → P n.
490 #n; #P; #H; 
491 ncut (∀q:nat. q ≤ n → P q);/2/;
492 nelim n; 
493  ##[#q; #HleO; (* applica male *) 
494     napply (le_n_O_elim ? HleO);
495     napply H; #p; #ltpO;
496     napply False_ind; /2/; (* 3 *)
497  ##|#p; #Hind; #q; #HleS; 
498     napply H; #a; #lta; napply Hind;
499     napply le_S_S_to_le;/2/;
500  ##]
501 nqed.
502
503 (* some properties of functions *)
504 (*
505 definition increasing \def \lambda f:nat \to nat. 
506 \forall n:nat. f n < f (S n).
507
508 theorem increasing_to_monotonic: \forall f:nat \to nat.
509 increasing f \to monotonic nat lt f.
510 unfold monotonic.unfold lt.unfold increasing.unfold lt.intros.elim H1.apply H.
511 apply (trans_le ? (f n1)).
512 assumption.apply (trans_le ? (S (f n1))).
513 apply le_n_Sn.
514 apply H.
515 qed.
516
517 theorem le_n_fn: \forall f:nat \to nat. (increasing f) 
518 \to \forall n:nat. n \le (f n).
519 intros.elim n.
520 apply le_O_n.
521 apply (trans_le ? (S (f n1))).
522 apply le_S_S.apply H1.
523 simplify in H. unfold increasing in H.unfold lt in H.apply H.
524 qed.
525
526 theorem increasing_to_le: \forall f:nat \to nat. (increasing f) 
527 \to \forall m:nat. \exists i. m \le (f i).
528 intros.elim m.
529 apply (ex_intro ? ? O).apply le_O_n.
530 elim H1.
531 apply (ex_intro ? ? (S a)).
532 apply (trans_le ? (S (f a))).
533 apply le_S_S.assumption.
534 simplify in H.unfold increasing in H.unfold lt in H.
535 apply H.
536 qed.
537
538 theorem increasing_to_le2: \forall f:nat \to nat. (increasing f) 
539 \to \forall m:nat. (f O) \le m \to 
540 \exists i. (f i) \le m \land m <(f (S i)).
541 intros.elim H1.
542 apply (ex_intro ? ? O).
543 split.apply le_n.apply H.
544 elim H3.elim H4.
545 cut ((S n1) < (f (S a)) \lor (S n1) = (f (S a))).
546 elim Hcut.
547 apply (ex_intro ? ? a).
548 split.apply le_S. assumption.assumption.
549 apply (ex_intro ? ? (S a)).
550 split.rewrite < H7.apply le_n.
551 rewrite > H7.
552 apply H.
553 apply le_to_or_lt_eq.apply H6.
554 qed.
555 *)
556
557 (*********************** monotonicity ***************************)
558 ntheorem monotonic_le_plus_r: 
559 ∀n:nat.monotonic nat le (λm.n + m).
560 #n; #a; #b; nelim n; nnormalize; //;
561 #m; #H; #leab;napply le_S_S; /2/; nqed.
562
563 (*
564 ntheorem le_plus_r: ∀p,n,m:nat. n ≤ m → p + n ≤ p + m
565 ≝ monotonic_le_plus_r. *)
566
567 ntheorem monotonic_le_plus_l: 
568 ∀m:nat.monotonic nat le (λn.n + m).
569 #m; #x; #y; #H; napplyS monotonic_le_plus_r;
570 /2/; nqed.
571
572 (*
573 ntheorem le_plus_l: \forall p,n,m:nat. n \le m \to n + p \le m + p
574 \def monotonic_le_plus_l. *)
575
576 ntheorem le_plus: ∀n1,n2,m1,m2:nat. n1 ≤ n2  → m1 ≤ m2 
577 → n1 + m1 ≤ n2 + m2.
578 #n1; #n2; #m1; #m2; #len; #lem; napply (transitive_le ? (n1+m2));
579 /2/; nqed.
580
581 ntheorem le_plus_n :∀n,m:nat. m ≤ n + m.
582 /2/; nqed. 
583
584 nlemma le_plus_a: ∀a,n,m. n ≤ m → n ≤ a + m.
585 /2/; nqed.
586
587 nlemma le_plus_b: ∀b,n,m. n + b ≤ m → n ≤ m.
588 /2/; nqed.
589
590 ntheorem le_plus_n_r :∀n,m:nat. m ≤ m + n.
591 /2/; nqed.
592
593 ntheorem eq_plus_to_le: ∀n,m,p:nat.n=m+p → m ≤ n.
594 //; nqed.
595
596 ntheorem le_plus_to_le: ∀a,n,m. a + n ≤ a + m → n ≤ m.
597 #a; nelim a; nnormalize; /3/; nqed. 
598
599 ntheorem le_plus_to_le_r: ∀a,n,m. n + a ≤ m +a → n ≤ m.
600 /2/; nqed. 
601
602 (* plus & lt *)
603
604 ntheorem monotonic_lt_plus_r: 
605 ∀n:nat.monotonic nat lt (λm.n+m).
606 /2/; nqed.
607
608 (*
609 variant lt_plus_r: \forall n,p,q:nat. p < q \to n + p < n + q \def
610 monotonic_lt_plus_r. *)
611
612 ntheorem monotonic_lt_plus_l: 
613 ∀n:nat.monotonic nat lt (λm.m+n).
614 /2/; nqed.
615
616 (*
617 variant lt_plus_l: \forall n,p,q:nat. p < q \to p + n < q + n \def
618 monotonic_lt_plus_l. *)
619
620 ntheorem lt_plus: ∀n,m,p,q:nat. n < m → p < q → n + p < m + q.
621 #n; #m; #p; #q; #ltnm; #ltpq;
622 napply (transitive_lt ? (n+q));/2/; nqed.
623
624 ntheorem lt_plus_to_lt_l :∀n,p,q:nat. p+n < q+n → p<q.
625 /2/; nqed.
626
627 ntheorem lt_plus_to_lt_r :∀n,p,q:nat. n+p < n+q → p<q.
628 /2/; nqed.
629
630 (*
631 ntheorem le_to_lt_to_lt_plus: ∀a,b,c,d:nat.
632 a ≤ c → b < d → a + b < c+d.
633 (* bello /2/ un po' lento *)
634 #a; #b; #c; #d; #leac; #lebd; 
635 nnormalize; napplyS le_plus; //; nqed.
636 *)
637
638 (* times *)
639 ntheorem monotonic_le_times_r: 
640 ∀n:nat.monotonic nat le (λm. n * m).
641 #n; #x; #y; #lexy; nelim n; nnormalize;//;(* lento /2/;*)
642 #a; #lea; napply le_plus; //;
643 nqed.
644
645 (*
646 ntheorem le_times_r: \forall p,n,m:nat. n \le m \to p*n \le p*m
647 \def monotonic_le_times_r. *)
648
649 (*
650 ntheorem monotonic_le_times_l: 
651 ∀m:nat.monotonic nat le (λn.n*m).
652 /2/; nqed.
653 *)
654
655 (*
656 theorem le_times_l: \forall p,n,m:nat. n \le m \to n*p \le m*p
657 \def monotonic_le_times_l. *)
658
659 ntheorem le_times: ∀n1,n2,m1,m2:nat. 
660 n1 ≤ n2  → m1 ≤ m2 → n1*m1 ≤ n2*m2.
661 #n1; #n2; #m1; #m2; #len; #lem; 
662 napply (transitive_le ? (n1*m2)); (* /2/ slow *)
663  ##[ napply monotonic_le_times_r;//; 
664  ##| napplyS monotonic_le_times_r;//;
665  ##]
666 nqed.
667
668 (* interesssante *)
669 ntheorem lt_times_n: ∀n,m:nat. O < n → m ≤ n*m.
670 #n; #m; #H; /2/; nqed.
671
672 ntheorem le_times_to_le: 
673 ∀a,n,m. O < a → a * n ≤ a * m → n ≤ m.
674 #a; napply nat_elim2; nnormalize;
675   ##[//;
676   ##|#n; #H1; #H2; 
677      napply (transitive_le ? (a*S n));/2/;
678   ##|#n; #m; #H; #lta; #le;
679      napply le_S_S; napply H; /2/;
680   ##]
681 nqed.
682
683 ntheorem le_S_times_2: ∀n,m.O < m → n ≤ m → S n ≤ 2*m.
684 #n; #m; #posm; #lenm;  (* interessante *)
685 napplyS (le_plus n m); //; nqed.
686
687 (* times & lt *)
688 (*
689 ntheorem lt_O_times_S_S: ∀n,m:nat.O < (S n)*(S m).
690 intros.simplify.unfold lt.apply le_S_S.apply le_O_n.
691 qed. *)
692
693 (*
694 ntheorem lt_times_eq_O: \forall a,b:nat.
695 O < a → a * b = O → b = O.
696 intros.
697 apply (nat_case1 b)
698 [ intros.
699   reflexivity
700 | intros.
701   rewrite > H2 in H1.
702   rewrite > (S_pred a) in H1
703   [ apply False_ind.
704     apply (eq_to_not_lt O ((S (pred a))*(S m)))
705     [ apply sym_eq.
706       assumption
707     | apply lt_O_times_S_S
708     ]
709   | assumption
710   ]
711 ]
712 qed. 
713
714 theorem O_lt_times_to_O_lt: \forall a,c:nat.
715 O \lt (a * c) \to O \lt a.
716 intros.
717 apply (nat_case1 a)
718 [ intros.
719   rewrite > H1 in H.
720   simplify in H.
721   assumption
722 | intros.
723   apply lt_O_S
724 ]
725 qed.
726
727 lemma lt_times_to_lt_O: \forall i,n,m:nat. i < n*m \to O < m.
728 intros.
729 elim (le_to_or_lt_eq O ? (le_O_n m))
730   [assumption
731   |apply False_ind.
732    rewrite < H1 in H.
733    rewrite < times_n_O in H.
734    apply (not_le_Sn_O ? H)
735   ]
736 qed. *)
737
738 (*
739 ntheorem monotonic_lt_times_r: 
740 ∀n:nat.monotonic nat lt (λm.(S n)*m).
741 /2/; 
742 simplify.
743 intros.elim n.
744 simplify.rewrite < plus_n_O.rewrite < plus_n_O.assumption.
745 apply lt_plus.assumption.assumption.
746 qed. *)
747
748 ntheorem monotonic_lt_times_l: 
749   ∀c:nat. O < c → monotonic nat lt (λt.(t*c)).
750 #c; #posc; #n; #m; #ltnm;
751 nelim ltnm; nnormalize;
752   ##[/2/; 
753   ##|#a; #_; #lt1; napply (transitive_le ??? lt1);//;
754   ##]
755 nqed.
756
757 ntheorem monotonic_lt_times_r: 
758   ∀c:nat. O < c → monotonic nat lt (λt.(c*t)).
759 /2/; nqed.
760
761 ntheorem lt_to_le_to_lt_times: 
762 ∀n,m,p,q:nat. n < m → p ≤ q → O < q → n*p < m*q.
763 #n; #m; #p; #q; #ltnm; #lepq; #posq;
764 napply (le_to_lt_to_lt ? (n*q));
765   ##[napply monotonic_le_times_r;//;
766   ##|napply monotonic_lt_times_l;//;
767   ##]
768 nqed.
769
770 ntheorem lt_times:∀n,m,p,q:nat. n<m → p<q → n*p < m*q.
771 #n; #m; #p; #q; #ltnm; #ltpq;
772 napply lt_to_le_to_lt_times;/2/;
773 nqed.
774
775 ntheorem lt_times_n_to_lt_l: 
776 ∀n,p,q:nat. p*n < q*n → p < q.
777 #n; #p; #q; #Hlt;
778 nelim (decidable_lt p q);//;
779 #nltpq; napply False_ind; 
780 napply (absurd ? ? (lt_to_not_le ? ? Hlt));
781 napplyS monotonic_le_times_r;/2/;
782 nqed.
783
784 ntheorem lt_times_n_to_lt_r: 
785 ∀n,p,q:nat. n*p < n*q → p < q.
786 /2/; nqed.
787
788 (*
789 theorem nat_compare_times_l : \forall n,p,q:nat. 
790 nat_compare p q = nat_compare ((S n) * p) ((S n) * q).
791 intros.apply nat_compare_elim.intro.
792 apply nat_compare_elim.
793 intro.reflexivity.
794 intro.absurd (p=q).
795 apply (inj_times_r n).assumption.
796 apply lt_to_not_eq. assumption.
797 intro.absurd (q<p).
798 apply (lt_times_to_lt_r n).assumption.
799 apply le_to_not_lt.apply lt_to_le.assumption.
800 intro.rewrite < H.rewrite > nat_compare_n_n.reflexivity.
801 intro.apply nat_compare_elim.intro.
802 absurd (p<q).
803 apply (lt_times_to_lt_r n).assumption.
804 apply le_to_not_lt.apply lt_to_le.assumption.
805 intro.absurd (q=p).
806 symmetry.
807 apply (inj_times_r n).assumption.
808 apply lt_to_not_eq.assumption.
809 intro.reflexivity.
810 qed. *)
811
812 (* times and plus 
813 theorem lt_times_plus_times: \forall a,b,n,m:nat. 
814 a < n \to b < m \to a*m + b < n*m.
815 intros 3.
816 apply (nat_case n)
817   [intros.apply False_ind.apply (not_le_Sn_O ? H)
818   |intros.simplify.
819    rewrite < sym_plus.
820    unfold.
821    change with (S b+a*m1 \leq m1+m*m1).
822    apply le_plus
823     [assumption
824     |apply le_times
825       [apply le_S_S_to_le.assumption
826       |apply le_n
827       ]
828     ]
829   ]
830 qed. *)
831
832 (************************** minus ******************************)
833
834 nlet rec minus n m ≝ 
835  match n with 
836  [ O ⇒ O
837  | S p ⇒ 
838         match m with
839           [ O ⇒ S p
840     | S q ⇒ minus p q ]].
841         
842 interpretation "natural minus" 'minus x y = (minus x y).
843
844 ntheorem minus_S_S: ∀n,m:nat.S n - S m = n -m.
845 //; nqed.
846
847 ntheorem minus_O_n: ∀n:nat.O=O-n.
848 #n; ncases n; //; nqed.
849
850 ntheorem minus_n_O: ∀n:nat.n=n-O.
851 #n; ncases n; //; nqed.
852
853 ntheorem minus_n_n: ∀n:nat.O=n-n.
854 #n; nelim n; //; nqed.
855
856 ntheorem minus_Sn_n: ∀n:nat. S O = (S n)-n.
857 #n; nelim n; nnormalize; //; nqed.
858
859 ntheorem minus_Sn_m: ∀m,n:nat. m ≤ n → S n -m = S (n-m).
860 (* qualcosa da capire qui 
861 #n; #m; #lenm; nelim lenm; napplyS refl_eq. *)
862 napply nat_elim2; 
863   ##[//
864   ##|#n; #abs; napply False_ind; /2/ 
865   ##|#n; #m; #Hind; #c; napplyS Hind; /2/;
866   ##]
867 nqed.
868
869 ntheorem not_eq_to_le_to_le_minus: 
870   ∀n,m.n ≠ m → n ≤ m → n ≤ m - 1.
871 #n; #m; ncases m;//; #m; nnormalize;
872 #H; #H1; napply le_S_S_to_le;
873 napplyS (not_eq_to_le_to_lt n (S m) H H1);
874 nqed.
875
876 ntheorem eq_minus_S_pred: ∀n,m. n - (S m) = pred(n -m).
877 napply nat_elim2; nnormalize; //; nqed.
878
879 ntheorem plus_minus:
880 ∀m,n,p:nat. m ≤ n → (n-m)+p = (n+p)-m.
881 napply nat_elim2; 
882   ##[//
883   ##|#n; #p; #abs; napply False_ind; /2/;
884   ##|nnormalize;/3/;
885   ##]
886 nqed.
887
888 ntheorem minus_plus_m_m: ∀n,m:nat.n = (n+m)-m.
889 #n; #m; napplyS (plus_minus m m n); //; nqed.
890
891 ntheorem plus_minus_m_m: ∀n,m:nat.
892   m ≤ n → n = (n-m)+m.
893 #n; #m; #lemn; napplyS symmetric_eq; 
894 napplyS (plus_minus m n m); //; nqed.
895
896 ntheorem le_plus_minus_m_m: ∀n,m:nat. n ≤ (n-m)+m.
897 #n; nelim n;
898   ##[//
899   ##|#a; #Hind; #m; ncases m;//;  
900      nnormalize; #n;/2/;  
901   ##]
902 nqed.
903
904 ntheorem minus_to_plus :∀n,m,p:nat.
905   m ≤ n → n-m = p → n = m+p.
906 #n; #m; #p; #lemn; #eqp; napplyS plus_minus_m_m; //;
907 nqed.
908
909 ntheorem plus_to_minus :∀n,m,p:nat.n = m+p → n-m = p.
910 (* /4/ done in 43.5 *)
911 #n; #m; #p; #eqp; 
912 napply symmetric_eq;
913 napplyS (minus_plus_m_m p m);
914 nqed.
915
916 ntheorem minus_pred_pred : ∀n,m:nat. O < n → O < m → 
917 pred n - pred m = n - m.
918 #n; #m; #posn; #posm;
919 napply (lt_O_n_elim n posn); 
920 napply (lt_O_n_elim m posm);//.
921 nqed.
922
923 (*
924 theorem eq_minus_n_m_O: \forall n,m:nat.
925 n \leq m \to n-m = O.
926 intros 2.
927 apply (nat_elim2 (\lambda n,m.n \leq m \to n-m = O)).
928 intros.simplify.reflexivity.
929 intros.apply False_ind.
930 apply not_le_Sn_O;
931 [2: apply H | skip].
932 intros.
933 simplify.apply H.apply le_S_S_to_le. apply H1.
934 qed.
935
936 theorem le_SO_minus: \forall n,m:nat.S n \leq m \to S O \leq m-n.
937 intros.elim H.elim (minus_Sn_n n).apply le_n.
938 rewrite > minus_Sn_m.
939 apply le_S.assumption.
940 apply lt_to_le.assumption.
941 qed.
942
943 theorem minus_le_S_minus_S: \forall n,m:nat. m-n \leq S (m-(S n)).
944 intros.
945 apply (nat_elim2 (\lambda n,m.m-n \leq S (m-(S n)))).
946 intro.elim n1.simplify.apply le_n_Sn.
947 simplify.rewrite < minus_n_O.apply le_n.
948 intros.simplify.apply le_n_Sn.
949 intros.simplify.apply H.
950 qed.
951
952 theorem lt_minus_S_n_to_le_minus_n : \forall n,m,p:nat. m-(S n) < p \to m-n \leq p. 
953 intros 3.intro.
954 (* autobatch *)
955 (* end auto($Revision: 9739 $) proof: TIME=1.33 SIZE=100 DEPTH=100 *)
956 apply (trans_le (m-n) (S (m-(S n))) p).
957 apply minus_le_S_minus_S.
958 assumption.
959 qed.
960
961 theorem le_minus_m: \forall n,m:nat. n-m \leq n.
962 intros.apply (nat_elim2 (\lambda m,n. n-m \leq n)).
963 intros.rewrite < minus_n_O.apply le_n.
964 intros.simplify.apply le_n.
965 intros.simplify.apply le_S.assumption.
966 qed.
967
968 theorem lt_minus_m: \forall n,m:nat. O < n \to O < m \to n-m \lt n.
969 intros.apply (lt_O_n_elim n H).intro.
970 apply (lt_O_n_elim m H1).intro.
971 simplify.unfold lt.apply le_S_S.apply le_minus_m.
972 qed.
973
974 theorem minus_le_O_to_le: \forall n,m:nat. n-m \leq O \to n \leq m.
975 intros 2.
976 apply (nat_elim2 (\lambda n,m:nat.n-m \leq O \to n \leq m)).
977 intros.apply le_O_n.
978 simplify.intros. assumption.
979 simplify.intros.apply le_S_S.apply H.assumption.
980 qed.
981 *)
982
983 (* monotonicity and galois *)
984
985 ntheorem monotonic_le_minus_l: 
986 ∀p,q,n:nat. q ≤ p → q-n ≤ p-n.
987 napply nat_elim2; #p; #q;
988   ##[#lePO; napply (le_n_O_elim ? lePO);//;
989   ##|//;
990   ##|#Hind; #n; ncases n;
991     ##[//;
992     ##|#a; #leSS; napply Hind; /2/;
993     ##]
994   ##]
995 nqed.
996
997 ntheorem le_minus_to_plus: ∀n,m,p. n-m ≤ p → n≤ p+m.
998 #n; #m; #p; #lep;
999 napply transitive_le;
1000   ##[##|napply le_plus_minus_m_m
1001   ##|napply monotonic_le_plus_l;//;
1002   ##]
1003 nqed.
1004
1005 ntheorem le_plus_to_minus: ∀n,m,p. n ≤ p+m → n-m ≤ p.
1006 #n; #m; #p; #lep;
1007 (* bello *)
1008 napplyS monotonic_le_minus_l;//;
1009 (* /2/; *)
1010 nqed.
1011
1012 ntheorem monotonic_le_minus_r: 
1013 ∀p,q,n:nat. q ≤ p → n-p ≤ n-q.
1014 #p; #q; #n; #lepq;
1015 napply le_plus_to_minus;
1016 napply (transitive_le ??? (le_plus_minus_m_m ? q));/2/;
1017 nqed.
1018
1019 (*********************** boolean arithmetics ********************) 
1020 include "basics/bool.ma".
1021
1022 nlet rec eqb n m ≝ 
1023 match n with 
1024   [ O ⇒ match m with [ O ⇒ true | S q ⇒ false] 
1025   | S p ⇒ match m with [ O ⇒ false | S q ⇒ eqb p q]
1026   ].
1027            
1028 (*
1029 ntheorem eqb_to_Prop: ∀n,m:nat. 
1030 match (eqb n m) with
1031 [ true  \Rightarrow n = m 
1032 | false \Rightarrow n \neq m].
1033 intros.
1034 apply (nat_elim2
1035 (\lambda n,m:nat.match (eqb n m) with
1036 [ true  \Rightarrow n = m 
1037 | false \Rightarrow n \neq m])).
1038 intro.elim n1.
1039 simplify.reflexivity.
1040 simplify.apply not_eq_O_S.
1041 intro.
1042 simplify.unfold Not.
1043 intro. apply (not_eq_O_S n1).apply sym_eq.assumption.
1044 intros.simplify.
1045 generalize in match H.
1046 elim ((eqb n1 m1)).
1047 simplify.apply eq_f.apply H1.
1048 simplify.unfold Not.intro.apply H1.apply inj_S.assumption.
1049 qed.
1050 *)
1051
1052 naxiom eqb_elim : ∀ n,m:nat.∀ P:bool → Prop.
1053 (n=m → (P true)) → (n ≠ m → (P false)) → (P (eqb n m)). 
1054 (*
1055 napply nat_elim2; 
1056   ##[#n; ncases n; nnormalize; /3/; 
1057   ##|nnormalize; /3/;
1058   ##|nnormalize; /4/; 
1059   ##] 
1060 nqed.*)
1061
1062 ntheorem eqb_n_n: ∀n. eqb n n = true.
1063 #n; nelim n; nnormalize; //.
1064 nqed. 
1065
1066 ntheorem eqb_true_to_eq: ∀n,m:nat. eqb n m = true → n = m.
1067 #n; #m; napply (eqb_elim n m);//;
1068 #_; #abs; napply False_ind; /2/;
1069 nqed.
1070
1071 ntheorem eqb_false_to_not_eq: ∀n,m:nat. eqb n m = false → n ≠ m.
1072 #n; #m; napply (eqb_elim n m);/2/;
1073 nqed.
1074
1075 ntheorem eq_to_eqb_true: ∀n,m:nat.
1076   n = m → eqb n m = true.
1077 //; nqed.
1078
1079 ntheorem not_eq_to_eqb_false: ∀n,m:nat.
1080   n ≠  m → eqb n m = false.
1081 #n; #m; #noteq; 
1082 napply eqb_elim;//;
1083 #Heq; napply False_ind; /2/; 
1084 nqed.
1085
1086 nlet rec leb n m ≝ 
1087 match n with 
1088     [ O ⇒ true
1089     | (S p) ⇒
1090         match m with 
1091         [ O ⇒ false
1092               | (S q) ⇒ leb p q]].
1093
1094 ntheorem leb_elim: ∀n,m:nat. ∀P:bool → Prop. 
1095 (n ≤ m → P true) → (n ≰ m → P false) → P (leb n m).
1096 napply nat_elim2; nnormalize;
1097   ##[/2/
1098   ##|/3/;
1099   ##|#n; #m; #Hind; #P; #Pt; #Pf; napply Hind;
1100     ##[#lenm; napply Pt; napply le_S_S;//;
1101     ##|#nlenm; napply Pf; /2/; 
1102     ##]
1103   ##]
1104 nqed.
1105
1106 ntheorem leb_true_to_le:∀n,m.leb n m = true → n ≤ m.
1107 #n; #m; napply leb_elim;
1108   ##[//;
1109   ##|#_; #abs; napply False_ind; /2/;
1110   ##]
1111 nqed.
1112
1113 ntheorem leb_false_to_not_le:∀n,m.
1114   leb n m = false → n ≰ m.
1115 #n; #m; napply leb_elim;
1116   ##[#_; #abs; napply False_ind; /2/;
1117   ##|//;
1118   ##]
1119 nqed.
1120
1121 ntheorem le_to_leb_true: ∀n,m. n ≤ m → leb n m = true.
1122 #n; #m; napply leb_elim; //;
1123 #H; #H1; napply False_ind; /2/;
1124 nqed.
1125
1126 ntheorem not_le_to_leb_false: ∀n,m. n ≰ m → leb n m = false.
1127 #n; #m; napply leb_elim; //;
1128 #H; #H1; napply False_ind; /2/;
1129 nqed.
1130
1131 ntheorem lt_to_leb_false: ∀n,m. m < n → leb n m = false.
1132 /3/; nqed.
1133
1134 (* serve anche ltb? 
1135 ndefinition ltb ≝λn,m. leb (S n) m.
1136
1137 ntheorem ltb_elim: ∀n,m:nat. ∀P:bool → Prop. 
1138 (n < m → P true) → (n ≮ m → P false) → P (ltb n m).
1139 #n; #m; #P; #Hlt; #Hnlt;
1140 napply leb_elim; /3/; nqed.
1141
1142 ntheorem ltb_true_to_lt:∀n,m.ltb n m = true → n < m.
1143 #n; #m; #Hltb; napply leb_true_to_le; nassumption;
1144 nqed.
1145
1146 ntheorem ltb_false_to_not_lt:∀n,m.
1147   ltb n m = false → n ≮ m.
1148 #n; #m; #Hltb; napply leb_false_to_not_le; nassumption;
1149 nqed.
1150
1151 ntheorem lt_to_ltb_true: ∀n,m. n < m → ltb n m = true.
1152 #n; #m; #Hltb; napply le_to_leb_true; nassumption;
1153 nqed.
1154
1155 ntheorem le_to_ltb_false: ∀n,m. m \le n → ltb n m = false.
1156 #n; #m; #Hltb; napply lt_to_leb_false; /2/;
1157 nqed. *)
1158
1159 ninductive compare : Type[0] ≝
1160 | LT : compare
1161 | EQ : compare
1162 | GT : compare.
1163
1164 ndefinition compare_invert: compare → compare ≝
1165   λc.match c with
1166       [ LT ⇒ GT
1167       | EQ ⇒ EQ
1168       | GT ⇒ LT ].
1169
1170 nlet rec nat_compare n m: compare ≝
1171 match n with
1172 [ O ⇒ match m with 
1173       [ O ⇒ EQ
1174       | (S q) ⇒ LT ]
1175 | S p ⇒ match m with 
1176       [ O ⇒ GT
1177       | S q ⇒ nat_compare p q]].
1178
1179 ntheorem nat_compare_n_n: ∀n. nat_compare n n = EQ.
1180 #n;nelim n
1181 ##[//
1182 ##|#m;#IH;nnormalize;//]
1183 nqed.
1184
1185 ntheorem nat_compare_S_S: ∀n,m:nat.nat_compare n m = nat_compare (S n) (S m).
1186 //;
1187 nqed.
1188
1189 ntheorem nat_compare_pred_pred: 
1190   ∀n,m.O < n → O < m → nat_compare n m = nat_compare (pred n) (pred m).
1191 #n;#m;#Hn;#Hm;
1192 napply (lt_O_n_elim n Hn);
1193 napply (lt_O_n_elim m Hm);
1194 #p;#q;//;
1195 nqed.
1196
1197 ntheorem nat_compare_to_Prop: 
1198   ∀n,m.match (nat_compare n m) with
1199     [ LT ⇒ n < m
1200     | EQ ⇒ n = m
1201     | GT ⇒ m < n ].
1202 #n;#m;
1203 napply (nat_elim2 (λn,m.match (nat_compare n m) with
1204   [ LT ⇒ n < m
1205   | EQ ⇒ n = m
1206   | GT ⇒ m < n ]) ?????) (* FIXME: don't want to put all these ?, especially when … does not work! *)
1207 ##[##1,2:#n1;ncases n1;//;
1208 ##|#n1;#m1;nnormalize;ncases (nat_compare n1 m1);
1209    ##[##1,3:nnormalize;#IH;napply le_S_S;//;
1210    ##|nnormalize;#IH;nrewrite > IH;//]
1211 nqed.
1212
1213 ntheorem nat_compare_n_m_m_n: 
1214   ∀n,m:nat.nat_compare n m = compare_invert (nat_compare m n).
1215 #n;#m;
1216 napply (nat_elim2 (λn,m. nat_compare n m = compare_invert (nat_compare m n)))
1217 ##[##1,2:#n1;ncases n1;//;
1218 ##|#n1;#m1;#IH;nnormalize;napply IH]
1219 nqed.
1220      
1221 ntheorem nat_compare_elim : 
1222   ∀n,m. ∀P:compare → Prop.
1223     (n < m → P LT) → (n=m → P EQ) → (m < n → P GT) → P (nat_compare n m).
1224 #n;#m;#P;#Hlt;#Heq;#Hgt;
1225 ncut (match (nat_compare n m) with
1226     [ LT ⇒ n < m
1227     | EQ ⇒ n=m
1228     | GT ⇒ m < n] →
1229   P (nat_compare n m))
1230 ##[ncases (nat_compare n m);
1231    ##[napply Hlt
1232    ##|napply Heq
1233    ##|napply Hgt]
1234 ##|#Hcut;napply Hcut;//;
1235 nqed.
1236
1237 ninductive cmp_cases (n,m:nat) : CProp[0] ≝
1238   | cmp_le : n ≤ m → cmp_cases n m
1239   | cmp_gt : m < n → cmp_cases n m.
1240
1241 ntheorem lt_to_le : ∀n,m:nat. n < m → n ≤ m.
1242 #n;#m;#H;nelim H
1243 ##[//
1244 ##|/2/]
1245 nqed.
1246
1247 nlemma cmp_nat: ∀n,m.cmp_cases n m.
1248 #n;#m; nlapply (nat_compare_to_Prop n m);
1249 ncases (nat_compare n m);#H
1250 ##[@;napply lt_to_le;//
1251 ##|@;//
1252 ##|@2;//]
1253 nqed.