From: Andrea Asperti Date: Thu, 23 Nov 2006 13:45:53 +0000 (+0000) Subject: Minor changes. X-Git-Tag: 0.4.95@7852~795 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=544444b7fbb3882577f55b7dbd16046f39c52031;p=helm.git Minor changes. -This line, and those below, will be ignored-- M library/nat/ord.ma M library/nat/gcd.ma M library/nat/factorization.ma M library/technicalities/setoids.ma --- diff --git a/matita/library/nat/factorization.ma b/matita/library/nat/factorization.ma index d0036e4a1..85351c06d 100644 --- a/matita/library/nat/factorization.ma +++ b/matita/library/nat/factorization.ma @@ -52,7 +52,7 @@ intros; apply divides_b_true_to_divides; apply (trans_lt ? (S O)); [ unfold lt; apply le_n; | assumption; ] *) ] ] - | letin x \def prime. auto new. + | auto. (* apply prime_to_nth_prime; apply prime_smallest_factor_n; @@ -70,8 +70,8 @@ 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). -auto new. -auto new. +auto. +auto. (* [ apply (transitive_divides ? n); [ apply divides_max_prime_factor_n. diff --git a/matita/library/nat/gcd.ma b/matita/library/nat/gcd.ma index fdb1e8d90..036167bd0 100644 --- a/matita/library/nat/gcd.ma +++ b/matita/library/nat/gcd.ma @@ -39,16 +39,18 @@ definition gcd : nat \to nat \to nat \def theorem divides_mod: \forall p,m,n:nat. O < n \to p \divides m \to p \divides n \to p \divides (m \mod n). intros.elim H1.elim H2. -apply (witness ? ? (n2 - n1*(m / n))). +(* apply (witness ? ? (n2 - n1*(m / n))). *) +apply witness[| rewrite > distr_times_minus. -rewrite < H3. +rewrite < H3 in \vdash (? ? ? (? % ?)). rewrite < assoc_times. -rewrite < H4. -apply sym_eq. -apply plus_to_minus. +rewrite < H4 in \vdash (? ? ? (? ? (? % ?))). +apply sym_eq.apply plus_to_minus. rewrite > sym_times. -apply div_mod. -assumption. +letin x \def div. +rewrite < (div_mod ? ? H). +reflexivity. +] qed. theorem divides_mod_to_divides: \forall p,m,n:nat. O < n \to diff --git a/matita/library/nat/ord.ma b/matita/library/nat/ord.ma index ba08a0dcf..b7d02f87f 100644 --- a/matita/library/nat/ord.ma +++ b/matita/library/nat/ord.ma @@ -215,7 +215,7 @@ apply (absurd ? ? H10 H5). apply (absurd ? ? H10 H7). (* rewrite > H6. rewrite > H8. *) -auto paramodulation library=yes. +auto paramodulation. unfold prime in H. elim H. assumption. qed. diff --git a/matita/library/technicalities/setoids.ma b/matita/library/technicalities/setoids.ma index c440795ff..4d23d7725 100644 --- a/matita/library/technicalities/setoids.ma +++ b/matita/library/technicalities/setoids.ma @@ -228,7 +228,7 @@ definition equality_morphism_of_symmetric_areflexive_transitive_relation: unfold transitive in H; unfold symmetric in sym; intro; - auto + auto new ]. qed.