]> matita.cs.unibo.it Git - helm.git/blob - matita/library/nat/propr_div_mod_lt_le_totient1_aux.ma
7a68a5bd8530f1dbfd8929df7038b89fd4e20779
[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 theorem n_gt_Uno_to_O_lt_pred_n: \forall n:nat.
74 (S O) \lt n \to O \lt (pred n).
75 intros.
76 apply (Sa_le_b_to_O_lt_b (pred (S O)) (pred n) ?).
77  apply (lt_pred (S O) n ? ?);
78  [ apply (lt_O_S O) 
79  | apply (H)
80  ]
81 qed.
82
83
84 theorem NdivM_times_M_to_N: \forall n,m:nat.
85 O \lt m \to m \divides n \to (n / m) * m = n.
86 intros.
87 rewrite > plus_n_O.
88 apply sym_eq.
89 cut (n \mod m = O)
90 [ rewrite < Hcut.
91   apply div_mod.
92   assumption
93 | apply divides_to_mod_O;
94     assumption.
95 ]
96 qed.
97
98 theorem lt_to_divides_to_div_le:  \forall a,c:nat.
99 O \lt c \to c \divides a \to a/c \le a.
100 intros.
101 apply (le_times_to_le c (a/c) a)
102 [ assumption
103 | rewrite > (sym_times c (a/c)).
104   rewrite > (NdivM_times_M_to_N a c) in \vdash (? % ?)
105   [ apply (le_times_n c a ?).   
106     assumption
107   | assumption
108   | assumption
109   ]
110 ]
111 qed.
112
113
114 theorem lt_O_a_lt_O_b_a_divides_b_to_lt_O_aDIVb:
115 \forall a,b:nat.
116 O \lt a \to O \lt b \to a \divides b \to O \lt (b/a).
117 intros.
118 elim H2.
119 rewrite > H3.
120 rewrite > (sym_times a n2).
121 rewrite > (div_times_ltO n2 a)
122 [ apply (divides_to_lt_O n2 b)
123   [ assumption
124   | apply (witness n2 b a).
125     rewrite > sym_times.
126     assumption
127   ]  
128 | assumption  
129 ]
130 qed.
131
132 (* some properties of div operator between natural numbers *)
133
134 theorem separazioneFattoriSuDivisione: \forall a,b,c:nat.
135 O \lt b \to c \divides b \to a * (b /c) = (a*b)/c.
136 intros.
137 elim H1.
138 rewrite > H2.
139 rewrite > (sym_times c n2).
140 cut(O \lt c)
141 [ rewrite > (div_times_ltO n2 c)
142   [ rewrite < assoc_times.
143     rewrite > (div_times_ltO (a *n2) c)
144     [ reflexivity
145     | assumption
146     ]
147   | assumption
148   ]  
149 | apply (divides_to_lt_O c b);
150     assumption.
151 ]
152 qed.
153
154
155 theorem div_mod_minus:
156 \forall a,b:nat. O \lt b \to
157 (a /b)*b = a - (a \mod b).
158 intros.
159 rewrite > (div_mod a b) in \vdash (? ? ? (? % ?))
160 [ apply (minus_plus_m_m (times (div a b) b) (mod a b))
161 | assumption
162 ]
163 qed.
164
165
166 (* A corollary to the division theorem (between natural numbers).
167  *
168  * Forall a,b,c: Nat, b > O,
169  *  a/b = c iff (b*c <= a && a < b*(c+1)
170  *
171  * two parts of the theorem are proved separately
172  *  - (=>) th_div_interi_2
173  *  - (<=) th_div_interi_1
174  *)
175
176 theorem th_div_interi_2: \forall a,b,c:nat.
177 O \lt b \to a/b = c \to (b*c \le a \land a \lt b*(S c)).
178 intros.
179 split
180 [ rewrite < H1.
181   rewrite > sym_times.
182   rewrite > div_mod_minus
183   [ apply (le_minus_m a (a \mod b))
184   | assumption
185   ]
186 | rewrite < (times_n_Sm b c).
187   rewrite < H1.
188   rewrite > sym_times.
189   rewrite > (div_mod a b) in \vdash (? % ?)
190   [ rewrite > (sym_plus b ((a/b)*b)).
191     apply lt_plus_r.
192     apply lt_mod_m_m.
193     assumption
194   | assumption
195   ]
196 ]
197 qed.
198   
199
200 theorem th_div_interi_1: \forall a,c,b:nat.
201 O \lt b \to (b*c) \le a \to a \lt (b*(S c)) \to a/b = c.
202 intros.
203 apply (le_to_le_to_eq)
204 [ apply (leb_elim (a/b) c);intros
205   [ assumption
206   | cut (c \lt (a/b))
207     [ apply False_ind.
208       apply (lt_to_not_le (a \mod b) O)
209       [ apply (lt_plus_to_lt_l ((a/b)*b)).
210         simplify.
211         rewrite < sym_plus.
212         rewrite < div_mod
213         [ apply (lt_to_le_to_lt ? (b*(S c)) ?)
214           [ assumption
215           | rewrite > (sym_times (a/b) b).
216             apply le_times_r.
217             assumption
218           ]
219         | assumption
220         ]
221       | apply le_O_n
222       ]
223     | apply not_le_to_lt.
224       assumption
225     ]
226   ]
227 | apply (leb_elim c (a/b));intros
228   [ assumption
229   | cut((a/b) \lt c) 
230     [ apply False_ind.
231       apply (lt_to_not_le (a \mod b) b)
232       [ apply (lt_mod_m_m).
233         assumption
234       | apply (le_plus_to_le ((a/b)*b)).
235         rewrite < (div_mod a b)
236         [ apply (trans_le ? (b*c) ?)
237           [ rewrite > (sym_times (a/b) b).
238             rewrite > (times_n_SO b) in \vdash (? (? ? %) ?).
239             rewrite < distr_times_plus.
240             rewrite > sym_plus.
241             simplify in \vdash (? (? ? %) ?).
242             apply le_times_r.
243             assumption
244           | assumption
245           ]
246         | assumption
247         ]
248       ]
249     | apply not_le_to_lt. 
250       assumption
251     ]
252   ]
253 ]
254 qed.
255
256
257 theorem times_numerator_denominator_aux: \forall a,b,c,d:nat.
258 O \lt c \to O \lt b \to d = (a/b) \to d= (a*c)/(b*c).
259 intros.
260 apply sym_eq.
261 cut (b*d \le a \land a \lt b*(S d))
262 [ elim Hcut.
263   apply th_div_interi_1
264   [ rewrite > (S_pred b)
265     [ rewrite > (S_pred c)
266       [ apply (lt_O_times_S_S)
267       | assumption
268       ]
269     | assumption
270     ]
271   | rewrite > assoc_times.
272     rewrite > (sym_times c d).
273     rewrite < assoc_times.
274     rewrite > (sym_times (b*d) c).
275     rewrite > (sym_times a c).
276     apply (le_times_r c (b*d) a).
277     assumption
278   | rewrite > (sym_times a c).
279     rewrite > (assoc_times ).
280     rewrite > (sym_times c (S d)).
281     rewrite < (assoc_times).
282     rewrite > (sym_times (b*(S d)) c).
283     apply (lt_times_r1 c a (b*(S d)));
284       assumption    
285   ]
286 | apply (th_div_interi_2)
287   [ assumption
288   | apply sym_eq.
289     assumption
290   ]
291 ]
292 qed.
293
294 theorem times_numerator_denominator: \forall a,b,c:nat. 
295 O \lt c \to O \lt b \to (a/b) = (a*c)/(b*c).
296 intros.
297 apply (times_numerator_denominator_aux a b c (a/b))
298 [ assumption
299 | assumption
300 | reflexivity
301 ]
302 qed.
303
304 theorem times_mod: \forall a,b,c:nat.
305 O \lt c \to O \lt b \to ((a*c) \mod (b*c)) = c*(a\mod b).
306 intros.
307 apply (div_mod_spec_to_eq2 (a*c) (b*c) (a/b) ((a*c) \mod (b*c)) (a/b) (c*(a \mod b)))
308 [ apply div_mod_spec_intro
309   [ apply (lt_mod_m_m (a*c) (b*c)).
310     rewrite > (S_pred b)
311     [ rewrite > (S_pred c)
312       [ apply lt_O_times_S_S
313       | assumption
314       ]
315     | assumption
316     ]
317   | rewrite > (times_numerator_denominator a b c)
318     [ apply (div_mod (a*c) (b*c)).
319       rewrite > (S_pred b)
320       [ rewrite > (S_pred c)
321         [ apply (lt_O_times_S_S)
322         | assumption
323         ]
324       | assumption
325       ]
326     | assumption
327     | assumption
328     ]
329   ]
330 | constructor 1
331   [ rewrite > (sym_times b c).
332     apply (lt_times_r1 c)
333     [ assumption
334     | apply (lt_mod_m_m).
335       assumption
336     ]
337   | rewrite < (assoc_times (a/b) b c).
338     rewrite > (sym_times a c).
339     rewrite > (sym_times ((a/b)*b) c).
340     rewrite < (distr_times_plus c ? ?).
341     apply eq_f.
342     apply (div_mod a b).
343     assumption
344   ]
345 ]
346 qed.
347
348
349
350
351
352
353
354
355
356
357