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 (* \ / Matita is distributed under the terms of the *)
11 (* v GNU Lesser General Public License Version 2.1 *)
13 (**************************************************************************)
15 set "baseuri" "cic:/matita/Z/moebius".
17 include "nat/factorization.ma".
18 include "Z/sigma_p.ma".
20 let rec moebius_aux p n : Z \def
24 match p_ord n (nth_prime p1) with
25 [ (pair q r) \Rightarrow
27 [ O \Rightarrow moebius_aux p1 r
30 [ O \Rightarrow Zopp (moebius_aux p1 r)
31 | (S q2) \Rightarrow OZ
38 definition moebius : nat \to Z \def
40 let p \def (max n (\lambda p:nat.eqb (n \mod (nth_prime p)) O)) in
44 theorem moebius_O : moebius O = pos O.
45 simplify. reflexivity.
48 theorem moebius_SO : moebius (S O) = pos O.
52 theorem moebius_SSO : moebius (S (S O)) = neg O.
56 theorem moebius_SSSO : moebius (S (S (S O))) = neg O.
60 theorem moebius_SSSSO : moebius (S (S (S (S O)))) = OZ.
65 theorem not_divides_to_eq_moebius_aux: \forall n,p,p1.p \le p1 \to
66 (\forall i. p \le i \to i \le p1 \to Not (divides (nth_prime i) n))
67 \to moebius_aux p n = moebius_aux p1 n.
72 rewrite > not_divides_to_p_ord_O
73 [simplify.apply H2.intros.
74 apply H3[assumption|apply le_S.assumption]
75 |apply H3[assumption|apply le_n_Sn]
80 theorem eq_moebius_moebius_aux: \forall n,p.
81 max_prime_factor n < p \to p \le n \to
82 moebius n = moebius_aux p n.
86 (moebius_aux (S (max n (\lambda p:nat.eqb (n\mod nth_prime p) O))) n
88 apply not_divides_to_eq_moebius_aux
91 apply divides_b_false_to_not_divides.
93 apply (lt_max_to_false ? n i)
95 |apply (trans_le ? p)[assumption|assumption]
100 theorem moebius_aux_SO: \forall p.moebius_aux p (S O) = pos O.
103 [simplify.reflexivity
106 apply not_divides_to_eq_moebius_aux
108 |intros.unfold.intro.
109 absurd (nth_prime i \le S O)
111 [apply le_n|assumption]
113 apply lt_SO_nth_prime_n.
119 theorem p_ord_SO_SO_to_moebius : \forall n,p.
121 p = max_prime_factor n \to
122 p_ord n (nth_prime p) = pair nat nat (S O) (S O) \to
123 moebius n = Zopp (pos O).
126 (moebius_aux (S (max_prime_factor n)) n = neg O).
127 rewrite < H1.simplify.
128 rewrite > H2.simplify.
129 rewrite > moebius_aux_SO.
133 theorem p_ord_SO_r_to_moebius1 : \forall n,p,r.
135 p = max_prime_factor n \to
137 p_ord n (nth_prime p) = pair nat nat (S O) r \to
138 moebius n = Zopp (moebius r).
141 (moebius_aux (S (max_prime_factor n)) n = -(moebius r)).
142 rewrite < H1.simplify.
143 rewrite > H3.simplify.
147 (moebius_aux (S (max_prime_factor r)) r = moebius_aux p r).
148 apply not_divides_to_eq_moebius_aux
149 [apply (p_ord_to_lt_max_prime_factor n p (S O) ? ? H1)
150 [apply (trans_lt ? (S O))[apply lt_O_S|assumption]
151 |apply sym_eq. assumption
155 lapply (decidable_le i r).
157 [apply divides_b_false_to_not_divides.
158 apply (lt_max_to_false ? r i)[assumption|assumption]
159 |unfold.intro.apply H6.
160 apply (trans_le ? (nth_prime i))
162 apply lt_n_nth_prime_n
164 [apply (trans_lt ? (S O))[apply lt_O_S|assumption]
172 theorem p_ord_SO_r_to_moebius : \forall n,p,r.
174 p = max_prime_factor n \to
175 p_ord n (nth_prime p) = pair nat nat (S O) r \to
176 moebius n = Zopp (moebius r).
178 apply (nat_case r);intro
180 apply (p_ord_to_not_eq_O ? ? ? ? H H2).
182 |apply (nat_case m);intros
183 [simplify.apply (p_ord_SO_SO_to_moebius ? ? H H1 H2)
184 |apply (p_ord_SO_r_to_moebius1 ? ? ? H H1 ? H2).
185 apply le_S_S.apply le_S_S.apply le_O_n
190 theorem p_ord_SSq_r_to_moebius : \forall n,p,q,r.
192 p = max_prime_factor n \to
193 p_ord n (nth_prime p) = pair nat nat (S (S q)) r \to
197 (moebius_aux (S (max n (\lambda p:nat.eqb (n\mod nth_prime p) O))) n =OZ).
198 rewrite < H1.simplify.
199 rewrite > H2.simplify.
203 theorem moebius_exp: \forall p,q,r:nat. O < r \to
204 r = (S O) \lor (max_prime_factor r) < p \to
205 (* r = (S O) \lor (max r (\lambda p:nat.eqb (r \mod (nth_prime p)) O)) < p \to *)
206 sigma_p (S (S q)) (\lambda y.true) (\lambda y.moebius (r*(exp (nth_prime p) y))) = O.
209 [rewrite > true_to_sigma_p_Sn
210 [rewrite > true_to_sigma_p_Sn
211 [rewrite > Zplus_z_OZ.
212 rewrite < times_n_SO.
213 rewrite > (p_ord_SO_r_to_moebius ? p r)
214 [rewrite > sym_Zplus.
215 rewrite > Zplus_Zopp.
219 rewrite > times_n_SO.
220 apply lt_to_le_to_lt_times
221 [apply lt_SO_nth_prime_n
226 [apply le_n|assumption|assumption]
228 [apply lt_O_nth_prime_n
229 |apply lt_max_prime_factor_to_not_divides;assumption
237 |rewrite > true_to_sigma_p_Sn
239 rewrite > Zplus_z_OZ.
240 apply (p_ord_SSq_r_to_moebius ? p n r)
241 [rewrite > times_n_SO.
242 rewrite > sym_times in \vdash (? ? %).
243 apply lt_to_le_to_lt_times
244 [apply (trans_lt ? (nth_prime p))
245 [apply lt_SO_nth_prime_n
246 |rewrite > exp_n_SO in \vdash (? % ?).
248 [apply lt_SO_nth_prime_n
249 |apply le_S_S.apply le_S_S.apply le_O_n
256 [apply le_S_S.apply le_O_n|assumption|assumption]
258 [apply lt_O_nth_prime_n
259 |apply lt_max_prime_factor_to_not_divides;assumption
268 theorem sigma_p_moebius1: \forall n,m,p:nat.O < n \to O < m \to
269 n = (S O) \lor max_prime_factor n < p \to
270 sigma_p (S (n*(exp (nth_prime p) m))) (\lambda y.divides_b y (n*(exp (nth_prime p) m))) moebius = O.
272 rewrite > sigma_p_divides_b
273 [apply (trans_eq ? ? (sigma_p (S n) (\lambda x:nat.divides_b x n) (\lambda x:nat.OZ)))
276 |apply (lt_O_n_elim m H1).
277 intros.apply moebius_exp
278 [apply (divides_b_true_to_lt_O ? ? ? H4).
281 [left.rewrite > H5 in H3.
282 elim (le_to_or_lt_eq ? ? (le_S_S_to_le ? ? H3))
284 apply (lt_to_not_le O x)
285 [apply (divides_b_true_to_lt_O n x H H4)
286 |apply le_S_S_to_le.assumption
291 apply (le_to_lt_to_lt ? ? ? ? H5).
292 apply (divides_to_max_prime_factor1 x n)
293 [apply (divides_b_true_to_lt_O ? ? H H4)
295 |apply divides_b_true_to_divides.
301 |generalize in match (\lambda x:nat.divides_b x n).
304 [simplify.elim (f O);reflexivity
305 |apply (bool_elim ? (f (S n1)))
307 rewrite > true_to_sigma_p_Sn
308 [rewrite > H3.reflexivity|assumption]
310 rewrite > false_to_sigma_p_Sn
311 [apply H3|assumption]
316 |apply prime_nth_prime
317 |apply lt_max_prime_factor_to_not_divides;assumption
321 theorem sigma_p_moebius: \forall n. (S O) < n \to
322 sigma_p (S n) (\lambda y.divides_b y n) moebius = O.
324 lapply (exp_ord (nth_prime (max_prime_factor n)) n)
325 [rewrite > sym_times in Hletin.
327 apply sigma_p_moebius1
329 [apply lt_SO_nth_prime_n
330 |apply lt_to_le.assumption
334 (fst ? ? (pair ? ? (S O) (S O)) \leq ord n (nth_prime (max_prime_factor n))).
335 rewrite < (p_ord_p (nth_prime (max_prime_factor n)))
336 [apply (divides_to_le_ord ? (nth_prime (max_prime_factor n)) n)
337 [apply lt_O_nth_prime_n
338 |apply lt_to_le.assumption
339 |apply prime_nth_prime
340 |apply divides_max_prime_factor_n.
343 |apply lt_SO_nth_prime_n
345 |lapply (lt_O_ord_rem (nth_prime (max_prime_factor n)) n)
346 [elim (le_to_or_lt_eq ? ? Hletin1)
348 apply (p_ord_to_lt_max_prime_factor1 n (max_prime_factor n)
349 (ord n (nth_prime (max_prime_factor n)))
350 (ord_rem n (nth_prime (max_prime_factor n))))
351 [apply lt_to_le.assumption
356 |left.apply sym_eq.assumption
358 |apply lt_SO_nth_prime_n
359 |apply lt_to_le.assumption
362 |apply lt_SO_nth_prime_n
363 |apply lt_to_le.assumption