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