]> matita.cs.unibo.it Git - helm.git/blob - matita/library/nat/totient.ma
2c6061e0b80c673158812a259d0e944280d0a8f6
[helm.git] / matita / library / nat / totient.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/totient".
16
17 include "nat/count.ma".
18 include "nat/chinese_reminder.ma".
19
20 definition totient : nat \to nat \def
21 \lambda n. count n (\lambda m. eqb (gcd m n) (S O)).
22
23 theorem totient3: totient (S(S(S O))) = (S(S O)).
24 reflexivity.
25 qed.
26
27 theorem totient6: totient (S(S(S(S(S(S O)))))) = (S(S O)).
28 reflexivity.
29 qed.
30
31 theorem totient_times: \forall n,m:nat. (gcd m n) = (S O) \to
32 totient (n*m) = (totient n)*(totient m).
33 intro.
34 apply (nat_case n).
35 intro.simplify.intro.reflexivity.
36 intros 2.apply (nat_case m1).
37 rewrite < sym_times.
38 rewrite < (sym_times (totient O)).
39 simplify.intro.reflexivity.
40 intros.
41 unfold totient.
42 apply (count_times m m2 ? ? ? 
43 (\lambda b,a. cr_pair (S m) (S m2) a b) (\lambda x. x \mod (S m)) (\lambda x. x \mod (S m2))).
44 intros.unfold cr_pair.
45 apply (le_to_lt_to_lt ? (pred ((S m)*(S m2)))).
46 unfold min.
47 apply le_min_aux_r.unfold lt.
48 apply (nat_case ((S m)*(S m2))).apply le_n.
49 intro.apply le_n.
50 intros.
51 generalize in match (mod_cr_pair (S m) (S m2) a b H1 H2 H).
52 intro.elim H3.
53 apply H4.
54 intros.
55 generalize in match (mod_cr_pair (S m) (S m2) a b H1 H2 H).
56 intro.elim H3.
57 apply H5.
58 intros.
59 generalize in match (mod_cr_pair (S m) (S m2) a b H1 H2 H).
60 intro.elim H3.
61 apply eqb_elim.
62 intro.
63 rewrite > eq_to_eqb_true.
64 rewrite > eq_to_eqb_true.
65 reflexivity.
66 rewrite < H4.
67 rewrite > sym_gcd.
68 rewrite > gcd_mod.
69 apply (gcd_times_SO_to_gcd_SO ? ? (S m2)).
70 unfold lt.apply le_S_S.apply le_O_n.
71 unfold lt.apply le_S_S.apply le_O_n.
72 assumption.
73 unfold lt.apply le_S_S.apply le_O_n.
74 rewrite < H5.
75 rewrite > sym_gcd.
76 rewrite > gcd_mod.
77 apply (gcd_times_SO_to_gcd_SO ? ? (S m)).
78 unfold lt.apply le_S_S.apply le_O_n.
79 unfold lt.apply le_S_S.apply le_O_n.
80 rewrite > sym_times.
81 assumption.
82 unfold lt.apply le_S_S.apply le_O_n.
83 intro.
84 apply eqb_elim.
85 intro.apply eqb_elim.
86 intro.apply False_ind.
87 apply H6.
88 apply eq_gcd_times_SO.
89 unfold lt.apply le_S_S.apply le_O_n.
90 unfold lt.apply le_S_S.apply le_O_n.
91 rewrite < gcd_mod.
92 rewrite > H4.
93 rewrite > sym_gcd.assumption.
94 unfold lt.apply le_S_S.apply le_O_n.
95 rewrite < gcd_mod.
96 rewrite > H5.
97 rewrite > sym_gcd.assumption.
98 unfold lt.apply le_S_S.apply le_O_n.
99 intro.reflexivity.
100 intro.reflexivity.
101 qed.