]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/nlibrary/arithmetics/nat.ma
4f1f6541164044c5b6be1e3a4117f4f73988b498
[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 "hints_declaration.ma".
16 include "basics/functions.ma".
17 include "basics/eq.ma".
18
19 ninductive nat : Type ≝
20   | O : nat
21   | S : nat → nat.
22   
23 interpretation "Natural numbers" 'N = nat.
24
25 alias num (instance 0) = "nnatural number".
26
27 (*
28 nrecord pos : Type ≝
29  {n:>nat; is_pos: n ≠ 0}.
30
31 ncoercion nat_to_pos: ∀n:nat. n ≠0 →pos ≝ mk_pos on 
32 *)
33
34 (* default "natural numbers" cic:/matita/ng/arithmetics/nat/nat.ind.
35 *)
36
37 ndefinition pred ≝
38  λn. match n with [ O ⇒ O | S p ⇒ p].
39
40 ntheorem pred_Sn : ∀n. n = pred (S n).
41 //; nqed.
42
43 ntheorem injective_S : injective nat nat S.
44 //; nqed.
45
46 (*
47 ntheorem inj_S : \forall n,m:nat.(S n)=(S m) \to n=m.
48 //. nqed. *)
49
50 ntheorem not_eq_S: ∀n,m:nat. n ≠ m → S n ≠ S m.
51 /3/; nqed.
52
53 ndefinition not_zero: nat → Prop ≝
54  λn: nat. match n with
55   [ O ⇒ False | (S p) ⇒ True ].
56
57 ntheorem not_eq_O_S : ∀n:nat. O ≠ S n.
58 #n; napply nmk; #eqOS; nchange with (not_zero O); nrewrite > eqOS; //.
59 nqed.
60
61 ntheorem not_eq_n_Sn: ∀n:nat. n ≠ S n.
62 #n; nelim n;/2/; nqed.
63
64 ntheorem nat_case:
65  ∀n:nat.∀P:nat → Prop. 
66   (n=O → P O) → (∀m:nat. (n=(S m) → P (S m))) → P n.
67 #n; #P; nelim n; /2/; nqed.
68
69 ntheorem nat_elim2 :
70  ∀R:nat → nat → Prop.
71   (∀n:nat. R O n) 
72   → (∀n:nat. R (S n) O)
73   → (∀n,m:nat. R n m → R (S n) (S m))
74   → ∀n,m:nat. R n m.
75 #R; #ROn; #RSO; #RSS; #n; nelim n;//;
76 #n0; #Rn0m; #m; ncases m;/2/; nqed.
77
78 ntheorem decidable_eq_nat : ∀n,m:nat.decidable (n=m).
79 napply nat_elim2; #n;
80  ##[ ncases n; /2/;
81  ##| /3/;
82  ##| #m; #Hind; ncases Hind;/3/;
83  (* 
84  ##[/2/; ##|#H; nwhd; 
85     napply or_introl;
86     napply eq_f2; /3/; *)
87  ##]
88 nqed. 
89
90 (*************************** plus ******************************)
91
92 nlet rec plus n m ≝ 
93  match n with 
94  [ O ⇒ m
95  | S p ⇒ S (plus p m) ].
96
97 interpretation "natural plus" 'plus x y = (plus x y).
98
99 ntheorem plus_O_n: ∀n:nat. n = 0+n.
100 //; nqed.
101
102 (*
103 ntheorem plus_Sn_m: ∀n,m:nat. S (n + m) = S n + m.
104 //; nqed.
105 *)
106
107 ntheorem plus_n_O: ∀n:nat. n = n+0.
108 #n; nelim n; nnormalize; //; nqed.
109
110 ntheorem plus_n_Sm : ∀n,m:nat. S (n+m) = n + S m.
111 #n; nelim n; nnormalize; //; nqed.
112
113 (*
114 ntheorem plus_Sn_m1: ∀n,m:nat. S m + n = n + S m.
115 #n; nelim n; nnormalize; //; nqed.
116 *)
117
118 (* deleterio?
119 ntheorem plus_n_1 : ∀n:nat. S n = n+1.
120 //; nqed.
121 *)
122
123 ntheorem symmetric_plus: symmetric ? plus.
124 #n; nelim n; nnormalize; //; nqed. 
125
126 ntheorem associative_plus : associative nat plus.
127 #n; nelim n; nnormalize; //; nqed.
128
129 ntheorem assoc_plus1: ∀a,b,c. c + (b + a) = b + c + a.
130 //; nqed.
131
132 ntheorem injective_plus_r: ∀n:nat.injective nat nat (λm.n+m).
133 #n; nelim n; nnormalize; /3/; nqed.
134
135 (* ntheorem inj_plus_r: \forall p,n,m:nat. p+n = p+m \to n=m
136 \def injective_plus_r. 
137
138 ntheorem injective_plus_l: ∀m:nat.injective nat nat (λn.n+m).
139 /2/; nqed. *)
140
141 (* ntheorem inj_plus_l: \forall p,n,m:nat. n+p = m+p \to n=m
142 \def injective_plus_l. *)
143
144 (*************************** times *****************************)
145
146 nlet rec times n m ≝ 
147  match n with 
148  [ O ⇒ O
149  | S p ⇒ m+(times p m) ].
150
151 interpretation "natural times" 'times x y = (times x y).
152
153 ntheorem times_Sn_m: ∀n,m:nat. m+n*m = S n*m.
154 //; nqed.
155
156 ntheorem times_O_n: ∀n:nat. O = O*n.
157 //; nqed.
158
159 ntheorem times_n_O: ∀n:nat. O = n*O.
160 #n; nelim n; //; nqed.
161
162 ntheorem times_n_Sm : ∀n,m:nat. n+(n*m) = n*(S m).
163 #n; nelim n; nnormalize; //; nqed.
164
165 ntheorem symmetric_times : symmetric nat times. 
166 #n; nelim n; nnormalize; //; nqed. 
167
168 (* variant sym_times : \forall n,m:nat. n*m = m*n \def
169 symmetric_times. *)
170
171 ntheorem distributive_times_plus : distributive nat times plus.
172 #n; nelim n; nnormalize; //; nqed.
173
174 ntheorem distributive_times_plus_r :
175   ∀a,b,c:nat. (b+c)*a = b*a + c*a.
176 //; nqed. 
177
178 ntheorem associative_times: associative nat times.
179 #n; nelim n; nnormalize; //; nqed.
180
181 nlemma times_times: ∀x,y,z. x*(y*z) = y*(x*z).
182 //; nqed. 
183
184 (* ci servono questi risultati? 
185 ntheorem times_O_to_O: ∀n,m:nat.n*m=O → n=O ∨ m=O.
186 napply nat_elim2; /2/; 
187 #n; #m; #H; nnormalize; #H1; napply False_ind;napply not_eq_O_S;
188 //; nqed.
189   
190 ntheorem times_n_SO : ∀n:nat. n = n * S O.
191 #n; //; nqed.
192
193 ntheorem times_SSO_n : ∀n:nat. n + n = (S(S O)) * n.
194 nnormalize; //; nqed.
195
196 nlemma times_SSO: \forall n.(S(S O))*(S n) = S(S((S(S O))*n)).
197 //; nqed.
198
199 ntheorem or_eq_eq_S: \forall n.\exists m. 
200 n = (S(S O))*m \lor n = S ((S(S O))*m).
201 #n; nelim n;
202   ##[@; /2/;
203   ##|#a; #H; nelim H; #b;#or;nelim or;#aeq;
204     ##[@ b; @ 2; //;
205     ##|@ (S b); @ 1; /2/;
206     ##]
207 nqed.
208 *)
209
210 (******************** ordering relations ************************)
211
212 ninductive le (n:nat) : nat → Prop ≝
213   | le_n : le n n
214   | le_S : ∀ m:nat. le n m → le n (S m).
215
216 interpretation "natural 'less or equal to'" 'leq x y = (le x y).
217
218 interpretation "natural 'neither less nor equal to'" 'nleq x y = (Not (le x y)).
219
220 ndefinition lt: nat → nat → Prop ≝
221 λn,m:nat. S n ≤ m.
222
223 interpretation "natural 'less than'" 'lt x y = (lt x y).
224
225 interpretation "natural 'not less than'" 'nless x y = (Not (lt x y)).
226
227 (* nlemma eq_lt: ∀n,m. (n < m) = (S n ≤ m).
228 //; nqed. *)
229
230 ndefinition ge: nat → nat → Prop ≝
231 λn,m:nat.m ≤ n.
232
233 interpretation "natural 'greater or equal to'" 'geq x y = (ge x y).
234
235 ndefinition gt: nat → nat → Prop ≝
236 λn,m:nat.m<n.
237
238 interpretation "natural 'greater than'" 'gt x y = (gt x y).
239
240 interpretation "natural 'not greater than'" 'ngtr x y = (Not (gt x y)).
241
242 ntheorem transitive_le : transitive nat le.
243 #a; #b; #c; #leab; #lebc;nelim lebc;/2/;
244 nqed.
245
246 (*
247 ntheorem trans_le: \forall n,m,p:nat. n \leq m \to m \leq p \to n \leq p
248 \def transitive_le. *)
249
250
251 ntheorem transitive_lt: transitive nat lt.
252 #a; #b; #c; #ltab; #ltbc;nelim ltbc;/2/;nqed.
253
254 (*
255 theorem trans_lt: \forall n,m,p:nat. lt n m \to lt m p \to lt n p
256 \def transitive_lt. *)
257
258 ntheorem le_S_S: ∀n,m:nat. n ≤ m → S n ≤ S m.
259 #n; #m; #lenm; nelim lenm; /2/; nqed.
260
261 ntheorem le_O_n : ∀n:nat. O ≤ n.
262 #n; nelim n; /2/; nqed.
263
264 ntheorem le_n_Sn : ∀n:nat. n ≤ S n.
265 /2/; nqed.
266
267 ntheorem le_pred_n : ∀n:nat. pred n ≤ n.
268 #n; nelim n; //; nqed.
269
270 (* XXX global problem 
271 nlemma my_trans_le : ∀x,y,z:nat.x ≤ y → y ≤ z → x ≤ z. 
272 napply transitive_le.
273 nqed. *)
274
275 ntheorem monotonic_pred: monotonic ? le pred.
276 #n; #m; #lenm; nelim lenm; /2/;nqed.
277
278 ntheorem le_S_S_to_le: ∀n,m:nat. S n ≤ S m → n ≤ m.
279 (* XXX *) nletin hint ≝ monotonic. 
280 #a; #b; #H; napplyS monotonic_pred;
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; 
457 ##[#n; #H1; #H2; /4/;
458  nqed. 
459
460 ntheorem lt_O_S : ∀n:nat. O < S n.
461 /2/; nqed.
462
463 (*
464 (* other abstract properties *)
465 theorem antisymmetric_le : antisymmetric nat le.
466 unfold antisymmetric.intros 2.
467 apply (nat_elim2 (\lambda n,m.(n \leq m \to m \leq n \to n=m))).
468 intros.apply le_n_O_to_eq.assumption.
469 intros.apply False_ind.apply (not_le_Sn_O ? H).
470 intros.apply eq_f.apply H.
471 apply le_S_S_to_le.assumption.
472 apply le_S_S_to_le.assumption.
473 qed.
474
475 theorem antisym_le: \forall n,m:nat. n \leq m \to m \leq n \to n=m
476 \def antisymmetric_le.
477
478 theorem le_n_m_to_lt_m_Sn_to_eq_n_m: ∀n,m. n ≤ m → m < S n → n=m.
479 intros;
480 unfold lt in H1;
481 generalize in match (le_S_S_to_le ? ? H1);
482 intro;
483 apply antisym_le;
484 assumption.
485 qed.
486 *)
487
488 (* well founded induction principles *)
489
490 ntheorem nat_elim1 : ∀n:nat.∀P:nat → Prop. 
491 (∀m.(∀p. p < m → P p) → P m) → P n.
492 #n; #P; #H; 
493 ncut (∀q:nat. q ≤ n → P q);/2/;
494 nelim n; 
495  ##[#q; #HleO; (* applica male *) 
496     napply (le_n_O_elim ? HleO);
497     napply H; #p; #ltpO;
498     napply False_ind; /2/; (* 3 *)
499  ##|#p; #Hind; #q; #HleS; 
500     napply H; #a; #lta; napply Hind;
501     napply le_S_S_to_le;/2/;
502  ##]
503 nqed.
504
505 (* some properties of functions *)
506 (*
507 definition increasing \def \lambda f:nat \to nat. 
508 \forall n:nat. f n < f (S n).
509
510 theorem increasing_to_monotonic: \forall f:nat \to nat.
511 increasing f \to monotonic nat lt f.
512 unfold monotonic.unfold lt.unfold increasing.unfold lt.intros.elim H1.apply H.
513 apply (trans_le ? (f n1)).
514 assumption.apply (trans_le ? (S (f n1))).
515 apply le_n_Sn.
516 apply H.
517 qed.
518
519 theorem le_n_fn: \forall f:nat \to nat. (increasing f) 
520 \to \forall n:nat. n \le (f n).
521 intros.elim n.
522 apply le_O_n.
523 apply (trans_le ? (S (f n1))).
524 apply le_S_S.apply H1.
525 simplify in H. unfold increasing in H.unfold lt in H.apply H.
526 qed.
527
528 theorem increasing_to_le: \forall f:nat \to nat. (increasing f) 
529 \to \forall m:nat. \exists i. m \le (f i).
530 intros.elim m.
531 apply (ex_intro ? ? O).apply le_O_n.
532 elim H1.
533 apply (ex_intro ? ? (S a)).
534 apply (trans_le ? (S (f a))).
535 apply le_S_S.assumption.
536 simplify in H.unfold increasing in H.unfold lt in H.
537 apply H.
538 qed.
539
540 theorem increasing_to_le2: \forall f:nat \to nat. (increasing f) 
541 \to \forall m:nat. (f O) \le m \to 
542 \exists i. (f i) \le m \land m <(f (S i)).
543 intros.elim H1.
544 apply (ex_intro ? ? O).
545 split.apply le_n.apply H.
546 elim H3.elim H4.
547 cut ((S n1) < (f (S a)) \lor (S n1) = (f (S a))).
548 elim Hcut.
549 apply (ex_intro ? ? a).
550 split.apply le_S. assumption.assumption.
551 apply (ex_intro ? ? (S a)).
552 split.rewrite < H7.apply le_n.
553 rewrite > H7.
554 apply H.
555 apply le_to_or_lt_eq.apply H6.
556 qed.
557 *)
558
559 (*********************** monotonicity ***************************)
560 ntheorem monotonic_le_plus_r: 
561 ∀n:nat.monotonic nat le (λm.n + m).
562 #n; #a; #b; nelim n; nnormalize; //;
563 #m; #H; #leab;napply le_S_S; /2/; nqed.
564
565 (*
566 ntheorem le_plus_r: ∀p,n,m:nat. n ≤ m → p + n ≤ p + m
567 ≝ monotonic_le_plus_r. *)
568
569 ntheorem monotonic_le_plus_l: 
570 ∀m:nat.monotonic nat le (λn.n + m).
571 /2/; nqed.
572
573 (*
574 ntheorem le_plus_l: \forall p,n,m:nat. n \le m \to n + p \le m + p
575 \def monotonic_le_plus_l. *)
576
577 ntheorem le_plus: ∀n1,n2,m1,m2:nat. n1 ≤ n2  → m1 ≤ m2 
578 → n1 + m1 ≤ n2 + m2.
579 #n1; #n2; #m1; #m2; #len; #lem; napply (transitive_le ? (n1+m2));
580 /2/; nqed.
581
582 ntheorem le_plus_n :∀n,m:nat. m ≤ n + m.
583 /2/; nqed. 
584
585 nlemma le_plus_a: ∀a,n,m. n ≤ m → n ≤ a + m.
586 /2/; nqed.
587
588 nlemma le_plus_b: ∀b,n,m. n + b ≤ m → n ≤ m.
589 /2/; nqed.
590
591 ntheorem le_plus_n_r :∀n,m:nat. m ≤ m + n.
592 /2/; nqed.
593
594 ntheorem eq_plus_to_le: ∀n,m,p:nat.n=m+p → m ≤ n.
595 //; nqed.
596
597 ntheorem le_plus_to_le: ∀a,n,m. a + n ≤ a + m → n ≤ m.
598 #a; nelim a; nnormalize; /3/; nqed. 
599
600 ntheorem le_plus_to_le_r: ∀a,n,m. n + a ≤ m +a → n ≤ m.
601 /2/; nqed. 
602
603 (* plus & lt *)
604
605 ntheorem monotonic_lt_plus_r: 
606 ∀n:nat.monotonic nat lt (λm.n+m).
607 /2/; nqed.
608
609 (*
610 variant lt_plus_r: \forall n,p,q:nat. p < q \to n + p < n + q \def
611 monotonic_lt_plus_r. *)
612
613 ntheorem monotonic_lt_plus_l: 
614 ∀n:nat.monotonic nat lt (λm.m+n).
615 /2/; nqed.
616
617 (*
618 variant lt_plus_l: \forall n,p,q:nat. p < q \to p + n < q + n \def
619 monotonic_lt_plus_l. *)
620
621 ntheorem lt_plus: ∀n,m,p,q:nat. n < m → p < q → n + p < m + q.
622 #n; #m; #p; #q; #ltnm; #ltpq;
623 napply (transitive_lt ? (n+q));/2/; nqed.
624
625 ntheorem lt_plus_to_lt_l :∀n,p,q:nat. p+n < q+n → p<q.
626 /2/; nqed.
627
628 ntheorem lt_plus_to_lt_r :∀n,p,q:nat. n+p < n+q → p<q.
629 /2/; nqed.
630
631 (*
632 ntheorem le_to_lt_to_lt_plus: ∀a,b,c,d:nat.
633 a ≤ c → b < d → a + b < c+d.
634 (* bello /2/ un po' lento *)
635 #a; #b; #c; #d; #leac; #lebd; 
636 nnormalize; napplyS le_plus; //; nqed.
637 *)
638
639 (* times *)
640 ntheorem monotonic_le_times_r: 
641 ∀n:nat.monotonic nat le (λm. n * m).
642 #n; #x; #y; #lexy; nelim n; nnormalize;//;(* lento /2/;*)
643 #a; #lea; napply le_plus; //;
644 nqed.
645
646 (*
647 ntheorem le_times_r: \forall p,n,m:nat. n \le m \to p*n \le p*m
648 \def monotonic_le_times_r. *)
649
650 (*
651 ntheorem monotonic_le_times_l: 
652 ∀m:nat.monotonic nat le (λn.n*m).
653 /2/; nqed.
654 *)
655
656 (*
657 theorem le_times_l: \forall p,n,m:nat. n \le m \to n*p \le m*p
658 \def monotonic_le_times_l. *)
659
660 ntheorem le_times: ∀n1,n2,m1,m2:nat. 
661 n1 ≤ n2  → m1 ≤ m2 → n1*m1 ≤ n2*m2.
662 #n1; #n2; #m1; #m2; #len; #lem; 
663 napply (transitive_le ? (n1*m2)); (* /2/ slow *)
664  ##[ napply monotonic_le_times_r;//; 
665  ##| napplyS monotonic_le_times_r;//;
666  ##]
667 nqed.
668
669 (* interesssante *)
670 ntheorem lt_times_n: ∀n,m:nat. O < n → m ≤ n*m.
671 #n; #m; #H; /2/; nqed.
672
673 ntheorem le_times_to_le: 
674 ∀a,n,m. O < a → a * n ≤ a * m → n ≤ m.
675 #a; napply nat_elim2; nnormalize;
676   ##[//;
677   ##|#n; #H1; #H2; 
678      napply (transitive_le ? (a*S n));/2/;
679   ##|#n; #m; #H; #lta; #le;
680      napply le_S_S; napply H; /2/;
681   ##]
682 nqed.
683
684 ntheorem le_S_times_2: ∀n,m.O < m → n ≤ m → S n ≤ 2*m.
685 #n; #m; #posm; #lenm;  (* interessante *)
686 napplyS (le_plus n m); //; nqed.
687
688 (* times & lt *)
689 (*
690 ntheorem lt_O_times_S_S: ∀n,m:nat.O < (S n)*(S m).
691 intros.simplify.unfold lt.apply le_S_S.apply le_O_n.
692 qed. *)
693
694 (*
695 ntheorem lt_times_eq_O: \forall a,b:nat.
696 O < a → a * b = O → b = O.
697 intros.
698 apply (nat_case1 b)
699 [ intros.
700   reflexivity
701 | intros.
702   rewrite > H2 in H1.
703   rewrite > (S_pred a) in H1
704   [ apply False_ind.
705     apply (eq_to_not_lt O ((S (pred a))*(S m)))
706     [ apply sym_eq.
707       assumption
708     | apply lt_O_times_S_S
709     ]
710   | assumption
711   ]
712 ]
713 qed. 
714
715 theorem O_lt_times_to_O_lt: \forall a,c:nat.
716 O \lt (a * c) \to O \lt a.
717 intros.
718 apply (nat_case1 a)
719 [ intros.
720   rewrite > H1 in H.
721   simplify in H.
722   assumption
723 | intros.
724   apply lt_O_S
725 ]
726 qed.
727
728 lemma lt_times_to_lt_O: \forall i,n,m:nat. i < n*m \to O < m.
729 intros.
730 elim (le_to_or_lt_eq O ? (le_O_n m))
731   [assumption
732   |apply False_ind.
733    rewrite < H1 in H.
734    rewrite < times_n_O in H.
735    apply (not_le_Sn_O ? H)
736   ]
737 qed. *)
738
739 (*
740 ntheorem monotonic_lt_times_r: 
741 ∀n:nat.monotonic nat lt (λm.(S n)*m).
742 /2/; 
743 simplify.
744 intros.elim n.
745 simplify.rewrite < plus_n_O.rewrite < plus_n_O.assumption.
746 apply lt_plus.assumption.assumption.
747 qed. *)
748
749 ntheorem monotonic_lt_times_l: 
750   ∀c:nat. O < c → monotonic nat lt (λt.(t*c)).
751 #c; #posc; #n; #m; #ltnm;
752 nelim ltnm; nnormalize;
753   ##[/2/; 
754   ##|#a; #_; #lt1; napply (transitive_le ??? lt1);//;
755   ##]
756 nqed.
757
758 ntheorem monotonic_lt_times_r: 
759   ∀c:nat. O < c → monotonic nat lt (λt.(c*t)).
760 /2/; nqed.
761
762 ntheorem lt_to_le_to_lt_times: 
763 ∀n,m,p,q:nat. n < m → p ≤ q → O < q → n*p < m*q.
764 #n; #m; #p; #q; #ltnm; #lepq; #posq;
765 napply (le_to_lt_to_lt ? (n*q));
766   ##[napply monotonic_le_times_r;//;
767   ##|napply monotonic_lt_times_l;//;
768   ##]
769 nqed.
770
771 ntheorem lt_times:∀n,m,p,q:nat. n<m → p<q → n*p < m*q.
772 #n; #m; #p; #q; #ltnm; #ltpq;
773 napply lt_to_le_to_lt_times;/2/;
774 nqed.
775
776 ntheorem lt_times_n_to_lt_l: 
777 ∀n,p,q:nat. p*n < q*n → p < q.
778 #n; #p; #q; #Hlt;
779 nelim (decidable_lt p q);//;
780 #nltpq; napply False_ind; 
781 napply (absurd ? ? (lt_to_not_le ? ? Hlt));
782 napplyS monotonic_le_times_r;/2/;
783 nqed.
784
785 ntheorem lt_times_n_to_lt_r: 
786 ∀n,p,q:nat. n*p < n*q → p < q.
787 /2/; nqed.
788
789 (*
790 theorem nat_compare_times_l : \forall n,p,q:nat. 
791 nat_compare p q = nat_compare ((S n) * p) ((S n) * q).
792 intros.apply nat_compare_elim.intro.
793 apply nat_compare_elim.
794 intro.reflexivity.
795 intro.absurd (p=q).
796 apply (inj_times_r n).assumption.
797 apply lt_to_not_eq. assumption.
798 intro.absurd (q<p).
799 apply (lt_times_to_lt_r n).assumption.
800 apply le_to_not_lt.apply lt_to_le.assumption.
801 intro.rewrite < H.rewrite > nat_compare_n_n.reflexivity.
802 intro.apply nat_compare_elim.intro.
803 absurd (p<q).
804 apply (lt_times_to_lt_r n).assumption.
805 apply le_to_not_lt.apply lt_to_le.assumption.
806 intro.absurd (q=p).
807 symmetry.
808 apply (inj_times_r n).assumption.
809 apply lt_to_not_eq.assumption.
810 intro.reflexivity.
811 qed. *)
812
813 (* times and plus 
814 theorem lt_times_plus_times: \forall a,b,n,m:nat. 
815 a < n \to b < m \to a*m + b < n*m.
816 intros 3.
817 apply (nat_case n)
818   [intros.apply False_ind.apply (not_le_Sn_O ? H)
819   |intros.simplify.
820    rewrite < sym_plus.
821    unfold.
822    change with (S b+a*m1 \leq m1+m*m1).
823    apply le_plus
824     [assumption
825     |apply le_times
826       [apply le_S_S_to_le.assumption
827       |apply le_n
828       ]
829     ]
830   ]
831 qed. *)
832
833 (************************** minus ******************************)
834
835 nlet rec minus n m ≝ 
836  match n with 
837  [ O ⇒ O
838  | S p ⇒ 
839         match m with
840           [ O ⇒ S p
841     | S q ⇒ minus p q ]].
842         
843 interpretation "natural minus" 'minus x y = (minus x y).
844
845 ntheorem minus_S_S: ∀n,m:nat.S n - S m = n -m.
846 //; nqed.
847
848 ntheorem minus_O_n: ∀n:nat.O=O-n.
849 #n; ncases n; //; nqed.
850
851 ntheorem minus_n_O: ∀n:nat.n=n-O.
852 #n; ncases n; //; nqed.
853
854 ntheorem minus_n_n: ∀n:nat.O=n-n.
855 #n; nelim n; //; nqed.
856
857 ntheorem minus_Sn_n: ∀n:nat. S O = (S n)-n.
858 #n; nelim n; nnormalize; //; nqed.
859
860 ntheorem minus_Sn_m: ∀m,n:nat. m ≤ n → S n -m = S (n-m).
861 (* qualcosa da capire qui 
862 #n; #m; #lenm; nelim lenm; napplyS refl_eq. *)
863 napply nat_elim2; 
864   ##[//
865   ##|#n; #abs; napply False_ind; /2/ 
866   ##|#n; #m; #Hind; #c; napplyS Hind; /2/;
867   ##]
868 nqed.
869
870 ntheorem not_eq_to_le_to_le_minus: 
871   ∀n,m.n ≠ m → n ≤ m → n ≤ m - 1.
872 #n; #m; ncases m;//; #m; nnormalize;
873 #H; #H1; napply le_S_S_to_le;
874 napplyS (not_eq_to_le_to_lt n (S m) H H1);
875 nqed.
876
877 ntheorem eq_minus_S_pred: ∀n,m. n - (S m) = pred(n -m).
878 napply nat_elim2; nnormalize; //; nqed.
879
880 ntheorem plus_minus:
881 ∀m,n,p:nat. m ≤ n → (n-m)+p = (n+p)-m.
882 napply nat_elim2; 
883   ##[//
884   ##|#n; #p; #abs; napply False_ind; /2/;
885   ##|nnormalize;/3/;
886   ##]
887 nqed.
888
889 ntheorem minus_plus_m_m: ∀n,m:nat.n = (n+m)-m.
890 #n; #m; napplyS (plus_minus m m n); //; nqed.
891
892 ntheorem plus_minus_m_m: ∀n,m:nat.
893   m ≤ n → n = (n-m)+m.
894 #n; #m; #lemn; napplyS sym_eq; 
895 napplyS (plus_minus m n m); //; nqed.
896
897 ntheorem le_plus_minus_m_m: ∀n,m:nat. n ≤ (n-m)+m.
898 #n; nelim n;
899   ##[//
900   ##|#a; #Hind; #m; ncases m;//;  
901      nnormalize; #n;/2/;  
902   ##]
903 nqed.
904
905 ntheorem minus_to_plus :∀n,m,p:nat.
906   m ≤ n → n-m = p → n = m+p.
907 #n; #m; #p; #lemn; #eqp; napplyS plus_minus_m_m; //;
908 nqed.
909
910 ntheorem plus_to_minus :∀n,m,p:nat.n = m+p → n-m = p.
911 (* /4/ done in 43.5 *)
912 #n; #m; #p; #eqp; 
913 napply sym_eq;
914 napplyS (minus_plus_m_m p m);
915 nqed.
916
917 ntheorem minus_pred_pred : ∀n,m:nat. O < n → O < m → 
918 pred n - pred m = n - m.
919 #n; #m; #posn; #posm;
920 napply (lt_O_n_elim n posn); 
921 napply (lt_O_n_elim m posm);//.
922 nqed.
923
924 (*
925 theorem eq_minus_n_m_O: \forall n,m:nat.
926 n \leq m \to n-m = O.
927 intros 2.
928 apply (nat_elim2 (\lambda n,m.n \leq m \to n-m = O)).
929 intros.simplify.reflexivity.
930 intros.apply False_ind.
931 apply not_le_Sn_O;
932 [2: apply H | skip].
933 intros.
934 simplify.apply H.apply le_S_S_to_le. apply H1.
935 qed.
936
937 theorem le_SO_minus: \forall n,m:nat.S n \leq m \to S O \leq m-n.
938 intros.elim H.elim (minus_Sn_n n).apply le_n.
939 rewrite > minus_Sn_m.
940 apply le_S.assumption.
941 apply lt_to_le.assumption.
942 qed.
943
944 theorem minus_le_S_minus_S: \forall n,m:nat. m-n \leq S (m-(S n)).
945 intros.
946 apply (nat_elim2 (\lambda n,m.m-n \leq S (m-(S n)))).
947 intro.elim n1.simplify.apply le_n_Sn.
948 simplify.rewrite < minus_n_O.apply le_n.
949 intros.simplify.apply le_n_Sn.
950 intros.simplify.apply H.
951 qed.
952
953 theorem lt_minus_S_n_to_le_minus_n : \forall n,m,p:nat. m-(S n) < p \to m-n \leq p. 
954 intros 3.intro.
955 (* autobatch *)
956 (* end auto($Revision: 9739 $) proof: TIME=1.33 SIZE=100 DEPTH=100 *)
957 apply (trans_le (m-n) (S (m-(S n))) p).
958 apply minus_le_S_minus_S.
959 assumption.
960 qed.
961
962 theorem le_minus_m: \forall n,m:nat. n-m \leq n.
963 intros.apply (nat_elim2 (\lambda m,n. n-m \leq n)).
964 intros.rewrite < minus_n_O.apply le_n.
965 intros.simplify.apply le_n.
966 intros.simplify.apply le_S.assumption.
967 qed.
968
969 theorem lt_minus_m: \forall n,m:nat. O < n \to O < m \to n-m \lt n.
970 intros.apply (lt_O_n_elim n H).intro.
971 apply (lt_O_n_elim m H1).intro.
972 simplify.unfold lt.apply le_S_S.apply le_minus_m.
973 qed.
974
975 theorem minus_le_O_to_le: \forall n,m:nat. n-m \leq O \to n \leq m.
976 intros 2.
977 apply (nat_elim2 (\lambda n,m:nat.n-m \leq O \to n \leq m)).
978 intros.apply le_O_n.
979 simplify.intros. assumption.
980 simplify.intros.apply le_S_S.apply H.assumption.
981 qed.
982 *)
983
984 (* monotonicity and galois *)
985
986 ntheorem monotonic_le_minus_l: 
987 ∀p,q,n:nat. q ≤ p → q-n ≤ p-n.
988 napply nat_elim2; #p; #q;
989   ##[#lePO; napply (le_n_O_elim ? lePO);//;
990   ##|//;
991   ##|#Hind; #n; ncases n;
992     ##[//;
993     ##|#a; #leSS; napply Hind; /2/;
994     ##]
995   ##]
996 nqed.
997
998 ntheorem le_minus_to_plus: ∀n,m,p. n-m ≤ p → n≤ p+m.
999 #n; #m; #p; #lep;
1000 napply transitive_le;
1001   ##[##|napply le_plus_minus_m_m
1002   ##|napply monotonic_le_plus_l;//;
1003   ##]
1004 nqed.
1005
1006 ntheorem le_plus_to_minus: ∀n,m,p. n ≤ p+m → n-m ≤ p.
1007 #n; #m; #p; #lep;
1008 (* bello *)
1009 napplyS monotonic_le_minus_l;//;
1010 (* /2/; *)
1011 nqed.
1012
1013 ntheorem monotonic_le_minus_r: 
1014 ∀p,q,n:nat. q ≤ p → n-p ≤ n-q.
1015 #p; #q; #n; #lepq;
1016 napply le_plus_to_minus;
1017 napply (transitive_le ??? (le_plus_minus_m_m ? q));/2/;
1018 nqed.
1019
1020 (*********************** boolean arithmetics ********************) 
1021 include "basics/bool.ma".
1022
1023 nlet rec eqb n m ≝ 
1024 match n with 
1025   [ O ⇒ match m with [ O ⇒ true | S q ⇒ false] 
1026   | S p ⇒ match m with [ O ⇒ false | S q ⇒ eqb p q]
1027   ].
1028            
1029 (*
1030 ntheorem eqb_to_Prop: ∀n,m:nat. 
1031 match (eqb n m) with
1032 [ true  \Rightarrow n = m 
1033 | false \Rightarrow n \neq m].
1034 intros.
1035 apply (nat_elim2
1036 (\lambda n,m:nat.match (eqb n m) with
1037 [ true  \Rightarrow n = m 
1038 | false \Rightarrow n \neq m])).
1039 intro.elim n1.
1040 simplify.reflexivity.
1041 simplify.apply not_eq_O_S.
1042 intro.
1043 simplify.unfold Not.
1044 intro. apply (not_eq_O_S n1).apply sym_eq.assumption.
1045 intros.simplify.
1046 generalize in match H.
1047 elim ((eqb n1 m1)).
1048 simplify.apply eq_f.apply H1.
1049 simplify.unfold Not.intro.apply H1.apply inj_S.assumption.
1050 qed.
1051 *)
1052
1053 ntheorem eqb_elim : ∀ n,m:nat.∀ P:bool → Prop.
1054 (n=m → (P true)) → (n ≠ m → (P false)) → (P (eqb n m)). 
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