]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/nlibrary/arithmetics/nat.ma
28764884064208238dd7966591bdfbd5ca53e56f
[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 : \forall 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 (*
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 ndefinition ge: nat \to nat \to Prop \def
224 \lambda n,m:nat.m \leq n.
225
226 interpretation "natural 'greater or equal to'" 'geq x y = (ge x y).
227
228 ndefinition gt: nat \to nat \to Prop \def
229 \lambda n,m:nat.m<n.
230
231 interpretation "natural 'greater than'" 'gt x y = (gt x y).
232
233 interpretation "natural 'not greater than'" 'ngtr x y = (Not (gt x y)).
234
235 ntheorem transitive_le : transitive nat le.
236 #a; #b; #c; #leab; #lebc;nelim lebc;/2/;
237 nqed.
238
239 (*
240 ntheorem trans_le: \forall n,m,p:nat. n \leq m \to m \leq p \to n \leq p
241 \def transitive_le. *)
242
243 ntheorem transitive_lt: transitive nat lt.
244 #a; #b; #c; #ltab; #ltbc;nelim ltbc;/2/;nqed.
245
246 (*
247 theorem trans_lt: \forall n,m,p:nat. lt n m \to lt m p \to lt n p
248 \def transitive_lt. *)
249
250 ntheorem le_S_S: ∀n,m:nat. n ≤ m → S n ≤ S m.
251 #n; #m; #lenm; nelim lenm; /2/; nqed.
252
253 ntheorem le_O_n : ∀n:nat. O ≤ n.
254 #n; nelim n; /2/; nqed.
255
256 ntheorem le_n_Sn : ∀n:nat. n ≤ S n.
257 /2/; nqed.
258
259 ntheorem le_pred_n : ∀n:nat. pred n ≤ n.
260 #n; nelim n; //; nqed.
261
262 ntheorem monotonic_pred: monotonic ? le pred.
263 #n; #m; #lenm; nelim lenm; /2/; nqed.
264
265 ntheorem le_S_S_to_le: ∀n,m:nat. S n ≤ S m → n ≤ m.
266 /2/; nqed.
267
268 ntheorem lt_S_S_to_lt: ∀n,m. S n < S m \to n < m.
269 /2/; nqed. 
270
271 ntheorem lt_to_lt_S_S: ∀n,m. n < m → S n < S m.
272 /2/; nqed.
273
274 ntheorem lt_to_not_zero : ∀n,m:nat. n < m → not_zero m.
275 #n; #m; #Hlt; nelim Hlt;//; nqed.
276
277 (* lt vs. le *)
278 ntheorem not_le_Sn_O: ∀ n:nat. S n ≰ O.
279 #n; #Hlen0; napply (lt_to_not_zero ?? Hlen0); nqed.
280
281 ntheorem not_le_to_not_le_S_S: ∀ n,m:nat. n ≰ m → S n ≰ S m.
282 /3/; nqed.
283
284 ntheorem not_le_S_S_to_not_le: ∀ n,m:nat. S n ≰ S m → n ≰ m.
285 /3/; nqed.
286
287 ntheorem decidable_le: ∀n,m. decidable (n≤m).
288 napply nat_elim2; #n; /2/;
289 #m; #dec; ncases dec;/3/; nqed.
290
291 ntheorem decidable_lt: ∀n,m. decidable (n < m).
292 #n; #m; napply decidable_le ; nqed.
293
294 ntheorem not_le_Sn_n: ∀n:nat. S n ≰ n.
295 #n; nelim n; /2/; nqed.
296
297 ntheorem lt_S_to_le: ∀n,m:nat. n < S m → n ≤ m.
298 /2/; nqed.
299
300 ntheorem not_le_to_lt: ∀n,m. n ≰ m → m < n.
301 napply nat_elim2; #n;
302  ##[#abs; napply False_ind;/2/;
303  ##|/2/;
304  ##|#m;#Hind;#HnotleSS; napply lt_to_lt_S_S;/3/;
305  ##]
306 nqed.
307
308 ntheorem lt_to_not_le: ∀n,m. n < m → m ≰ n.
309 #n; #m; #Hltnm; nelim Hltnm;/3/; nqed.
310
311 ntheorem not_lt_to_le: ∀n,m:nat. n ≮ m → m ≤ n.
312 #n; #m; #Hnlt; napply lt_S_to_le;
313 (* something strange here: /2/ fails *)
314 napply not_le_to_lt; napply Hnlt; nqed. 
315
316 ntheorem le_to_not_lt: ∀n,m:nat. n ≤ m → m ≮ n.
317 /2/; nqed.
318
319 (* lt and le trans *)
320
321 ntheorem lt_to_le_to_lt: ∀n,m,p:nat. n < m → m ≤ p → n < p.
322 #n; #m; #p; #H; #H1; nelim H1; /2/; nqed.
323
324 ntheorem le_to_lt_to_lt: ∀n,m,p:nat. n ≤ m → m < p → n < p.
325 #n; #m; #p; #H; nelim H; /3/; nqed.
326
327 ntheorem lt_S_to_lt: ∀n,m. S n < m → n < m.
328 /2/; nqed.
329
330 ntheorem ltn_to_ltO: ∀n,m:nat. n < m → O < m.
331 /2/; nqed.
332
333 (*
334 theorem lt_SO_n_to_lt_O_pred_n: \forall n:nat.
335 (S O) \lt n \to O \lt (pred n).
336 intros.
337 apply (ltn_to_ltO (pred (S O)) (pred n) ?).
338  apply (lt_pred (S O) n);
339  [ apply (lt_O_S O) 
340  | assumption
341  ]
342 qed. *)
343
344 ntheorem lt_O_n_elim: ∀n:nat. O < n → 
345   ∀P:nat → Prop.(∀m:nat.P (S m)) → P n.
346 #n; nelim n; //; #abs; napply False_ind; /2/; nqed.
347
348 (*
349 theorem lt_pred: \forall n,m. 
350   O < n \to n < m \to pred n < pred m. 
351 apply nat_elim2
352   [intros.apply False_ind.apply (not_le_Sn_O ? H)
353   |intros.apply False_ind.apply (not_le_Sn_O ? H1)
354   |intros.simplify.unfold.apply le_S_S_to_le.assumption
355   ]
356 qed.
357
358 theorem S_pred: \forall n:nat.lt O n \to eq nat n (S (pred n)).
359 intro.elim n.apply False_ind.exact (not_le_Sn_O O H).
360 apply eq_f.apply pred_Sn.
361 qed.
362
363 theorem le_pred_to_le:
364  ∀n,m. O < m → pred n ≤ pred m → n ≤ m.
365 intros 2;
366 elim n;
367 [ apply le_O_n
368 | simplify in H2;
369   rewrite > (S_pred m);
370   [ apply le_S_S;
371     assumption
372   | assumption
373   ]
374 ].
375 qed.
376
377 *)
378
379 (* le to lt or eq *)
380 ntheorem le_to_or_lt_eq: ∀n,m:nat. n ≤ m → n < m ∨ n = m.
381 #n; #m; #lenm; nelim lenm; /3/; nqed.
382
383 (* not eq *)
384 ntheorem lt_to_not_eq : ∀n,m:nat. n < m → n ≠ m.
385 /2/; nqed.
386
387 (*not lt 
388 ntheorem eq_to_not_lt: ∀a,b:nat. a = b → a ≮ b.
389 intros.
390 unfold Not.
391 intros.
392 rewrite > H in H1.
393 apply (lt_to_not_eq b b)
394 [ assumption
395 | reflexivity
396 ]
397 qed. 
398
399 theorem lt_n_m_to_not_lt_m_Sn: ∀n,m. n < m → m ≮ S n.
400 intros;
401 unfold Not;
402 intro;
403 unfold lt in H;
404 unfold lt in H1;
405 generalize in match (le_S_S ? ? H);
406 intro;
407 generalize in match (transitive_le ? ? ? H2 H1);
408 intro;
409 apply (not_le_Sn_n ? H3).
410 qed. *)
411
412 ntheorem not_eq_to_le_to_lt: ∀n,m. n≠m → n≤m → n<m.
413 #n; #m; #Hneq; #Hle; ncases (le_to_or_lt_eq ?? Hle); //;
414 #Heq; nelim (Hneq Heq); nqed.
415
416 (* le elimination *)
417 ntheorem le_n_O_to_eq : ∀n:nat. n ≤ O → O=n.
418 #n; ncases n; //; #a ; #abs; nelim (not_le_Sn_O ? abs); nqed.
419
420 ntheorem le_n_O_elim: ∀n:nat. n ≤ O → ∀P: nat →Prop. P O → P n.
421 #n; ncases n; //; #a; #abs; nelim (not_le_Sn_O ? abs); nqed. 
422
423 ntheorem le_n_Sm_elim : ∀n,m:nat.n ≤ S m → 
424 ∀P:Prop. (S n ≤ S m → P) → (n=S m → P) → P.
425 #n; #m; #Hle; #P; nelim Hle; /3/; nqed.
426
427 (* le and eq *)
428
429 ntheorem le_to_le_to_eq: ∀n,m. n ≤ m → m ≤ n → n = m.
430 napply nat_elim2; /3/; nqed.
431
432 ntheorem lt_O_S : \forall n:nat. O < S n.
433 /2/; nqed.
434
435 (*
436 (* other abstract properties *)
437 theorem antisymmetric_le : antisymmetric nat le.
438 unfold antisymmetric.intros 2.
439 apply (nat_elim2 (\lambda n,m.(n \leq m \to m \leq n \to n=m))).
440 intros.apply le_n_O_to_eq.assumption.
441 intros.apply False_ind.apply (not_le_Sn_O ? H).
442 intros.apply eq_f.apply H.
443 apply le_S_S_to_le.assumption.
444 apply le_S_S_to_le.assumption.
445 qed.
446
447 theorem antisym_le: \forall n,m:nat. n \leq m \to m \leq n \to n=m
448 \def antisymmetric_le.
449
450 theorem le_n_m_to_lt_m_Sn_to_eq_n_m: ∀n,m. n ≤ m → m < S n → n=m.
451 intros;
452 unfold lt in H1;
453 generalize in match (le_S_S_to_le ? ? H1);
454 intro;
455 apply antisym_le;
456 assumption.
457 qed.
458 *)
459
460 (* well founded induction principles *)
461
462 ntheorem nat_elim1 : ∀n:nat.∀P:nat → Prop. 
463 (∀m.(∀p. p < m → P p) → P m) → P n.
464 #n; #P; #H; 
465 ncut (∀q:nat. q ≤ n → P q);/2/;
466 nelim n; 
467  ##[#q; #HleO; (* applica male *) 
468     napply (le_n_O_elim ? HleO);
469     napply H; #p; #ltpO;
470     napply False_ind; /2/; 
471  ##|#p; #Hind; #q; #HleS; 
472     napply H; #a; #lta; napply Hind;
473     napply le_S_S_to_le;/2/;
474  ##]
475 nqed.
476
477 (* some properties of functions *)
478 (*
479 definition increasing \def \lambda f:nat \to nat. 
480 \forall n:nat. f n < f (S n).
481
482 theorem increasing_to_monotonic: \forall f:nat \to nat.
483 increasing f \to monotonic nat lt f.
484 unfold monotonic.unfold lt.unfold increasing.unfold lt.intros.elim H1.apply H.
485 apply (trans_le ? (f n1)).
486 assumption.apply (trans_le ? (S (f n1))).
487 apply le_n_Sn.
488 apply H.
489 qed.
490
491 theorem le_n_fn: \forall f:nat \to nat. (increasing f) 
492 \to \forall n:nat. n \le (f n).
493 intros.elim n.
494 apply le_O_n.
495 apply (trans_le ? (S (f n1))).
496 apply le_S_S.apply H1.
497 simplify in H. unfold increasing in H.unfold lt in H.apply H.
498 qed.
499
500 theorem increasing_to_le: \forall f:nat \to nat. (increasing f) 
501 \to \forall m:nat. \exists i. m \le (f i).
502 intros.elim m.
503 apply (ex_intro ? ? O).apply le_O_n.
504 elim H1.
505 apply (ex_intro ? ? (S a)).
506 apply (trans_le ? (S (f a))).
507 apply le_S_S.assumption.
508 simplify in H.unfold increasing in H.unfold lt in H.
509 apply H.
510 qed.
511
512 theorem increasing_to_le2: \forall f:nat \to nat. (increasing f) 
513 \to \forall m:nat. (f O) \le m \to 
514 \exists i. (f i) \le m \land m <(f (S i)).
515 intros.elim H1.
516 apply (ex_intro ? ? O).
517 split.apply le_n.apply H.
518 elim H3.elim H4.
519 cut ((S n1) < (f (S a)) \lor (S n1) = (f (S a))).
520 elim Hcut.
521 apply (ex_intro ? ? a).
522 split.apply le_S. assumption.assumption.
523 apply (ex_intro ? ? (S a)).
524 split.rewrite < H7.apply le_n.
525 rewrite > H7.
526 apply H.
527 apply le_to_or_lt_eq.apply H6.
528 qed.
529 *)
530
531 (******************* monotonicity ******************************)
532 ntheorem monotonic_le_plus_r: 
533 ∀n:nat.monotonic nat le (λm.n + m).
534 #n; #a; #b; nelim n; nnormalize; //;
535 #m; #H; #leab;napply le_S_S; /2/; nqed.
536
537 ntheorem le_plus_r: ∀p,n,m:nat. n ≤ m → p + n ≤ p + m
538 ≝ monotonic_le_plus_r.
539
540 ntheorem monotonic_le_plus_l: 
541 ∀m:nat.monotonic nat le (λn.n + m).
542 /2/; nqed.
543
544 ntheorem le_plus_l: \forall p,n,m:nat. n \le m \to n + p \le m + p
545 \def monotonic_le_plus_l. 
546
547 ntheorem le_plus: ∀n1,n2,m1,m2:nat. n1 ≤ n2  \to m1 ≤ m2 
548 → n1 + m1 ≤ n2 + m2.
549 #n1; #n2; #m1; #m2; #len; #lem; napply transitive_le;
550 /2/; nqed.
551
552 ntheorem le_plus_n :∀n,m:nat. m ≤ n + m.
553 /2/; nqed. 
554
555 ntheorem le_plus_n_r :∀n,m:nat. m ≤ m + n.
556 /2/; nqed.
557
558 ntheorem eq_plus_to_le: ∀n,m,p:nat.n=m+p → m ≤ n.
559 //; nqed.
560
561 ntheorem le_plus_to_le: ∀a,n,m. a + n ≤ a + m → n ≤ m.
562 #a; nelim a; /3/; nqed. 
563
564 (* times *)
565 ntheorem monotonic_le_times_r: 
566 ∀n:nat.monotonic nat le (λm. n * m).
567 #n; #x; #y; #lexy; nelim n; nnormalize;//;
568 #a; #lea; napply le_plus;//; (* lentissimo /2/ *)
569 nqed.
570
571 (*
572 ntheorem le_times_r: \forall p,n,m:nat. n \le m \to p*n \le p*m
573 \def monotonic_le_times_r. *)
574
575 ntheorem monotonic_le_times_l: 
576 ∀m:nat.monotonic nat le (λn.n*m).
577 /2/; nqed.
578
579 (*
580 theorem le_times_l: \forall p,n,m:nat. n \le m \to n*p \le m*p
581 \def monotonic_le_times_l. *)
582
583 ntheorem le_times: ∀n1,n2,m1,m2:nat. 
584 n1 ≤ n2  → m1 ≤ m2 → n1*m1 ≤ n2*m2.
585 #n1; #n2; #m1; #m2; #len; #lem; 
586 napply transitive_le; (* too slow *)
587  ##[ ##| napply monotonic_le_times_l;//; 
588      ##| napply monotonic_le_times_r;//;
589  ##]
590 nqed.
591
592 ntheorem lt_times_n: ∀n,m:nat. O < n → m ≤ n*m.
593 (* bello *)
594 /2/; nqed.
595
596 ntheorem le_times_to_le: 
597 ∀a,n,m. O < a → a * n ≤ a * m → n ≤ m.
598 #a; napply nat_elim2; nnormalize;
599   ##[//;
600   ##|#n; #H1; #H2; napply False_ind;
601      ngeneralize in match H2;
602      napply lt_to_not_le;
603      napply (transitive_le ? (S n));/2/;
604   ##|#n; #m; #H; #lta; #le;
605      napply le_S_S; napply H; /2/;
606   ##]
607 nqed.
608
609 ntheorem le_S_times_2: ∀n,m.O < m → n ≤ m → n < 2*m.
610 #n; #m; #posm; #lenm; (* interessante *)
611 nnormalize; napplyS (le_plus n); //; nqed.