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