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