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 "nat/map_iter_p.ma".
16 include "nat/totient.ma".
19 lemma count_card: \forall p.\forall n.
20 p O = false \to count (S n) p = card n p.
22 [simplify.rewrite > H. reflexivity
29 lemma count_card1: \forall p.\forall n.
30 p O = false \to p n = false \to count n p = card n p.
31 intros 3.apply (nat_case n)
32 [intro.simplify.rewrite > H. reflexivity
33 |intros.rewrite > (count_card ? ? H).
34 simplify.rewrite > H1.reflexivity
39 ( a reformulation of totient using card insted of count )
41 lemma totient_card: \forall n.
42 totient n = card n (\lambda i.eqb (gcd i n) (S O)).
43 intro.apply (nat_case n)
45 |intro.apply (nat_case m)
47 |intro.apply count_card1
49 |rewrite > gcd_n_n.reflexivity
56 (*this obvious property is useful because simplify, sometimes,
57 "simplifies too much", and doesn't allow to obtain this simple result.
59 theorem card_Sn: \forall n:nat. \forall p:nat \to bool.
60 card (S n) p = (bool_to_nat (p (S n))) + (card n p).
66 (* a reformulation of totient using card insted of sigma_p *)
68 theorem totient_card_aux: \forall n,m: nat.
70 sigma_p (S (S n)) (\lambda m:nat.eqb (gcd m (S (S n))) (S O)) (\lambda x:nat. (S O))
71 = card (S n) (\lambda m:nat.eqb (gcd m (S (S n))) (S O)).
73 rewrite < H in \vdash (? ? (? % ? ?) ?).
74 rewrite < H in \vdash (? ? ? (? % ?)).
77 cut ((eqb (gcd (S O)(S (S n))) (S O) ) = true)
79 simplify in \vdash (? ? ? %).
80 rewrite > true_to_sigma_p_Sn
81 [ rewrite > false_to_sigma_p_Sn in \vdash (? ? (? ? %) ?)
82 [ simplify in \vdash (? ? % ?).
85 apply not_eq_to_eqb_false.
86 apply cic:/matita/nat/nat/not_eq_S.con.
90 [ apply (not_le_Sn_n n ?).
91 apply (transitive_le (S n) O n ? ?);
101 | apply eq_to_eqb_true.
105 | cut ((eqb (gcd (S (S n1)) (S (S n))) (S O)) = true
106 \lor (eqb (gcd (S (S n1)) (S (S n))) (S O)) = false)
109 rewrite > true_to_sigma_p_Sn
111 simplify in \vdash (? ? ? (? % ?)).
117 rewrite > false_to_sigma_p_Sn
119 simplify in \vdash (? ? ? (? % ?)).
126 | elim (eqb (gcd (S (S n1)) (S (S n))) (S O))
136 lemma totient_card: \forall n.
137 totient n = card n (\lambda i.eqb (gcd i n) (S O)).
148 (* here simplify creates problems: it seems it simplifies too much. I had to
149 * introduce the obvious theorem card_Sn.
152 rewrite > (gcd_n_n (S (S n2))).
153 cut ((eqb (S (S n2)) (S O)) = false)
155 simplify in \vdash (? ? ? (? % ?)).
157 rewrite < (plus_n_O).
159 apply (totient_card_aux n2 n2).
161 | apply not_eq_to_eqb_false.
162 apply cic:/matita/nat/nat/not_eq_S.con.
166 [ apply (not_le_Sn_n n2 ?).
167 apply (transitive_le (S n2) O n2 ? ?);
179 theorem gcd_pi_p: \forall n,k. O < k \to k \le n \to
180 gcd n (pi_p (\lambda i.eqb (gcd i n) (S O)) k) = (S O).
183 cut (eqb (gcd (S O) n) (S O) = true)
185 change with ((gcd n (S O)) = (S O)).
186 apply (transitive_eq nat (gcd n (S O)) (gcd (S O) n) (S O) ? ?);
187 [ apply (sym_eq nat (gcd (S O) n) (gcd n (S O)) ?).
188 apply (symmetric_gcd (S O) n).
189 | apply (gcd_SO_n n).
191 |apply eq_to_eqb_true.
198 ((gcd n ((S n1)*(pi_p (\lambda i.eqb (gcd i n) (S O)) n1))) = (S O)).
199 apply eq_gcd_times_SO
200 [unfold.apply le_S.assumption
202 |rewrite > sym_gcd. assumption.
204 apply (trans_le ? (S n1))[apply le_n_Sn|assumption]
208 (gcd n (pi_p (\lambda i.eqb (gcd i n) (S O)) n1) = (S O)).
210 apply (trans_le ? (S n1))[apply le_n_Sn|assumption]
215 theorem congruent_map_iter_p_times:\forall f:nat \to nat. \forall a,n:nat.
218 (map_iter_p n (\lambda i.eqb (gcd i a) (S O)) (\lambda x.f x) (S O) times)
219 (map_iter_p n (\lambda i.eqb (gcd i a) (S O))
220 (\lambda x.f x \mod a) (S O) times) a.
223 [rewrite > map_iter_p_O.
224 apply (congruent_n_n ? a)
225 |apply (eqb_elim (gcd (S n1) a) (S O))
227 rewrite > map_iter_p_S_true
228 [rewrite > map_iter_p_S_true
229 [apply congruent_times
231 |apply congruent_n_mod_n.assumption
234 |apply eq_to_eqb_true.assumption
236 |apply eq_to_eqb_true.assumption
239 rewrite > map_iter_p_S_false
240 [rewrite > map_iter_p_S_false
242 |apply not_eq_to_eqb_false.assumption
244 |apply not_eq_to_eqb_false.assumption
250 theorem permut_p_mod: \forall a,n. S O < n \to O < a \to gcd a n = (S O) \to
251 permut_p (\lambda x:nat.a*x \mod n) (\lambda i:nat.eqb (gcd i n) (S O)) n.
253 lapply (lt_S_to_lt ? ? H) as H3.
264 [apply eq_to_eqb_true.
266 apply eq_gcd_times_SO
268 |apply (gcd_SO_to_lt_O i n H).
269 apply eqb_true_to_eq.
271 |rewrite > sym_gcd.assumption
272 |rewrite > sym_gcd.apply eqb_true_to_eq.
279 lapply (gcd_SO_to_lt_n ? ? H H4 (eqb_true_to_eq ? ? H5)) as H9.
280 lapply (gcd_SO_to_lt_n ? ? H H7 (eqb_true_to_eq ? ? H6)) as H10.
281 lapply (gcd_SO_to_lt_O ? ? H (eqb_true_to_eq ? ? H5)) as H11.
282 lapply (gcd_SO_to_lt_O ? ? H (eqb_true_to_eq ? ? H6)) as H12.
285 apply (nat_compare_elim i j)
290 apply (trans_le ? (j -i))
293 apply (lt_plus_to_lt_l i).
295 rewrite < (plus_minus_m_m)
296 [assumption|apply lt_to_le.assumption]
297 |apply (gcd_SO_to_divides_times_to_divides a)
299 |rewrite > sym_gcd.assumption
300 |apply mod_O_to_divides
302 |rewrite > distr_times_minus.
303 apply (divides_to_mod_O n (minus (times a j) (times a i)) ? ?);
305 | apply (eq_mod_to_divides (times a j) (times a i) n ? ?);
307 |apply (sym_eq nat (mod (times a i) n) (mod (times a j) n) ?).
314 | apply (le_minus_m j i).
322 apply (trans_le ? (i -j))
325 apply (lt_plus_to_lt_l j).
327 rewrite < (plus_minus_m_m)
328 [assumption|apply lt_to_le.assumption]
329 |apply (gcd_SO_to_divides_times_to_divides a)
331 |rewrite > sym_gcd.assumption
332 |apply mod_O_to_divides
334 |rewrite > distr_times_minus.
335 apply (divides_to_mod_O n (minus (times a i) (times a j)) ? ?);
337 | apply (eq_mod_to_divides (times a i) (times a j) n ? ?);
345 | apply (le_minus_m i j).
352 theorem congruent_exp_totient_SO: \forall n,a:nat. (S O) < n \to
353 gcd a n = (S O) \to congruent (exp a (totient n)) (S O) n.
356 [ apply divides_to_congruent
357 [apply (trans_lt ? (S O)).apply lt_O_S. assumption
358 |change with (O < exp a (totient n)).apply lt_O_exp.assumption
359 |apply (gcd_SO_to_divides_times_to_divides (pi_p (\lambda i.eqb (gcd i n) (S O)) n))
360 [apply (trans_lt ? (S O)).apply lt_O_S. assumption
362 [apply (trans_lt ? (S O)).apply lt_O_S. assumption
365 |rewrite < sym_times.
366 rewrite > times_minus_l.
367 rewrite > (sym_times (S O)).
368 rewrite < times_n_SO.
369 rewrite > totient_card.
370 rewrite > a_times_pi_p.
371 apply congruent_to_divides
372 [apply (trans_lt ? (S O)).apply lt_O_S. assumption
373 | apply (transitive_congruent n ?
374 (map_iter_p n (\lambda i.eqb (gcd i n) (S O)) (\lambda x.a*x \mod n) (S O) times))
375 [apply (congruent_map_iter_p_times ? n n).
376 apply (trans_lt ? (S O))
377 [apply lt_O_S|assumption]
379 cut ( (map_iter_p n (\lambda i:nat.eqb (gcd i n) (S O)) (\lambda n:nat.n) (S O) times)
380 = (map_iter_p n (\lambda i:nat.eqb (gcd i n) (S O)) (\lambda x:nat.a*x\mod n) (S O) times))
381 [rewrite < Hcut1.apply congruent_n_n
382 |apply (eq_map_iter_p_permut ? ? ? ? ? (λm.m))
385 |apply (permut_p_mod ? ? H Hcut H1)
387 apply not_eq_to_eqb_false.
389 apply (lt_to_not_eq (S O) n)
390 [assumption|apply sym_eq.assumption]
397 |elim (le_to_or_lt_eq O a (le_O_n a))
399 |absurd (gcd a n = S O)
404 apply (lt_to_not_eq (S O) n)
405 [assumption|apply sym_eq.assumption]