X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Farithmetics%2Fprimes.ma;h=68e511e982617f293eedfc320aaa6be1ec888447;hb=81fc94f4f091ec35d41e2711207218d255b75273;hp=7c5df29f35151cebe5d3c6664368c8ee9f2ece73;hpb=ddc80515997a3f56085c6234d4db326141e189aa;p=helm.git diff --git a/matita/matita/lib/arithmetics/primes.ma b/matita/matita/lib/arithmetics/primes.ma index 7c5df29f3..68e511e98 100644 --- a/matita/matita/lib/arithmetics/primes.ma +++ b/matita/matita/lib/arithmetics/primes.ma @@ -247,8 +247,7 @@ qed. (* smallest factor *) definition smallest_factor : nat → nat ≝ -λn:nat. if_then_else ? (leb n 1) n - (min n 2 (λm.(eqb (n \mod m) O))). +λn:nat. if leb n 1 then n else min n 2 (λm.(eqb (n \mod m) O)). theorem smallest_factor_to_min : ∀n. 1 < n → smallest_factor n = (min n 2 (λm.(eqb (n \mod m) O))). @@ -439,9 +438,9 @@ let rec nth_prime n ≝ match n with min upper_bound (S previous_prime) primeb]. lemma nth_primeS: ∀n.nth_prime (S n) = - let previous_prime ≝ nth_prime n in + (let previous_prime ≝ nth_prime n in let upper_bound ≝ S previous_prime! in - min upper_bound (S previous_prime) primeb. + min upper_bound (S previous_prime) primeb). // qed. (* it works *)