X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Fnat%2Ffactorization.ma;h=bfdf947b4ba1ec437d1520b4095bd01a1320e54c;hb=7288b45eacf9f7dcd118b3b89b81ff19ae9d6ce5;hp=14696ca2891d7322622f2b7aabd2a1031fc36572;hpb=d8b17e4c77989c669c9db3847bd6b9e7c236e2c3;p=helm.git diff --git a/helm/software/matita/library/nat/factorization.ma b/helm/software/matita/library/nat/factorization.ma index 14696ca28..bfdf947b4 100644 --- a/helm/software/matita/library/nat/factorization.ma +++ b/helm/software/matita/library/nat/factorization.ma @@ -12,17 +12,33 @@ (* *) (**************************************************************************) -set "baseuri" "cic:/matita/nat/factorization". - 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 (n \mod (nth_prime p)) O)). +theorem lt_SO_max_prime: \forall m. S O < m \to +S O < max m (λi:nat.primeb i∧divides_b i m). +intros. +apply (lt_to_le_to_lt ? (smallest_factor m)) + [apply lt_SO_smallest_factor.assumption + |apply f_m_to_le_max + [apply le_smallest_factor_n + |apply true_to_true_to_andb_true + [apply prime_to_primeb_true. + apply prime_smallest_factor_n. + assumption + |apply divides_to_divides_b_true + [apply lt_O_smallest_factor.apply lt_to_le.assumption + |apply divides_smallest_factor_n. + apply lt_to_le.assumption + ] + ] + ] + ] +qed. (* max_prime_factor is indeed a factor *) theorem divides_max_prime_factor_n: \forall n:nat. (S O) < n @@ -40,25 +56,33 @@ cut (\exists i. nth_prime i = smallest_factor n); | rewrite > H1; apply le_smallest_factor_n; ] | rewrite > H1; - (*CSC: simplify here does something nasty! *) change with (divides_b (smallest_factor n) n = true); apply divides_to_divides_b_true; [ apply (trans_lt ? (S O)); [ unfold lt; apply le_n; | apply lt_SO_smallest_factor; assumption; ] - | letin x \def le.autobatch new. - (* - apply divides_smallest_factor_n; - apply (trans_lt ? (S O)); - [ unfold lt; apply le_n; - | assumption; ] *) ] ] - | autobatch. + | apply divides_smallest_factor_n; + autobatch; ] ] + | apply prime_to_nth_prime; + autobatch. (* apply prime_to_nth_prime; apply prime_smallest_factor_n; assumption; *) ] qed. +lemma divides_to_prime_divides : \forall n,m.1 < m \to m < n \to m \divides n \to + \exists p.p \leq m \land prime p \land p \divides n. +intros;apply (ex_intro ? ? (nth_prime (max_prime_factor m)));split + [split + [apply divides_to_le + [apply lt_to_le;assumption + |apply divides_max_prime_factor_n;assumption] + |apply prime_nth_prime;] + |apply (transitive_divides ? ? ? ? H2);apply divides_max_prime_factor_n; + assumption] +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.unfold max_prime_factor. @@ -129,27 +153,22 @@ apply divides_max_prime_factor_n. assumption.unfold Not. intro. cut (r \mod (nth_prime (max_prime_factor n)) \neq O); - [unfold Not in Hcut1.autobatch new. + [ apply Hcut1; autobatch depth=2; (* apply Hcut1.apply divides_to_mod_O; [ apply lt_O_nth_prime_n. | assumption. ] *) - |letin z \def le. - cut(pair nat nat q r=p_ord_aux n n (nth_prime (max_prime_factor n))); - [2: rewrite < H1.assumption.|letin x \def le.autobatch width = 4 depth = 2] - (* CERCA COME MAI le_n non lo applica se lo trova come Const e non Rel *) - ]. -(* - apply (p_ord_aux_to_not_mod_O n n ? q r); + | unfold p_ord in H2; lapply depth=0 le_n; autobatch width = 4 depth = 2; + (*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 divides_to_max_prime_factor. assumption.assumption. @@ -356,7 +375,7 @@ theorem defactorize_factorize: \forall n:nat.defactorize (factorize n) = n. intro. apply (nat_case n).reflexivity. intro.apply (nat_case m).reflexivity. -intro.(*CSC: simplify here does something really nasty *) +intro. change 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 @@ -377,7 +396,6 @@ simplify. cut ((S(S m1)) = (nth_prime p) \sup q *r). cut (O defactorize_aux_factorize_aux. -(*CSC: simplify here does something really nasty *) change with (r*(nth_prime p) \sup (S (pred q)) = (S(S m1))). cut ((S (pred q)) = q). rewrite > Hcut2. @@ -395,7 +413,6 @@ 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 (? (? ? %)). -(*CSC: simplify here does something really nasty *) change with (nth_prime p \divides r \to False). intro. apply (p_ord_aux_to_not_mod_O (S(S m1)) (S(S m1)) (nth_prime p) q r). @@ -472,30 +489,6 @@ apply (witness ? ? (n2* (nth_prime i) \sup n)). reflexivity. qed. -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 divides_SO_n. -cut (p \divides n \lor p \divides n \sup n1). -elim Hcut.assumption. -apply H.assumption.assumption. -apply divides_times_to_divides.assumption. -exact H2. -qed. - -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. -unfold prime in H1. -elim H1.apply H4. -apply (divides_exp_to_divides p q m). -assumption.assumption. -unfold prime in H.elim H.assumption. -qed. - lemma eq_p_max: \forall n,p,r:nat. O < n \to O < r \to r = (S O) \lor (max r (\lambda p:nat.eqb (r \mod (nth_prime p)) O)) < p \to @@ -525,7 +518,7 @@ split [apply lt_O_nth_prime_n |apply (lt_O_n_elim ? H). intro. - apply (witness ? ? (r*(nth_prime p \sup m))). + apply (witness ? ? ((r*(nth_prime p) \sup m))). rewrite < assoc_times. rewrite < sym_times in \vdash (? ? ? (? % ?)). rewrite > exp_n_SO in \vdash (? ? ? (? (? ? %) ?)). @@ -647,21 +640,19 @@ 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. elim f. -generalize in match H. -elim g. +elim g in H ⊢ %. apply eq_f. apply inj_S. apply (inj_exp_r (nth_prime i)). apply lt_SO_nth_prime_n. assumption. apply False_ind. -apply (not_eq_nf_last_nf_cons n2 n n1 i H2). -generalize in match H1. -elim g. +apply (not_eq_nf_last_nf_cons n2 n n1 i H1). +elim g in H1 ⊢ %. apply False_ind. apply (not_eq_nf_last_nf_cons n1 n2 n i). apply sym_eq. assumption. -simplify in H3. -generalize in match H3. +simplify in H2. +generalize in match H2. 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 @@ -693,7 +684,7 @@ change with [ (nf_last m) \Rightarrow m | (nf_cons m g) \Rightarrow m ] = m). rewrite > Hcut.simplify.reflexivity. -apply H4.simplify in H5. +apply H3.simplify in H4. apply (inj_times_r1 (nth_prime i)). apply lt_O_nth_prime_n. rewrite < assoc_times.rewrite < assoc_times.assumption. @@ -711,22 +702,21 @@ injective nat_fact_all nat defactorize. unfold injective. change with (\forall f,g.defactorize f = defactorize g \to f=g). intro.elim f. -generalize in match H.elim g. +elim g in H ⊢ %. (* zero - zero *) reflexivity. (* zero - one *) simplify in H1. apply False_ind. -apply (not_eq_O_S O H1). +apply (not_eq_O_S O H). (* zero - proper *) simplify in H1. apply False_ind. apply (not_le_Sn_n O). -rewrite > H1 in \vdash (? ? %). +rewrite > H in \vdash (? ? %). change with (O < defactorize_aux n O). apply lt_O_defactorize_aux. -generalize in match H. -elim g. +elim g in H ⊢ %. (* one - zero *) simplify in H1. apply False_ind. @@ -737,32 +727,32 @@ reflexivity. simplify in H1. apply False_ind. apply (not_le_Sn_n (S O)). -rewrite > H1 in \vdash (? ? %). +rewrite > H in \vdash (? ? %). change with ((S O) < defactorize_aux n O). apply lt_SO_defactorize_aux. -generalize in match H.elim g. +elim g in H ⊢ %. (* proper - zero *) simplify in H1. apply False_ind. apply (not_le_Sn_n O). -rewrite < H1 in \vdash (? ? %). +rewrite < H 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 (? ? %). +rewrite < H 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. +exact H. qed. theorem factorize_defactorize: -\forall f,g: nat_fact_all. factorize (defactorize f) = f. +\forall f: nat_fact_all. factorize (defactorize f) = f. intros. apply injective_defactorize. apply defactorize_factorize.