X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Farithmetics%2Fprimes.ma;h=68e511e982617f293eedfc320aaa6be1ec888447;hb=5535cd4e08fd8d1e7e6e067eac1bb6c1bf8fcbbf;hp=3898105c6aaac49345f32c6ea925e52daf0029ed;hpb=31e8729021717072f88d250ef41527da3488289e;p=helm.git diff --git a/matita/matita/lib/arithmetics/primes.ma b/matita/matita/lib/arithmetics/primes.ma index 3898105c6..68e511e98 100644 --- a/matita/matita/lib/arithmetics/primes.ma +++ b/matita/matita/lib/arithmetics/primes.ma @@ -247,26 +247,24 @@ 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))). #n #lt1n normalize >lt_to_leb_false // qed. -(* it works ! -theorem example1 : smallest_factor 3 = 3. +example example1 : smallest_factor 3 = 3. normalize // qed. -theorem example2: smallest_factor 4 = 2. +example example2: smallest_factor 4 = 2. normalize // qed. -theorem example3 : smallest_factor 7 = 7. +example example3 : smallest_factor 7 = 7. normalize // -qed. *) +qed. theorem le_SO_smallest_factor: ∀n:nat. n ≤ 1 → smallest_factor n = n. @@ -291,7 +289,7 @@ theorem divides_smallest_factor_n : ∀n:nat. 0 < n → #n #posn (cases (le_to_or_lt_eq … posn)) [#lt1n @mod_O_to_divides [@lt_O_smallest_factor //] >smallest_factor_to_min // @eqb_true_to_eq @f_min_true - @(ex_intro … n) % /2/ @eq_to_eqb_true /2/ + @(ex_intro … n) % /2/ |#H nth_primeS @primeb_true_to_prime check f_min_true @(f_min_true primeb) + |#m >nth_primeS @primeb_true_to_prime @(f_min_true primeb) @(ex_intro nat ? (smallest_factor (S (nth_prime m)!))) % [% // @le_S_S @(transitive_le … (le_smallest_factor_n …))