]> matita.cs.unibo.it Git - helm.git/blob - matita/library/nat/lt_arith.ma
a568ca408be94412a0203a82abf170709edfaa05
[helm.git] / matita / library / nat / lt_arith.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 set "baseuri" "cic:/matita/nat/lt_arith".
16
17 include "nat/div_and_mod.ma".
18
19 (* plus *)
20 theorem monotonic_lt_plus_r: 
21 \forall n:nat.monotonic nat lt (\lambda m.n+m).
22 simplify.intros.
23 elim n.simplify.assumption.
24 simplify.unfold lt.
25 apply le_S_S.assumption.
26 qed.
27
28 variant lt_plus_r: \forall n,p,q:nat. p < q \to n + p < n + q \def
29 monotonic_lt_plus_r.
30
31 theorem monotonic_lt_plus_l: 
32 \forall n:nat.monotonic nat lt (\lambda m.m+n).
33 simplify.
34 intros.
35 rewrite < sym_plus. rewrite < (sym_plus n).
36 apply lt_plus_r.assumption.
37 qed.
38
39 variant lt_plus_l: \forall n,p,q:nat. p < q \to p + n < q + n \def
40 monotonic_lt_plus_l.
41
42 theorem lt_plus: \forall n,m,p,q:nat. n < m \to p < q \to n + p < m + q.
43 intros.
44 apply (trans_lt ? (n + q)).
45 apply lt_plus_r.assumption.
46 apply lt_plus_l.assumption.
47 qed.
48
49 theorem lt_plus_to_lt_l :\forall n,p,q:nat. p+n < q+n \to p<q.
50 intro.elim n.
51 rewrite > plus_n_O.
52 rewrite > (plus_n_O q).assumption.
53 apply H.
54 unfold lt.apply le_S_S_to_le.
55 rewrite > plus_n_Sm.
56 rewrite > (plus_n_Sm q).
57 exact H1.
58 qed.
59
60 theorem lt_plus_to_lt_r :\forall n,p,q:nat. n+p < n+q \to p<q.
61 intros.apply (lt_plus_to_lt_l n). 
62 rewrite > sym_plus.
63 rewrite > (sym_plus q).assumption.
64 qed.
65
66 theorem le_to_lt_to_plus_lt: \forall a,b,c,d:nat.
67 a \le c \to b \lt d \to (a + b) \lt (c+d).
68 intros.
69 cut (a \lt c \lor a = c)
70 [ elim Hcut
71   [ apply (lt_plus );
72       assumption
73   | rewrite > H2.
74     apply (lt_plus_r c b d).
75     assumption
76   ]
77 | apply le_to_or_lt_eq.
78   assumption
79 ]
80 qed.
81
82
83 (* times and zero *)
84 theorem lt_O_times_S_S: \forall n,m:nat.O < (S n)*(S m).
85 intros.simplify.unfold lt.apply le_S_S.apply le_O_n.
86 qed.
87
88 theorem lt_times_eq_O: \forall a,b:nat.
89 O \lt a \to (a * b) = O \to b = O.
90 intros.
91 apply (nat_case1 b)
92 [ intros.
93   reflexivity
94 | intros.
95   rewrite > H2 in H1.
96   rewrite > (S_pred a) in H1
97   [ apply False_ind.
98     apply (eq_to_not_lt O ((S (pred a))*(S m)))
99     [ apply sym_eq.
100       assumption
101     | apply lt_O_times_S_S
102     ]
103   | assumption
104   ]
105 ]
106 qed.
107
108 theorem O_lt_times_to_O_lt: \forall a,c:nat.
109 O \lt (a * c) \to O \lt a.
110 intros.
111 apply (nat_case1 a)
112 [ intros.
113   rewrite > H1 in H.
114   simplify in H.
115   assumption
116 | intros.
117   apply lt_O_S
118 ]
119 qed.
120
121 (* times *)
122 theorem monotonic_lt_times_r: 
123 \forall n:nat.monotonic nat lt (\lambda m.(S n)*m).
124 simplify.
125 intros.elim n.
126 simplify.rewrite < plus_n_O.rewrite < plus_n_O.assumption.
127 apply lt_plus.assumption.assumption.
128 qed.
129
130 (* a simple variant of the previus monotionic_lt_times *)
131 theorem monotonic_lt_times_variant: \forall c:nat.
132 O \lt c \to monotonic nat lt (\lambda t.(t*c)).
133 intros.
134 apply (increasing_to_monotonic).
135 unfold increasing.
136 intros.
137 simplify.
138 rewrite > sym_plus.
139 rewrite > plus_n_O in \vdash (? % ?).
140 apply lt_plus_r.
141 assumption.
142 qed.
143
144 theorem lt_times_r: \forall n,p,q:nat. p < q \to (S n) * p < (S n) * q
145 \def monotonic_lt_times_r.
146
147 theorem monotonic_lt_times_l: 
148 \forall m:nat.monotonic nat lt (\lambda n.n * (S m)).
149 simplify.
150 intros.
151 rewrite < sym_times.rewrite < (sym_times (S m)).
152 apply lt_times_r.assumption.
153 qed.
154
155 variant lt_times_l: \forall n,p,q:nat. p<q \to p*(S n) < q*(S n)
156 \def monotonic_lt_times_l.
157
158 theorem lt_times:\forall n,m,p,q:nat. n<m \to p<q \to n*p < m*q.
159 intro.
160 elim n.
161 apply (lt_O_n_elim m H).
162 intro.
163 cut (lt O q).
164 apply (lt_O_n_elim q Hcut).
165 intro.change with (O < (S m1)*(S m2)).
166 apply lt_O_times_S_S.
167 apply (ltn_to_ltO p q H1).
168 apply (trans_lt ? ((S n1)*q)).
169 apply lt_times_r.assumption.
170 cut (lt O q).
171 apply (lt_O_n_elim q Hcut).
172 intro.
173 apply lt_times_l.
174 assumption.
175 apply (ltn_to_ltO p q H2).
176 qed.
177
178 theorem lt_times_r1: 
179 \forall n,m,p. O < n \to m < p \to n*m < n*p.
180 intros.
181 elim H;apply lt_times_r;assumption.
182 qed.
183
184 theorem lt_times_l1: 
185 \forall n,m,p. O < n \to m < p \to m*n < p*n.
186 intros.
187 elim H;apply lt_times_l;assumption.
188 qed.
189
190 theorem lt_to_le_to_lt_times : 
191 \forall n,n1,m,m1. n < n1 \to m \le m1 \to O < m1 \to n*m < n1*m1.
192 intros.
193 apply (le_to_lt_to_lt ? (n*m1))
194   [apply le_times_r.assumption
195   |apply lt_times_l1
196     [assumption|assumption]
197   ]
198 qed.
199
200 theorem lt_times_to_lt_l: 
201 \forall n,p,q:nat. p*(S n) < q*(S n) \to p < q.
202 intros.
203 cut (p < q \lor p \nlt q).
204 elim Hcut.
205 assumption.
206 absurd (p * (S n) < q * (S n)).
207 assumption.
208 apply le_to_not_lt.
209 apply le_times_l.
210 apply not_lt_to_le.
211 assumption.
212 exact (decidable_lt p q).
213 qed.
214
215 theorem lt_times_n_to_lt: 
216 \forall n,p,q:nat. O < n \to p*n < q*n \to p < q.
217 intro.
218 apply (nat_case n)
219   [intros.apply False_ind.apply (not_le_Sn_n ? H)
220   |intros 4.apply lt_times_to_lt_l
221   ]
222 qed.
223
224 theorem lt_times_to_lt_r: 
225 \forall n,p,q:nat. (S n)*p < (S n)*q \to lt p q.
226 intros.
227 apply (lt_times_to_lt_l n).
228 rewrite < sym_times.
229 rewrite < (sym_times (S n)).
230 assumption.
231 qed.
232
233 theorem lt_times_n_to_lt_r: 
234 \forall n,p,q:nat. O < n \to n*p < n*q \to lt p q.
235 intro.
236 apply (nat_case n)
237   [intros.apply False_ind.apply (not_le_Sn_n ? H)
238   |intros 4.apply lt_times_to_lt_r
239   ]
240 qed.
241
242 theorem nat_compare_times_l : \forall n,p,q:nat. 
243 nat_compare p q = nat_compare ((S n) * p) ((S n) * q).
244 intros.apply nat_compare_elim.intro.
245 apply nat_compare_elim.
246 intro.reflexivity.
247 intro.absurd (p=q).
248 apply (inj_times_r n).assumption.
249 apply lt_to_not_eq. assumption.
250 intro.absurd (q<p).
251 apply (lt_times_to_lt_r n).assumption.
252 apply le_to_not_lt.apply lt_to_le.assumption.
253 intro.rewrite < H.rewrite > nat_compare_n_n.reflexivity.
254 intro.apply nat_compare_elim.intro.
255 absurd (p<q).
256 apply (lt_times_to_lt_r n).assumption.
257 apply le_to_not_lt.apply lt_to_le.assumption.
258 intro.absurd (q=p).
259 symmetry.
260 apply (inj_times_r n).assumption.
261 apply lt_to_not_eq.assumption.
262 intro.reflexivity.
263 qed.
264
265 (* times and plus *)
266 theorem lt_times_plus_times: \forall a,b,n,m:nat. 
267 a < n \to b < m \to a*m + b < n*m.
268 intros 3.
269 apply (nat_case n)
270   [intros.apply False_ind.apply (not_le_Sn_O ? H)
271   |intros.simplify.
272    rewrite < sym_plus.
273    unfold.
274    change with (S b+a*m1 \leq m1+m*m1).
275    apply le_plus
276     [assumption
277     |apply le_times
278       [apply le_S_S_to_le.assumption
279       |apply le_n
280       ]
281     ]
282   ]
283 qed.
284
285 (* div *) 
286
287 theorem eq_mod_O_to_lt_O_div: \forall n,m:nat. O < m \to O < n\to n \mod m = O \to O < n / m. 
288 intros 4.apply (lt_O_n_elim m H).intros.
289 apply (lt_times_to_lt_r m1).
290 rewrite < times_n_O.
291 rewrite > (plus_n_O ((S m1)*(n / (S m1)))).
292 rewrite < H2.
293 rewrite < sym_times.
294 rewrite < div_mod.
295 rewrite > H2.
296 assumption.
297 unfold lt.apply le_S_S.apply le_O_n.
298 qed.
299
300 theorem lt_div_n_m_n: \forall n,m:nat. (S O) < m \to O < n \to n / m \lt n.
301 intros.
302 apply (nat_case1 (n / m)).intro.
303 assumption.intros.rewrite < H2.
304 rewrite > (div_mod n m) in \vdash (? ? %).
305 apply (lt_to_le_to_lt ? ((n / m)*m)).
306 apply (lt_to_le_to_lt ? ((n / m)*(S (S O)))).
307 rewrite < sym_times.
308 rewrite > H2.
309 simplify.unfold lt.
310 rewrite < plus_n_O.
311 rewrite < plus_n_Sm.
312 apply le_S_S.
313 apply le_S_S.
314 apply le_plus_n.
315 apply le_times_r.
316 assumption.
317 rewrite < sym_plus.
318 apply le_plus_n.
319 apply (trans_lt ? (S O)).
320 unfold lt. apply le_n.assumption.
321 qed.
322
323
324 (* Forall a,b : N. 0 < b \to b * (a/b) <= a < b * (a/b +1) *)
325 (* The theorem is shown in two different parts: *)
326
327 theorem lt_to_div_to_and_le_times_lt_S: \forall a,b,c:nat.
328 O \lt b \to a/b = c \to (b*c \le a \land a \lt b*(S c)).
329 intros.
330 split
331 [ rewrite < H1.
332   rewrite > sym_times.
333   rewrite > eq_times_div_minus_mod
334   [ apply (le_minus_m a (a \mod b))
335   | assumption
336   ]
337 | rewrite < (times_n_Sm b c).
338   rewrite < H1.
339   rewrite > sym_times.
340   rewrite > (div_mod a b) in \vdash (? % ?)
341   [ rewrite > (sym_plus b ((a/b)*b)).
342     apply lt_plus_r.
343     apply lt_mod_m_m.
344     assumption
345   | assumption
346   ]
347 ]
348 qed.
349
350 theorem lt_to_le_times_to_lt_S_to_div: \forall a,c,b:nat.
351 O \lt b \to (b*c) \le a \to a \lt (b*(S c)) \to a/b = c.
352 intros.
353 apply (le_to_le_to_eq)
354 [ apply (leb_elim (a/b) c);intros
355   [ assumption
356   | cut (c \lt (a/b))
357     [ apply False_ind.
358       apply (lt_to_not_le (a \mod b) O)
359       [ apply (lt_plus_to_lt_l ((a/b)*b)).
360         simplify.
361         rewrite < sym_plus.
362         rewrite < div_mod
363         [ apply (lt_to_le_to_lt ? (b*(S c)) ?)
364           [ assumption
365           | rewrite > (sym_times (a/b) b).
366             apply le_times_r.
367             assumption
368           ]
369         | assumption
370         ]
371       | apply le_O_n
372       ]
373     | apply not_le_to_lt.
374       assumption
375     ]
376   ]
377 | apply (leb_elim c (a/b));intros
378   [ assumption
379   | cut((a/b) \lt c) 
380     [ apply False_ind.
381       apply (lt_to_not_le (a \mod b) b)
382       [ apply (lt_mod_m_m).
383         assumption
384       | apply (le_plus_to_le ((a/b)*b)).
385         rewrite < (div_mod a b)
386         [ apply (trans_le ? (b*c) ?)
387           [ rewrite > (sym_times (a/b) b).
388             rewrite > (times_n_SO b) in \vdash (? (? ? %) ?).
389             rewrite < distr_times_plus.
390             rewrite > sym_plus.
391             simplify in \vdash (? (? ? %) ?).
392             apply le_times_r.
393             assumption
394           | assumption
395           ]
396         | assumption
397         ]
398       ]
399     | apply not_le_to_lt. 
400       assumption
401     ]
402   ]
403 ]
404 qed.
405
406
407 theorem lt_to_lt_to_eq_div_div_times_times: \forall a,b,c:nat. 
408 O \lt c \to O \lt b \to (a/b) = (a*c)/(b*c).
409 intros.
410 apply sym_eq.
411 cut (b*(a/b) \le a \land a \lt b*(S (a/b)))
412 [ elim Hcut.
413   apply lt_to_le_times_to_lt_S_to_div
414   [ rewrite > (S_pred b)
415     [ rewrite > (S_pred c)
416       [ apply (lt_O_times_S_S)
417       | assumption
418       ]
419     | assumption
420     ]
421   | rewrite > assoc_times.
422     rewrite > (sym_times c (a/b)).
423     rewrite < assoc_times.
424     rewrite > (sym_times (b*(a/b)) c).
425     rewrite > (sym_times a c).
426     apply (le_times_r c (b*(a/b)) a).
427     assumption
428   | rewrite > (sym_times a c).
429     rewrite > (assoc_times ).
430     rewrite > (sym_times c (S (a/b))).
431     rewrite < (assoc_times).
432     rewrite > (sym_times (b*(S (a/b))) c).
433     apply (lt_times_r1 c a (b*(S (a/b))));
434       assumption    
435   ]
436 | apply (lt_to_div_to_and_le_times_lt_S)
437   [ assumption
438   | reflexivity
439   ]
440 ]
441 qed.
442
443 theorem times_mod: \forall a,b,c:nat.
444 O \lt c \to O \lt b \to ((a*c) \mod (b*c)) = c*(a\mod b).
445 intros.
446 apply (div_mod_spec_to_eq2 (a*c) (b*c) (a/b) ((a*c) \mod (b*c)) (a/b) (c*(a \mod b)))
447 [ rewrite > (lt_to_lt_to_eq_div_div_times_times a b c)
448   [ apply div_mod_spec_div_mod.
449     rewrite > (S_pred b)
450     [ rewrite > (S_pred c)
451       [ apply lt_O_times_S_S
452       | assumption
453       ]
454     | assumption
455     ]
456   | assumption
457   | assumption
458   ]
459 | apply div_mod_spec_intro
460   [ rewrite > (sym_times b c).
461     apply (lt_times_r1 c)
462     [ assumption
463     | apply (lt_mod_m_m).
464       assumption
465     ]
466   | rewrite < (assoc_times (a/b) b c).
467     rewrite > (sym_times a c).
468     rewrite > (sym_times ((a/b)*b) c).
469     rewrite < (distr_times_plus c ? ?).
470     apply eq_f.
471     apply (div_mod a b).
472     assumption
473   ]
474 ]
475 qed.
476
477
478
479
480 (* general properties of functions *)
481 theorem monotonic_to_injective: \forall f:nat\to nat.
482 monotonic nat lt f \to injective nat nat f.
483 unfold injective.intros.
484 apply (nat_compare_elim x y).
485 intro.apply False_ind.apply (not_le_Sn_n (f x)).
486 rewrite > H1 in \vdash (? ? %).
487 change with (f x < f y).
488 apply H.apply H2.
489 intros.assumption.
490 intro.apply False_ind.apply (not_le_Sn_n (f y)).
491 rewrite < H1 in \vdash (? ? %).
492 change with (f y < f x).
493 apply H.apply H2.
494 qed.
495
496 theorem increasing_to_injective: \forall f:nat\to nat.
497 increasing f \to injective nat nat f.
498 intros.apply monotonic_to_injective.
499 apply increasing_to_monotonic.assumption.
500 qed.
501