1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| A.Asperti, C.Sacerdoti Coen, *)
8 (* ||A|| E.Tassi, S.Zacchiroli *)
10 (* \ / Matita is distributed under the terms of the *)
11 (* v GNU Lesser General Public License Version 2.1 *)
13 (**************************************************************************)
17 (* the following factorization algorithm looks for the largest prime
19 definition max_prime_factor \def \lambda n:nat.
20 (max n (\lambda p:nat.eqb (n \mod (nth_prime p)) O)).
22 theorem lt_SO_max_prime: \forall m. S O < m \to
23 S O < max m (λi:nat.primeb i∧divides_b i m).
25 apply (lt_to_le_to_lt ? (smallest_factor m))
26 [apply lt_SO_smallest_factor.assumption
28 [apply le_smallest_factor_n
29 |apply true_to_true_to_andb_true
30 [apply prime_to_primeb_true.
31 apply prime_smallest_factor_n.
33 |apply divides_to_divides_b_true
34 [apply lt_O_smallest_factor.apply lt_to_le.assumption
35 |apply divides_smallest_factor_n.
36 apply lt_to_le.assumption
42 (* max_prime_factor is indeed a factor *)
43 theorem divides_max_prime_factor_n:
44 \forall n:nat. (S O) < n
45 \to nth_prime (max_prime_factor n) \divides n.
47 apply divides_b_true_to_divides.
48 apply (f_max_true (\lambda p:nat.eqb (n \mod (nth_prime p)) O) n);
49 cut (\exists i. nth_prime i = smallest_factor n);
51 apply (ex_intro nat ? a);
53 [ apply (trans_le a (nth_prime a));
55 exact lt_nth_prime_n_nth_prime_Sn;
57 apply le_smallest_factor_n; ]
59 change with (divides_b (smallest_factor n) n = true);
60 apply divides_to_divides_b_true;
61 [ apply (trans_lt ? (S O));
62 [ unfold lt; apply le_n;
63 | apply lt_SO_smallest_factor; assumption; ]
64 | letin x \def le.autobatch.
66 apply divides_smallest_factor_n;
67 apply (trans_lt ? (S O));
68 [ unfold lt; apply le_n;
69 | assumption; ] *) ] ]
72 apply prime_to_nth_prime;
73 apply prime_smallest_factor_n;
77 lemma divides_to_prime_divides : \forall n,m.1 < m \to m < n \to m \divides n \to
78 \exists p.p \leq m \land prime p \land p \divides n.
79 intros;apply (ex_intro ? ? (nth_prime (max_prime_factor m)));split
82 [apply lt_to_le;assumption
83 |apply divides_max_prime_factor_n;assumption]
84 |apply prime_nth_prime;]
85 |apply (transitive_divides ? ? ? ? H2);apply divides_max_prime_factor_n;
89 theorem divides_to_max_prime_factor : \forall n,m. (S O) < n \to O < m \to n \divides m \to
90 max_prime_factor n \le max_prime_factor m.
91 intros.unfold max_prime_factor.
94 apply le_max_n.apply divides_to_le.assumption.assumption.
95 change with (divides_b (nth_prime (max_prime_factor n)) m = true).
96 apply divides_to_divides_b_true.
97 cut (prime (nth_prime (max_prime_factor n))).
98 apply lt_O_nth_prime_n.apply prime_nth_prime.
99 cut (nth_prime (max_prime_factor n) \divides n).
103 [ apply (transitive_divides ? n);
104 [ apply divides_max_prime_factor_n.
108 | apply divides_b_true_to_divides;
109 [ apply lt_O_nth_prime_n.
110 | apply divides_to_divides_b_true;
111 [ apply lt_O_nth_prime_n.
112 | apply divides_max_prime_factor_n.
120 theorem divides_to_max_prime_factor1 : \forall n,m. O < n \to O < m \to n \divides m \to
121 max_prime_factor n \le max_prime_factor m.
123 elim (le_to_or_lt_eq ? ? H)
124 [apply divides_to_max_prime_factor
125 [assumption|assumption|assumption]
127 simplify.apply le_O_n.
131 theorem max_prime_factor_to_not_p_ord_O : \forall n,p,r.
133 p = max_prime_factor n \to
134 p_ord n (nth_prime p) \neq pair nat nat O r.
135 intros.unfold Not.intro.
136 apply (p_ord_O_to_not_divides ? ? ? ? H2)
137 [apply (trans_lt ? (S O))[apply lt_O_S|assumption]
139 apply divides_max_prime_factor_n.
144 theorem p_ord_to_lt_max_prime_factor: \forall n,p,q,r. O < n \to
145 p = max_prime_factor n \to
146 (pair nat nat q r) = p_ord n (nth_prime p) \to
147 (S O) < r \to max_prime_factor r < p.
150 cut (max_prime_factor r \lt max_prime_factor n \lor
151 max_prime_factor r = max_prime_factor n).
152 elim Hcut.assumption.
153 absurd (nth_prime (max_prime_factor n) \divides r).
155 apply divides_max_prime_factor_n.
156 assumption.unfold Not.
158 cut (r \mod (nth_prime (max_prime_factor n)) \neq O);
159 [unfold Not in Hcut1.autobatch.
161 apply Hcut1.apply divides_to_mod_O;
162 [ apply lt_O_nth_prime_n.
167 cut(pair nat nat q r=p_ord_aux n n (nth_prime (max_prime_factor n)));
168 [2: rewrite < H1.assumption.|letin x \def le.autobatch width = 4 depth = 2]
169 (* CERCA COME MAI le_n non lo applica se lo trova come Const e non Rel *)
172 apply (p_ord_aux_to_not_mod_O n n ? q r);
173 [ apply lt_SO_nth_prime_n.
176 | rewrite < H1.assumption.
180 apply (le_to_or_lt_eq (max_prime_factor r) (max_prime_factor n)).
181 apply divides_to_max_prime_factor.
182 assumption.assumption.
183 apply (witness r n ((nth_prime p) \sup q)).
185 apply (p_ord_aux_to_exp n n ? q r).
186 apply lt_O_nth_prime_n.assumption.
189 theorem p_ord_to_lt_max_prime_factor1: \forall n,p,q,r. O < n \to
190 max_prime_factor n \le p \to
191 (pair nat nat q r) = p_ord n (nth_prime p) \to
192 (S O) < r \to max_prime_factor r < p.
194 cut (max_prime_factor n < p \lor max_prime_factor n = p).
195 elim Hcut.apply (le_to_lt_to_lt ? (max_prime_factor n)).
196 apply divides_to_max_prime_factor.assumption.assumption.
197 apply (witness r n ((nth_prime p) \sup q)).
199 apply (p_ord_aux_to_exp n n).
200 apply lt_O_nth_prime_n.
201 assumption.assumption.
202 apply (p_ord_to_lt_max_prime_factor n ? q).
203 assumption.apply sym_eq.assumption.assumption.assumption.
204 apply (le_to_or_lt_eq ? p H1).
207 lemma lt_max_prime_factor_to_not_divides: \forall n,p:nat.
208 O < n \to n=S O \lor max_prime_factor n < p \to
209 (nth_prime p \ndivides n).
210 intros.unfold Not.intro.
213 apply (le_to_not_lt (nth_prime p) (S O))
214 [apply divides_to_le[apply le_n|assumption]
215 |apply lt_SO_nth_prime_n
217 |apply (not_le_Sn_n p).
219 apply (le_to_lt_to_lt ? ? ? ? H3).
220 unfold max_prime_factor.
222 [apply (trans_le ? (nth_prime p))
224 apply lt_n_nth_prime_n
225 |apply divides_to_le;assumption
227 |apply eq_to_eqb_true.
228 apply divides_to_mod_O
229 [apply lt_O_nth_prime_n|assumption]
234 (* datatypes and functions *)
236 inductive nat_fact : Set \def
237 nf_last : nat \to nat_fact
238 | nf_cons : nat \to nat_fact \to nat_fact.
240 inductive nat_fact_all : Set \def
241 nfa_zero : nat_fact_all
242 | nfa_one : nat_fact_all
243 | nfa_proper : nat_fact \to nat_fact_all.
245 let rec factorize_aux p n acc \def
249 match p_ord n (nth_prime p1) with
250 [ (pair q r) \Rightarrow
251 factorize_aux p1 r (nf_cons q acc)]].
253 definition factorize : nat \to nat_fact_all \def \lambda n:nat.
255 [ O \Rightarrow nfa_zero
258 [ O \Rightarrow nfa_one
260 let p \def (max (S(S n2)) (\lambda p:nat.eqb ((S(S n2)) \mod (nth_prime p)) O)) in
261 match p_ord (S(S n2)) (nth_prime p) with
262 [ (pair q r) \Rightarrow
263 nfa_proper (factorize_aux p r (nf_last (pred q)))]]].
265 let rec defactorize_aux f i \def
267 [ (nf_last n) \Rightarrow (nth_prime i) \sup (S n)
268 | (nf_cons n g) \Rightarrow
269 (nth_prime i) \sup n *(defactorize_aux g (S i))].
271 definition defactorize : nat_fact_all \to nat \def
272 \lambda f : nat_fact_all.
274 [ nfa_zero \Rightarrow O
275 | nfa_one \Rightarrow (S O)
276 | (nfa_proper g) \Rightarrow defactorize_aux g O].
278 theorem lt_O_defactorize_aux:
281 O < defactorize_aux f i.
285 rewrite > times_n_SO;
287 [ change with (O < nth_prime i);
288 apply lt_O_nth_prime_n;
290 change with (O < exp (nth_prime i) n);
292 apply lt_O_nth_prime_n;
293 | change with (O < defactorize_aux n1 (S i));
297 theorem lt_SO_defactorize_aux: \forall f:nat_fact.\forall i:nat.
298 S O < defactorize_aux f i.
299 intro.elim f.simplify.unfold lt.
300 rewrite > times_n_SO.
302 change with (S O < nth_prime i).
303 apply lt_SO_nth_prime_n.
304 change with (O < exp (nth_prime i) n).
306 apply lt_O_nth_prime_n.
308 rewrite > times_n_SO.
311 change with (O < exp (nth_prime i) n).
313 apply lt_O_nth_prime_n.
314 change with (S O < defactorize_aux n1 (S i)).
318 theorem defactorize_aux_factorize_aux :
319 \forall p,n:nat.\forall acc:nat_fact.O < n \to
320 ((n=(S O) \land p=O) \lor max_prime_factor n < p) \to
321 defactorize_aux (factorize_aux p n acc) O = n*(defactorize_aux acc p).
322 intro.elim p.simplify.
323 elim H1.elim H2.rewrite > H3.
324 rewrite > sym_times. apply times_n_SO.
325 apply False_ind.apply (not_le_Sn_O (max_prime_factor n) H2).
327 (* generalizing the goal: I guess there exists a better way *)
328 cut (\forall q,r.(pair nat nat q r) = (p_ord_aux n1 n1 (nth_prime n)) \to
329 defactorize_aux match (p_ord_aux n1 n1 (nth_prime n)) with
330 [(pair q r) \Rightarrow (factorize_aux n r (nf_cons q acc))] O =
331 n1*defactorize_aux acc (S n)).
332 apply (Hcut (fst ? ? (p_ord_aux n1 n1 (nth_prime n)))
333 (snd ? ? (p_ord_aux n1 n1 (nth_prime n)))).
334 apply sym_eq.apply eq_pair_fst_snd.
338 cut (n1 = r * (nth_prime n) \sup q).
340 simplify.rewrite < assoc_times.
341 rewrite < Hcut.reflexivity.
342 cut (O < r \lor O = r).
343 elim Hcut1.assumption.absurd (n1 = O).
344 rewrite > Hcut.rewrite < H4.reflexivity.
345 unfold Not. intro.apply (not_le_Sn_O O).
346 rewrite < H5 in \vdash (? ? %).assumption.
347 apply le_to_or_lt_eq.apply le_O_n.
348 cut ((S O) < r \lor (S O) \nlt r).
351 apply (p_ord_to_lt_max_prime_factor1 n1 ? q r).
355 apply (not_eq_O_S n).apply sym_eq.assumption.
358 assumption.assumption.
361 left.split.assumption.reflexivity.
362 intro.right.rewrite > Hcut2.
363 simplify.unfold lt.apply le_S_S.apply le_O_n.
364 cut (r < (S O) ∨ r=(S O)).
365 elim Hcut2.absurd (O=r).
366 apply le_n_O_to_eq.apply le_S_S_to_le.exact H5.
369 apply (not_le_Sn_O O).
370 rewrite > Hcut3 in ⊢ (? ? %).
371 assumption.rewrite > Hcut.
372 rewrite < H6.reflexivity.
374 apply (le_to_or_lt_eq r (S O)).
375 apply not_lt_to_le.assumption.
376 apply (decidable_lt (S O) r).
378 apply (p_ord_aux_to_exp n1 n1).
379 apply lt_O_nth_prime_n.assumption.
382 theorem defactorize_factorize: \forall n:nat.defactorize (factorize n) = n.
384 apply (nat_case n).reflexivity.
385 intro.apply (nat_case m).reflexivity.
388 (let p \def (max (S(S m1)) (\lambda p:nat.eqb ((S(S m1)) \mod (nth_prime p)) O)) in
389 defactorize (match p_ord (S(S m1)) (nth_prime p) with
390 [ (pair q r) \Rightarrow
391 nfa_proper (factorize_aux p r (nf_last (pred q)))])=(S(S m1))).
393 (* generalizing the goal; find a better way *)
394 cut (\forall q,r.(pair nat nat q r) = (p_ord (S(S m1)) (nth_prime p)) \to
395 defactorize (match p_ord (S(S m1)) (nth_prime p) with
396 [ (pair q r) \Rightarrow
397 nfa_proper (factorize_aux p r (nf_last (pred q)))])=(S(S m1))).
398 apply (Hcut (fst ? ? (p_ord (S(S m1)) (nth_prime p)))
399 (snd ? ? (p_ord (S(S m1)) (nth_prime p)))).
400 apply sym_eq.apply eq_pair_fst_snd.
404 cut ((S(S m1)) = (nth_prime p) \sup q *r).
406 rewrite > defactorize_aux_factorize_aux.
407 change with (r*(nth_prime p) \sup (S (pred q)) = (S(S m1))).
408 cut ((S (pred q)) = q).
412 apply (p_ord_aux_to_exp (S(S m1))).
413 apply lt_O_nth_prime_n.
416 apply sym_eq. apply S_pred.
417 cut (O < q \lor O = q).
418 elim Hcut2.assumption.
419 absurd (nth_prime p \divides S (S m1)).
420 apply (divides_max_prime_factor_n (S (S m1))).
421 unfold lt.apply le_S_S.apply le_S_S. apply le_O_n.
423 rewrite > Hcut3 in \vdash (? (? ? %)).
424 change with (nth_prime p \divides r \to False).
426 apply (p_ord_aux_to_not_mod_O (S(S m1)) (S(S m1)) (nth_prime p) q r).
427 apply lt_SO_nth_prime_n.
428 unfold lt.apply le_S_S.apply le_O_n.apply le_n.
430 apply divides_to_mod_O.apply lt_O_nth_prime_n.assumption.
431 rewrite > times_n_SO in \vdash (? ? ? %).
433 rewrite > (exp_n_O (nth_prime p)).
434 rewrite > H1 in \vdash (? ? ? (? (? ? %) ?)).
436 apply le_to_or_lt_eq.apply le_O_n.assumption.
437 (* e adesso l'ultimo goal. TASSI: che ora non e' piu' l'ultimo :P *)
438 cut ((S O) < r \lor S O \nlt r).
441 apply (p_ord_to_lt_max_prime_factor1 (S(S m1)) ? q r).
442 unfold lt.apply le_S_S. apply le_O_n.
444 assumption.assumption.
447 left.split.assumption.reflexivity.
448 intro.right.rewrite > Hcut3.
449 simplify.unfold lt.apply le_S_S.apply le_O_n.
450 cut (r \lt (S O) \or r=(S O)).
451 elim Hcut3.absurd (O=r).
452 apply le_n_O_to_eq.apply le_S_S_to_le.exact H2.
454 apply (not_le_Sn_O O).
455 rewrite > H3 in \vdash (? ? %).assumption.assumption.
456 apply (le_to_or_lt_eq r (S O)).
457 apply not_lt_to_le.assumption.
458 apply (decidable_lt (S O) r).
460 cut (O < r \lor O = r).
461 elim Hcut1.assumption.
463 apply (not_eq_O_S (S m1)).
464 rewrite > Hcut.rewrite < H1.rewrite < times_n_O.reflexivity.
465 apply le_to_or_lt_eq.apply le_O_n.
467 apply (p_ord_aux_to_exp (S(S m1))).
468 apply lt_O_nth_prime_n.
475 [ (nf_last n) \Rightarrow O
476 | (nf_cons n g) \Rightarrow S (max_p g)].
478 let rec max_p_exponent f \def
480 [ (nf_last n) \Rightarrow n
481 | (nf_cons n g) \Rightarrow max_p_exponent g].
483 theorem divides_max_p_defactorize: \forall f:nat_fact.\forall i:nat.
484 nth_prime ((max_p f)+i) \divides defactorize_aux f i.
486 elim f.simplify.apply (witness ? ? ((nth_prime i) \sup n)).
489 (nth_prime (S(max_p n1)+i) \divides
490 (nth_prime i) \sup n *(defactorize_aux n1 (S i))).
494 rewrite > assoc_times.
496 apply (witness ? ? (n2* (nth_prime i) \sup n)).
500 lemma eq_p_max: \forall n,p,r:nat. O < n \to
502 r = (S O) \lor (max r (\lambda p:nat.eqb (r \mod (nth_prime p)) O)) < p \to
503 p = max_prime_factor (r*(nth_prime p)\sup n).
506 unfold max_prime_factor.
507 apply max_spec_to_max.
510 [rewrite > times_n_SO in \vdash (? % ?).
515 apply (lt_to_le_to_lt ? (nth_prime p))
516 [apply lt_n_nth_prime_n
517 |rewrite > exp_n_SO in \vdash (? % ?).
519 [apply lt_O_nth_prime_n
524 |apply eq_to_eqb_true.
525 apply divides_to_mod_O
526 [apply lt_O_nth_prime_n
527 |apply (lt_O_n_elim ? H).
529 apply (witness ? ? (r*(nth_prime p \sup m))).
530 rewrite < assoc_times.
531 rewrite < sym_times in \vdash (? ? ? (? % ?)).
532 rewrite > exp_n_SO in \vdash (? ? ? (? (? ? %) ?)).
533 rewrite > assoc_times.
534 rewrite < exp_plus_times.
539 apply not_eq_to_eqb_false.
541 lapply (mod_O_to_divides ? ? ? H5)
542 [apply lt_O_nth_prime_n
543 |cut (Not (divides (nth_prime i) ((nth_prime p)\sup n)))
545 [rewrite > H6 in Hletin.
547 rewrite < plus_n_O in Hletin.
548 apply Hcut.assumption
549 |elim (divides_times_to_divides ? ? ? ? Hletin)
550 [apply (lt_to_not_le ? ? H3).
552 apply (le_to_lt_to_lt ? ? ? ? H6).
554 [apply (trans_le ? (nth_prime i))
556 apply lt_n_nth_prime_n
557 |apply divides_to_le[assumption|assumption]
559 |apply eq_to_eqb_true.
560 apply divides_to_mod_O
561 [apply lt_O_nth_prime_n|assumption]
563 |apply prime_nth_prime
564 |apply Hcut.assumption
568 apply (lt_to_not_eq ? ? H3).
570 elim (prime_nth_prime p).
571 apply injective_nth_prime.
573 [apply (divides_exp_to_divides ? ? ? ? H6).
574 apply prime_nth_prime.
575 |apply lt_SO_nth_prime_n
582 theorem not_divides_defactorize_aux: \forall f:nat_fact. \forall i,j:nat.
583 i < j \to nth_prime i \ndivides defactorize_aux f j.
586 (nth_prime i \divides (nth_prime j) \sup (S n) \to False).
587 intro.absurd ((nth_prime i) = (nth_prime j)).
588 apply (divides_exp_to_eq ? ? (S n)).
589 apply prime_nth_prime.apply prime_nth_prime.
590 assumption.unfold Not.
592 apply (not_le_Sn_n i).rewrite > Hcut in \vdash (? ? %).assumption.
593 apply (injective_nth_prime ? ? H2).
596 cut (nth_prime i \divides (nth_prime j) \sup n
597 \lor nth_prime i \divides defactorize_aux n1 (S j)).
599 absurd ((nth_prime i) = (nth_prime j)).
600 apply (divides_exp_to_eq ? ? n).
601 apply prime_nth_prime.apply prime_nth_prime.
602 assumption.unfold Not.
605 apply (not_le_Sn_n i).rewrite > Hcut1 in \vdash (? ? %).assumption.
606 apply (injective_nth_prime ? ? H4).
608 apply (trans_lt ? j).assumption.unfold lt.apply le_n.
610 apply divides_times_to_divides.
611 apply prime_nth_prime.assumption.
614 lemma not_eq_nf_last_nf_cons: \forall g:nat_fact.\forall n,m,i:nat.
615 \lnot (defactorize_aux (nf_last n) i= defactorize_aux (nf_cons m g) i).
618 (exp (nth_prime i) (S n) = defactorize_aux (nf_cons m g) i \to False).
620 cut (S(max_p g)+i= i).
621 apply (not_le_Sn_n i).
622 rewrite < Hcut in \vdash (? ? %).
623 simplify.apply le_S_S.
625 apply injective_nth_prime.
626 apply (divides_exp_to_eq ? ? (S n)).
627 apply prime_nth_prime.apply prime_nth_prime.
629 change with (divides (nth_prime ((max_p (nf_cons m g))+i))
630 (defactorize_aux (nf_cons m g) i)).
631 apply divides_max_p_defactorize.
634 lemma not_eq_nf_cons_O_nf_cons: \forall f,g:nat_fact.\forall n,i:nat.
635 \lnot (defactorize_aux (nf_cons O f) i= defactorize_aux (nf_cons (S n) g) i).
637 simplify.unfold Not.rewrite < plus_n_O.
639 apply (not_divides_defactorize_aux f i (S i) ?).
640 unfold lt.apply le_n.
642 rewrite > assoc_times.
643 apply (witness ? ? ((exp (nth_prime i) n)*(defactorize_aux g (S i)))).
647 theorem eq_defactorize_aux_to_eq: \forall f,g:nat_fact.\forall i:nat.
648 defactorize_aux f i = defactorize_aux g i \to f = g.
653 apply inj_S. apply (inj_exp_r (nth_prime i)).
654 apply lt_SO_nth_prime_n.
657 apply (not_eq_nf_last_nf_cons n2 n n1 i H1).
660 apply (not_eq_nf_last_nf_cons n1 n2 n i).
661 apply sym_eq. assumption.
663 generalize in match H2.
664 apply (nat_elim2 (\lambda n,n2.
665 ((nth_prime i) \sup n)*(defactorize_aux n1 (S i)) =
666 ((nth_prime i) \sup n2)*(defactorize_aux n3 (S i)) \to
667 nf_cons n n1 = nf_cons n2 n3)).
673 rewrite > (plus_n_O (defactorize_aux n3 (S i))).assumption.
675 apply (not_eq_nf_cons_O_nf_cons n1 n3 n5 i).assumption.
678 apply (not_eq_nf_cons_O_nf_cons n3 n1 n4 i).
679 apply sym_eq.assumption.
681 cut (nf_cons n4 n1 = nf_cons m n3).
684 rewrite > Hcut1.rewrite > Hcut2.reflexivity.
686 (match nf_cons n4 n1 with
687 [ (nf_last m) \Rightarrow n1
688 | (nf_cons m g) \Rightarrow g ] = n3).
689 rewrite > Hcut.simplify.reflexivity.
691 (match nf_cons n4 n1 with
692 [ (nf_last m) \Rightarrow m
693 | (nf_cons m g) \Rightarrow m ] = m).
694 rewrite > Hcut.simplify.reflexivity.
695 apply H3.simplify in H4.
696 apply (inj_times_r1 (nth_prime i)).
697 apply lt_O_nth_prime_n.
698 rewrite < assoc_times.rewrite < assoc_times.assumption.
701 theorem injective_defactorize_aux: \forall i:nat.
702 injective nat_fact nat (\lambda f.defactorize_aux f i).
705 apply (eq_defactorize_aux_to_eq x y i H).
708 theorem injective_defactorize:
709 injective nat_fact_all nat defactorize.
711 change with (\forall f,g.defactorize f = defactorize g \to f=g).
719 apply (not_eq_O_S O H).
723 apply (not_le_Sn_n O).
724 rewrite > H in \vdash (? ? %).
725 change with (O < defactorize_aux n O).
726 apply lt_O_defactorize_aux.
731 apply (not_eq_O_S O).apply sym_eq. assumption.
737 apply (not_le_Sn_n (S O)).
738 rewrite > H in \vdash (? ? %).
739 change with ((S O) < defactorize_aux n O).
740 apply lt_SO_defactorize_aux.
745 apply (not_le_Sn_n O).
746 rewrite < H in \vdash (? ? %).
747 change with (O < defactorize_aux n O).
748 apply lt_O_defactorize_aux.
752 apply (not_le_Sn_n (S O)).
753 rewrite < H in \vdash (? ? %).
754 change with ((S O) < defactorize_aux n O).
755 apply lt_SO_defactorize_aux.
756 (* proper - proper *)
758 apply (injective_defactorize_aux O).
762 theorem factorize_defactorize:
763 \forall f: nat_fact_all. factorize (defactorize f) = f.
765 apply injective_defactorize.
766 apply defactorize_factorize.