X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Fnat%2Ffactorization.ma;h=8c50d1d7db1e5612e160776b39a59e55ab21fece;hb=10f29fdd78ee089a9a94446207b543d33d6c851c;hp=37b5ea1ddb8e5d82d19ff44acc362d07f0fbb8ce;hpb=55b82bd235d82ff7f0a40d980effe1efde1f5073;p=helm.git diff --git a/helm/software/matita/library/nat/factorization.ma b/helm/software/matita/library/nat/factorization.ma index 37b5ea1dd..8c50d1d7d 100644 --- a/helm/software/matita/library/nat/factorization.ma +++ b/helm/software/matita/library/nat/factorization.ma @@ -15,22 +15,40 @@ 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 \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 (n \mod (nth_prime p)) O) n); - cut (\exists i. nth_prime i = smallest_factor n); +intros. +apply divides_b_true_to_divides. +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); split; @@ -45,20 +63,22 @@ intros; apply divides_b_true_to_divides; [ apply (trans_lt ? (S O)); [ unfold lt; apply le_n; | apply lt_SO_smallest_factor; assumption; ] - | apply divides_smallest_factor_n; + | letin x \def le.autobatch new. + (* + apply divides_smallest_factor_n; apply (trans_lt ? (S O)); [ unfold lt; apply le_n; - | assumption; ] ] ] - | apply prime_to_nth_prime; + | assumption; ] *) ] ] + | autobatch. + (* + apply prime_to_nth_prime; apply prime_smallest_factor_n; - assumption; ] ] + 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.change with -((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))). +intros.unfold max_prime_factor. apply f_m_to_le_max. apply (trans_le ? n). apply le_max_n.apply divides_to_le.assumption.assumption. @@ -67,15 +87,48 @@ apply divides_to_divides_b_true. 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). -apply divides_max_prime_factor_n. -assumption.assumption. -apply divides_b_true_to_divides. -apply lt_O_nth_prime_n. -apply divides_to_divides_b_true. -apply lt_O_nth_prime_n. -apply divides_max_prime_factor_n. -assumption. +autobatch. +autobatch. +(* + [ apply (transitive_divides ? n); + [ apply divides_max_prime_factor_n. + assumption. + | assumption. + ] + | apply divides_b_true_to_divides; + [ apply lt_O_nth_prime_n. + | apply divides_to_divides_b_true; + [ apply lt_O_nth_prime_n. + | apply divides_max_prime_factor_n. + assumption. + ] + ] + ] +*) +qed. + +theorem divides_to_max_prime_factor1 : \forall n,m. O < n \to O < m \to n \divides m \to +max_prime_factor n \le max_prime_factor m. +intros 3. +elim (le_to_or_lt_eq ? ? H) + [apply divides_to_max_prime_factor + [assumption|assumption|assumption] + |rewrite < H1. + simplify.apply le_O_n. + ] +qed. + +theorem max_prime_factor_to_not_p_ord_O : \forall n,p,r. + (S O) < n \to + p = max_prime_factor n \to + p_ord n (nth_prime p) \neq pair nat nat O r. +intros.unfold Not.intro. +apply (p_ord_O_to_not_divides ? ? ? ? H2) + [apply (trans_lt ? (S O))[apply lt_O_S|assumption] + |rewrite > H1. + apply divides_max_prime_factor_n. + assumption + ] qed. theorem p_ord_to_lt_max_prime_factor: \forall n,p,q,r. O < n \to @@ -90,16 +143,30 @@ elim Hcut.assumption. 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). +assumption.unfold Not. intro. -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 (p_ord_aux_to_not_mod_O n n ? q r). -apply lt_SO_nth_prime_n.assumption. -apply le_n. -rewrite < H1.assumption. +cut (r \mod (nth_prime (max_prime_factor n)) \neq O); + [unfold Not in Hcut1.autobatch new. + (* + 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); + [ 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. @@ -127,6 +194,33 @@ assumption.apply sym_eq.assumption.assumption.assumption. apply (le_to_or_lt_eq ? p H1). qed. +lemma lt_max_prime_factor_to_not_divides: \forall n,p:nat. +O < n \to n=S O \lor max_prime_factor n < p \to +(nth_prime p \ndivides n). +intros.unfold Not.intro. +elim H1 + [rewrite > H3 in H2. + apply (le_to_not_lt (nth_prime p) (S O)) + [apply divides_to_le[apply le_n|assumption] + |apply lt_SO_nth_prime_n + ] + |apply (not_le_Sn_n p). + change with (p < p). + apply (le_to_lt_to_lt ? ? ? ? H3). + unfold max_prime_factor. + apply f_m_to_le_max + [apply (trans_le ? (nth_prime p)) + [apply lt_to_le. + apply lt_n_nth_prime_n + |apply divides_to_le;assumption + ] + |apply eq_to_eqb_true. + apply divides_to_mod_O + [apply lt_O_nth_prime_n|assumption] + ] + ] +qed. + (* datatypes and functions *) inductive nat_fact : Set \def @@ -171,24 +265,23 @@ 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. +theorem lt_O_defactorize_aux: + \forall f:nat_fact. + \forall i:nat. + O < defactorize_aux f i. +intro; elim f; +[1,2: + simplify; unfold lt; + rewrite > times_n_SO; + apply le_times; + [ change with (O < nth_prime i); + apply lt_O_nth_prime_n; + |2,3: + 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. @@ -258,13 +351,13 @@ apply (nat_case n). left.split.assumption.reflexivity. intro.right.rewrite > Hcut2. simplify.unfold lt.apply le_S_S.apply le_O_n. -cut (r \lt (S O) \or r=(S O)). +cut (r < (S O) ∨ r=(S O)). elim Hcut2.absurd (O=r). apply le_n_O_to_eq.apply le_S_S_to_le.exact H5. unfold Not.intro. cut (O=n1). apply (not_le_Sn_O O). -rewrite > Hcut3 in \vdash (? ? %). +rewrite > Hcut3 in ⊢ (? ? %). assumption.rewrite > Hcut. rewrite < H6.reflexivity. assumption. @@ -280,7 +373,8 @@ theorem defactorize_factorize: \forall n:nat.defactorize (factorize n) = n. intro. apply (nat_case n).reflexivity. intro.apply (nat_case m).reflexivity. -intro.change with +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 [ (pair q r) \Rightarrow @@ -296,8 +390,7 @@ apply (Hcut (fst ? ? (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))). +simplify. cut ((S(S m1)) = (nth_prime p) \sup q *r). cut (O defactorize_aux_factorize_aux. @@ -361,7 +454,6 @@ 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 (p_ord_aux_to_exp (S(S m1))). apply lt_O_nth_prime_n. assumption. @@ -395,28 +487,86 @@ 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. +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 +p = max_prime_factor (r*(nth_prime p)\sup n). 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. +apply sym_eq. +unfold max_prime_factor. +apply max_spec_to_max. +split + [split + [rewrite > times_n_SO in \vdash (? % ?). + rewrite > sym_times. + apply le_times + [assumption + |apply lt_to_le. + apply (lt_to_le_to_lt ? (nth_prime p)) + [apply lt_n_nth_prime_n + |rewrite > exp_n_SO in \vdash (? % ?). + apply le_exp + [apply lt_O_nth_prime_n + |assumption + ] + ] + ] + |apply eq_to_eqb_true. + apply divides_to_mod_O + [apply lt_O_nth_prime_n + |apply (lt_O_n_elim ? H). + intro. + apply (witness ? ? (r*(nth_prime p \sup m))). + rewrite < assoc_times. + rewrite < sym_times in \vdash (? ? ? (? % ?)). + rewrite > exp_n_SO in \vdash (? ? ? (? (? ? %) ?)). + rewrite > assoc_times. + rewrite < exp_plus_times. + reflexivity + ] + ] + |intros. + apply not_eq_to_eqb_false. + unfold Not.intro. + lapply (mod_O_to_divides ? ? ? H5) + [apply lt_O_nth_prime_n + |cut (Not (divides (nth_prime i) ((nth_prime p)\sup n))) + [elim H2 + [rewrite > H6 in Hletin. + simplify in Hletin. + rewrite < plus_n_O in Hletin. + apply Hcut.assumption + |elim (divides_times_to_divides ? ? ? ? Hletin) + [apply (lt_to_not_le ? ? H3). + apply lt_to_le. + apply (le_to_lt_to_lt ? ? ? ? H6). + apply f_m_to_le_max + [apply (trans_le ? (nth_prime i)) + [apply lt_to_le. + apply lt_n_nth_prime_n + |apply divides_to_le[assumption|assumption] + ] + |apply eq_to_eqb_true. + apply divides_to_mod_O + [apply lt_O_nth_prime_n|assumption] + ] + |apply prime_nth_prime + |apply Hcut.assumption + ] + ] + |unfold Not.intro. + apply (lt_to_not_eq ? ? H3). + apply sym_eq. + elim (prime_nth_prime p). + apply injective_nth_prime. + apply H8 + [apply (divides_exp_to_divides ? ? ? ? H6). + apply prime_nth_prime. + |apply lt_SO_nth_prime_n + ] + ] + ] + ] qed. theorem not_divides_defactorize_aux: \forall f:nat_fact. \forall i,j:nat. @@ -427,13 +577,11 @@ change with 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). +assumption.unfold Not. 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). +unfold Not.simplify. intro. cut (nth_prime i \divides (nth_prime j) \sup n \lor nth_prime i \divides defactorize_aux n1 (S j)). @@ -441,8 +589,7 @@ elim Hcut. 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). +assumption.unfold Not. intro. cut (i = j). apply (not_le_Sn_n i).rewrite > Hcut1 in \vdash (? ? %).assumption. @@ -466,8 +613,6 @@ 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. @@ -547,16 +692,15 @@ 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). +simplify. intros. -apply (eq_defactorize_aux_to_eq f g i H). +apply (eq_defactorize_aux_to_eq x y 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). +unfold injective. +change with (\forall f,g.defactorize f = defactorize g \to f=g). intro.elim f. generalize in match H.elim g. (* zero - zero *) @@ -609,11 +753,8 @@ exact H1. 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. -(* uffa: perche' semplifica ??? *) -change with (defactorize(factorize (defactorize f)) = (defactorize f)). apply defactorize_factorize. qed. -