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 new.
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 theorem divides_to_max_prime_factor : \forall n,m. (S O) < n \to O < m \to n \divides m \to
78 max_prime_factor n \le max_prime_factor m.
79 intros.unfold max_prime_factor.
82 apply le_max_n.apply divides_to_le.assumption.assumption.
83 change with (divides_b (nth_prime (max_prime_factor n)) m = true).
84 apply divides_to_divides_b_true.
85 cut (prime (nth_prime (max_prime_factor n))).
86 apply lt_O_nth_prime_n.apply prime_nth_prime.
87 cut (nth_prime (max_prime_factor n) \divides n).
91 [ apply (transitive_divides ? n);
92 [ apply divides_max_prime_factor_n.
96 | apply divides_b_true_to_divides;
97 [ apply lt_O_nth_prime_n.
98 | apply divides_to_divides_b_true;
99 [ apply lt_O_nth_prime_n.
100 | apply divides_max_prime_factor_n.
108 theorem divides_to_max_prime_factor1 : \forall n,m. O < n \to O < m \to n \divides m \to
109 max_prime_factor n \le max_prime_factor m.
111 elim (le_to_or_lt_eq ? ? H)
112 [apply divides_to_max_prime_factor
113 [assumption|assumption|assumption]
115 simplify.apply le_O_n.
119 theorem max_prime_factor_to_not_p_ord_O : \forall n,p,r.
121 p = max_prime_factor n \to
122 p_ord n (nth_prime p) \neq pair nat nat O r.
123 intros.unfold Not.intro.
124 apply (p_ord_O_to_not_divides ? ? ? ? H2)
125 [apply (trans_lt ? (S O))[apply lt_O_S|assumption]
127 apply divides_max_prime_factor_n.
132 theorem p_ord_to_lt_max_prime_factor: \forall n,p,q,r. O < n \to
133 p = max_prime_factor n \to
134 (pair nat nat q r) = p_ord n (nth_prime p) \to
135 (S O) < r \to max_prime_factor r < p.
138 cut (max_prime_factor r \lt max_prime_factor n \lor
139 max_prime_factor r = max_prime_factor n).
140 elim Hcut.assumption.
141 absurd (nth_prime (max_prime_factor n) \divides r).
143 apply divides_max_prime_factor_n.
144 assumption.unfold Not.
146 cut (r \mod (nth_prime (max_prime_factor n)) \neq O);
147 [unfold Not in Hcut1.autobatch new.
149 apply Hcut1.apply divides_to_mod_O;
150 [ apply lt_O_nth_prime_n.
155 cut(pair nat nat q r=p_ord_aux n n (nth_prime (max_prime_factor n)));
156 [2: rewrite < H1.assumption.|letin x \def le.autobatch width = 4 depth = 2]
157 (* CERCA COME MAI le_n non lo applica se lo trova come Const e non Rel *)
160 apply (p_ord_aux_to_not_mod_O n n ? q r);
161 [ apply lt_SO_nth_prime_n.
164 | rewrite < H1.assumption.
168 apply (le_to_or_lt_eq (max_prime_factor r) (max_prime_factor n)).
169 apply divides_to_max_prime_factor.
170 assumption.assumption.
171 apply (witness r n ((nth_prime p) \sup q)).
173 apply (p_ord_aux_to_exp n n ? q r).
174 apply lt_O_nth_prime_n.assumption.
177 theorem p_ord_to_lt_max_prime_factor1: \forall n,p,q,r. O < n \to
178 max_prime_factor n \le p \to
179 (pair nat nat q r) = p_ord n (nth_prime p) \to
180 (S O) < r \to max_prime_factor r < p.
182 cut (max_prime_factor n < p \lor max_prime_factor n = p).
183 elim Hcut.apply (le_to_lt_to_lt ? (max_prime_factor n)).
184 apply divides_to_max_prime_factor.assumption.assumption.
185 apply (witness r n ((nth_prime p) \sup q)).
187 apply (p_ord_aux_to_exp n n).
188 apply lt_O_nth_prime_n.
189 assumption.assumption.
190 apply (p_ord_to_lt_max_prime_factor n ? q).
191 assumption.apply sym_eq.assumption.assumption.assumption.
192 apply (le_to_or_lt_eq ? p H1).
195 lemma lt_max_prime_factor_to_not_divides: \forall n,p:nat.
196 O < n \to n=S O \lor max_prime_factor n < p \to
197 (nth_prime p \ndivides n).
198 intros.unfold Not.intro.
201 apply (le_to_not_lt (nth_prime p) (S O))
202 [apply divides_to_le[apply le_n|assumption]
203 |apply lt_SO_nth_prime_n
205 |apply (not_le_Sn_n p).
207 apply (le_to_lt_to_lt ? ? ? ? H3).
208 unfold max_prime_factor.
210 [apply (trans_le ? (nth_prime p))
212 apply lt_n_nth_prime_n
213 |apply divides_to_le;assumption
215 |apply eq_to_eqb_true.
216 apply divides_to_mod_O
217 [apply lt_O_nth_prime_n|assumption]
222 (* datatypes and functions *)
224 inductive nat_fact : Set \def
225 nf_last : nat \to nat_fact
226 | nf_cons : nat \to nat_fact \to nat_fact.
228 inductive nat_fact_all : Set \def
229 nfa_zero : nat_fact_all
230 | nfa_one : nat_fact_all
231 | nfa_proper : nat_fact \to nat_fact_all.
233 let rec factorize_aux p n acc \def
237 match p_ord n (nth_prime p1) with
238 [ (pair q r) \Rightarrow
239 factorize_aux p1 r (nf_cons q acc)]].
241 definition factorize : nat \to nat_fact_all \def \lambda n:nat.
243 [ O \Rightarrow nfa_zero
246 [ O \Rightarrow nfa_one
248 let p \def (max (S(S n2)) (\lambda p:nat.eqb ((S(S n2)) \mod (nth_prime p)) O)) in
249 match p_ord (S(S n2)) (nth_prime p) with
250 [ (pair q r) \Rightarrow
251 nfa_proper (factorize_aux p r (nf_last (pred q)))]]].
253 let rec defactorize_aux f i \def
255 [ (nf_last n) \Rightarrow (nth_prime i) \sup (S n)
256 | (nf_cons n g) \Rightarrow
257 (nth_prime i) \sup n *(defactorize_aux g (S i))].
259 definition defactorize : nat_fact_all \to nat \def
260 \lambda f : nat_fact_all.
262 [ nfa_zero \Rightarrow O
263 | nfa_one \Rightarrow (S O)
264 | (nfa_proper g) \Rightarrow defactorize_aux g O].
266 theorem lt_O_defactorize_aux:
269 O < defactorize_aux f i.
273 rewrite > times_n_SO;
275 [ change with (O < nth_prime i);
276 apply lt_O_nth_prime_n;
278 change with (O < exp (nth_prime i) n);
280 apply lt_O_nth_prime_n;
281 | change with (O < defactorize_aux n1 (S i));
285 theorem lt_SO_defactorize_aux: \forall f:nat_fact.\forall i:nat.
286 S O < defactorize_aux f i.
287 intro.elim f.simplify.unfold lt.
288 rewrite > times_n_SO.
290 change with (S O < nth_prime i).
291 apply lt_SO_nth_prime_n.
292 change with (O < exp (nth_prime i) n).
294 apply lt_O_nth_prime_n.
296 rewrite > times_n_SO.
299 change with (O < exp (nth_prime i) n).
301 apply lt_O_nth_prime_n.
302 change with (S O < defactorize_aux n1 (S i)).
306 theorem defactorize_aux_factorize_aux :
307 \forall p,n:nat.\forall acc:nat_fact.O < n \to
308 ((n=(S O) \land p=O) \lor max_prime_factor n < p) \to
309 defactorize_aux (factorize_aux p n acc) O = n*(defactorize_aux acc p).
310 intro.elim p.simplify.
311 elim H1.elim H2.rewrite > H3.
312 rewrite > sym_times. apply times_n_SO.
313 apply False_ind.apply (not_le_Sn_O (max_prime_factor n) H2).
315 (* generalizing the goal: I guess there exists a better way *)
316 cut (\forall q,r.(pair nat nat q r) = (p_ord_aux n1 n1 (nth_prime n)) \to
317 defactorize_aux match (p_ord_aux n1 n1 (nth_prime n)) with
318 [(pair q r) \Rightarrow (factorize_aux n r (nf_cons q acc))] O =
319 n1*defactorize_aux acc (S n)).
320 apply (Hcut (fst ? ? (p_ord_aux n1 n1 (nth_prime n)))
321 (snd ? ? (p_ord_aux n1 n1 (nth_prime n)))).
322 apply sym_eq.apply eq_pair_fst_snd.
326 cut (n1 = r * (nth_prime n) \sup q).
328 simplify.rewrite < assoc_times.
329 rewrite < Hcut.reflexivity.
330 cut (O < r \lor O = r).
331 elim Hcut1.assumption.absurd (n1 = O).
332 rewrite > Hcut.rewrite < H4.reflexivity.
333 unfold Not. intro.apply (not_le_Sn_O O).
334 rewrite < H5 in \vdash (? ? %).assumption.
335 apply le_to_or_lt_eq.apply le_O_n.
336 cut ((S O) < r \lor (S O) \nlt r).
339 apply (p_ord_to_lt_max_prime_factor1 n1 ? q r).
343 apply (not_eq_O_S n).apply sym_eq.assumption.
346 assumption.assumption.
349 left.split.assumption.reflexivity.
350 intro.right.rewrite > Hcut2.
351 simplify.unfold lt.apply le_S_S.apply le_O_n.
352 cut (r < (S O) ∨ r=(S O)).
353 elim Hcut2.absurd (O=r).
354 apply le_n_O_to_eq.apply le_S_S_to_le.exact H5.
357 apply (not_le_Sn_O O).
358 rewrite > Hcut3 in ⊢ (? ? %).
359 assumption.rewrite > Hcut.
360 rewrite < H6.reflexivity.
362 apply (le_to_or_lt_eq r (S O)).
363 apply not_lt_to_le.assumption.
364 apply (decidable_lt (S O) r).
366 apply (p_ord_aux_to_exp n1 n1).
367 apply lt_O_nth_prime_n.assumption.
370 theorem defactorize_factorize: \forall n:nat.defactorize (factorize n) = n.
372 apply (nat_case n).reflexivity.
373 intro.apply (nat_case m).reflexivity.
376 (let p \def (max (S(S m1)) (\lambda p:nat.eqb ((S(S m1)) \mod (nth_prime p)) O)) in
377 defactorize (match p_ord (S(S m1)) (nth_prime p) with
378 [ (pair q r) \Rightarrow
379 nfa_proper (factorize_aux p r (nf_last (pred q)))])=(S(S m1))).
381 (* generalizing the goal; find a better way *)
382 cut (\forall q,r.(pair nat nat q r) = (p_ord (S(S m1)) (nth_prime p)) \to
383 defactorize (match p_ord (S(S m1)) (nth_prime p) with
384 [ (pair q r) \Rightarrow
385 nfa_proper (factorize_aux p r (nf_last (pred q)))])=(S(S m1))).
386 apply (Hcut (fst ? ? (p_ord (S(S m1)) (nth_prime p)))
387 (snd ? ? (p_ord (S(S m1)) (nth_prime p)))).
388 apply sym_eq.apply eq_pair_fst_snd.
392 cut ((S(S m1)) = (nth_prime p) \sup q *r).
394 rewrite > defactorize_aux_factorize_aux.
395 change with (r*(nth_prime p) \sup (S (pred q)) = (S(S m1))).
396 cut ((S (pred q)) = q).
400 apply (p_ord_aux_to_exp (S(S m1))).
401 apply lt_O_nth_prime_n.
404 apply sym_eq. apply S_pred.
405 cut (O < q \lor O = q).
406 elim Hcut2.assumption.
407 absurd (nth_prime p \divides S (S m1)).
408 apply (divides_max_prime_factor_n (S (S m1))).
409 unfold lt.apply le_S_S.apply le_S_S. apply le_O_n.
411 rewrite > Hcut3 in \vdash (? (? ? %)).
412 change with (nth_prime p \divides r \to False).
414 apply (p_ord_aux_to_not_mod_O (S(S m1)) (S(S m1)) (nth_prime p) q r).
415 apply lt_SO_nth_prime_n.
416 unfold lt.apply le_S_S.apply le_O_n.apply le_n.
418 apply divides_to_mod_O.apply lt_O_nth_prime_n.assumption.
419 rewrite > times_n_SO in \vdash (? ? ? %).
421 rewrite > (exp_n_O (nth_prime p)).
422 rewrite > H1 in \vdash (? ? ? (? (? ? %) ?)).
424 apply le_to_or_lt_eq.apply le_O_n.assumption.
425 (* e adesso l'ultimo goal. TASSI: che ora non e' piu' l'ultimo :P *)
426 cut ((S O) < r \lor S O \nlt r).
429 apply (p_ord_to_lt_max_prime_factor1 (S(S m1)) ? q r).
430 unfold lt.apply le_S_S. apply le_O_n.
432 assumption.assumption.
435 left.split.assumption.reflexivity.
436 intro.right.rewrite > Hcut3.
437 simplify.unfold lt.apply le_S_S.apply le_O_n.
438 cut (r \lt (S O) \or r=(S O)).
439 elim Hcut3.absurd (O=r).
440 apply le_n_O_to_eq.apply le_S_S_to_le.exact H2.
442 apply (not_le_Sn_O O).
443 rewrite > H3 in \vdash (? ? %).assumption.assumption.
444 apply (le_to_or_lt_eq r (S O)).
445 apply not_lt_to_le.assumption.
446 apply (decidable_lt (S O) r).
448 cut (O < r \lor O = r).
449 elim Hcut1.assumption.
451 apply (not_eq_O_S (S m1)).
452 rewrite > Hcut.rewrite < H1.rewrite < times_n_O.reflexivity.
453 apply le_to_or_lt_eq.apply le_O_n.
455 apply (p_ord_aux_to_exp (S(S m1))).
456 apply lt_O_nth_prime_n.
463 [ (nf_last n) \Rightarrow O
464 | (nf_cons n g) \Rightarrow S (max_p g)].
466 let rec max_p_exponent f \def
468 [ (nf_last n) \Rightarrow n
469 | (nf_cons n g) \Rightarrow max_p_exponent g].
471 theorem divides_max_p_defactorize: \forall f:nat_fact.\forall i:nat.
472 nth_prime ((max_p f)+i) \divides defactorize_aux f i.
474 elim f.simplify.apply (witness ? ? ((nth_prime i) \sup n)).
477 (nth_prime (S(max_p n1)+i) \divides
478 (nth_prime i) \sup n *(defactorize_aux n1 (S i))).
482 rewrite > assoc_times.
484 apply (witness ? ? (n2* (nth_prime i) \sup n)).
488 lemma eq_p_max: \forall n,p,r:nat. O < n \to
490 r = (S O) \lor (max r (\lambda p:nat.eqb (r \mod (nth_prime p)) O)) < p \to
491 p = max_prime_factor (r*(nth_prime p)\sup n).
494 unfold max_prime_factor.
495 apply max_spec_to_max.
498 [rewrite > times_n_SO in \vdash (? % ?).
503 apply (lt_to_le_to_lt ? (nth_prime p))
504 [apply lt_n_nth_prime_n
505 |rewrite > exp_n_SO in \vdash (? % ?).
507 [apply lt_O_nth_prime_n
512 |apply eq_to_eqb_true.
513 apply divides_to_mod_O
514 [apply lt_O_nth_prime_n
515 |apply (lt_O_n_elim ? H).
517 apply (witness ? ? (r*(nth_prime p \sup m))).
518 rewrite < assoc_times.
519 rewrite < sym_times in \vdash (? ? ? (? % ?)).
520 rewrite > exp_n_SO in \vdash (? ? ? (? (? ? %) ?)).
521 rewrite > assoc_times.
522 rewrite < exp_plus_times.
527 apply not_eq_to_eqb_false.
529 lapply (mod_O_to_divides ? ? ? H5)
530 [apply lt_O_nth_prime_n
531 |cut (Not (divides (nth_prime i) ((nth_prime p)\sup n)))
533 [rewrite > H6 in Hletin.
535 rewrite < plus_n_O in Hletin.
536 apply Hcut.assumption
537 |elim (divides_times_to_divides ? ? ? ? Hletin)
538 [apply (lt_to_not_le ? ? H3).
540 apply (le_to_lt_to_lt ? ? ? ? H6).
542 [apply (trans_le ? (nth_prime i))
544 apply lt_n_nth_prime_n
545 |apply divides_to_le[assumption|assumption]
547 |apply eq_to_eqb_true.
548 apply divides_to_mod_O
549 [apply lt_O_nth_prime_n|assumption]
551 |apply prime_nth_prime
552 |apply Hcut.assumption
556 apply (lt_to_not_eq ? ? H3).
558 elim (prime_nth_prime p).
559 apply injective_nth_prime.
561 [apply (divides_exp_to_divides ? ? ? ? H6).
562 apply prime_nth_prime.
563 |apply lt_SO_nth_prime_n
570 theorem not_divides_defactorize_aux: \forall f:nat_fact. \forall i,j:nat.
571 i < j \to nth_prime i \ndivides defactorize_aux f j.
574 (nth_prime i \divides (nth_prime j) \sup (S n) \to False).
575 intro.absurd ((nth_prime i) = (nth_prime j)).
576 apply (divides_exp_to_eq ? ? (S n)).
577 apply prime_nth_prime.apply prime_nth_prime.
578 assumption.unfold Not.
580 apply (not_le_Sn_n i).rewrite > Hcut in \vdash (? ? %).assumption.
581 apply (injective_nth_prime ? ? H2).
584 cut (nth_prime i \divides (nth_prime j) \sup n
585 \lor nth_prime i \divides defactorize_aux n1 (S j)).
587 absurd ((nth_prime i) = (nth_prime j)).
588 apply (divides_exp_to_eq ? ? n).
589 apply prime_nth_prime.apply prime_nth_prime.
590 assumption.unfold Not.
593 apply (not_le_Sn_n i).rewrite > Hcut1 in \vdash (? ? %).assumption.
594 apply (injective_nth_prime ? ? H4).
596 apply (trans_lt ? j).assumption.unfold lt.apply le_n.
598 apply divides_times_to_divides.
599 apply prime_nth_prime.assumption.
602 lemma not_eq_nf_last_nf_cons: \forall g:nat_fact.\forall n,m,i:nat.
603 \lnot (defactorize_aux (nf_last n) i= defactorize_aux (nf_cons m g) i).
606 (exp (nth_prime i) (S n) = defactorize_aux (nf_cons m g) i \to False).
608 cut (S(max_p g)+i= i).
609 apply (not_le_Sn_n i).
610 rewrite < Hcut in \vdash (? ? %).
611 simplify.apply le_S_S.
613 apply injective_nth_prime.
614 apply (divides_exp_to_eq ? ? (S n)).
615 apply prime_nth_prime.apply prime_nth_prime.
617 change with (divides (nth_prime ((max_p (nf_cons m g))+i))
618 (defactorize_aux (nf_cons m g) i)).
619 apply divides_max_p_defactorize.
622 lemma not_eq_nf_cons_O_nf_cons: \forall f,g:nat_fact.\forall n,i:nat.
623 \lnot (defactorize_aux (nf_cons O f) i= defactorize_aux (nf_cons (S n) g) i).
625 simplify.unfold Not.rewrite < plus_n_O.
627 apply (not_divides_defactorize_aux f i (S i) ?).
628 unfold lt.apply le_n.
630 rewrite > assoc_times.
631 apply (witness ? ? ((exp (nth_prime i) n)*(defactorize_aux g (S i)))).
635 theorem eq_defactorize_aux_to_eq: \forall f,g:nat_fact.\forall i:nat.
636 defactorize_aux f i = defactorize_aux g i \to f = g.
639 generalize in match H.
642 apply inj_S. apply (inj_exp_r (nth_prime i)).
643 apply lt_SO_nth_prime_n.
646 apply (not_eq_nf_last_nf_cons n2 n n1 i H2).
647 generalize in match H1.
650 apply (not_eq_nf_last_nf_cons n1 n2 n i).
651 apply sym_eq. assumption.
653 generalize in match H3.
654 apply (nat_elim2 (\lambda n,n2.
655 ((nth_prime i) \sup n)*(defactorize_aux n1 (S i)) =
656 ((nth_prime i) \sup n2)*(defactorize_aux n3 (S i)) \to
657 nf_cons n n1 = nf_cons n2 n3)).
663 rewrite > (plus_n_O (defactorize_aux n3 (S i))).assumption.
665 apply (not_eq_nf_cons_O_nf_cons n1 n3 n5 i).assumption.
668 apply (not_eq_nf_cons_O_nf_cons n3 n1 n4 i).
669 apply sym_eq.assumption.
671 cut (nf_cons n4 n1 = nf_cons m n3).
674 rewrite > Hcut1.rewrite > Hcut2.reflexivity.
676 (match nf_cons n4 n1 with
677 [ (nf_last m) \Rightarrow n1
678 | (nf_cons m g) \Rightarrow g ] = n3).
679 rewrite > Hcut.simplify.reflexivity.
681 (match nf_cons n4 n1 with
682 [ (nf_last m) \Rightarrow m
683 | (nf_cons m g) \Rightarrow m ] = m).
684 rewrite > Hcut.simplify.reflexivity.
685 apply H4.simplify in H5.
686 apply (inj_times_r1 (nth_prime i)).
687 apply lt_O_nth_prime_n.
688 rewrite < assoc_times.rewrite < assoc_times.assumption.
691 theorem injective_defactorize_aux: \forall i:nat.
692 injective nat_fact nat (\lambda f.defactorize_aux f i).
695 apply (eq_defactorize_aux_to_eq x y i H).
698 theorem injective_defactorize:
699 injective nat_fact_all nat defactorize.
701 change with (\forall f,g.defactorize f = defactorize g \to f=g).
703 generalize in match H.elim g.
709 apply (not_eq_O_S O H1).
713 apply (not_le_Sn_n O).
714 rewrite > H1 in \vdash (? ? %).
715 change with (O < defactorize_aux n O).
716 apply lt_O_defactorize_aux.
717 generalize in match H.
722 apply (not_eq_O_S O).apply sym_eq. assumption.
728 apply (not_le_Sn_n (S O)).
729 rewrite > H1 in \vdash (? ? %).
730 change with ((S O) < defactorize_aux n O).
731 apply lt_SO_defactorize_aux.
732 generalize in match H.elim g.
736 apply (not_le_Sn_n O).
737 rewrite < H1 in \vdash (? ? %).
738 change with (O < defactorize_aux n O).
739 apply lt_O_defactorize_aux.
743 apply (not_le_Sn_n (S O)).
744 rewrite < H1 in \vdash (? ? %).
745 change with ((S O) < defactorize_aux n O).
746 apply lt_SO_defactorize_aux.
747 (* proper - proper *)
749 apply (injective_defactorize_aux O).
753 theorem factorize_defactorize:
754 \forall f: nat_fact_all. factorize (defactorize f) = f.
756 apply injective_defactorize.
757 apply defactorize_factorize.