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