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