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