]> matita.cs.unibo.it Git - helm.git/blob - matita/library/nat/propr_div_mod_lt_le_totient1_aux.ma
f26f0ea004f5ce5f5c8d8df6d9c3c56d40ef5fae
[helm.git] / matita / library / nat / propr_div_mod_lt_le_totient1_aux.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 set "baseuri" "cic:/matita/nat/propr_div_mod_lt_le_totient1_aux/".
16
17 include "nat/iteration2.ma".
18
19 (*this part of the library contains some properties useful to prove
20   the main theorem in totient1.ma, and some new properties about gcd
21   (see gcd_properties1.ma).
22   These theorems are saved separately from the other part of the library
23   in order to avoid to create circular dependences in it.
24  *)
25
26 (* some basic properties of and - or*)
27 (*theorem andb_sym: \forall A,B:bool.
28 (A \land B) = (B \land A).
29 intros.
30 elim A;
31   elim B;
32     simplify;
33     reflexivity.
34 qed.
35
36 theorem andb_assoc: \forall A,B,C:bool.
37 (A \land (B \land C)) = ((A \land B) \land C).
38 intros.
39 elim A;
40   elim B;
41     elim C;
42       simplify;
43       reflexivity.
44 qed.
45
46 theorem orb_sym: \forall A,B:bool.
47 (A \lor B) = (B \lor A).
48 intros.
49 elim A;
50   elim B;
51     simplify;
52     reflexivity.
53 qed.
54
55 theorem andb_t_t_t: \forall A,B,C:bool.
56 A = true \to B = true \to C = true \to (A \land (B \land C)) = true.
57 intros.
58 rewrite > H.
59 rewrite > H1.
60 rewrite > H2.
61 reflexivity.
62 qed.
63 *)
64 (*properties about relational operators*)
65
66 theorem Sa_le_b_to_O_lt_b: \forall a,b:nat.
67 (S a) \le b \to O \lt b.
68 intros.
69 elim H;
70   apply lt_O_S.
71 qed.
72
73
74 theorem n_gt_Uno_to_O_lt_pred_n: \forall n:nat.
75 (S O) \lt n \to O \lt (pred n).
76 intros.
77 apply (Sa_le_b_to_O_lt_b (pred (S O)) (pred n) ?).
78  apply (lt_pred (S O) n ? ?);
79  [ apply (lt_O_S O) 
80  | apply (H)
81  ]
82 qed.
83
84
85 theorem NdivM_times_M_to_N: \forall n,m:nat.
86 O \lt m \to m \divides n \to (n / m) * m = n.
87 intros.
88 rewrite > plus_n_O.
89 apply sym_eq.
90 cut (n \mod m = O)
91 [ rewrite < Hcut.
92   apply div_mod.
93   assumption
94 | apply divides_to_mod_O;
95     assumption.
96 ]
97 qed.
98
99 theorem lt_to_divides_to_div_le:  \forall a,c:nat.
100 O \lt c \to c \divides a \to a/c \le a.
101 intros.
102 apply (le_times_to_le c (a/c) a)
103 [ assumption
104 | rewrite > (sym_times c (a/c)).
105   rewrite > (NdivM_times_M_to_N a c) in \vdash (? % ?)
106   [ apply (le_times_n c a ?).   
107     assumption
108   | assumption
109   | assumption
110   ]
111 ]
112 qed.
113
114
115 theorem lt_O_a_lt_O_b_a_divides_b_to_lt_O_aDIVb:
116 \forall a,b:nat.
117 O \lt a \to O \lt b \to a \divides b \to O \lt (b/a).
118 intros.
119 elim H2.
120 rewrite > H3.
121 rewrite > (sym_times a n2).
122 rewrite > (div_times_ltO n2 a)
123 [ apply (divides_to_lt_O n2 b)
124   [ assumption
125   | apply (witness n2 b a).
126     rewrite > sym_times.
127     assumption
128   ]  
129 | assumption  
130 ]
131 qed.
132
133 (* some properties of div operator between natural numbers *)
134
135 theorem separazioneFattoriSuDivisione: \forall a,b,c:nat.
136 O \lt b \to c \divides b \to a * (b /c) = (a*b)/c.
137 intros.
138 elim H1.
139 rewrite > H2.
140 rewrite > (sym_times c n2).
141 cut(O \lt c)
142 [ rewrite > (div_times_ltO n2 c)
143   [ rewrite < assoc_times.
144     rewrite > (div_times_ltO (a *n2) c)
145     [ reflexivity
146     | assumption
147     ]
148   | assumption
149   ]  
150 | apply (divides_to_lt_O c b);
151     assumption.
152 ]
153 qed.
154
155
156 theorem div_mod_minus:
157 \forall a,b:nat. O \lt b \to
158 (a /b)*b = a - (a \mod b).
159 intros.
160 rewrite > (div_mod a b) in \vdash (? ? ? (? % ?))
161 [ apply (minus_plus_m_m (times (div a b) b) (mod a b))
162 | assumption
163 ]
164 qed.
165
166
167 (* A corollary to the division theorem (between natural numbers).
168  *
169  * Forall a,b,c: Nat, b > O,
170  *  a/b = c iff (b*c <= a && a < b*(c+1)
171  *
172  * two parts of the theorem are proved separately
173  *  - (=>) th_div_interi_2
174  *  - (<=) th_div_interi_1
175  *)
176
177 theorem th_div_interi_2: \forall a,b,c:nat.
178 O \lt b \to a/b = c \to (b*c \le a \land a \lt b*(S c)).
179 intros.
180 split
181 [ rewrite < H1.
182   rewrite > sym_times.
183   rewrite > div_mod_minus
184   [ apply (le_minus_m a (a \mod b))
185   | assumption
186   ]
187 | rewrite < (times_n_Sm b c).
188   rewrite < H1.
189   rewrite > sym_times.
190   rewrite > (div_mod a b) in \vdash (? % ?)
191   [ rewrite > (sym_plus b ((a/b)*b)).
192     apply lt_plus_r.
193     apply lt_mod_m_m.
194     assumption
195   | assumption
196   ]
197 ]
198 qed.
199   
200
201 theorem th_div_interi_1: \forall a,c,b:nat.
202 O \lt b \to (b*c) \le a \to a \lt (b*(S c)) \to a/b = c.
203 intros.
204 apply (le_to_le_to_eq)
205 [ apply (leb_elim (a/b) c);intros
206   [ assumption
207   | cut (c \lt (a/b))
208     [ apply False_ind.
209       apply (lt_to_not_le (a \mod b) O)
210       [ apply (lt_plus_to_lt_l ((a/b)*b)).
211         simplify.
212         rewrite < sym_plus.
213         rewrite < div_mod
214         [ apply (lt_to_le_to_lt ? (b*(S c)) ?)
215           [ assumption
216           | rewrite > (sym_times (a/b) b).
217             apply le_times_r.
218             assumption
219           ]
220         | assumption
221         ]
222       | apply le_O_n
223       ]
224     | apply not_le_to_lt.
225       assumption
226     ]
227   ]
228 | apply (leb_elim c (a/b));intros
229   [ assumption
230   | cut((a/b) \lt c) 
231     [ apply False_ind.
232       apply (lt_to_not_le (a \mod b) b)
233       [ apply (lt_mod_m_m).
234         assumption
235       | apply (le_plus_to_le ((a/b)*b)).
236         rewrite < (div_mod a b)
237         [ apply (trans_le ? (b*c) ?)
238           [ rewrite > (sym_times (a/b) b).
239             rewrite > (times_n_SO b) in \vdash (? (? ? %) ?).
240             rewrite < distr_times_plus.
241             rewrite > sym_plus.
242             simplify in \vdash (? (? ? %) ?).
243             apply le_times_r.
244             assumption
245           | assumption
246           ]
247         | assumption
248         ]
249       ]
250     | apply not_le_to_lt. 
251       assumption
252     ]
253   ]
254 ]
255 qed.
256
257
258 theorem times_numerator_denominator_aux: \forall a,b,c,d:nat.
259 O \lt c \to O \lt b \to d = (a/b) \to d= (a*c)/(b*c).
260 intros.
261 apply sym_eq.
262 cut (b*d \le a \land a \lt b*(S d))
263 [ elim Hcut.
264   apply th_div_interi_1
265   [ rewrite > (S_pred b)
266     [ rewrite > (S_pred c)
267       [ apply (lt_O_times_S_S)
268       | assumption
269       ]
270     | assumption
271     ]
272   | rewrite > assoc_times.
273     rewrite > (sym_times c d).
274     rewrite < assoc_times.
275     rewrite > (sym_times (b*d) c).
276     rewrite > (sym_times a c).
277     apply (le_times_r c (b*d) a).
278     assumption
279   | rewrite > (sym_times a c).
280     rewrite > (assoc_times ).
281     rewrite > (sym_times c (S d)).
282     rewrite < (assoc_times).
283     rewrite > (sym_times (b*(S d)) c).
284     apply (lt_times_r1 c a (b*(S d)));
285       assumption    
286   ]
287 | apply (th_div_interi_2)
288   [ assumption
289   | apply sym_eq.
290     assumption
291   ]
292 ]
293 qed.
294
295 theorem times_numerator_denominator: \forall a,b,c:nat. 
296 O \lt c \to O \lt b \to (a/b) = (a*c)/(b*c).
297 intros.
298 apply (times_numerator_denominator_aux a b c (a/b))
299 [ assumption
300 | assumption
301 | reflexivity
302 ]
303 qed.
304
305 theorem times_mod: \forall a,b,c:nat.
306 O \lt c \to O \lt b \to ((a*c) \mod (b*c)) = c*(a\mod b).
307 intros.
308 apply (div_mod_spec_to_eq2 (a*c) (b*c) (a/b) ((a*c) \mod (b*c)) (a/b) (c*(a \mod b)))
309 [ apply div_mod_spec_intro
310   [ apply (lt_mod_m_m (a*c) (b*c)).
311     rewrite > (S_pred b)
312     [ rewrite > (S_pred c)
313       [ apply lt_O_times_S_S
314       | assumption
315       ]
316     | assumption
317     ]
318   | rewrite > (times_numerator_denominator a b c)
319     [ apply (div_mod (a*c) (b*c)).
320       rewrite > (S_pred b)
321       [ rewrite > (S_pred c)
322         [ apply (lt_O_times_S_S)
323         | assumption
324         ]
325       | assumption
326       ]
327     | assumption
328     | assumption
329     ]
330   ]
331 | constructor 1
332   [ rewrite > (sym_times b c).
333     apply (lt_times_r1 c)
334     [ assumption
335     | apply (lt_mod_m_m).
336       assumption
337     ]
338   | rewrite < (assoc_times (a/b) b c).
339     rewrite > (sym_times a c).
340     rewrite > (sym_times ((a/b)*b) c).
341     rewrite < (distr_times_plus c ? ?).
342     apply eq_f.
343     apply (div_mod a b).
344     assumption
345   ]
346 ]
347 qed.
348
349
350
351
352
353
354
355
356
357
358