X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fmatita%2Flibrary%2Fnat%2Ffactorization.ma;h=6a309430302ee18b576e2def76303d3333a6fdfd;hb=4167cea65ca58897d1a3dbb81ff95de5074700cc;hp=1b946be522c143b6be9e1082dab079c9b82a86dd;hpb=30b7e65d641fe7243c4f36ed448f56360a1c5e1c;p=helm.git diff --git a/helm/matita/library/nat/factorization.ma b/helm/matita/library/nat/factorization.ma index 1b946be52..6a3094303 100644 --- a/helm/matita/library/nat/factorization.ma +++ b/helm/matita/library/nat/factorization.ma @@ -14,36 +14,36 @@ set "baseuri" "cic:/matita/nat/factorization". -include "nat/log.ma". +include "nat/ord.ma". include "nat/gcd.ma". include "nat/nth_prime.ma". (* the following factorization algorithm looks for the largest prime factor. *) definition max_prime_factor \def \lambda n:nat. -(max n (\lambda p:nat.eqb (mod n (nth_prime p)) O)). +(max n (\lambda p:nat.eqb (n \mod (nth_prime p)) O)). (* max_prime_factor is indeed a factor *) theorem divides_max_prime_factor_n: \forall n:nat. (S O) < n \to nth_prime (max_prime_factor n) \divides n. intros.apply divides_b_true_to_divides. apply lt_O_nth_prime_n. -apply f_max_true (\lambda p:nat.eqb (mod n (nth_prime p)) O) n. -cut \exists i. nth_prime i = smallest_factor n. +apply (f_max_true (\lambda p:nat.eqb (n \mod (nth_prime p)) O) n). +cut (\exists i. nth_prime i = smallest_factor n). elim Hcut. -apply ex_intro nat ? a. +apply (ex_intro nat ? a). split. -apply trans_le a (nth_prime a). +apply (trans_le a (nth_prime a)). apply le_n_fn. exact lt_nth_prime_n_nth_prime_Sn. rewrite > H1. apply le_smallest_factor_n. rewrite > H1. -change with divides_b (smallest_factor n) n = true. +change with (divides_b (smallest_factor n) n = true). apply divides_to_divides_b_true. -apply trans_lt ? (S O).simplify. apply le_n. +apply (trans_lt ? (S O)).unfold lt. apply le_n. apply lt_SO_smallest_factor.assumption. apply divides_smallest_factor_n. -apply trans_lt ? (S O). simplify. apply le_n. assumption. +apply (trans_lt ? (S O)). unfold lt. apply le_n. assumption. apply prime_to_nth_prime. apply prime_smallest_factor_n.assumption. qed. @@ -51,17 +51,17 @@ qed. theorem divides_to_max_prime_factor : \forall n,m. (S O) < n \to O < m \to n \divides m \to max_prime_factor n \le max_prime_factor m. intros.change with -(max n (\lambda p:nat.eqb (mod n (nth_prime p)) O)) \le -(max m (\lambda p:nat.eqb (mod m (nth_prime p)) O)). +((max n (\lambda p:nat.eqb (n \mod (nth_prime p)) O)) \le +(max m (\lambda p:nat.eqb (m \mod (nth_prime p)) O))). apply f_m_to_le_max. -apply trans_le ? n. +apply (trans_le ? n). apply le_max_n.apply divides_to_le.assumption.assumption. -change with divides_b (nth_prime (max_prime_factor n)) m = true. +change with (divides_b (nth_prime (max_prime_factor n)) m = true). apply divides_to_divides_b_true. -cut prime (nth_prime (max_prime_factor n)). +cut (prime (nth_prime (max_prime_factor n))). apply lt_O_nth_prime_n.apply prime_nth_prime. -cut nth_prime (max_prime_factor n) \divides n. -apply transitive_divides ? n. +cut (nth_prime (max_prime_factor n) \divides n). +apply (transitive_divides ? n). apply divides_max_prime_factor_n. assumption.assumption. apply divides_b_true_to_divides. @@ -72,53 +72,53 @@ apply divides_max_prime_factor_n. assumption. qed. -theorem plog_to_lt_max_prime_factor: \forall n,p,q,r. O < n \to +theorem p_ord_to_lt_max_prime_factor: \forall n,p,q,r. O < n \to p = max_prime_factor n \to -(pair nat nat q r) = plog n (nth_prime p) \to +(pair nat nat q r) = p_ord n (nth_prime p) \to (S O) < r \to max_prime_factor r < p. intros. rewrite > H1. -cut max_prime_factor r \lt max_prime_factor n \lor - max_prime_factor r = max_prime_factor n. +cut (max_prime_factor r \lt max_prime_factor n \lor + max_prime_factor r = max_prime_factor n). elim Hcut.assumption. -absurd nth_prime (max_prime_factor n) \divides r. +absurd (nth_prime (max_prime_factor n) \divides r). rewrite < H4. apply divides_max_prime_factor_n. assumption. -change with nth_prime (max_prime_factor n) \divides r \to False. +change with (nth_prime (max_prime_factor n) \divides r \to False). intro. -cut \lnot (mod r (nth_prime (max_prime_factor n))) = O. +cut (r \mod (nth_prime (max_prime_factor n)) \neq O). apply Hcut1.apply divides_to_mod_O. apply lt_O_nth_prime_n.assumption. -apply plog_aux_to_not_mod_O n n ? q r. +apply (p_ord_aux_to_not_mod_O n n ? q r). apply lt_SO_nth_prime_n.assumption. apply le_n. rewrite < H1.assumption. -apply le_to_or_lt_eq (max_prime_factor r) (max_prime_factor n). +apply (le_to_or_lt_eq (max_prime_factor r) (max_prime_factor n)). apply divides_to_max_prime_factor. assumption.assumption. -apply witness r n ((nth_prime p) \sup q). +apply (witness r n ((nth_prime p) \sup q)). rewrite < sym_times. -apply plog_aux_to_exp n n ? q r. +apply (p_ord_aux_to_exp n n ? q r). apply lt_O_nth_prime_n.assumption. qed. -theorem plog_to_lt_max_prime_factor1: \forall n,p,q,r. O < n \to +theorem p_ord_to_lt_max_prime_factor1: \forall n,p,q,r. O < n \to max_prime_factor n \le p \to -(pair nat nat q r) = plog n (nth_prime p) \to +(pair nat nat q r) = p_ord n (nth_prime p) \to (S O) < r \to max_prime_factor r < p. intros. -cut max_prime_factor n < p \lor max_prime_factor n = p. -elim Hcut.apply le_to_lt_to_lt ? (max_prime_factor n). +cut (max_prime_factor n < p \lor max_prime_factor n = p). +elim Hcut.apply (le_to_lt_to_lt ? (max_prime_factor n)). apply divides_to_max_prime_factor.assumption.assumption. -apply witness r n ((nth_prime p) \sup q). +apply (witness r n ((nth_prime p) \sup q)). rewrite > sym_times. -apply plog_aux_to_exp n n. +apply (p_ord_aux_to_exp n n). apply lt_O_nth_prime_n. assumption.assumption. -apply plog_to_lt_max_prime_factor n ? q. +apply (p_ord_to_lt_max_prime_factor n ? q). assumption.apply sym_eq.assumption.assumption.assumption. -apply le_to_or_lt_eq ? p H1. +apply (le_to_or_lt_eq ? p H1). qed. (* datatypes and functions *) @@ -136,7 +136,7 @@ let rec factorize_aux p n acc \def match p with [ O \Rightarrow acc | (S p1) \Rightarrow - match plog n (nth_prime p1) with + match p_ord n (nth_prime p1) with [ (pair q r) \Rightarrow factorize_aux p1 r (nf_cons q acc)]]. @@ -147,8 +147,8 @@ definition factorize : nat \to nat_fact_all \def \lambda n:nat. match n1 with [ O \Rightarrow nfa_one | (S n2) \Rightarrow - let p \def (max (S(S n2)) (\lambda p:nat.eqb (mod (S(S n2)) (nth_prime p)) O)) in - match plog (S(S n2)) (nth_prime p) with + let p \def (max (S(S n2)) (\lambda p:nat.eqb ((S(S n2)) \mod (nth_prime p)) O)) in + match p_ord (S(S n2)) (nth_prime p) with [ (pair q r) \Rightarrow nfa_proper (factorize_aux p r (nf_last (pred q)))]]]. @@ -165,6 +165,47 @@ match f with | nfa_one \Rightarrow (S O) | (nfa_proper g) \Rightarrow defactorize_aux g O]. +theorem lt_O_defactorize_aux: \forall f:nat_fact.\forall i:nat. +O < defactorize_aux f i. +intro.elim f.simplify.unfold lt. +rewrite > times_n_SO. +apply le_times. +change with (O < nth_prime i). +apply lt_O_nth_prime_n. +change with (O < exp (nth_prime i) n). +apply lt_O_exp. +apply lt_O_nth_prime_n. +simplify.unfold lt. +rewrite > times_n_SO. +apply le_times. +change with (O < exp (nth_prime i) n). +apply lt_O_exp. +apply lt_O_nth_prime_n. +change with (O < defactorize_aux n1 (S i)). +apply H. +qed. + +theorem lt_SO_defactorize_aux: \forall f:nat_fact.\forall i:nat. +S O < defactorize_aux f i. +intro.elim f.simplify.unfold lt. +rewrite > times_n_SO. +apply le_times. +change with (S O < nth_prime i). +apply lt_SO_nth_prime_n. +change with (O < exp (nth_prime i) n). +apply lt_O_exp. +apply lt_O_nth_prime_n. +simplify.unfold lt. +rewrite > times_n_SO. +rewrite > sym_times. +apply le_times. +change with (O < exp (nth_prime i) n). +apply lt_O_exp. +apply lt_O_nth_prime_n. +change with (S O < defactorize_aux n1 (S i)). +apply H. +qed. + theorem defactorize_aux_factorize_aux : \forall p,n:nat.\forall acc:nat_fact.O < n \to ((n=(S O) \land p=O) \lor max_prime_factor n < p) \to @@ -172,150 +213,150 @@ defactorize_aux (factorize_aux p n acc) O = n*(defactorize_aux acc p). intro.elim p.simplify. elim H1.elim H2.rewrite > H3. rewrite > sym_times. apply times_n_SO. -apply False_ind.apply not_le_Sn_O (max_prime_factor n) H2. +apply False_ind.apply (not_le_Sn_O (max_prime_factor n) H2). simplify. (* generalizing the goal: I guess there exists a better way *) -cut \forall q,r.(pair nat nat q r) = (plog_aux n1 n1 (nth_prime n)) \to -defactorize_aux match (plog_aux n1 n1 (nth_prime n)) with +cut (\forall q,r.(pair nat nat q r) = (p_ord_aux n1 n1 (nth_prime n)) \to +defactorize_aux match (p_ord_aux n1 n1 (nth_prime n)) with [(pair q r) \Rightarrow (factorize_aux n r (nf_cons q acc))] O = -n1*defactorize_aux acc (S n). -apply Hcut (fst ? ? (plog_aux n1 n1 (nth_prime n))) -(snd ? ? (plog_aux n1 n1 (nth_prime n))). +n1*defactorize_aux acc (S n)). +apply (Hcut (fst ? ? (p_ord_aux n1 n1 (nth_prime n))) +(snd ? ? (p_ord_aux n1 n1 (nth_prime n)))). apply sym_eq.apply eq_pair_fst_snd. intros. rewrite < H3. simplify. -cut n1 = r * (nth_prime n) \sup q. +cut (n1 = r * (nth_prime n) \sup q). rewrite > H. simplify.rewrite < assoc_times. rewrite < Hcut.reflexivity. -cut O < r \lor O = r. -elim Hcut1.assumption.absurd n1 = O. +cut (O < r \lor O = r). +elim Hcut1.assumption.absurd (n1 = O). rewrite > Hcut.rewrite < H4.reflexivity. -simplify. intro.apply not_le_Sn_O O. +unfold Not. intro.apply (not_le_Sn_O O). rewrite < H5 in \vdash (? ? %).assumption. apply le_to_or_lt_eq.apply le_O_n. -cut (S O) < r \lor \lnot (S O) < r. +cut ((S O) < r \lor (S O) \nlt r). elim Hcut1. right. -apply plog_to_lt_max_prime_factor1 n1 ? q r. +apply (p_ord_to_lt_max_prime_factor1 n1 ? q r). assumption.elim H2. elim H5. apply False_ind. -apply not_eq_O_S n.apply sym_eq.assumption. +apply (not_eq_O_S n).apply sym_eq.assumption. apply le_S_S_to_le. exact H5. assumption.assumption. -cut r=(S O). -apply nat_case n. +cut (r=(S O)). +apply (nat_case n). left.split.assumption.reflexivity. intro.right.rewrite > Hcut2. -simplify.apply le_S_S.apply le_O_n. -cut r \lt (S O) \or r=(S O). -elim Hcut2.absurd O=r. +simplify.unfold lt.apply le_S_S.apply le_O_n. +cut (r \lt (S O) \or r=(S O)). +elim Hcut2.absurd (O=r). apply le_n_O_to_eq.apply le_S_S_to_le.exact H5. -simplify.intro. -cut O=n1. -apply not_le_Sn_O O. +unfold Not.intro. +cut (O=n1). +apply (not_le_Sn_O O). rewrite > Hcut3 in \vdash (? ? %). assumption.rewrite > Hcut. rewrite < H6.reflexivity. assumption. -apply le_to_or_lt_eq r (S O). +apply (le_to_or_lt_eq r (S O)). apply not_lt_to_le.assumption. -apply decidable_lt (S O) r. +apply (decidable_lt (S O) r). rewrite > sym_times. -apply plog_aux_to_exp n1 n1. +apply (p_ord_aux_to_exp n1 n1). apply lt_O_nth_prime_n.assumption. qed. theorem defactorize_factorize: \forall n:nat.defactorize (factorize n) = n. intro. -apply nat_case n.reflexivity. -intro.apply nat_case m.reflexivity. +apply (nat_case n).reflexivity. +intro.apply (nat_case m).reflexivity. intro.change with -let p \def (max (S(S m1)) (\lambda p:nat.eqb (mod (S(S m1)) (nth_prime p)) O)) in -defactorize (match plog (S(S m1)) (nth_prime p) with +(let p \def (max (S(S m1)) (\lambda p:nat.eqb ((S(S m1)) \mod (nth_prime p)) O)) in +defactorize (match p_ord (S(S m1)) (nth_prime p) with [ (pair q r) \Rightarrow - nfa_proper (factorize_aux p r (nf_last (pred q)))])=(S(S m1)). + nfa_proper (factorize_aux p r (nf_last (pred q)))])=(S(S m1))). intro. (* generalizing the goal; find a better way *) -cut \forall q,r.(pair nat nat q r) = (plog (S(S m1)) (nth_prime p)) \to -defactorize (match plog (S(S m1)) (nth_prime p) with +cut (\forall q,r.(pair nat nat q r) = (p_ord (S(S m1)) (nth_prime p)) \to +defactorize (match p_ord (S(S m1)) (nth_prime p) with [ (pair q r) \Rightarrow - nfa_proper (factorize_aux p r (nf_last (pred q)))])=(S(S m1)). -apply Hcut (fst ? ? (plog (S(S m1)) (nth_prime p))) -(snd ? ? (plog (S(S m1)) (nth_prime p))). + nfa_proper (factorize_aux p r (nf_last (pred q)))])=(S(S m1))). +apply (Hcut (fst ? ? (p_ord (S(S m1)) (nth_prime p))) +(snd ? ? (p_ord (S(S m1)) (nth_prime p)))). apply sym_eq.apply eq_pair_fst_snd. intros. rewrite < H. change with -defactorize_aux (factorize_aux p r (nf_last (pred q))) O = (S(S m1)). -cut (S(S m1)) = (nth_prime p) \sup q *r. -cut O defactorize_aux_factorize_aux. -change with r*(nth_prime p) \sup (S (pred q)) = (S(S m1)). -cut (S (pred q)) = q. +change with (r*(nth_prime p) \sup (S (pred q)) = (S(S m1))). +cut ((S (pred q)) = q). rewrite > Hcut2. rewrite > sym_times. apply sym_eq. -apply plog_aux_to_exp (S(S m1)). +apply (p_ord_aux_to_exp (S(S m1))). apply lt_O_nth_prime_n. assumption. (* O < q *) apply sym_eq. apply S_pred. -cut O < q \lor O = q. +cut (O < q \lor O = q). elim Hcut2.assumption. -absurd nth_prime p \divides S (S m1). -apply divides_max_prime_factor_n (S (S m1)). -simplify.apply le_S_S.apply le_S_S. apply le_O_n. -cut (S(S m1)) = r. +absurd (nth_prime p \divides S (S m1)). +apply (divides_max_prime_factor_n (S (S m1))). +unfold lt.apply le_S_S.apply le_S_S. apply le_O_n. +cut ((S(S m1)) = r). rewrite > Hcut3 in \vdash (? (? ? %)). -change with nth_prime p \divides r \to False. +change with (nth_prime p \divides r \to False). intro. -apply plog_aux_to_not_mod_O (S(S m1)) (S(S m1)) (nth_prime p) q r. +apply (p_ord_aux_to_not_mod_O (S(S m1)) (S(S m1)) (nth_prime p) q r). apply lt_SO_nth_prime_n. -simplify.apply le_S_S.apply le_O_n.apply le_n. +unfold lt.apply le_S_S.apply le_O_n.apply le_n. assumption. apply divides_to_mod_O.apply lt_O_nth_prime_n.assumption. rewrite > times_n_SO in \vdash (? ? ? %). rewrite < sym_times. -rewrite > exp_n_O (nth_prime p). +rewrite > (exp_n_O (nth_prime p)). rewrite > H1 in \vdash (? ? ? (? (? ? %) ?)). assumption. apply le_to_or_lt_eq.apply le_O_n.assumption. (* e adesso l'ultimo goal. TASSI: che ora non e' piu' l'ultimo :P *) -cut (S O) < r \lor \lnot (S O) < r. +cut ((S O) < r \lor S O \nlt r). elim Hcut2. right. -apply plog_to_lt_max_prime_factor1 (S(S m1)) ? q r. -simplify.apply le_S_S. apply le_O_n. +apply (p_ord_to_lt_max_prime_factor1 (S(S m1)) ? q r). +unfold lt.apply le_S_S. apply le_O_n. apply le_n. assumption.assumption. -cut r=(S O). -apply nat_case p. +cut (r=(S O)). +apply (nat_case p). left.split.assumption.reflexivity. intro.right.rewrite > Hcut3. -simplify.apply le_S_S.apply le_O_n. -cut r \lt (S O) \or r=(S O). -elim Hcut3.absurd O=r. +simplify.unfold lt.apply le_S_S.apply le_O_n. +cut (r \lt (S O) \or r=(S O)). +elim Hcut3.absurd (O=r). apply le_n_O_to_eq.apply le_S_S_to_le.exact H2. -simplify.intro. -apply not_le_Sn_O O. +unfold Not.intro. +apply (not_le_Sn_O O). rewrite > H3 in \vdash (? ? %).assumption.assumption. -apply le_to_or_lt_eq r (S O). +apply (le_to_or_lt_eq r (S O)). apply not_lt_to_le.assumption. -apply decidable_lt (S O) r. +apply (decidable_lt (S O) r). (* O < r *) -cut O < r \lor O = r. +cut (O < r \lor O = r). elim Hcut1.assumption. apply False_ind. -apply not_eq_O_S (S m1). +apply (not_eq_O_S (S m1)). rewrite > Hcut.rewrite < H1.rewrite < times_n_O.reflexivity. apply le_to_or_lt_eq.apply le_O_n. (* prova del cut *) goal 20. -apply plog_aux_to_exp (S(S m1)). +apply (p_ord_aux_to_exp (S(S m1))). apply lt_O_nth_prime_n. assumption. (* fine prova cut *) @@ -334,17 +375,17 @@ match f with theorem divides_max_p_defactorize: \forall f:nat_fact.\forall i:nat. nth_prime ((max_p f)+i) \divides defactorize_aux f i. intro. -elim f.simplify.apply witness ? ? ((nth_prime i) \sup n). +elim f.simplify.apply (witness ? ? ((nth_prime i) \sup n)). reflexivity. change with -nth_prime (S(max_p n1)+i) \divides -(nth_prime i) \sup n *(defactorize_aux n1 (S i)). -elim H (S i). +(nth_prime (S(max_p n1)+i) \divides +(nth_prime i) \sup n *(defactorize_aux n1 (S i))). +elim (H (S i)). rewrite > H1. rewrite < sym_times. rewrite > assoc_times. rewrite < plus_n_Sm. -apply witness ? ? (n2* (nth_prime i) \sup n). +apply (witness ? ? (n2* (nth_prime i) \sup n)). reflexivity. qed. @@ -352,9 +393,9 @@ theorem divides_exp_to_divides: \forall p,n,m:nat. prime p \to p \divides n \sup m \to p \divides n. intros 3.elim m.simplify in H1. -apply transitive_divides p (S O).assumption. +apply (transitive_divides p (S O)).assumption. apply divides_SO_n. -cut p \divides n \lor p \divides n \sup n1. +cut (p \divides n \lor p \divides n \sup n1). elim Hcut.assumption. apply H.assumption.assumption. apply divides_times_to_divides.assumption. @@ -365,48 +406,83 @@ theorem divides_exp_to_eq: \forall p,q,m:nat. prime p \to prime q \to p \divides q \sup m \to p = q. intros. -simplify in H1. +unfold prime in H1. elim H1.apply H4. -apply divides_exp_to_divides p q m. +apply (divides_exp_to_divides p q m). assumption.assumption. -simplify in H.elim H.assumption. +unfold prime in H.elim H.assumption. qed. theorem not_divides_defactorize_aux: \forall f:nat_fact. \forall i,j:nat. i < j \to nth_prime i \ndivides defactorize_aux f j. intro.elim f. change with -nth_prime i \divides (nth_prime j) \sup (S n) \to False. -intro.absurd (nth_prime i) = (nth_prime j). -apply divides_exp_to_eq ? ? (S n). +(nth_prime i \divides (nth_prime j) \sup (S n) \to False). +intro.absurd ((nth_prime i) = (nth_prime j)). +apply (divides_exp_to_eq ? ? (S n)). apply prime_nth_prime.apply prime_nth_prime. assumption. -change with (nth_prime i) = (nth_prime j) \to False. -intro.cut i = j. -apply not_le_Sn_n i.rewrite > Hcut in \vdash (? ? %).assumption. -apply injective_nth_prime ? ? H2. +change with ((nth_prime i) = (nth_prime j) \to False). +intro.cut (i = j). +apply (not_le_Sn_n i).rewrite > Hcut in \vdash (? ? %).assumption. +apply (injective_nth_prime ? ? H2). change with -nth_prime i \divides (nth_prime j) \sup n *(defactorize_aux n1 (S j)) \to False. +(nth_prime i \divides (nth_prime j) \sup n *(defactorize_aux n1 (S j)) \to False). intro. -cut nth_prime i \divides (nth_prime j) \sup n -\lor nth_prime i \divides defactorize_aux n1 (S j). +cut (nth_prime i \divides (nth_prime j) \sup n +\lor nth_prime i \divides defactorize_aux n1 (S j)). elim Hcut. -absurd (nth_prime i) = (nth_prime j). -apply divides_exp_to_eq ? ? n. +absurd ((nth_prime i) = (nth_prime j)). +apply (divides_exp_to_eq ? ? n). apply prime_nth_prime.apply prime_nth_prime. assumption. -change with (nth_prime i) = (nth_prime j) \to False. +change with ((nth_prime i) = (nth_prime j) \to False). intro. -cut i = j. -apply not_le_Sn_n i.rewrite > Hcut1 in \vdash (? ? %).assumption. -apply injective_nth_prime ? ? H4. -apply H i (S j). -apply trans_lt ? j.assumption.simplify.apply le_n. +cut (i = j). +apply (not_le_Sn_n i).rewrite > Hcut1 in \vdash (? ? %).assumption. +apply (injective_nth_prime ? ? H4). +apply (H i (S j)). +apply (trans_lt ? j).assumption.unfold lt.apply le_n. assumption. apply divides_times_to_divides. apply prime_nth_prime.assumption. qed. +lemma not_eq_nf_last_nf_cons: \forall g:nat_fact.\forall n,m,i:nat. +\lnot (defactorize_aux (nf_last n) i= defactorize_aux (nf_cons m g) i). +intros. +change with +(exp (nth_prime i) (S n) = defactorize_aux (nf_cons m g) i \to False). +intro. +cut (S(max_p g)+i= i). +apply (not_le_Sn_n i). +rewrite < Hcut in \vdash (? ? %). +simplify.apply le_S_S. +apply le_plus_n. +apply injective_nth_prime. +(* uffa, perche' semplifica ? *) +change with (nth_prime (S(max_p g)+i)= nth_prime i). +apply (divides_exp_to_eq ? ? (S n)). +apply prime_nth_prime.apply prime_nth_prime. +rewrite > H. +change with (divides (nth_prime ((max_p (nf_cons m g))+i)) +(defactorize_aux (nf_cons m g) i)). +apply divides_max_p_defactorize. +qed. + +lemma not_eq_nf_cons_O_nf_cons: \forall f,g:nat_fact.\forall n,i:nat. +\lnot (defactorize_aux (nf_cons O f) i= defactorize_aux (nf_cons (S n) g) i). +intros. +simplify.unfold Not.rewrite < plus_n_O. +intro. +apply (not_divides_defactorize_aux f i (S i) ?). +unfold lt.apply le_n. +rewrite > H. +rewrite > assoc_times. +apply (witness ? ? ((exp (nth_prime i) n)*(defactorize_aux g (S i)))). +reflexivity. +qed. + theorem eq_defactorize_aux_to_eq: \forall f,g:nat_fact.\forall i:nat. defactorize_aux f i = defactorize_aux g i \to f = g. intro. @@ -414,116 +490,124 @@ elim f. generalize in match H. elim g. apply eq_f. -apply inj_S. apply inj_exp_r (nth_prime i). +apply inj_S. apply (inj_exp_r (nth_prime i)). apply lt_SO_nth_prime_n. assumption. -absurd defactorize_aux (nf_last n) i = -defactorize_aux (nf_cons n1 n2) i. -rewrite > H2.reflexivity. -absurd nth_prime (S(max_p n2)+i) \divides defactorize_aux (nf_cons n1 n2) i. -apply divides_max_p_defactorize. -rewrite < H2. -change with -(nth_prime (S(max_p n2)+i) \divides (nth_prime i) \sup (S n)) \to False. -intro. -absurd nth_prime (S (max_p n2) + i) = nth_prime i. -apply divides_exp_to_eq ? ? (S n). -apply prime_nth_prime.apply prime_nth_prime.assumption. -change with nth_prime (S (max_p n2) + i) = nth_prime i \to False. -intro.apply not_le_Sn_n i. -cut S(max_p n2)+i= i. -rewrite < Hcut in \vdash (? ? %). -simplify.apply le_S_S. -apply le_plus_n. -apply injective_nth_prime ? ? H4. +apply False_ind. +apply (not_eq_nf_last_nf_cons n2 n n1 i H2). generalize in match H1. elim g. -absurd defactorize_aux (nf_last n2) i = -defactorize_aux (nf_cons n n1) i. +apply False_ind. +apply (not_eq_nf_last_nf_cons n1 n2 n i). apply sym_eq. assumption. -absurd nth_prime (S(max_p n1)+i) \divides defactorize_aux (nf_cons n n1) i. -apply divides_max_p_defactorize. -rewrite > H2. -change with -(nth_prime (S(max_p n1)+i) \divides (nth_prime i) \sup (S n2)) \to False. -intro. -absurd nth_prime (S (max_p n1) + i) = nth_prime i. -apply divides_exp_to_eq ? ? (S n2). -apply prime_nth_prime.apply prime_nth_prime.assumption. -change with nth_prime (S (max_p n1) + i) = nth_prime i \to False. -intro.apply not_le_Sn_n i. -cut S(max_p n1)+i= i. -rewrite < Hcut in \vdash (? ? %). -simplify.apply le_S_S. -apply le_plus_n. -apply injective_nth_prime ? ? H4. simplify in H3. generalize in match H3. -apply nat_elim2 (\lambda n,n2. +apply (nat_elim2 (\lambda n,n2. ((nth_prime i) \sup n)*(defactorize_aux n1 (S i)) = ((nth_prime i) \sup n2)*(defactorize_aux n3 (S i)) \to -nf_cons n n1 = nf_cons n2 n3). +nf_cons n n1 = nf_cons n2 n3)). intro. elim n4. apply eq_f. -apply H n3 (S i). +apply (H n3 (S i)). simplify in H4. rewrite > plus_n_O. -rewrite > plus_n_O (defactorize_aux n3 (S i)).assumption. +rewrite > (plus_n_O (defactorize_aux n3 (S i))).assumption. apply False_ind. -apply not_divides_defactorize_aux n1 i (S i). -simplify. apply le_n. -simplify in H5. -rewrite > plus_n_O (defactorize_aux n1 (S i)). -rewrite > H5. -rewrite > assoc_times. -apply witness ? ? (((nth_prime i) \sup n5)*(defactorize_aux n3 (S i))). -reflexivity. +apply (not_eq_nf_cons_O_nf_cons n1 n3 n5 i).assumption. intros. apply False_ind. -apply not_divides_defactorize_aux n3 i (S i). -simplify. apply le_n. -simplify in H4. -rewrite > plus_n_O (defactorize_aux n3 (S i)). -rewrite < H4. -rewrite > assoc_times. -apply witness ? ? (((nth_prime i) \sup n4)*(defactorize_aux n1 (S i))). -reflexivity. +apply (not_eq_nf_cons_O_nf_cons n3 n1 n4 i). +apply sym_eq.assumption. intros. -cut nf_cons n4 n1 = nf_cons m n3. -cut n4=m. -cut n1=n3. +cut (nf_cons n4 n1 = nf_cons m n3). +cut (n4=m). +cut (n1=n3). rewrite > Hcut1.rewrite > Hcut2.reflexivity. change with -match nf_cons n4 n1 with +(match nf_cons n4 n1 with [ (nf_last m) \Rightarrow n1 -| (nf_cons m g) \Rightarrow g ] = n3. +| (nf_cons m g) \Rightarrow g ] = n3). rewrite > Hcut.simplify.reflexivity. change with -match nf_cons n4 n1 with +(match nf_cons n4 n1 with [ (nf_last m) \Rightarrow m -| (nf_cons m g) \Rightarrow m ] = m. +| (nf_cons m g) \Rightarrow m ] = m). rewrite > Hcut.simplify.reflexivity. apply H4.simplify in H5. -apply inj_times_r1 (nth_prime i). +apply (inj_times_r1 (nth_prime i)). apply lt_O_nth_prime_n. rewrite < assoc_times.rewrite < assoc_times.assumption. qed. theorem injective_defactorize_aux: \forall i:nat. injective nat_fact nat (\lambda f.defactorize_aux f i). -change with \forall i:nat.\forall f,g:nat_fact. -defactorize_aux f i = defactorize_aux g i \to f = g. +change with (\forall i:nat.\forall f,g:nat_fact. +defactorize_aux f i = defactorize_aux g i \to f = g). intros. -apply eq_defactorize_aux_to_eq f g i H. +apply (eq_defactorize_aux_to_eq f g i H). qed. -(* theorem injective_defactorize: injective nat_fact_all nat defactorize. -change with \forall f,g:nat_fact_all. -defactorize f = defactorize g \to f = g. +change with (\forall f,g:nat_fact_all. +defactorize f = defactorize g \to f = g). intro.elim f. generalize in match H.elim g. -reflexivity.simplify in H1. -*) +(* zero - zero *) +reflexivity. +(* zero - one *) +simplify in H1. +apply False_ind. +apply (not_eq_O_S O H1). +(* zero - proper *) +simplify in H1. +apply False_ind. +apply (not_le_Sn_n O). +rewrite > H1 in \vdash (? ? %). +change with (O < defactorize_aux n O). +apply lt_O_defactorize_aux. +generalize in match H. +elim g. +(* one - zero *) +simplify in H1. +apply False_ind. +apply (not_eq_O_S O).apply sym_eq. assumption. +(* one - one *) +reflexivity. +(* one - proper *) +simplify in H1. +apply False_ind. +apply (not_le_Sn_n (S O)). +rewrite > H1 in \vdash (? ? %). +change with ((S O) < defactorize_aux n O). +apply lt_SO_defactorize_aux. +generalize in match H.elim g. +(* proper - zero *) +simplify in H1. +apply False_ind. +apply (not_le_Sn_n O). +rewrite < H1 in \vdash (? ? %). +change with (O < defactorize_aux n O). +apply lt_O_defactorize_aux. +(* proper - one *) +simplify in H1. +apply False_ind. +apply (not_le_Sn_n (S O)). +rewrite < H1 in \vdash (? ? %). +change with ((S O) < defactorize_aux n O). +apply lt_SO_defactorize_aux. +(* proper - proper *) +apply eq_f. +apply (injective_defactorize_aux O). +exact H1. +qed. + +theorem factorize_defactorize: +\forall f,g: nat_fact_all. factorize (defactorize f) = f. +intros. +apply injective_defactorize. +(* uffa: perche' semplifica ??? *) +change with (defactorize(factorize (defactorize f)) = (defactorize f)). +apply defactorize_factorize. +qed.