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/nat/factorization".
19 include "nat/nth_prime.ma".
21 (* the following factorization algorithm looks for the largest prime
23 definition max_prime_factor \def \lambda n:nat.
24 (max n (\lambda p:nat.eqb (n \mod (nth_prime p)) O)).
26 (* max_prime_factor is indeed a factor *)
27 theorem divides_max_prime_factor_n:
28 \forall n:nat. (S O) < n
29 \to nth_prime (max_prime_factor n) \divides n.
30 intros; apply divides_b_true_to_divides;
31 [ apply lt_O_nth_prime_n;
32 | apply (f_max_true (\lambda p:nat.eqb (n \mod (nth_prime p)) O) n);
33 cut (\exists i. nth_prime i = smallest_factor n);
35 apply (ex_intro nat ? a);
37 [ apply (trans_le a (nth_prime a));
39 exact lt_nth_prime_n_nth_prime_Sn;
41 apply le_smallest_factor_n; ]
43 (*CSC: simplify here does something nasty! *)
44 change with (divides_b (smallest_factor n) n = true);
45 apply divides_to_divides_b_true;
46 [ apply (trans_lt ? (S O));
47 [ unfold lt; apply le_n;
48 | apply lt_SO_smallest_factor; assumption; ]
49 | apply divides_smallest_factor_n;
50 apply (trans_lt ? (S O));
51 [ unfold lt; apply le_n;
53 | apply prime_to_nth_prime;
54 apply prime_smallest_factor_n;
58 theorem divides_to_max_prime_factor : \forall n,m. (S O) < n \to O < m \to n \divides m \to
59 max_prime_factor n \le max_prime_factor m.
60 intros.unfold max_prime_factor.
63 apply le_max_n.apply divides_to_le.assumption.assumption.
64 change with (divides_b (nth_prime (max_prime_factor n)) m = true).
65 apply divides_to_divides_b_true.
66 cut (prime (nth_prime (max_prime_factor n))).
67 apply lt_O_nth_prime_n.apply prime_nth_prime.
68 cut (nth_prime (max_prime_factor n) \divides n).
69 apply (transitive_divides ? n).
70 apply divides_max_prime_factor_n.
71 assumption.assumption.
72 apply divides_b_true_to_divides.
73 apply lt_O_nth_prime_n.
74 apply divides_to_divides_b_true.
75 apply lt_O_nth_prime_n.
76 apply divides_max_prime_factor_n.
80 theorem p_ord_to_lt_max_prime_factor: \forall n,p,q,r. O < n \to
81 p = max_prime_factor n \to
82 (pair nat nat q r) = p_ord n (nth_prime p) \to
83 (S O) < r \to max_prime_factor r < p.
86 cut (max_prime_factor r \lt max_prime_factor n \lor
87 max_prime_factor r = max_prime_factor n).
89 absurd (nth_prime (max_prime_factor n) \divides r).
91 apply divides_max_prime_factor_n.
92 assumption.unfold Not.
94 cut (r \mod (nth_prime (max_prime_factor n)) \neq O).
95 apply Hcut1.apply divides_to_mod_O.
96 apply lt_O_nth_prime_n.assumption.
97 apply (p_ord_aux_to_not_mod_O n n ? q r).
98 apply lt_SO_nth_prime_n.assumption.
100 rewrite < H1.assumption.
101 apply (le_to_or_lt_eq (max_prime_factor r) (max_prime_factor n)).
102 apply divides_to_max_prime_factor.
103 assumption.assumption.
104 apply (witness r n ((nth_prime p) \sup q)).
106 apply (p_ord_aux_to_exp n n ? q r).
107 apply lt_O_nth_prime_n.assumption.
110 theorem p_ord_to_lt_max_prime_factor1: \forall n,p,q,r. O < n \to
111 max_prime_factor n \le p \to
112 (pair nat nat q r) = p_ord n (nth_prime p) \to
113 (S O) < r \to max_prime_factor r < p.
115 cut (max_prime_factor n < p \lor max_prime_factor n = p).
116 elim Hcut.apply (le_to_lt_to_lt ? (max_prime_factor n)).
117 apply divides_to_max_prime_factor.assumption.assumption.
118 apply (witness r n ((nth_prime p) \sup q)).
120 apply (p_ord_aux_to_exp n n).
121 apply lt_O_nth_prime_n.
122 assumption.assumption.
123 apply (p_ord_to_lt_max_prime_factor n ? q).
124 assumption.apply sym_eq.assumption.assumption.assumption.
125 apply (le_to_or_lt_eq ? p H1).
128 (* datatypes and functions *)
130 inductive nat_fact : Set \def
131 nf_last : nat \to nat_fact
132 | nf_cons : nat \to nat_fact \to nat_fact.
134 inductive nat_fact_all : Set \def
135 nfa_zero : nat_fact_all
136 | nfa_one : nat_fact_all
137 | nfa_proper : nat_fact \to nat_fact_all.
139 let rec factorize_aux p n acc \def
143 match p_ord n (nth_prime p1) with
144 [ (pair q r) \Rightarrow
145 factorize_aux p1 r (nf_cons q acc)]].
147 definition factorize : nat \to nat_fact_all \def \lambda n:nat.
149 [ O \Rightarrow nfa_zero
152 [ O \Rightarrow nfa_one
154 let p \def (max (S(S n2)) (\lambda p:nat.eqb ((S(S n2)) \mod (nth_prime p)) O)) in
155 match p_ord (S(S n2)) (nth_prime p) with
156 [ (pair q r) \Rightarrow
157 nfa_proper (factorize_aux p r (nf_last (pred q)))]]].
159 let rec defactorize_aux f i \def
161 [ (nf_last n) \Rightarrow (nth_prime i) \sup (S n)
162 | (nf_cons n g) \Rightarrow
163 (nth_prime i) \sup n *(defactorize_aux g (S i))].
165 definition defactorize : nat_fact_all \to nat \def
166 \lambda f : nat_fact_all.
168 [ nfa_zero \Rightarrow O
169 | nfa_one \Rightarrow (S O)
170 | (nfa_proper g) \Rightarrow defactorize_aux g O].
172 theorem lt_O_defactorize_aux: \forall f:nat_fact.\forall i:nat.
173 O < defactorize_aux f i.
174 intro.elim f.simplify.unfold lt.
175 rewrite > times_n_SO.
177 change with (O < nth_prime i).
178 apply lt_O_nth_prime_n.
179 change with (O < exp (nth_prime i) n).
181 apply lt_O_nth_prime_n.
183 rewrite > times_n_SO.
185 change with (O < exp (nth_prime i) n).
187 apply lt_O_nth_prime_n.
188 change with (O < defactorize_aux n1 (S i)).
192 theorem lt_SO_defactorize_aux: \forall f:nat_fact.\forall i:nat.
193 S O < defactorize_aux f i.
194 intro.elim f.simplify.unfold lt.
195 rewrite > times_n_SO.
197 change with (S O < nth_prime i).
198 apply lt_SO_nth_prime_n.
199 change with (O < exp (nth_prime i) n).
201 apply lt_O_nth_prime_n.
203 rewrite > times_n_SO.
206 change with (O < exp (nth_prime i) n).
208 apply lt_O_nth_prime_n.
209 change with (S O < defactorize_aux n1 (S i)).
213 theorem defactorize_aux_factorize_aux :
214 \forall p,n:nat.\forall acc:nat_fact.O < n \to
215 ((n=(S O) \land p=O) \lor max_prime_factor n < p) \to
216 defactorize_aux (factorize_aux p n acc) O = n*(defactorize_aux acc p).
217 intro.elim p.simplify.
218 elim H1.elim H2.rewrite > H3.
219 rewrite > sym_times. apply times_n_SO.
220 apply False_ind.apply (not_le_Sn_O (max_prime_factor n) H2).
222 (* generalizing the goal: I guess there exists a better way *)
223 cut (\forall q,r.(pair nat nat q r) = (p_ord_aux n1 n1 (nth_prime n)) \to
224 defactorize_aux match (p_ord_aux n1 n1 (nth_prime n)) with
225 [(pair q r) \Rightarrow (factorize_aux n r (nf_cons q acc))] O =
226 n1*defactorize_aux acc (S n)).
227 apply (Hcut (fst ? ? (p_ord_aux n1 n1 (nth_prime n)))
228 (snd ? ? (p_ord_aux n1 n1 (nth_prime n)))).
229 apply sym_eq.apply eq_pair_fst_snd.
233 cut (n1 = r * (nth_prime n) \sup q).
235 simplify.rewrite < assoc_times.
236 rewrite < Hcut.reflexivity.
237 cut (O < r \lor O = r).
238 elim Hcut1.assumption.absurd (n1 = O).
239 rewrite > Hcut.rewrite < H4.reflexivity.
240 unfold Not. intro.apply (not_le_Sn_O O).
241 rewrite < H5 in \vdash (? ? %).assumption.
242 apply le_to_or_lt_eq.apply le_O_n.
243 cut ((S O) < r \lor (S O) \nlt r).
246 apply (p_ord_to_lt_max_prime_factor1 n1 ? q r).
250 apply (not_eq_O_S n).apply sym_eq.assumption.
253 assumption.assumption.
256 left.split.assumption.reflexivity.
257 intro.right.rewrite > Hcut2.
258 simplify.unfold lt.apply le_S_S.apply le_O_n.
259 cut (r \lt (S O) \or r=(S O)).
260 elim Hcut2.absurd (O=r).
261 apply le_n_O_to_eq.apply le_S_S_to_le.exact H5.
264 apply (not_le_Sn_O O).
265 rewrite > Hcut3 in \vdash (? ? %).
266 assumption.rewrite > Hcut.
267 rewrite < H6.reflexivity.
269 apply (le_to_or_lt_eq r (S O)).
270 apply not_lt_to_le.assumption.
271 apply (decidable_lt (S O) r).
273 apply (p_ord_aux_to_exp n1 n1).
274 apply lt_O_nth_prime_n.assumption.
277 theorem defactorize_factorize: \forall n:nat.defactorize (factorize n) = n.
279 apply (nat_case n).reflexivity.
280 intro.apply (nat_case m).reflexivity.
281 intro.(*CSC: simplify here does something really nasty *)
283 (let p \def (max (S(S m1)) (\lambda p:nat.eqb ((S(S m1)) \mod (nth_prime p)) O)) in
284 defactorize (match p_ord (S(S m1)) (nth_prime p) with
285 [ (pair q r) \Rightarrow
286 nfa_proper (factorize_aux p r (nf_last (pred q)))])=(S(S m1))).
288 (* generalizing the goal; find a better way *)
289 cut (\forall q,r.(pair nat nat q r) = (p_ord (S(S m1)) (nth_prime p)) \to
290 defactorize (match p_ord (S(S m1)) (nth_prime p) with
291 [ (pair q r) \Rightarrow
292 nfa_proper (factorize_aux p r (nf_last (pred q)))])=(S(S m1))).
293 apply (Hcut (fst ? ? (p_ord (S(S m1)) (nth_prime p)))
294 (snd ? ? (p_ord (S(S m1)) (nth_prime p)))).
295 apply sym_eq.apply eq_pair_fst_snd.
299 cut ((S(S m1)) = (nth_prime p) \sup q *r).
301 rewrite > defactorize_aux_factorize_aux.
302 (*CSC: simplify here does something really nasty *)
303 change with (r*(nth_prime p) \sup (S (pred q)) = (S(S m1))).
304 cut ((S (pred q)) = q).
308 apply (p_ord_aux_to_exp (S(S m1))).
309 apply lt_O_nth_prime_n.
312 apply sym_eq. apply S_pred.
313 cut (O < q \lor O = q).
314 elim Hcut2.assumption.
315 absurd (nth_prime p \divides S (S m1)).
316 apply (divides_max_prime_factor_n (S (S m1))).
317 unfold lt.apply le_S_S.apply le_S_S. apply le_O_n.
319 rewrite > Hcut3 in \vdash (? (? ? %)).
320 (*CSC: simplify here does something really nasty *)
321 change with (nth_prime p \divides r \to False).
323 apply (p_ord_aux_to_not_mod_O (S(S m1)) (S(S m1)) (nth_prime p) q r).
324 apply lt_SO_nth_prime_n.
325 unfold lt.apply le_S_S.apply le_O_n.apply le_n.
327 apply divides_to_mod_O.apply lt_O_nth_prime_n.assumption.
328 rewrite > times_n_SO in \vdash (? ? ? %).
330 rewrite > (exp_n_O (nth_prime p)).
331 rewrite > H1 in \vdash (? ? ? (? (? ? %) ?)).
333 apply le_to_or_lt_eq.apply le_O_n.assumption.
334 (* e adesso l'ultimo goal. TASSI: che ora non e' piu' l'ultimo :P *)
335 cut ((S O) < r \lor S O \nlt r).
338 apply (p_ord_to_lt_max_prime_factor1 (S(S m1)) ? q r).
339 unfold lt.apply le_S_S. apply le_O_n.
341 assumption.assumption.
344 left.split.assumption.reflexivity.
345 intro.right.rewrite > Hcut3.
346 simplify.unfold lt.apply le_S_S.apply le_O_n.
347 cut (r \lt (S O) \or r=(S O)).
348 elim Hcut3.absurd (O=r).
349 apply le_n_O_to_eq.apply le_S_S_to_le.exact H2.
351 apply (not_le_Sn_O O).
352 rewrite > H3 in \vdash (? ? %).assumption.assumption.
353 apply (le_to_or_lt_eq r (S O)).
354 apply not_lt_to_le.assumption.
355 apply (decidable_lt (S O) r).
357 cut (O < r \lor O = r).
358 elim Hcut1.assumption.
360 apply (not_eq_O_S (S m1)).
361 rewrite > Hcut.rewrite < H1.rewrite < times_n_O.reflexivity.
362 apply le_to_or_lt_eq.apply le_O_n.
365 apply (p_ord_aux_to_exp (S(S m1))).
366 apply lt_O_nth_prime_n.
373 [ (nf_last n) \Rightarrow O
374 | (nf_cons n g) \Rightarrow S (max_p g)].
376 let rec max_p_exponent f \def
378 [ (nf_last n) \Rightarrow n
379 | (nf_cons n g) \Rightarrow max_p_exponent g].
381 theorem divides_max_p_defactorize: \forall f:nat_fact.\forall i:nat.
382 nth_prime ((max_p f)+i) \divides defactorize_aux f i.
384 elim f.simplify.apply (witness ? ? ((nth_prime i) \sup n)).
387 (nth_prime (S(max_p n1)+i) \divides
388 (nth_prime i) \sup n *(defactorize_aux n1 (S i))).
392 rewrite > assoc_times.
394 apply (witness ? ? (n2* (nth_prime i) \sup n)).
398 theorem divides_exp_to_divides:
399 \forall p,n,m:nat. prime p \to
400 p \divides n \sup m \to p \divides n.
401 intros 3.elim m.simplify in H1.
402 apply (transitive_divides p (S O)).assumption.
404 cut (p \divides n \lor p \divides n \sup n1).
405 elim Hcut.assumption.
406 apply H.assumption.assumption.
407 apply divides_times_to_divides.assumption.
411 theorem divides_exp_to_eq:
412 \forall p,q,m:nat. prime p \to prime q \to
413 p \divides q \sup m \to p = q.
417 apply (divides_exp_to_divides p q m).
418 assumption.assumption.
419 unfold prime in H.elim H.assumption.
422 theorem not_divides_defactorize_aux: \forall f:nat_fact. \forall i,j:nat.
423 i < j \to nth_prime i \ndivides defactorize_aux f j.
426 (nth_prime i \divides (nth_prime j) \sup (S n) \to False).
427 intro.absurd ((nth_prime i) = (nth_prime j)).
428 apply (divides_exp_to_eq ? ? (S n)).
429 apply prime_nth_prime.apply prime_nth_prime.
430 assumption.unfold Not.
432 apply (not_le_Sn_n i).rewrite > Hcut in \vdash (? ? %).assumption.
433 apply (injective_nth_prime ? ? H2).
436 cut (nth_prime i \divides (nth_prime j) \sup n
437 \lor nth_prime i \divides defactorize_aux n1 (S j)).
439 absurd ((nth_prime i) = (nth_prime j)).
440 apply (divides_exp_to_eq ? ? n).
441 apply prime_nth_prime.apply prime_nth_prime.
442 assumption.unfold Not.
445 apply (not_le_Sn_n i).rewrite > Hcut1 in \vdash (? ? %).assumption.
446 apply (injective_nth_prime ? ? H4).
448 apply (trans_lt ? j).assumption.unfold lt.apply le_n.
450 apply divides_times_to_divides.
451 apply prime_nth_prime.assumption.
454 lemma not_eq_nf_last_nf_cons: \forall g:nat_fact.\forall n,m,i:nat.
455 \lnot (defactorize_aux (nf_last n) i= defactorize_aux (nf_cons m g) i).
458 (exp (nth_prime i) (S n) = defactorize_aux (nf_cons m g) i \to False).
460 cut (S(max_p g)+i= i).
461 apply (not_le_Sn_n i).
462 rewrite < Hcut in \vdash (? ? %).
463 simplify.apply le_S_S.
465 apply injective_nth_prime.
466 apply (divides_exp_to_eq ? ? (S n)).
467 apply prime_nth_prime.apply prime_nth_prime.
469 change with (divides (nth_prime ((max_p (nf_cons m g))+i))
470 (defactorize_aux (nf_cons m g) i)).
471 apply divides_max_p_defactorize.
474 lemma not_eq_nf_cons_O_nf_cons: \forall f,g:nat_fact.\forall n,i:nat.
475 \lnot (defactorize_aux (nf_cons O f) i= defactorize_aux (nf_cons (S n) g) i).
477 simplify.unfold Not.rewrite < plus_n_O.
479 apply (not_divides_defactorize_aux f i (S i) ?).
480 unfold lt.apply le_n.
482 rewrite > assoc_times.
483 apply (witness ? ? ((exp (nth_prime i) n)*(defactorize_aux g (S i)))).
487 theorem eq_defactorize_aux_to_eq: \forall f,g:nat_fact.\forall i:nat.
488 defactorize_aux f i = defactorize_aux g i \to f = g.
491 generalize in match H.
494 apply inj_S. apply (inj_exp_r (nth_prime i)).
495 apply lt_SO_nth_prime_n.
498 apply (not_eq_nf_last_nf_cons n2 n n1 i H2).
499 generalize in match H1.
502 apply (not_eq_nf_last_nf_cons n1 n2 n i).
503 apply sym_eq. assumption.
505 generalize in match H3.
506 apply (nat_elim2 (\lambda n,n2.
507 ((nth_prime i) \sup n)*(defactorize_aux n1 (S i)) =
508 ((nth_prime i) \sup n2)*(defactorize_aux n3 (S i)) \to
509 nf_cons n n1 = nf_cons n2 n3)).
515 rewrite > (plus_n_O (defactorize_aux n3 (S i))).assumption.
517 apply (not_eq_nf_cons_O_nf_cons n1 n3 n5 i).assumption.
520 apply (not_eq_nf_cons_O_nf_cons n3 n1 n4 i).
521 apply sym_eq.assumption.
523 cut (nf_cons n4 n1 = nf_cons m n3).
526 rewrite > Hcut1.rewrite > Hcut2.reflexivity.
528 (match nf_cons n4 n1 with
529 [ (nf_last m) \Rightarrow n1
530 | (nf_cons m g) \Rightarrow g ] = n3).
531 rewrite > Hcut.simplify.reflexivity.
533 (match nf_cons n4 n1 with
534 [ (nf_last m) \Rightarrow m
535 | (nf_cons m g) \Rightarrow m ] = m).
536 rewrite > Hcut.simplify.reflexivity.
537 apply H4.simplify in H5.
538 apply (inj_times_r1 (nth_prime i)).
539 apply lt_O_nth_prime_n.
540 rewrite < assoc_times.rewrite < assoc_times.assumption.
543 theorem injective_defactorize_aux: \forall i:nat.
544 injective nat_fact nat (\lambda f.defactorize_aux f i).
547 apply (eq_defactorize_aux_to_eq x y i H).
550 theorem injective_defactorize:
551 injective nat_fact_all nat defactorize.
553 change with (\forall f,g.defactorize f = defactorize g \to f=g).
555 generalize in match H.elim g.
561 apply (not_eq_O_S O H1).
565 apply (not_le_Sn_n O).
566 rewrite > H1 in \vdash (? ? %).
567 change with (O < defactorize_aux n O).
568 apply lt_O_defactorize_aux.
569 generalize in match H.
574 apply (not_eq_O_S O).apply sym_eq. assumption.
580 apply (not_le_Sn_n (S O)).
581 rewrite > H1 in \vdash (? ? %).
582 change with ((S O) < defactorize_aux n O).
583 apply lt_SO_defactorize_aux.
584 generalize in match H.elim g.
588 apply (not_le_Sn_n O).
589 rewrite < H1 in \vdash (? ? %).
590 change with (O < defactorize_aux n O).
591 apply lt_O_defactorize_aux.
595 apply (not_le_Sn_n (S O)).
596 rewrite < H1 in \vdash (? ? %).
597 change with ((S O) < defactorize_aux n O).
598 apply lt_SO_defactorize_aux.
599 (* proper - proper *)
601 apply (injective_defactorize_aux O).
605 theorem factorize_defactorize:
606 \forall f,g: nat_fact_all. factorize (defactorize f) = f.
608 apply injective_defactorize.
609 apply defactorize_factorize.