]> matita.cs.unibo.it Git - helm.git/blob - helm/matita/library/nat/factorization.ma
factorization.ma
[helm.git] / helm / matita / library / nat / factorization.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 (*       \ /        Matita 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/factorization".
16
17 include "nat/log.ma".
18 include "nat/gcd.ma".
19 include "nat/nth_prime.ma".
20
21 (* the following factorization algorithm looks for the largest prime
22    factor. *)
23 definition max_prime_factor \def \lambda n:nat.
24 (max n (\lambda p:nat.eqb (mod n (nth_prime p)) O)).
25
26 (* max_prime_factor is indeed a factor *)
27 theorem divides_max_prime_factor_n: \forall n:nat. (S O) < n \to
28 divides (nth_prime (max_prime_factor n)) n.
29 intros.apply divides_b_true_to_divides.
30 apply lt_O_nth_prime_n.
31 apply f_max_true  (\lambda p:nat.eqb (mod n (nth_prime p)) O) n.
32 cut ex nat (\lambda i. nth_prime i = smallest_factor n).
33 elim Hcut.
34 apply ex_intro nat ? a.
35 split.
36 apply trans_le a (nth_prime a).
37 apply le_n_fn.
38 exact lt_nth_prime_n_nth_prime_Sn.
39 rewrite > H1. apply le_smallest_factor_n.
40 rewrite > H1.
41 change with divides_b (smallest_factor n) n = true.
42 apply divides_to_divides_b_true.
43 apply trans_lt ? (S O).simplify. apply le_n.
44 apply lt_SO_smallest_factor.assumption.
45 apply divides_smallest_factor_n.
46 apply trans_lt ? (S O). simplify. apply le_n. assumption.
47 apply prime_to_nth_prime.
48 apply prime_smallest_factor_n.assumption.
49 qed.
50
51 theorem divides_to_max_prime_factor : \forall n,m. (S O) < n \to O < m \to divides n m \to 
52 max_prime_factor n \le max_prime_factor m.
53 intros.change with
54 (max n (\lambda p:nat.eqb (mod n (nth_prime p)) O)) \le
55 (max m (\lambda p:nat.eqb (mod m (nth_prime p)) O)).
56 apply f_m_to_le_max.
57 apply trans_le ? n.
58 apply le_max_n.apply divides_to_le.assumption.assumption.
59 change with divides_b (nth_prime (max_prime_factor n)) m = true.
60 apply divides_to_divides_b_true.
61 cut prime (nth_prime (max_prime_factor n)).
62 apply lt_O_nth_prime_n.apply prime_nth_prime.
63 cut divides (nth_prime (max_prime_factor n)) n.
64 apply transitive_divides ? n.
65 apply divides_max_prime_factor_n.
66 assumption.assumption.
67 apply divides_b_true_to_divides.
68 apply lt_O_nth_prime_n.
69 apply divides_to_divides_b_true.
70 apply lt_O_nth_prime_n.
71 apply divides_max_prime_factor_n.
72 assumption.
73 qed.
74
75 theorem plog_to_lt_max_prime_factor: \forall n,p,q,r. O < n \to
76 p = max_prime_factor n \to 
77 (pair nat nat q r) = plog n (nth_prime p) \to
78 (S O) < r \to max_prime_factor r < p.
79 intros.
80 rewrite > H1.
81 cut max_prime_factor r \lt max_prime_factor n \lor
82     max_prime_factor r = max_prime_factor n.
83 elim Hcut.assumption.
84 absurd (divides (nth_prime (max_prime_factor n)) r).
85 rewrite < H4.
86 apply divides_max_prime_factor_n.
87 assumption.
88 change with (divides (nth_prime (max_prime_factor n)) r) \to False.
89 intro.
90 cut \not (mod r (nth_prime (max_prime_factor n))) = O.
91 apply Hcut1.apply divides_to_mod_O.
92 apply lt_O_nth_prime_n.assumption.
93 apply plog_aux_to_not_mod_O n n ? q r.
94 apply lt_SO_nth_prime_n.assumption.
95 apply le_n.
96 rewrite < H1.assumption.
97 apply le_to_or_lt_eq (max_prime_factor r)  (max_prime_factor n).
98 apply divides_to_max_prime_factor.
99 assumption.assumption.
100 apply witness r n (exp (nth_prime p) q).
101 rewrite < sym_times.
102 apply plog_aux_to_exp n n ? q r.
103 apply lt_O_nth_prime_n.assumption.
104 qed.
105
106 theorem plog_to_lt_max_prime_factor1: \forall n,p,q,r. O < n \to
107 max_prime_factor n \le p \to 
108 (pair nat nat q r) = plog n (nth_prime p) \to
109 (S O) < r \to max_prime_factor r < p.
110 intros.
111 cut max_prime_factor n < p \lor max_prime_factor n = p.
112 elim Hcut.apply le_to_lt_to_lt ? (max_prime_factor n).
113 apply divides_to_max_prime_factor.assumption.assumption.
114 apply witness r n (exp (nth_prime p) q).
115 rewrite > sym_times.
116 apply plog_aux_to_exp n n.
117 apply lt_O_nth_prime_n.
118 assumption.assumption.
119 apply plog_to_lt_max_prime_factor n ? q.
120 assumption.apply sym_eq.assumption.assumption.assumption.
121 apply le_to_or_lt_eq ? p H1.
122 qed.
123
124 (* datatypes and functions *)
125
126 inductive nat_fact : Set \def
127     nf_last : nat \to nat_fact   
128   | nf_cons : nat \to nat_fact \to nat_fact.
129
130 inductive nat_fact_all : Set \def
131     nfa_zero : nat_fact_all
132   | nfa_one : nat_fact_all
133   | nfa_proper : nat_fact \to nat_fact_all.
134
135 let rec factorize_aux p n acc \def
136   match p with 
137   [ O \Rightarrow acc
138   | (S p1) \Rightarrow 
139     match plog n (nth_prime p1) with
140     [ (pair q r) \Rightarrow 
141       factorize_aux p1 r (nf_cons q acc)]].
142   
143 definition factorize : nat \to nat_fact_all \def \lambda n:nat.
144   match n with
145     [ O \Rightarrow nfa_zero
146     | (S n1) \Rightarrow
147       match n1 with
148       [ O \Rightarrow nfa_one
149     | (S n2) \Rightarrow 
150       let p \def (max (S(S n2)) (\lambda p:nat.eqb (mod (S(S n2)) (nth_prime p)) O)) in
151       match plog (S(S n2)) (nth_prime p) with
152       [ (pair q r) \Rightarrow 
153            nfa_proper (factorize_aux p r (nf_last (pred q)))]]].
154            
155 let rec defactorize_aux f i \def
156   match f with
157   [ (nf_last n) \Rightarrow exp (nth_prime i) (S n)
158   | (nf_cons n g) \Rightarrow 
159       (exp (nth_prime i) n)*(defactorize_aux g (S i))].
160       
161 definition defactorize : nat_fact_all \to nat \def
162 \lambda f : nat_fact_all. 
163 match f with 
164 [ nfa_zero \Rightarrow O
165 | nfa_one \Rightarrow (S O)
166 | (nfa_proper g) \Rightarrow defactorize_aux g O]. 
167
168 theorem defactorize_aux_factorize_aux : 
169 \forall p,n:nat.\forall acc:nat_fact.O < n \to
170 ((n=(S O) \land p=O) \lor max_prime_factor n < p) \to
171 defactorize_aux (factorize_aux p n acc) O = n*(defactorize_aux acc p).
172 intro.elim p.simplify.
173 elim H1.elim H2.rewrite > H3.
174 rewrite > sym_times. apply times_n_SO.
175 apply False_ind.apply not_le_Sn_O (max_prime_factor n) H2.
176 simplify.
177 (* generalizing the goal: I guess there exists a better way *)
178 cut \forall q,r.(pair nat nat q r) = (plog_aux n1 n1 (nth_prime n)) \to
179 defactorize_aux match (plog_aux n1 n1 (nth_prime n)) with
180 [(pair q r)  \Rightarrow (factorize_aux n r (nf_cons q acc))] O =
181 n1*defactorize_aux acc (S n).
182 apply Hcut (fst ? ? (plog_aux n1 n1 (nth_prime n)))
183 (snd ? ? (plog_aux n1 n1 (nth_prime n))).
184 apply sym_eq.apply eq_pair_fst_snd.
185 intros.
186 rewrite < H3.
187 simplify.
188 cut n1 = r*(exp (nth_prime n) q).
189 rewrite > H.
190 simplify.rewrite < assoc_times.
191 rewrite < Hcut.reflexivity.
192 rewrite > sym_times.
193 apply plog_aux_to_exp n1 n1.
194 apply lt_O_nth_prime_n.assumption.
195 cut O < r \lor O = r.
196 elim Hcut1.assumption.absurd n1 = O.
197 rewrite > Hcut.rewrite < H4.reflexivity.
198 simplify. intro.apply not_le_Sn_O O.
199 rewrite < H5 in \vdash (? ? %).assumption.
200 apply le_to_or_lt_eq.apply le_O_n.
201 cut (S O) < r \lor \lnot (S O) < r.
202 elim Hcut1.
203 right.
204 apply plog_to_lt_max_prime_factor1 n1 ? q r.
205 assumption.elim H2.
206 elim H5.
207 apply False_ind.
208 apply not_eq_O_S n.apply sym_eq.assumption.
209 apply le_S_S_to_le.
210 exact H5.
211 assumption.assumption.
212 cut r=(S O).
213 apply nat_case n.
214 left.split.assumption.reflexivity.
215 intro.right.rewrite > Hcut2.
216 simplify.apply le_S_S.apply le_O_n.
217 cut r \lt (S O) \or r=(S O).
218 elim Hcut2.absurd O=r.
219 apply le_n_O_to_eq.apply le_S_S_to_le.exact H5.
220 simplify.intro.
221 cut O=n1.
222 apply not_le_Sn_O O ?.
223 rewrite > Hcut3 in \vdash (? ? %).
224 assumption.rewrite > Hcut. 
225 rewrite < H6.reflexivity.
226 assumption.
227 apply le_to_or_lt_eq r (S O).
228 apply not_lt_to_le.assumption.
229 apply decidable_lt (S O) r.
230 qed.
231
232 theorem defactorize_factorize: \forall n:nat.defactorize (factorize n) = n.
233 intro.
234 apply nat_case n.reflexivity.
235 intro.apply nat_case m.reflexivity.
236 intro.change with  
237 let p \def (max (S(S m1)) (\lambda p:nat.eqb (mod (S(S m1)) (nth_prime p)) O)) in
238 defactorize (match plog (S(S m1)) (nth_prime p) with
239 [ (pair q r) \Rightarrow 
240    nfa_proper (factorize_aux p r (nf_last (pred q)))])=(S(S m1)).
241 intro.
242 (* generalizing the goal; find a better way *)
243 cut \forall q,r.(pair nat nat q r) = (plog (S(S m1)) (nth_prime p)) \to
244 defactorize (match plog (S(S m1)) (nth_prime p) with
245 [ (pair q r) \Rightarrow 
246    nfa_proper (factorize_aux p r (nf_last (pred q)))])=(S(S m1)).
247 apply Hcut (fst ? ? (plog (S(S m1)) (nth_prime p)))
248 (snd ? ? (plog (S(S m1)) (nth_prime p))).
249 apply sym_eq.apply eq_pair_fst_snd.
250 intros.
251 rewrite < H.
252 change with 
253 defactorize_aux (factorize_aux p r (nf_last (pred q))) O = (S(S m1)).
254 cut (S(S m1)) = (exp (nth_prime p) q)*r.
255 cut O <r.
256 rewrite > defactorize_aux_factorize_aux.
257 change with r*(exp (nth_prime p) (S (pred q))) = (S(S m1)).
258 cut (S (pred q)) = q.
259 rewrite > Hcut2.
260 rewrite > sym_times.
261 apply sym_eq.
262 apply plog_aux_to_exp (S(S m1)).
263 apply lt_O_nth_prime_n.
264 assumption.
265 (* O < q *)
266 apply sym_eq. apply S_pred.
267 cut O < q \lor O = q.
268 elim Hcut2.assumption.
269 absurd divides (nth_prime p) (S (S m1)).
270 apply divides_max_prime_factor_n (S (S m1)).
271 simplify.apply le_S_S.apply le_S_S. apply le_O_n.
272 cut (S(S m1)) = r.
273 rewrite > Hcut3 in \vdash (? (? ? %)).
274 change with divides (nth_prime p) r \to False.
275 intro.
276 apply plog_aux_to_not_mod_O (S(S m1)) (S(S m1)) (nth_prime p) q r ? ? ? ?.
277 apply divides_to_mod_O.
278 apply lt_O_nth_prime_n.
279 assumption.
280 apply lt_SO_nth_prime_n.
281 simplify.apply le_S_S.apply le_O_n.
282 apply le_n.
283 assumption.rewrite > times_n_SO in \vdash (? ? ? %).
284 rewrite < sym_times.
285 rewrite > exp_n_O (nth_prime p).
286 rewrite > H1 in \vdash (? ? ? (? (? ? %) ?)).
287 assumption.
288 apply le_to_or_lt_eq.apply le_O_n.
289 (* O < r *)
290 cut O < r \lor O = r.
291 elim Hcut1.assumption.
292 apply False_ind.
293 apply not_eq_O_S (S m1).
294 rewrite > Hcut.rewrite < H1.rewrite < times_n_O.reflexivity.
295 apply le_to_or_lt_eq.apply le_O_n.
296 (* prova del cut *)
297 apply plog_aux_to_exp (S(S m1)).
298 apply lt_O_nth_prime_n.
299 assumption.
300 (* fine prova cut *)
301 assumption.
302 (* e adesso l'ultimo goal *)
303 cut (S O) < r \lor \lnot (S O) < r.
304 elim Hcut2.
305 right.
306 apply plog_to_lt_max_prime_factor1 (S(S m1)) ? q r.
307 simplify.apply le_S_S. apply le_O_n.
308 apply le_n.
309 assumption.assumption.
310 cut r=(S O).
311 apply nat_case p.
312 left.split.assumption.reflexivity.
313 intro.right.rewrite > Hcut3.
314 simplify.apply le_S_S.apply le_O_n.
315 cut r \lt (S O) \or r=(S O).
316 elim Hcut3.absurd O=r.
317 apply le_n_O_to_eq.apply le_S_S_to_le.exact H2.
318 simplify.intro.
319 apply not_le_Sn_O O ?.
320 rewrite > H3 in \vdash (? ? %).assumption.assumption.
321 apply le_to_or_lt_eq r (S O).
322 apply not_lt_to_le.assumption.
323 apply decidable_lt (S O) r.
324 qed.
325
326 let rec max_p f \def
327 match f with
328 [ (nf_last n) \Rightarrow O
329 | (nf_cons n g) \Rightarrow S (max_p g)].
330
331 let rec max_p_exponent f \def
332 match f with
333 [ (nf_last n) \Rightarrow n
334 | (nf_cons n g) \Rightarrow max_p_exponent g].
335
336 theorem divides_max_p_defactorize: \forall f:nat_fact.\forall i:nat. 
337 divides (nth_prime ((max_p f)+i)) (defactorize_aux f i).
338 intro.
339 elim f.simplify.apply witness ? ? (exp (nth_prime i) n).
340 reflexivity.
341 change with 
342 divides (nth_prime (S(max_p n1)+i)) 
343 ((exp (nth_prime i) n)*(defactorize_aux n1 (S i))).
344 elim H (S i).
345 rewrite > H1.
346 rewrite < sym_times.
347 rewrite > assoc_times.
348 rewrite < plus_n_Sm.
349 apply witness ? ? (n2*(exp (nth_prime i) n)).
350 reflexivity.
351 qed.
352
353 theorem divides_exp_to_divides: 
354 \forall p,n,m:nat. prime p \to 
355 divides p (exp n m) \to divides p n.
356 intros 3.elim m.simplify in H1.
357 apply transitive_divides p (S O).assumption.
358 apply divides_SO_n.
359 cut divides p n \lor divides p (exp n n1).
360 elim Hcut.assumption.
361 apply H.assumption.assumption.
362 apply divides_times_to_divides.assumption.
363 exact H2.
364 qed.
365
366 theorem divides_exp_to_eq: 
367 \forall p,q,m:nat. prime p \to prime q \to
368 divides p (exp q m) \to p = q.
369 intros.
370 simplify in H1.
371 elim H1.apply H4.
372 apply divides_exp_to_divides p q m.
373 assumption.assumption.
374 simplify in H.elim H.assumption.
375 qed.
376
377 theorem  not_divides_defactorize_aux: \forall f:nat_fact. \forall i,j:nat.
378 i < j \to \not divides (nth_prime i) (defactorize_aux f j).
379 intro.elim f.
380 change with
381 divides (nth_prime i) (exp (nth_prime j) (S n)) \to False.
382 intro.absurd (nth_prime i) = (nth_prime j).
383 apply divides_exp_to_eq ? ? (S n).
384 apply prime_nth_prime.apply prime_nth_prime.
385 assumption.
386 change with (nth_prime i) = (nth_prime j) \to False.
387 intro.cut i = j.
388 apply not_le_Sn_n i.rewrite > Hcut in \vdash (? ? %).assumption.
389 apply injective_nth_prime ? ? H2.
390 change with 
391 divides (nth_prime i) ((exp (nth_prime j) n)*(defactorize_aux n1 (S j))) \to False.
392 intro.
393 cut divides (nth_prime i) (exp (nth_prime j) n)
394 \lor divides (nth_prime i) (defactorize_aux n1 (S j)).
395 elim Hcut.
396 absurd (nth_prime i) = (nth_prime j).
397 apply divides_exp_to_eq ? ? n.
398 apply prime_nth_prime.apply prime_nth_prime.
399 assumption.
400 change with (nth_prime i) = (nth_prime j) \to False.
401 intro.
402 cut i = j.
403 apply not_le_Sn_n i.rewrite > Hcut1 in \vdash (? ? %).assumption.
404 apply injective_nth_prime ? ? H4.
405 apply H i (S j) ?.
406 assumption.apply trans_lt ? j.assumption.simplify.apply le_n.
407 apply divides_times_to_divides.
408 apply prime_nth_prime.assumption.
409 qed.
410
411 theorem eq_defactorize_aux_to_eq: \forall f,g:nat_fact.\forall i:nat.
412 defactorize_aux f i = defactorize_aux g i \to f = g.
413 intro.
414 elim f.
415 generalize in match H.
416 elim g.
417 apply eq_f.
418 apply inj_S. apply inj_exp_r (nth_prime i).
419 apply lt_SO_nth_prime_n.
420 assumption.
421 absurd defactorize_aux (nf_last n) i =
422 defactorize_aux (nf_cons n1 n2) i.
423 rewrite > H2.reflexivity.
424 absurd divides (nth_prime (S(max_p n2)+i)) (defactorize_aux (nf_cons n1 n2) i).
425 apply divides_max_p_defactorize.
426 rewrite < H2.
427 change with 
428 (divides (nth_prime (S(max_p n2)+i)) (exp (nth_prime i) (S n))) \to False.
429 intro.
430 absurd nth_prime (S (max_p n2) + i) = nth_prime i.
431 apply divides_exp_to_eq ? ? (S n).
432 apply prime_nth_prime.apply prime_nth_prime.assumption.
433 change with nth_prime (S (max_p n2) + i) = nth_prime i \to False.
434 intro.apply not_le_Sn_n i.
435 cut S(max_p n2)+i= i.
436 rewrite < Hcut in \vdash (? ? %).
437 simplify.apply le_S_S.
438 apply le_plus_n.
439 apply injective_nth_prime ? ? H4.
440 generalize in match H1.
441 elim g.
442 absurd defactorize_aux (nf_last n2) i =
443 defactorize_aux (nf_cons n n1) i.
444 apply sym_eq. assumption.
445 absurd divides (nth_prime (S(max_p n1)+i)) (defactorize_aux (nf_cons n n1) i).
446 apply divides_max_p_defactorize.
447 rewrite > H2.
448 change with 
449 (divides (nth_prime (S(max_p n1)+i)) (exp (nth_prime i) (S n2))) \to False.
450 intro.
451 absurd nth_prime (S (max_p n1) + i) = nth_prime i.
452 apply divides_exp_to_eq ? ? (S n2).
453 apply prime_nth_prime.apply prime_nth_prime.assumption.
454 change with nth_prime (S (max_p n1) + i) = nth_prime i \to False.
455 intro.apply not_le_Sn_n i.
456 cut S(max_p n1)+i= i.
457 rewrite < Hcut in \vdash (? ? %).
458 simplify.apply le_S_S.
459 apply le_plus_n.
460 apply injective_nth_prime ? ? H4.
461 simplify in H3.
462 generalize in match H3.
463 apply nat_elim2 (\lambda n,n2.
464 (exp (nth_prime i) n)*(defactorize_aux n1 (S i)) =
465 (exp (nth_prime i) n2)*(defactorize_aux n3 (S i)) \to
466 nf_cons n n1 = nf_cons n2 n3).
467 intro.
468 elim n4. apply eq_f.
469 apply H n3 (S i).
470 simplify in H4.
471 rewrite > plus_n_O.
472 rewrite > plus_n_O (defactorize_aux n3 (S i)).assumption.
473 apply False_ind.
474 apply not_divides_defactorize_aux n1 i (S i) ?.
475 simplify in H5.
476 rewrite > plus_n_O (defactorize_aux n1 (S i)).
477 rewrite > H5.
478 rewrite > assoc_times.
479 apply witness ? ? ((exp (nth_prime i) n5)*(defactorize_aux n3 (S i))).
480 reflexivity.
481 simplify. apply le_n.
482 intros.
483 apply False_ind.
484 apply not_divides_defactorize_aux n3 i (S i) ?.
485 simplify in H4.
486 rewrite > plus_n_O (defactorize_aux n3 (S i)).
487 rewrite < H4.
488 rewrite > assoc_times.
489 apply witness ? ? ((exp (nth_prime i) n4)*(defactorize_aux n1 (S i))).
490 reflexivity.
491 simplify. apply le_n.
492 intros.
493 cut nf_cons n4 n1 = nf_cons m n3.
494 cut n4=m.
495 cut n1=n3.
496 rewrite > Hcut1.rewrite > Hcut2.reflexivity.
497 change with 
498 match nf_cons n4 n1 with
499 [ (nf_last m) \Rightarrow n1
500 | (nf_cons m g) \Rightarrow g ] = n3.
501 rewrite > Hcut.simplify.reflexivity.
502 change with 
503 match nf_cons n4 n1 with
504 [ (nf_last m) \Rightarrow m
505 | (nf_cons m g) \Rightarrow m ] = m.
506 rewrite > Hcut.simplify.reflexivity.
507 apply H4.simplify in H5.
508 apply inj_times_r1 (nth_prime i).
509 apply lt_O_nth_prime_n.
510 rewrite < assoc_times.rewrite < assoc_times.assumption.
511 qed.
512
513 theorem injective_defactorize_aux: \forall i:nat.
514 injective nat_fact nat (\lambda f.defactorize_aux f i).
515 change with \forall i:nat.\forall f,g:nat_fact.
516 defactorize_aux f i = defactorize_aux g i \to f = g.
517 intros.
518 apply eq_defactorize_aux_to_eq f g i H.
519 qed.
520
521 (*
522 theorem injective_defactorize: 
523 injective nat_fact_all nat defactorize.
524 change with \forall f,g:nat_fact_all.
525 defactorize f = defactorize g \to f = g.
526 intro.elim f.
527 generalize in match H.elim g.
528 reflexivity.simplify in H1.
529 *)
530