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