diff -r d7c525fd68b2 -r 1cf4f1e930af src/HOL/Computational_Algebra/Primes.thy --- a/src/HOL/Computational_Algebra/Primes.thy Sun Nov 02 19:47:30 2025 +0100 +++ b/src/HOL/Computational_Algebra/Primes.thy Sun Nov 02 19:47:30 2025 +0100 @@ -1061,7 +1061,26 @@ lemmas prime_exp = prime_elem_power_iff[where ?'a = nat] text \Code generation\ - + +lemma divisor_less_eq_half_nat: + \m \ n div 2\ if \m dvd n\ \m < n\ for m n :: nat + using that by (auto simp add: less_eq_div_iff_mult_less_eq) + +lemma divisor_less_eq_half_int: + \k \ l div 2\ if \k dvd l\ \k < l\ \l \ 0\ \k \ 0\ for k l :: int +proof - + define m n where \m = nat \k\\ \n = nat \l\\ + with \l \ 0\ \k \ 0\ have \k = int m\ \l = int n\ + by simp_all + with that show ?thesis + using divisor_less_eq_half_nat [of m n] by simp +qed + +lemma + \prime n \ 1 < n \ (\n\{1<..n div 2}. \ n dvd p)\ + +thm prime_nat_iff prime_int_iff [no_vars] + context begin