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