]> matita.cs.unibo.it Git - helm.git/blob - matita/library/nat/factorization.ma
- renamed ocaml/ to components/
[helm.git] / 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/ord.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 (n \mod (nth_prime p)) O)).
25
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);
34   [ elim Hcut.
35     apply (ex_intro nat ? a);
36     split;
37     [ apply (trans_le a (nth_prime a));
38       [ apply le_n_fn;
39         exact lt_nth_prime_n_nth_prime_Sn;
40       | rewrite > H1;
41         apply le_smallest_factor_n; ]
42     | rewrite > H1;
43       change with (divides_b (smallest_factor n) n = true);
44       apply divides_to_divides_b_true;
45       [ apply (trans_lt ? (S O));
46         [ unfold lt; apply le_n;
47         | apply lt_SO_smallest_factor; assumption; ]
48       | apply divides_smallest_factor_n;
49         apply (trans_lt ? (S O));
50         [ unfold lt; apply le_n;
51         | assumption; ] ] ]
52   | apply prime_to_nth_prime;
53     apply prime_smallest_factor_n;
54     assumption; ] ]
55 qed.
56
57 theorem divides_to_max_prime_factor : \forall n,m. (S O) < n \to O < m \to n \divides m \to 
58 max_prime_factor n \le max_prime_factor m.
59 intros.change with
60 ((max n (\lambda p:nat.eqb (n \mod (nth_prime p)) O)) \le
61 (max m (\lambda p:nat.eqb (m \mod (nth_prime p)) O))).
62 apply f_m_to_le_max.
63 apply (trans_le ? n).
64 apply le_max_n.apply divides_to_le.assumption.assumption.
65 change with (divides_b (nth_prime (max_prime_factor n)) m = true).
66 apply divides_to_divides_b_true.
67 cut (prime (nth_prime (max_prime_factor n))).
68 apply lt_O_nth_prime_n.apply prime_nth_prime.
69 cut (nth_prime (max_prime_factor n) \divides n).
70 apply (transitive_divides ? n).
71 apply divides_max_prime_factor_n.
72 assumption.assumption.
73 apply divides_b_true_to_divides.
74 apply lt_O_nth_prime_n.
75 apply divides_to_divides_b_true.
76 apply lt_O_nth_prime_n.
77 apply divides_max_prime_factor_n.
78 assumption.
79 qed.
80
81 theorem p_ord_to_lt_max_prime_factor: \forall n,p,q,r. O < n \to
82 p = max_prime_factor n \to 
83 (pair nat nat q r) = p_ord n (nth_prime p) \to
84 (S O) < r \to max_prime_factor r < p.
85 intros.
86 rewrite > H1.
87 cut (max_prime_factor r \lt max_prime_factor n \lor
88     max_prime_factor r = max_prime_factor n).
89 elim Hcut.assumption.
90 absurd (nth_prime (max_prime_factor n) \divides r).
91 rewrite < H4.
92 apply divides_max_prime_factor_n.
93 assumption.
94 change with (nth_prime (max_prime_factor n) \divides r \to False).
95 intro.
96 cut (r \mod (nth_prime (max_prime_factor n)) \neq O).
97 apply Hcut1.apply divides_to_mod_O.
98 apply lt_O_nth_prime_n.assumption.
99 apply (p_ord_aux_to_not_mod_O n n ? q r).
100 apply lt_SO_nth_prime_n.assumption.
101 apply le_n.
102 rewrite < H1.assumption.
103 apply (le_to_or_lt_eq (max_prime_factor r)  (max_prime_factor n)).
104 apply divides_to_max_prime_factor.
105 assumption.assumption.
106 apply (witness r n ((nth_prime p) \sup q)).
107 rewrite < sym_times.
108 apply (p_ord_aux_to_exp n n ? q r).
109 apply lt_O_nth_prime_n.assumption.
110 qed.
111
112 theorem p_ord_to_lt_max_prime_factor1: \forall n,p,q,r. O < n \to
113 max_prime_factor n \le p \to 
114 (pair nat nat q r) = p_ord n (nth_prime p) \to
115 (S O) < r \to max_prime_factor r < p.
116 intros.
117 cut (max_prime_factor n < p \lor max_prime_factor n = p).
118 elim Hcut.apply (le_to_lt_to_lt ? (max_prime_factor n)).
119 apply divides_to_max_prime_factor.assumption.assumption.
120 apply (witness r n ((nth_prime p) \sup q)).
121 rewrite > sym_times.
122 apply (p_ord_aux_to_exp n n).
123 apply lt_O_nth_prime_n.
124 assumption.assumption.
125 apply (p_ord_to_lt_max_prime_factor n ? q).
126 assumption.apply sym_eq.assumption.assumption.assumption.
127 apply (le_to_or_lt_eq ? p H1).
128 qed.
129
130 (* datatypes and functions *)
131
132 inductive nat_fact : Set \def
133     nf_last : nat \to nat_fact   
134   | nf_cons : nat \to nat_fact \to nat_fact.
135
136 inductive nat_fact_all : Set \def
137     nfa_zero : nat_fact_all
138   | nfa_one : nat_fact_all
139   | nfa_proper : nat_fact \to nat_fact_all.
140
141 let rec factorize_aux p n acc \def
142   match p with 
143   [ O \Rightarrow acc
144   | (S p1) \Rightarrow 
145     match p_ord n (nth_prime p1) with
146     [ (pair q r) \Rightarrow 
147       factorize_aux p1 r (nf_cons q acc)]].
148   
149 definition factorize : nat \to nat_fact_all \def \lambda n:nat.
150   match n with
151     [ O \Rightarrow nfa_zero
152     | (S n1) \Rightarrow
153       match n1 with
154       [ O \Rightarrow nfa_one
155     | (S n2) \Rightarrow 
156       let p \def (max (S(S n2)) (\lambda p:nat.eqb ((S(S n2)) \mod (nth_prime p)) O)) in
157       match p_ord (S(S n2)) (nth_prime p) with
158       [ (pair q r) \Rightarrow 
159            nfa_proper (factorize_aux p r (nf_last (pred q)))]]].
160            
161 let rec defactorize_aux f i \def
162   match f with
163   [ (nf_last n) \Rightarrow (nth_prime i) \sup (S n)
164   | (nf_cons n g) \Rightarrow 
165       (nth_prime i) \sup n *(defactorize_aux g (S i))].
166       
167 definition defactorize : nat_fact_all \to nat \def
168 \lambda f : nat_fact_all. 
169 match f with 
170 [ nfa_zero \Rightarrow O
171 | nfa_one \Rightarrow (S O)
172 | (nfa_proper g) \Rightarrow defactorize_aux g O]. 
173
174 theorem lt_O_defactorize_aux: \forall f:nat_fact.\forall i:nat.
175 O < defactorize_aux f i.
176 intro.elim f.simplify.unfold lt. 
177 rewrite > times_n_SO.
178 apply le_times.
179 change with (O < nth_prime i).
180 apply lt_O_nth_prime_n.
181 change with (O < exp (nth_prime i) n).
182 apply lt_O_exp.
183 apply lt_O_nth_prime_n.
184 simplify.unfold lt.
185 rewrite > times_n_SO.
186 apply le_times.
187 change with (O < exp (nth_prime i) n).
188 apply lt_O_exp.
189 apply lt_O_nth_prime_n.
190 change with (O < defactorize_aux n1 (S i)).
191 apply H.
192 qed.
193
194 theorem lt_SO_defactorize_aux: \forall f:nat_fact.\forall i:nat.
195 S O < defactorize_aux f i.
196 intro.elim f.simplify.unfold lt.
197 rewrite > times_n_SO.
198 apply le_times.
199 change with (S O < nth_prime i).
200 apply lt_SO_nth_prime_n.
201 change with (O < exp (nth_prime i) n).
202 apply lt_O_exp.
203 apply lt_O_nth_prime_n.
204 simplify.unfold lt.
205 rewrite > times_n_SO.
206 rewrite > sym_times.
207 apply le_times.
208 change with (O < exp (nth_prime i) n).
209 apply lt_O_exp.
210 apply lt_O_nth_prime_n.
211 change with (S O < defactorize_aux n1 (S i)).
212 apply H.
213 qed.
214
215 theorem defactorize_aux_factorize_aux : 
216 \forall p,n:nat.\forall acc:nat_fact.O < n \to
217 ((n=(S O) \land p=O) \lor max_prime_factor n < p) \to
218 defactorize_aux (factorize_aux p n acc) O = n*(defactorize_aux acc p).
219 intro.elim p.simplify.
220 elim H1.elim H2.rewrite > H3.
221 rewrite > sym_times. apply times_n_SO.
222 apply False_ind.apply (not_le_Sn_O (max_prime_factor n) H2).
223 simplify.
224 (* generalizing the goal: I guess there exists a better way *)
225 cut (\forall q,r.(pair nat nat q r) = (p_ord_aux n1 n1 (nth_prime n)) \to
226 defactorize_aux match (p_ord_aux n1 n1 (nth_prime n)) with
227 [(pair q r)  \Rightarrow (factorize_aux n r (nf_cons q acc))] O =
228 n1*defactorize_aux acc (S n)).
229 apply (Hcut (fst ? ? (p_ord_aux n1 n1 (nth_prime n)))
230 (snd ? ? (p_ord_aux n1 n1 (nth_prime n)))).
231 apply sym_eq.apply eq_pair_fst_snd.
232 intros.
233 rewrite < H3.
234 simplify.
235 cut (n1 = r * (nth_prime n) \sup q).
236 rewrite > H.
237 simplify.rewrite < assoc_times.
238 rewrite < Hcut.reflexivity.
239 cut (O < r \lor O = r).
240 elim Hcut1.assumption.absurd (n1 = O).
241 rewrite > Hcut.rewrite < H4.reflexivity.
242 unfold Not. intro.apply (not_le_Sn_O O).
243 rewrite < H5 in \vdash (? ? %).assumption.
244 apply le_to_or_lt_eq.apply le_O_n.
245 cut ((S O) < r \lor (S O) \nlt r).
246 elim Hcut1.
247 right.
248 apply (p_ord_to_lt_max_prime_factor1 n1 ? q r).
249 assumption.elim H2.
250 elim H5.
251 apply False_ind.
252 apply (not_eq_O_S n).apply sym_eq.assumption.
253 apply le_S_S_to_le.
254 exact H5.
255 assumption.assumption.
256 cut (r=(S O)).
257 apply (nat_case n).
258 left.split.assumption.reflexivity.
259 intro.right.rewrite > Hcut2.
260 simplify.unfold lt.apply le_S_S.apply le_O_n.
261 cut (r \lt (S O) \or r=(S O)).
262 elim Hcut2.absurd (O=r).
263 apply le_n_O_to_eq.apply le_S_S_to_le.exact H5.
264 unfold Not.intro.
265 cut (O=n1).
266 apply (not_le_Sn_O O).
267 rewrite > Hcut3 in \vdash (? ? %).
268 assumption.rewrite > Hcut. 
269 rewrite < H6.reflexivity.
270 assumption.
271 apply (le_to_or_lt_eq r (S O)).
272 apply not_lt_to_le.assumption.
273 apply (decidable_lt (S O) r).
274 rewrite > sym_times.
275 apply (p_ord_aux_to_exp n1 n1).
276 apply lt_O_nth_prime_n.assumption.
277 qed.
278
279 theorem defactorize_factorize: \forall n:nat.defactorize (factorize n) = n.
280 intro.
281 apply (nat_case n).reflexivity.
282 intro.apply (nat_case m).reflexivity.
283 intro.change with  
284 (let p \def (max (S(S m1)) (\lambda p:nat.eqb ((S(S m1)) \mod (nth_prime p)) O)) in
285 defactorize (match p_ord (S(S m1)) (nth_prime p) with
286 [ (pair q r) \Rightarrow 
287    nfa_proper (factorize_aux p r (nf_last (pred q)))])=(S(S m1))).
288 intro.
289 (* generalizing the goal; find a better way *)
290 cut (\forall q,r.(pair nat nat q r) = (p_ord (S(S m1)) (nth_prime p)) \to
291 defactorize (match p_ord (S(S m1)) (nth_prime p) with
292 [ (pair q r) \Rightarrow 
293    nfa_proper (factorize_aux p r (nf_last (pred q)))])=(S(S m1))).
294 apply (Hcut (fst ? ? (p_ord (S(S m1)) (nth_prime p)))
295 (snd ? ? (p_ord (S(S m1)) (nth_prime p)))).
296 apply sym_eq.apply eq_pair_fst_snd.
297 intros.
298 rewrite < H.
299 change with 
300 (defactorize_aux (factorize_aux p r (nf_last (pred q))) O = (S(S m1))).
301 cut ((S(S m1)) = (nth_prime p) \sup q *r).
302 cut (O<r).
303 rewrite > defactorize_aux_factorize_aux.
304 change with (r*(nth_prime p) \sup (S (pred q)) = (S(S m1))).
305 cut ((S (pred q)) = q).
306 rewrite > Hcut2.
307 rewrite > sym_times.
308 apply sym_eq.
309 apply (p_ord_aux_to_exp (S(S m1))).
310 apply lt_O_nth_prime_n.
311 assumption.
312 (* O < q *)
313 apply sym_eq. apply S_pred.
314 cut (O < q \lor O = q).
315 elim Hcut2.assumption.
316 absurd (nth_prime p \divides S (S m1)).
317 apply (divides_max_prime_factor_n (S (S m1))).
318 unfold lt.apply le_S_S.apply le_S_S. apply le_O_n.
319 cut ((S(S m1)) = r).
320 rewrite > Hcut3 in \vdash (? (? ? %)).
321 change with (nth_prime p \divides r \to False).
322 intro.
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.
326 assumption.
327 apply divides_to_mod_O.apply lt_O_nth_prime_n.assumption.
328 rewrite > times_n_SO in \vdash (? ? ? %).
329 rewrite < sym_times.
330 rewrite > (exp_n_O (nth_prime p)).
331 rewrite > H1 in \vdash (? ? ? (? (? ? %) ?)).
332 assumption.
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).
336 elim Hcut2.
337 right. 
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.
340 apply le_n.
341 assumption.assumption.
342 cut (r=(S O)).
343 apply (nat_case p).
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.
350 unfold Not.intro.
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).
356 (* O < r *)
357 cut (O < r \lor O = r).
358 elim Hcut1.assumption. 
359 apply False_ind.
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.
363 (* prova del cut *)
364 goal 20.
365 apply (p_ord_aux_to_exp (S(S m1))).
366 apply lt_O_nth_prime_n.
367 assumption.
368 (* fine prova cut *)
369 qed.
370
371 let rec max_p f \def
372 match f with
373 [ (nf_last n) \Rightarrow O
374 | (nf_cons n g) \Rightarrow S (max_p g)].
375
376 let rec max_p_exponent f \def
377 match f with
378 [ (nf_last n) \Rightarrow n
379 | (nf_cons n g) \Rightarrow max_p_exponent g].
380
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.
383 intro.
384 elim f.simplify.apply (witness ? ? ((nth_prime i) \sup n)).
385 reflexivity.
386 change with 
387 (nth_prime (S(max_p n1)+i) \divides
388 (nth_prime i) \sup n *(defactorize_aux n1 (S i))).
389 elim (H (S i)).
390 rewrite > H1.
391 rewrite < sym_times.
392 rewrite > assoc_times.
393 rewrite < plus_n_Sm.
394 apply (witness ? ? (n2* (nth_prime i) \sup n)).
395 reflexivity.
396 qed.
397
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.
403 apply divides_SO_n.
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.
408 exact H2.
409 qed.
410
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.
414 intros.
415 unfold prime in H1.
416 elim H1.apply H4.
417 apply (divides_exp_to_divides p q m).
418 assumption.assumption.
419 unfold prime in H.elim H.assumption.
420 qed.
421
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.
424 intro.elim f.
425 change with
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.
431 change with ((nth_prime i) = (nth_prime j) \to False).
432 intro.cut (i = j).
433 apply (not_le_Sn_n i).rewrite > Hcut in \vdash (? ? %).assumption.
434 apply (injective_nth_prime ? ? H2).
435 change with 
436 (nth_prime i \divides (nth_prime j) \sup n *(defactorize_aux n1 (S j)) \to False).
437 intro.
438 cut (nth_prime i \divides (nth_prime j) \sup n
439 \lor nth_prime i \divides defactorize_aux n1 (S j)).
440 elim Hcut.
441 absurd ((nth_prime i) = (nth_prime j)).
442 apply (divides_exp_to_eq ? ? n).
443 apply prime_nth_prime.apply prime_nth_prime.
444 assumption.
445 change with ((nth_prime i) = (nth_prime j) \to False).
446 intro.
447 cut (i = j).
448 apply (not_le_Sn_n i).rewrite > Hcut1 in \vdash (? ? %).assumption.
449 apply (injective_nth_prime ? ? H4).
450 apply (H i (S j)).
451 apply (trans_lt ? j).assumption.unfold lt.apply le_n.
452 assumption.
453 apply divides_times_to_divides.
454 apply prime_nth_prime.assumption.
455 qed.
456
457 lemma not_eq_nf_last_nf_cons: \forall g:nat_fact.\forall n,m,i:nat.
458 \lnot (defactorize_aux (nf_last n) i= defactorize_aux (nf_cons m g) i).
459 intros.
460 change with 
461 (exp (nth_prime i) (S n) = defactorize_aux (nf_cons m g) i \to False).
462 intro.
463 cut (S(max_p g)+i= i).
464 apply (not_le_Sn_n i).
465 rewrite < Hcut in \vdash (? ? %).
466 simplify.apply le_S_S.
467 apply le_plus_n.
468 apply injective_nth_prime.
469 (* uffa, perche' semplifica ? *)
470 change with (nth_prime (S(max_p g)+i)= nth_prime i).
471 apply (divides_exp_to_eq ? ? (S n)).
472 apply prime_nth_prime.apply prime_nth_prime.
473 rewrite > H.
474 change with (divides (nth_prime ((max_p (nf_cons m g))+i)) 
475 (defactorize_aux (nf_cons m g) i)).
476 apply divides_max_p_defactorize.
477 qed.
478
479 lemma not_eq_nf_cons_O_nf_cons: \forall f,g:nat_fact.\forall n,i:nat.
480 \lnot (defactorize_aux (nf_cons O f) i= defactorize_aux (nf_cons (S n) g) i).
481 intros.
482 simplify.unfold Not.rewrite < plus_n_O.
483 intro.
484 apply (not_divides_defactorize_aux f i (S i) ?).
485 unfold lt.apply le_n.
486 rewrite > H.
487 rewrite > assoc_times.
488 apply (witness ? ? ((exp (nth_prime i) n)*(defactorize_aux g (S i)))).
489 reflexivity.
490 qed.
491
492 theorem eq_defactorize_aux_to_eq: \forall f,g:nat_fact.\forall i:nat.
493 defactorize_aux f i = defactorize_aux g i \to f = g.
494 intro.
495 elim f.
496 generalize in match H.
497 elim g.
498 apply eq_f.
499 apply inj_S. apply (inj_exp_r (nth_prime i)).
500 apply lt_SO_nth_prime_n.
501 assumption.
502 apply False_ind.
503 apply (not_eq_nf_last_nf_cons n2 n n1 i H2).
504 generalize in match H1.
505 elim g.
506 apply False_ind.
507 apply (not_eq_nf_last_nf_cons n1 n2 n i).
508 apply sym_eq. assumption.
509 simplify in H3.
510 generalize in match H3.
511 apply (nat_elim2 (\lambda n,n2.
512 ((nth_prime i) \sup n)*(defactorize_aux n1 (S i)) =
513 ((nth_prime i) \sup n2)*(defactorize_aux n3 (S i)) \to
514 nf_cons n n1 = nf_cons n2 n3)).
515 intro.
516 elim n4. apply eq_f.
517 apply (H n3 (S i)).
518 simplify in H4.
519 rewrite > plus_n_O.
520 rewrite > (plus_n_O (defactorize_aux n3 (S i))).assumption.
521 apply False_ind.
522 apply (not_eq_nf_cons_O_nf_cons n1 n3 n5 i).assumption.
523 intros.
524 apply False_ind.
525 apply (not_eq_nf_cons_O_nf_cons n3 n1 n4 i).
526 apply sym_eq.assumption.
527 intros.
528 cut (nf_cons n4 n1 = nf_cons m n3).
529 cut (n4=m).
530 cut (n1=n3).
531 rewrite > Hcut1.rewrite > Hcut2.reflexivity.
532 change with 
533 (match nf_cons n4 n1 with
534 [ (nf_last m) \Rightarrow n1
535 | (nf_cons m g) \Rightarrow g ] = n3).
536 rewrite > Hcut.simplify.reflexivity.
537 change with 
538 (match nf_cons n4 n1 with
539 [ (nf_last m) \Rightarrow m
540 | (nf_cons m g) \Rightarrow m ] = m).
541 rewrite > Hcut.simplify.reflexivity.
542 apply H4.simplify in H5.
543 apply (inj_times_r1 (nth_prime i)).
544 apply lt_O_nth_prime_n.
545 rewrite < assoc_times.rewrite < assoc_times.assumption.
546 qed.
547
548 theorem injective_defactorize_aux: \forall i:nat.
549 injective nat_fact nat (\lambda f.defactorize_aux f i).
550 change with (\forall i:nat.\forall f,g:nat_fact.
551 defactorize_aux f i = defactorize_aux g i \to f = g).
552 intros.
553 apply (eq_defactorize_aux_to_eq f g i H).
554 qed.
555
556 theorem injective_defactorize: 
557 injective nat_fact_all nat defactorize.
558 change with (\forall f,g:nat_fact_all.
559 defactorize f = defactorize g \to f = g).
560 intro.elim f.
561 generalize in match H.elim g.
562 (* zero - zero *)
563 reflexivity.
564 (* zero - one *)
565 simplify in H1.
566 apply False_ind.
567 apply (not_eq_O_S O H1).
568 (* zero - proper *)
569 simplify in H1.
570 apply False_ind.
571 apply (not_le_Sn_n O).
572 rewrite > H1 in \vdash (? ? %).
573 change with (O < defactorize_aux n O).
574 apply lt_O_defactorize_aux.
575 generalize in match H.
576 elim g.
577 (* one - zero *)
578 simplify in H1.
579 apply False_ind.
580 apply (not_eq_O_S O).apply sym_eq. assumption.
581 (* one - one *)
582 reflexivity.
583 (* one - proper *)
584 simplify in H1.
585 apply False_ind.
586 apply (not_le_Sn_n (S O)).
587 rewrite > H1 in \vdash (? ? %).
588 change with ((S O) < defactorize_aux n O).
589 apply lt_SO_defactorize_aux.
590 generalize in match H.elim g.
591 (* proper - zero *)
592 simplify in H1.
593 apply False_ind.
594 apply (not_le_Sn_n O).
595 rewrite < H1 in \vdash (? ? %).
596 change with (O < defactorize_aux n O).
597 apply lt_O_defactorize_aux.
598 (* proper - one *)
599 simplify in H1.
600 apply False_ind.
601 apply (not_le_Sn_n (S O)).
602 rewrite < H1 in \vdash (? ? %).
603 change with ((S O) < defactorize_aux n O).
604 apply lt_SO_defactorize_aux.
605 (* proper - proper *)
606 apply eq_f.
607 apply (injective_defactorize_aux O).
608 exact H1.
609 qed.
610
611 theorem factorize_defactorize: 
612 \forall f,g: nat_fact_all. factorize (defactorize f) = f.
613 intros.
614 apply injective_defactorize.
615 (* uffa: perche' semplifica ??? *)
616 change with (defactorize(factorize (defactorize f)) = (defactorize f)).
617 apply defactorize_factorize.
618 qed.
619