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