1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| A.Asperti, C.Sacerdoti Coen, *)
8 (* ||A|| E.Tassi, S.Zacchiroli *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU Lesser General Public License Version 2.1 *)
13 (**************************************************************************)
15 include "hints_declaration.ma".
16 include "basics/functions.ma".
17 include "basics/eq.ma".
19 ninductive nat : Type ≝
23 interpretation "Natural numbers" 'N = nat.
25 alias num (instance 0) = "nnatural number".
29 {n:>nat; is_pos: n ≠ 0}.
31 ncoercion nat_to_pos: ∀n:nat. n ≠0 →pos ≝ mk_pos on
34 (* default "natural numbers" cic:/matita/ng/arithmetics/nat/nat.ind.
38 λn. match n with [ O ⇒ O | S p ⇒ p].
40 ntheorem pred_Sn : ∀n. n = pred (S n).
43 ntheorem injective_S : injective nat nat S.
47 ntheorem inj_S : \forall n,m:nat.(S n)=(S m) \to n=m.
50 ntheorem not_eq_S: ∀n,m:nat. n ≠ m → S n ≠ S m.
53 ndefinition not_zero: nat → Prop ≝
55 [ O ⇒ False | (S p) ⇒ True ].
57 ntheorem not_eq_O_S : ∀n:nat. O ≠ S n.
58 #n; napply nmk; #eqOS; nchange with (not_zero O); nrewrite > eqOS; //.
61 ntheorem not_eq_n_Sn: ∀n:nat. n ≠ S n.
62 #n; nelim n;/2/; nqed.
66 (n=O → P O) → (∀m:nat. (n=(S m) → P (S m))) → P n.
67 #n; #P; nelim n; /2/; nqed.
73 → (∀n,m:nat. R n m → R (S n) (S m))
75 #R; #ROn; #RSO; #RSS; #n; nelim n;//;
76 #n0; #Rn0m; #m; ncases m;/2/; nqed.
78 ntheorem decidable_eq_nat : ∀n,m:nat.decidable (n=m).
82 ##| #m; #Hind; ncases Hind;/3/;
86 (*************************** plus ******************************)
91 | S p ⇒ S (plus p m) ].
93 interpretation "natural plus" 'plus x y = (plus x y).
95 ntheorem plus_O_n: ∀n:nat. n = 0+n.
99 ntheorem plus_Sn_m: ∀n,m:nat. S (n + m) = S n + m.
103 ntheorem plus_n_O: ∀n:nat. n = n+0.
104 #n; nelim n; nnormalize; //; nqed.
106 ntheorem plus_n_Sm : ∀n,m:nat. S (n+m) = n + S m.
107 #n; nelim n; nnormalize; //; nqed.
110 ntheorem plus_Sn_m1: ∀n,m:nat. S m + n = n + S m.
111 #n; nelim n; nnormalize; //; nqed.
115 ntheorem plus_n_1 : ∀n:nat. S n = n+1.
119 ntheorem symmetric_plus: symmetric ? plus.
120 #n; nelim n; nnormalize; //; nqed.
122 ntheorem associative_plus : associative nat plus.
123 #n; nelim n; nnormalize; //; nqed.
125 ntheorem assoc_plus1: ∀a,b,c. c + (b + a) = b + c + a.
128 ntheorem injective_plus_r: ∀n:nat.injective nat nat (λm.n+m).
129 #n; nelim n; nnormalize; /3/; nqed.
131 (* ntheorem inj_plus_r: \forall p,n,m:nat. p+n = p+m \to n=m
132 \def injective_plus_r.
134 ntheorem injective_plus_l: ∀m:nat.injective nat nat (λn.n+m).
137 (* ntheorem inj_plus_l: \forall p,n,m:nat. n+p = m+p \to n=m
138 \def injective_plus_l. *)
140 (*************************** times *****************************)
145 | S p ⇒ m+(times p m) ].
147 interpretation "natural times" 'times x y = (times x y).
149 ntheorem times_Sn_m: ∀n,m:nat. m+n*m = S n*m.
152 ntheorem times_O_n: ∀n:nat. O = O*n.
155 ntheorem times_n_O: ∀n:nat. O = n*O.
156 #n; nelim n; //; nqed.
158 ntheorem times_n_Sm : ∀n,m:nat. n+(n*m) = n*(S m).
159 #n; nelim n; nnormalize; //; nqed.
161 ntheorem symmetric_times : symmetric nat times.
162 #n; nelim n; nnormalize; //; nqed.
164 (* variant sym_times : \forall n,m:nat. n*m = m*n \def
167 ntheorem distributive_times_plus : distributive nat times plus.
168 #n; nelim n; nnormalize; //; nqed.
170 ntheorem distributive_times_plus_r :
171 ∀a,b,c:nat. (b+c)*a = b*a + c*a.
174 ntheorem associative_times: associative nat times.
175 #n; nelim n; nnormalize; //; nqed.
177 nlemma times_times: ∀x,y,z. x*(y*z) = y*(x*z).
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;
186 ntheorem times_n_SO : ∀n:nat. n = n * S O.
189 ntheorem times_SSO_n : ∀n:nat. n + n = (S(S O)) * n.
190 nnormalize; //; nqed.
192 nlemma times_SSO: \forall n.(S(S O))*(S n) = S(S((S(S O))*n)).
195 ntheorem or_eq_eq_S: \forall n.\exists m.
196 n = (S(S O))*m \lor n = S ((S(S O))*m).
199 ##|#a; #H; nelim H; #b;#or;nelim or;#aeq;
201 ##|@ (S b); @ 1; /2/;
206 (******************** ordering relations ************************)
208 ninductive le (n:nat) : nat → Prop ≝
210 | le_S : ∀ m:nat. le n m → le n (S m).
212 interpretation "natural 'less or equal to'" 'leq x y = (le x y).
214 interpretation "natural 'neither less nor equal to'" 'nleq x y = (Not (le x y)).
216 ndefinition lt: nat → nat → Prop ≝
219 interpretation "natural 'less than'" 'lt x y = (lt x y).
221 interpretation "natural 'not less than'" 'nless x y = (Not (lt x y)).
223 (* nlemma eq_lt: ∀n,m. (n < m) = (S n ≤ m).
226 ndefinition ge: nat → nat → Prop ≝
229 interpretation "natural 'greater or equal to'" 'geq x y = (ge x y).
231 ndefinition gt: nat → nat → Prop ≝
234 interpretation "natural 'greater than'" 'gt x y = (gt x y).
236 interpretation "natural 'not greater than'" 'ngtr x y = (Not (gt x y)).
238 ntheorem transitive_le : transitive nat le.
239 #a; #b; #c; #leab; #lebc;nelim lebc;/2/;
243 ntheorem trans_le: \forall n,m,p:nat. n \leq m \to m \leq p \to n \leq p
244 \def transitive_le. *)
247 ntheorem transitive_lt: transitive nat lt.
248 #a; #b; #c; #ltab; #ltbc;nelim ltbc;/2/;nqed.
251 theorem trans_lt: \forall n,m,p:nat. lt n m \to lt m p \to lt n p
252 \def transitive_lt. *)
254 ntheorem le_S_S: ∀n,m:nat. n ≤ m → S n ≤ S m.
255 #n; #m; #lenm; nelim lenm; /2/; nqed.
257 ntheorem le_O_n : ∀n:nat. O ≤ n.
258 #n; nelim n; /2/; nqed.
260 ntheorem le_n_Sn : ∀n:nat. n ≤ S n.
263 ntheorem le_pred_n : ∀n:nat. pred n ≤ n.
264 #n; nelim n; //; nqed.
266 (* XXX global problem
267 nlemma my_trans_le : ∀x,y,z:nat.x ≤ y → y ≤ z → x ≤ z.
268 napply transitive_le.
271 ntheorem monotonic_pred: monotonic ? le pred.
272 #n; #m; #lenm; nelim lenm; /2/;nqed.
274 ntheorem le_S_S_to_le: ∀n,m:nat. S n ≤ S m → n ≤ m.
277 (* this are instances of the le versions
278 ntheorem lt_S_S_to_lt: ∀n,m. S n < S m → n < m.
281 ntheorem lt_to_lt_S_S: ∀n,m. n < m → S n < S m.
284 ntheorem lt_to_not_zero : ∀n,m:nat. n < m → not_zero m.
285 #n; #m; #Hlt; nelim Hlt;//; nqed.
288 ntheorem not_le_Sn_O: ∀ n:nat. S n ≰ O.
289 #n; napply nmk; #Hlen0; napply (lt_to_not_zero ?? Hlen0); nqed.
291 ntheorem not_le_to_not_le_S_S: ∀ n,m:nat. n ≰ m → S n ≰ S m.
294 ntheorem not_le_S_S_to_not_le: ∀ n,m:nat. S n ≰ S m → n ≰ m.
297 ntheorem decidable_le: ∀n,m. decidable (n≤m).
298 napply nat_elim2; #n; /2/;
301 ntheorem decidable_lt: ∀n,m. decidable (n < m).
302 #n; #m; napply decidable_le ; nqed.
304 ntheorem not_le_Sn_n: ∀n:nat. S n ≰ n.
305 #n; nelim n; /2/; nqed.
307 (* this is le_S_S_to_le
308 ntheorem lt_S_to_le: ∀n,m:nat. n < S m → n ≤ m.
312 ntheorem not_le_to_lt: ∀n,m. n ≰ m → m < n.
313 napply nat_elim2; #n;
314 ##[#abs; napply False_ind;/2/;
316 ##|#m;#Hind;#HnotleSS; napply le_S_S;/3/;
320 ntheorem lt_to_not_le: ∀n,m. n < m → m ≰ n.
321 #n; #m; #Hltnm; nelim Hltnm;/3/; nqed.
323 ntheorem not_lt_to_le: ∀n,m:nat. n ≮ m → m ≤ n.
327 #n; #m; #Hnlt; napply le_S_S_to_le;/2/;
328 (* something strange here: /2/ fails *)
329 napply not_le_to_lt; napply Hnlt; nqed. *)
331 ntheorem le_to_not_lt: ∀n,m:nat. n ≤ m → m ≮ n.
332 #n; #m; #H;napply lt_to_not_le; /2/; (* /3/ *) nqed.
334 (* lt and le trans *)
336 ntheorem lt_to_le_to_lt: ∀n,m,p:nat. n < m → m ≤ p → n < p.
337 #n; #m; #p; #H; #H1; nelim H1; /2/; nqed.
339 ntheorem le_to_lt_to_lt: ∀n,m,p:nat. n ≤ m → m < p → n < p.
340 #n; #m; #p; #H; nelim H; /3/; nqed.
342 ntheorem lt_S_to_lt: ∀n,m. S n < m → n < m.
345 ntheorem ltn_to_ltO: ∀n,m:nat. n < m → O < m.
349 theorem lt_SO_n_to_lt_O_pred_n: \forall n:nat.
350 (S O) \lt n \to O \lt (pred n).
352 apply (ltn_to_ltO (pred (S O)) (pred n) ?).
353 apply (lt_pred (S O) n);
359 ntheorem lt_O_n_elim: ∀n:nat. O < n →
360 ∀P:nat → Prop.(∀m:nat.P (S m)) → P n.
361 #n; nelim n; //; #abs; napply False_ind;/2/;
365 theorem lt_pred: \forall n,m.
366 O < n \to n < m \to pred n < pred m.
368 [intros.apply False_ind.apply (not_le_Sn_O ? H)
369 |intros.apply False_ind.apply (not_le_Sn_O ? H1)
370 |intros.simplify.unfold.apply le_S_S_to_le.assumption
374 theorem S_pred: \forall n:nat.lt O n \to eq nat n (S (pred n)).
375 intro.elim n.apply False_ind.exact (not_le_Sn_O O H).
376 apply eq_f.apply pred_Sn.
379 theorem le_pred_to_le:
380 ∀n,m. O < m → pred n ≤ pred m → n ≤ m.
385 rewrite > (S_pred m);
396 ntheorem le_to_or_lt_eq: ∀n,m:nat. n ≤ m → n < m ∨ n = m.
397 #n; #m; #lenm; nelim lenm; /3/; nqed.
400 ntheorem lt_to_not_eq : ∀n,m:nat. n < m → n ≠ m.
401 #n; #m; #H; napply not_to_not;/2/; nqed.
404 ntheorem eq_to_not_lt: ∀a,b:nat. a = b → a ≮ b.
409 apply (lt_to_not_eq b b)
415 theorem lt_n_m_to_not_lt_m_Sn: ∀n,m. n < m → m ≮ S n.
421 generalize in match (le_S_S ? ? H);
423 generalize in match (transitive_le ? ? ? H2 H1);
425 apply (not_le_Sn_n ? H3).
428 ntheorem not_eq_to_le_to_lt: ∀n,m. n≠m → n≤m → n<m.
429 #n; #m; #Hneq; #Hle; ncases (le_to_or_lt_eq ?? Hle); //;
432 nelim (Hneq Heq); nqed. *)
435 ntheorem le_n_O_to_eq : ∀n:nat. n ≤ O → O=n.
436 #n; ncases n; //; #a ; #abs;
437 napply False_ind; /2/;nqed.
439 ntheorem le_n_O_elim: ∀n:nat. n ≤ O → ∀P: nat →Prop. P O → P n.
440 #n; ncases n; //; #a; #abs;
441 napply False_ind; /2/; nqed.
443 ntheorem le_n_Sm_elim : ∀n,m:nat.n ≤ S m →
444 ∀P:Prop. (S n ≤ S m → P) → (n=S m → P) → P.
445 #n; #m; #Hle; #P; nelim Hle; /3/; nqed.
449 ntheorem le_to_le_to_eq: ∀n,m. n ≤ m → m ≤ n → n = m.
450 napply nat_elim2; /4/;
453 ntheorem lt_O_S : ∀n:nat. O < S n.
457 (* other abstract properties *)
458 theorem antisymmetric_le : antisymmetric nat le.
459 unfold antisymmetric.intros 2.
460 apply (nat_elim2 (\lambda n,m.(n \leq m \to m \leq n \to n=m))).
461 intros.apply le_n_O_to_eq.assumption.
462 intros.apply False_ind.apply (not_le_Sn_O ? H).
463 intros.apply eq_f.apply H.
464 apply le_S_S_to_le.assumption.
465 apply le_S_S_to_le.assumption.
468 theorem antisym_le: \forall n,m:nat. n \leq m \to m \leq n \to n=m
469 \def antisymmetric_le.
471 theorem le_n_m_to_lt_m_Sn_to_eq_n_m: ∀n,m. n ≤ m → m < S n → n=m.
474 generalize in match (le_S_S_to_le ? ? H1);
481 (* well founded induction principles *)
483 ntheorem nat_elim1 : ∀n:nat.∀P:nat → Prop.
484 (∀m.(∀p. p < m → P p) → P m) → P n.
486 ncut (∀q:nat. q ≤ n → P q);/2/;
488 ##[#q; #HleO; (* applica male *)
489 napply (le_n_O_elim ? HleO);
491 napply False_ind; /2/; (* 3 *)
492 ##|#p; #Hind; #q; #HleS;
493 napply H; #a; #lta; napply Hind;
494 napply le_S_S_to_le;/2/;
498 (* some properties of functions *)
500 definition increasing \def \lambda f:nat \to nat.
501 \forall n:nat. f n < f (S n).
503 theorem increasing_to_monotonic: \forall f:nat \to nat.
504 increasing f \to monotonic nat lt f.
505 unfold monotonic.unfold lt.unfold increasing.unfold lt.intros.elim H1.apply H.
506 apply (trans_le ? (f n1)).
507 assumption.apply (trans_le ? (S (f n1))).
512 theorem le_n_fn: \forall f:nat \to nat. (increasing f)
513 \to \forall n:nat. n \le (f n).
516 apply (trans_le ? (S (f n1))).
517 apply le_S_S.apply H1.
518 simplify in H. unfold increasing in H.unfold lt in H.apply H.
521 theorem increasing_to_le: \forall f:nat \to nat. (increasing f)
522 \to \forall m:nat. \exists i. m \le (f i).
524 apply (ex_intro ? ? O).apply le_O_n.
526 apply (ex_intro ? ? (S a)).
527 apply (trans_le ? (S (f a))).
528 apply le_S_S.assumption.
529 simplify in H.unfold increasing in H.unfold lt in H.
533 theorem increasing_to_le2: \forall f:nat \to nat. (increasing f)
534 \to \forall m:nat. (f O) \le m \to
535 \exists i. (f i) \le m \land m <(f (S i)).
537 apply (ex_intro ? ? O).
538 split.apply le_n.apply H.
540 cut ((S n1) < (f (S a)) \lor (S n1) = (f (S a))).
542 apply (ex_intro ? ? a).
543 split.apply le_S. assumption.assumption.
544 apply (ex_intro ? ? (S a)).
545 split.rewrite < H7.apply le_n.
548 apply le_to_or_lt_eq.apply H6.
552 (*********************** monotonicity ***************************)
553 ntheorem monotonic_le_plus_r:
554 ∀n:nat.monotonic nat le (λm.n + m).
555 #n; #a; #b; nelim n; nnormalize; //;
556 #m; #H; #leab;napply le_S_S; /2/; nqed.
559 ntheorem le_plus_r: ∀p,n,m:nat. n ≤ m → p + n ≤ p + m
560 ≝ monotonic_le_plus_r. *)
562 ntheorem monotonic_le_plus_l:
563 ∀m:nat.monotonic nat le (λn.n + m).
567 ntheorem le_plus_l: \forall p,n,m:nat. n \le m \to n + p \le m + p
568 \def monotonic_le_plus_l. *)
570 ntheorem le_plus: ∀n1,n2,m1,m2:nat. n1 ≤ n2 → m1 ≤ m2
572 #n1; #n2; #m1; #m2; #len; #lem; napply (transitive_le ? (n1+m2));
575 ntheorem le_plus_n :∀n,m:nat. m ≤ n + m.
578 nlemma le_plus_a: ∀a,n,m. n ≤ m → n ≤ a + m.
581 nlemma le_plus_b: ∀b,n,m. n + b ≤ m → n ≤ m.
584 ntheorem le_plus_n_r :∀n,m:nat. m ≤ m + n.
587 ntheorem eq_plus_to_le: ∀n,m,p:nat.n=m+p → m ≤ n.
590 ntheorem le_plus_to_le: ∀a,n,m. a + n ≤ a + m → n ≤ m.
591 #a; nelim a; nnormalize; /3/; nqed.
593 ntheorem le_plus_to_le_r: ∀a,n,m. n + a ≤ m +a → n ≤ m.
598 ntheorem monotonic_lt_plus_r:
599 ∀n:nat.monotonic nat lt (λm.n+m).
603 variant lt_plus_r: \forall n,p,q:nat. p < q \to n + p < n + q \def
604 monotonic_lt_plus_r. *)
606 ntheorem monotonic_lt_plus_l:
607 ∀n:nat.monotonic nat lt (λm.m+n).
611 variant lt_plus_l: \forall n,p,q:nat. p < q \to p + n < q + n \def
612 monotonic_lt_plus_l. *)
614 ntheorem lt_plus: ∀n,m,p,q:nat. n < m → p < q → n + p < m + q.
615 #n; #m; #p; #q; #ltnm; #ltpq;
616 napply (transitive_lt ? (n+q));/2/; nqed.
618 ntheorem lt_plus_to_lt_l :∀n,p,q:nat. p+n < q+n → p<q.
621 ntheorem lt_plus_to_lt_r :∀n,p,q:nat. n+p < n+q → p<q.
625 ntheorem le_to_lt_to_lt_plus: ∀a,b,c,d:nat.
626 a ≤ c → b < d → a + b < c+d.
627 (* bello /2/ un po' lento *)
628 #a; #b; #c; #d; #leac; #lebd;
629 nnormalize; napplyS le_plus; //; nqed.
633 ntheorem monotonic_le_times_r:
634 ∀n:nat.monotonic nat le (λm. n * m).
635 #n; #x; #y; #lexy; nelim n; nnormalize;//;(* lento /2/;*)
636 #a; #lea; napply le_plus; //;
640 ntheorem le_times_r: \forall p,n,m:nat. n \le m \to p*n \le p*m
641 \def monotonic_le_times_r. *)
644 ntheorem monotonic_le_times_l:
645 ∀m:nat.monotonic nat le (λn.n*m).
650 theorem le_times_l: \forall p,n,m:nat. n \le m \to n*p \le m*p
651 \def monotonic_le_times_l. *)
653 ntheorem le_times: ∀n1,n2,m1,m2:nat.
654 n1 ≤ n2 → m1 ≤ m2 → n1*m1 ≤ n2*m2.
655 #n1; #n2; #m1; #m2; #len; #lem;
656 napply (transitive_le ? (n1*m2)); (* /2/ slow *)
657 ##[ napply monotonic_le_times_r;//;
658 ##| napplyS monotonic_le_times_r;//;
663 ntheorem lt_times_n: ∀n,m:nat. O < n → m ≤ n*m.
664 #n; #m; #H; /2/; nqed.
666 ntheorem le_times_to_le:
667 ∀a,n,m. O < a → a * n ≤ a * m → n ≤ m.
668 #a; napply nat_elim2; nnormalize;
671 napply (transitive_le ? (a*S n));/2/;
672 ##|#n; #m; #H; #lta; #le;
673 napply le_S_S; napply H; /2/;
677 ntheorem le_S_times_2: ∀n,m.O < m → n ≤ m → S n ≤ 2*m.
678 #n; #m; #posm; #lenm; (* interessante *)
679 napplyS (le_plus n m); //; nqed.
683 ntheorem lt_O_times_S_S: ∀n,m:nat.O < (S n)*(S m).
684 intros.simplify.unfold lt.apply le_S_S.apply le_O_n.
688 ntheorem lt_times_eq_O: \forall a,b:nat.
689 O < a → a * b = O → b = O.
696 rewrite > (S_pred a) in H1
698 apply (eq_to_not_lt O ((S (pred a))*(S m)))
701 | apply lt_O_times_S_S
708 theorem O_lt_times_to_O_lt: \forall a,c:nat.
709 O \lt (a * c) \to O \lt a.
721 lemma lt_times_to_lt_O: \forall i,n,m:nat. i < n*m \to O < m.
723 elim (le_to_or_lt_eq O ? (le_O_n m))
727 rewrite < times_n_O in H.
728 apply (not_le_Sn_O ? H)
733 ntheorem monotonic_lt_times_r:
734 ∀n:nat.monotonic nat lt (λm.(S n)*m).
738 simplify.rewrite < plus_n_O.rewrite < plus_n_O.assumption.
739 apply lt_plus.assumption.assumption.
742 ntheorem monotonic_lt_times_l:
743 ∀c:nat. O < c → monotonic nat lt (λt.(t*c)).
744 #c; #posc; #n; #m; #ltnm;
745 nelim ltnm; nnormalize;
747 ##|#a; #_; #lt1; napply (transitive_le ??? lt1);//;
751 ntheorem monotonic_lt_times_r:
752 ∀c:nat. O < c → monotonic nat lt (λt.(c*t)).
755 ntheorem lt_to_le_to_lt_times:
756 ∀n,m,p,q:nat. n < m → p ≤ q → O < q → n*p < m*q.
757 #n; #m; #p; #q; #ltnm; #lepq; #posq;
758 napply (le_to_lt_to_lt ? (n*q));
759 ##[napply monotonic_le_times_r;//;
760 ##|napply monotonic_lt_times_l;//;
764 ntheorem lt_times:∀n,m,p,q:nat. n<m → p<q → n*p < m*q.
765 #n; #m; #p; #q; #ltnm; #ltpq;
766 napply lt_to_le_to_lt_times;/2/;
769 ntheorem lt_times_n_to_lt_l:
770 ∀n,p,q:nat. p*n < q*n → p < q.
772 nelim (decidable_lt p q);//;
773 #nltpq; napply False_ind;
774 napply (absurd ? ? (lt_to_not_le ? ? Hlt));
775 napplyS monotonic_le_times_r;/2/;
778 ntheorem lt_times_n_to_lt_r:
779 ∀n,p,q:nat. n*p < n*q → p < q.
783 theorem nat_compare_times_l : \forall n,p,q:nat.
784 nat_compare p q = nat_compare ((S n) * p) ((S n) * q).
785 intros.apply nat_compare_elim.intro.
786 apply nat_compare_elim.
789 apply (inj_times_r n).assumption.
790 apply lt_to_not_eq. assumption.
792 apply (lt_times_to_lt_r n).assumption.
793 apply le_to_not_lt.apply lt_to_le.assumption.
794 intro.rewrite < H.rewrite > nat_compare_n_n.reflexivity.
795 intro.apply nat_compare_elim.intro.
797 apply (lt_times_to_lt_r n).assumption.
798 apply le_to_not_lt.apply lt_to_le.assumption.
801 apply (inj_times_r n).assumption.
802 apply lt_to_not_eq.assumption.
807 theorem lt_times_plus_times: \forall a,b,n,m:nat.
808 a < n \to b < m \to a*m + b < n*m.
811 [intros.apply False_ind.apply (not_le_Sn_O ? H)
815 change with (S b+a*m1 \leq m1+m*m1).
819 [apply le_S_S_to_le.assumption
826 (************************** minus ******************************)
834 | S q ⇒ minus p q ]].
836 interpretation "natural minus" 'minus x y = (minus x y).
838 ntheorem minus_S_S: ∀n,m:nat.S n - S m = n -m.
841 ntheorem minus_O_n: ∀n:nat.O=O-n.
842 #n; ncases n; //; nqed.
844 ntheorem minus_n_O: ∀n:nat.n=n-O.
845 #n; ncases n; //; nqed.
847 ntheorem minus_n_n: ∀n:nat.O=n-n.
848 #n; nelim n; //; nqed.
850 ntheorem minus_Sn_n: ∀n:nat. S O = (S n)-n.
851 #n; nelim n; nnormalize; //; nqed.
853 ntheorem minus_Sn_m: ∀m,n:nat. m ≤ n → S n -m = S (n-m).
854 (* qualcosa da capire qui
855 #n; #m; #lenm; nelim lenm; napplyS refl_eq. *)
858 ##|#n; #abs; napply False_ind; /2/
859 ##|#n; #m; #Hind; #c; napplyS Hind; /2/;
863 ntheorem not_eq_to_le_to_le_minus:
864 ∀n,m.n ≠ m → n ≤ m → n ≤ m - 1.
865 #n; #m; ncases m;//; #m; nnormalize;
866 #H; #H1; napply le_S_S_to_le;
867 napplyS (not_eq_to_le_to_lt n (S m) H H1);
870 ntheorem eq_minus_S_pred: ∀n,m. n - (S m) = pred(n -m).
871 napply nat_elim2; nnormalize; //; nqed.
874 ∀m,n,p:nat. m ≤ n → (n-m)+p = (n+p)-m.
877 ##|#n; #p; #abs; napply False_ind; /2/;
882 ntheorem minus_plus_m_m: ∀n,m:nat.n = (n+m)-m.
883 #n; #m; napplyS (plus_minus m m n); //; nqed.
885 ntheorem plus_minus_m_m: ∀n,m:nat.
887 #n; #m; #lemn; napplyS sym_eq;
888 napplyS (plus_minus m n m); //; nqed.
890 ntheorem le_plus_minus_m_m: ∀n,m:nat. n ≤ (n-m)+m.
893 ##|#a; #Hind; #m; ncases m;//;
898 ntheorem minus_to_plus :∀n,m,p:nat.
899 m ≤ n → n-m = p → n = m+p.
900 #n; #m; #p; #lemn; #eqp; napplyS plus_minus_m_m; //;
903 ntheorem plus_to_minus :∀n,m,p:nat.n = m+p → n-m = p.
904 (* /4/ done in 43.5 *)
907 napplyS (minus_plus_m_m p m);
910 ntheorem minus_pred_pred : ∀n,m:nat. O < n → O < m →
911 pred n - pred m = n - m.
912 #n; #m; #posn; #posm;
913 napply (lt_O_n_elim n posn);
914 napply (lt_O_n_elim m posm);//.
918 theorem eq_minus_n_m_O: \forall n,m:nat.
919 n \leq m \to n-m = O.
921 apply (nat_elim2 (\lambda n,m.n \leq m \to n-m = O)).
922 intros.simplify.reflexivity.
923 intros.apply False_ind.
927 simplify.apply H.apply le_S_S_to_le. apply H1.
930 theorem le_SO_minus: \forall n,m:nat.S n \leq m \to S O \leq m-n.
931 intros.elim H.elim (minus_Sn_n n).apply le_n.
932 rewrite > minus_Sn_m.
933 apply le_S.assumption.
934 apply lt_to_le.assumption.
937 theorem minus_le_S_minus_S: \forall n,m:nat. m-n \leq S (m-(S n)).
939 apply (nat_elim2 (\lambda n,m.m-n \leq S (m-(S n)))).
940 intro.elim n1.simplify.apply le_n_Sn.
941 simplify.rewrite < minus_n_O.apply le_n.
942 intros.simplify.apply le_n_Sn.
943 intros.simplify.apply H.
946 theorem lt_minus_S_n_to_le_minus_n : \forall n,m,p:nat. m-(S n) < p \to m-n \leq p.
949 (* end auto($Revision: 9739 $) proof: TIME=1.33 SIZE=100 DEPTH=100 *)
950 apply (trans_le (m-n) (S (m-(S n))) p).
951 apply minus_le_S_minus_S.
955 theorem le_minus_m: \forall n,m:nat. n-m \leq n.
956 intros.apply (nat_elim2 (\lambda m,n. n-m \leq n)).
957 intros.rewrite < minus_n_O.apply le_n.
958 intros.simplify.apply le_n.
959 intros.simplify.apply le_S.assumption.
962 theorem lt_minus_m: \forall n,m:nat. O < n \to O < m \to n-m \lt n.
963 intros.apply (lt_O_n_elim n H).intro.
964 apply (lt_O_n_elim m H1).intro.
965 simplify.unfold lt.apply le_S_S.apply le_minus_m.
968 theorem minus_le_O_to_le: \forall n,m:nat. n-m \leq O \to n \leq m.
970 apply (nat_elim2 (\lambda n,m:nat.n-m \leq O \to n \leq m)).
972 simplify.intros. assumption.
973 simplify.intros.apply le_S_S.apply H.assumption.
977 (* monotonicity and galois *)
979 ntheorem monotonic_le_minus_l:
980 ∀p,q,n:nat. q ≤ p → q-n ≤ p-n.
981 napply nat_elim2; #p; #q;
982 ##[#lePO; napply (le_n_O_elim ? lePO);//;
984 ##|#Hind; #n; ncases n;
986 ##|#a; #leSS; napply Hind; /2/;
991 ntheorem le_minus_to_plus: ∀n,m,p. n-m ≤ p → n≤ p+m.
993 napply transitive_le;
994 ##[##|napply le_plus_minus_m_m
995 ##|napply monotonic_le_plus_l;//;
999 ntheorem le_plus_to_minus: ∀n,m,p. n ≤ p+m → n-m ≤ p.
1002 napplyS monotonic_le_minus_l;//;
1006 ntheorem monotonic_le_minus_r:
1007 ∀p,q,n:nat. q ≤ p → n-p ≤ n-q.
1009 napply le_plus_to_minus;
1010 napply (transitive_le ??? (le_plus_minus_m_m ? q));/2/;
1013 (*********************** boolean arithmetics ********************)
1014 include "basics/bool.ma".
1018 [ O ⇒ match m with [ O ⇒ true | S q ⇒ false]
1019 | S p ⇒ match m with [ O ⇒ false | S q ⇒ eqb p q]
1023 ntheorem eqb_to_Prop: ∀n,m:nat.
1024 match (eqb n m) with
1025 [ true \Rightarrow n = m
1026 | false \Rightarrow n \neq m].
1029 (\lambda n,m:nat.match (eqb n m) with
1030 [ true \Rightarrow n = m
1031 | false \Rightarrow n \neq m])).
1033 simplify.reflexivity.
1034 simplify.apply not_eq_O_S.
1036 simplify.unfold Not.
1037 intro. apply (not_eq_O_S n1).apply sym_eq.assumption.
1039 generalize in match H.
1041 simplify.apply eq_f.apply H1.
1042 simplify.unfold Not.intro.apply H1.apply inj_S.assumption.
1046 ntheorem eqb_elim : ∀ n,m:nat.∀ P:bool → Prop.
1047 (n=m → (P true)) → (n ≠ m → (P false)) → (P (eqb n m)).
1049 ##[#n; ncases n; nnormalize; /3/;
1055 ntheorem eqb_n_n: ∀n. eqb n n = true.
1056 #n; nelim n; nnormalize; //.
1059 ntheorem eqb_true_to_eq: ∀n,m:nat. eqb n m = true → n = m.
1060 #n; #m; napply (eqb_elim n m);//;
1061 #_; #abs; napply False_ind; /2/;
1064 ntheorem eqb_false_to_not_eq: ∀n,m:nat. eqb n m = false → n ≠ m.
1065 #n; #m; napply (eqb_elim n m);/2/;
1068 ntheorem eq_to_eqb_true: ∀n,m:nat.
1069 n = m → eqb n m = true.
1072 ntheorem not_eq_to_eqb_false: ∀n,m:nat.
1073 n ≠ m → eqb n m = false.
1076 #Heq; napply False_ind; /2/;
1085 | (S q) ⇒ leb p q]].
1087 ntheorem leb_elim: ∀n,m:nat. ∀P:bool → Prop.
1088 (n ≤ m → P true) → (n ≰ m → P false) → P (leb n m).
1089 napply nat_elim2; nnormalize;
1092 ##|#n; #m; #Hind; #P; #Pt; #Pf; napply Hind;
1093 ##[#lenm; napply Pt; napply le_S_S;//;
1094 ##|#nlenm; napply Pf; /2/;
1099 ntheorem leb_true_to_le:∀n,m.leb n m = true → n ≤ m.
1100 #n; #m; napply leb_elim;
1102 ##|#_; #abs; napply False_ind; /2/;
1106 ntheorem leb_false_to_not_le:∀n,m.
1107 leb n m = false → n ≰ m.
1108 #n; #m; napply leb_elim;
1109 ##[#_; #abs; napply False_ind; /2/;
1114 ntheorem le_to_leb_true: ∀n,m. n ≤ m → leb n m = true.
1115 #n; #m; napply leb_elim; //;
1116 #H; #H1; napply False_ind; /2/;
1119 ntheorem not_le_to_leb_false: ∀n,m. n ≰ m → leb n m = false.
1120 #n; #m; napply leb_elim; //;
1121 #H; #H1; napply False_ind; /2/;
1124 ntheorem lt_to_leb_false: ∀n,m. m < n → leb n m = false.
1128 ndefinition ltb ≝λn,m. leb (S n) m.
1130 ntheorem ltb_elim: ∀n,m:nat. ∀P:bool → Prop.
1131 (n < m → P true) → (n ≮ m → P false) → P (ltb n m).
1132 #n; #m; #P; #Hlt; #Hnlt;
1133 napply leb_elim; /3/; nqed.
1135 ntheorem ltb_true_to_lt:∀n,m.ltb n m = true → n < m.
1136 #n; #m; #Hltb; napply leb_true_to_le; nassumption;
1139 ntheorem ltb_false_to_not_lt:∀n,m.
1140 ltb n m = false → n ≮ m.
1141 #n; #m; #Hltb; napply leb_false_to_not_le; nassumption;
1144 ntheorem lt_to_ltb_true: ∀n,m. n < m → ltb n m = true.
1145 #n; #m; #Hltb; napply le_to_leb_true; nassumption;
1148 ntheorem le_to_ltb_false: ∀n,m. m \le n → ltb n m = false.
1149 #n; #m; #Hltb; napply lt_to_leb_false; /2/;