]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/library/nat/factorization.ma
- libraryObjects: new default "natural numbers" with the uri of nat.
[helm.git] / helm / software / matita / library / nat / factorization.ma
index 8c50d1d7db1e5612e160776b39a59e55ab21fece..bfdf947b4ba1ec437d1520b4095bd01a1320e54c 100644 (file)
@@ -12,8 +12,6 @@
 (*                                                                        *)
 (**************************************************************************)
 
-set "baseuri" "cic:/matita/nat/factorization".
-
 include "nat/ord.ma".
 
 (* the following factorization algorithm looks for the largest prime
@@ -63,19 +61,28 @@ cut (\exists i. nth_prime i = smallest_factor n);
       [ 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.
@@ -146,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.
@@ -516,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 (? ? ? (? (? ? %) ?)).
@@ -638,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
@@ -684,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.
@@ -702,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.
@@ -728,28 +727,28 @@ 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: