wenzelm@38159: (* Title: HOL/Old_Number_Theory/Pocklington.thy huffman@30488: Author: Amine Chaieb chaieb@26126: *) chaieb@26126: chaieb@26126: header {* Pocklington's Theorem for Primes *} chaieb@26126: chaieb@26126: theory Pocklington wenzelm@38159: imports Primes chaieb@26126: begin chaieb@26126: chaieb@26126: definition modeq:: "nat => nat => nat => bool" ("(1[_ = _] '(mod _'))") chaieb@26126: where "[a = b] (mod p) == ((a mod p) = (b mod p))" chaieb@26126: chaieb@26126: definition modneq:: "nat => nat => nat => bool" ("(1[_ \ _] '(mod _'))") chaieb@26126: where "[a \ b] (mod p) == ((a mod p) \ (b mod p))" chaieb@26126: chaieb@26126: lemma modeq_trans: chaieb@26126: "\ [a = b] (mod p); [b = c] (mod p) \ \ [a = c] (mod p)" chaieb@26126: by (simp add:modeq_def) chaieb@26126: hoelzl@57129: lemma modeq_sym[sym]: hoelzl@57129: "[a = b] (mod p) \ [b = a] (mod p)" hoelzl@57129: unfolding modeq_def by simp hoelzl@57129: hoelzl@57129: lemma modneq_sym[sym]: hoelzl@57129: "[a \ b] (mod p) \ [b \ a] (mod p)" hoelzl@57129: by (simp add: modneq_def) chaieb@26126: chaieb@26126: lemma nat_mod_lemma: assumes xyn: "[x = y] (mod n)" and xy:"y \ x" chaieb@26126: shows "\q. x = y + n * q" chaieb@27668: using xyn xy unfolding modeq_def using nat_mod_eq_lemma by blast chaieb@26126: huffman@30488: lemma nat_mod[algebra]: "[x = y] (mod n) \ (\q1 q2. x + n * q1 = y + n * q2)" chaieb@27668: unfolding modeq_def nat_mod_eq_iff .. chaieb@26126: chaieb@26126: (* Lemmas about previously defined terms. *) chaieb@26126: huffman@30488: lemma prime: "prime p \ p \ 0 \ p\1 \ (\m. 0 < m \ m < p \ coprime p m)" huffman@30488: (is "?lhs \ ?rhs") chaieb@26126: proof- chaieb@26126: {assume "p=0 \ p=1" hence ?thesis using prime_0 prime_1 by (cases "p=0", simp_all)} chaieb@26126: moreover chaieb@26126: {assume p0: "p\0" "p\1" chaieb@26126: {assume H: "?lhs" chaieb@26126: {fix m assume m: "m > 0" "m < p" wenzelm@32960: {assume "m=1" hence "coprime p m" by simp} wenzelm@32960: moreover wenzelm@32960: {assume "p dvd m" hence "p \ m" using dvd_imp_le m by blast with m(2) wenzelm@32960: have "coprime p m" by simp} wenzelm@32960: ultimately have "coprime p m" using prime_coprime[OF H, of m] by blast} chaieb@26126: hence ?rhs using p0 by auto} chaieb@26126: moreover chaieb@26126: { assume H: "\m. 0 < m \ m < p \ coprime p m" chaieb@26126: from prime_factor[OF p0(2)] obtain q where q: "prime q" "q dvd p" by blast chaieb@26126: from prime_ge_2[OF q(1)] have q0: "q > 0" by arith chaieb@26126: from dvd_imp_le[OF q(2)] p0 have qp: "q \ p" by arith chaieb@26126: {assume "q = p" hence ?lhs using q(1) by blast} chaieb@26126: moreover chaieb@26126: {assume "q\p" with qp have qplt: "q < p" by arith wenzelm@32960: from H[rule_format, of q] qplt q0 have "coprime p q" by arith wenzelm@32960: with coprime_prime[of p q q] q have False by simp hence ?lhs by blast} chaieb@26126: ultimately have ?lhs by blast} chaieb@26126: ultimately have ?thesis by blast} chaieb@26126: ultimately show ?thesis by (cases"p=0 \ p=1", auto) chaieb@26126: qed chaieb@26126: chaieb@26126: lemma finite_number_segment: "card { m. 0 < m \ m < n } = n - 1" chaieb@26126: proof- chaieb@26126: have "{ m. 0 < m \ m < n } = {1.. 0" shows "coprime (a mod n) n \ coprime a n" chaieb@26126: using n dvd_mod_iff[of _ n a] by (auto simp add: coprime) chaieb@26126: chaieb@26126: (* Congruences. *) chaieb@26126: huffman@30488: lemma cong_mod_01[simp,presburger]: chaieb@26126: "[x = y] (mod 0) \ x = y" "[x = y] (mod 1)" "[x = 0] (mod n) \ n dvd x" chaieb@26126: by (simp_all add: modeq_def, presburger) chaieb@26126: huffman@30488: lemma cong_sub_cases: chaieb@26126: "[x = y] (mod n) \ (if x <= y then [y - x = 0] (mod n) else [x - y = 0] (mod n))" chaieb@26126: apply (auto simp add: nat_mod) chaieb@26126: apply (rule_tac x="q2" in exI) chaieb@26126: apply (rule_tac x="q1" in exI, simp) chaieb@26126: apply (rule_tac x="q2" in exI) chaieb@26126: apply (rule_tac x="q1" in exI, simp) chaieb@26126: apply (rule_tac x="q1" in exI) chaieb@26126: apply (rule_tac x="q2" in exI, simp) chaieb@26126: apply (rule_tac x="q1" in exI) chaieb@26126: apply (rule_tac x="q2" in exI, simp) chaieb@26126: done chaieb@26126: chaieb@26126: lemma cong_mult_lcancel: assumes an: "coprime a n" and axy:"[a * x = a * y] (mod n)" chaieb@26126: shows "[x = y] (mod n)" chaieb@26126: proof- chaieb@26126: {assume "a = 0" with an axy coprime_0'[of n] have ?thesis by (simp add: modeq_def) } chaieb@26126: moreover chaieb@26126: {assume az: "a\0" chaieb@26126: {assume xy: "x \ y" hence axy': "a*x \ a*y" by simp chaieb@26126: with axy cong_sub_cases[of "a*x" "a*y" n] have "[a*(y - x) = 0] (mod n)" wenzelm@32960: by (simp only: if_True diff_mult_distrib2) huffman@30488: hence th: "n dvd a*(y -x)" by simp chaieb@26126: from coprime_divprod[OF th] an have "n dvd y - x" wenzelm@32960: by (simp add: coprime_commute) chaieb@26126: hence ?thesis using xy cong_sub_cases[of x y n] by simp} chaieb@26126: moreover huffman@30488: {assume H: "\x \ y" hence xy: "y \ x" by arith chaieb@26126: from H az have axy': "\ a*x \ a*y" by auto chaieb@26126: with axy H cong_sub_cases[of "a*x" "a*y" n] have "[a*(x - y) = 0] (mod n)" wenzelm@32960: by (simp only: if_False diff_mult_distrib2) huffman@30488: hence th: "n dvd a*(x - y)" by simp chaieb@26126: from coprime_divprod[OF th] an have "n dvd x - y" wenzelm@32960: by (simp add: coprime_commute) chaieb@26126: hence ?thesis using xy cong_sub_cases[of x y n] by simp} chaieb@26126: ultimately have ?thesis by blast} chaieb@26126: ultimately show ?thesis by blast chaieb@26126: qed chaieb@26126: chaieb@26126: lemma cong_mult_rcancel: assumes an: "coprime a n" and axy:"[x*a = y*a] (mod n)" chaieb@26126: shows "[x = y] (mod n)" haftmann@57512: using cong_mult_lcancel[OF an axy[unfolded mult.commute[of _a]]] . chaieb@26126: chaieb@26126: lemma cong_refl: "[x = x] (mod n)" by (simp add: modeq_def) chaieb@26126: chaieb@26126: lemma eq_imp_cong: "a = b \ [a = b] (mod n)" by (simp add: cong_refl) chaieb@26126: huffman@30488: lemma cong_commute: "[x = y] (mod n) \ [y = x] (mod n)" chaieb@26126: by (auto simp add: modeq_def) chaieb@26126: chaieb@26126: lemma cong_trans[trans]: "[x = y] (mod n) \ [y = z] (mod n) \ [x = z] (mod n)" chaieb@26126: by (simp add: modeq_def) chaieb@26126: chaieb@26126: lemma cong_add: assumes xx': "[x = x'] (mod n)" and yy':"[y = y'] (mod n)" chaieb@26126: shows "[x + y = x' + y'] (mod n)" chaieb@26126: proof- chaieb@26126: have "(x + y) mod n = (x mod n + y mod n) mod n" chaieb@26126: by (simp add: mod_add_left_eq[of x y n] mod_add_right_eq[of "x mod n" y n]) huffman@30488: also have "\ = (x' mod n + y' mod n) mod n" using xx' yy' modeq_def by simp chaieb@26126: also have "\ = (x' + y') mod n" chaieb@26126: by (simp add: mod_add_left_eq[of x' y' n] mod_add_right_eq[of "x' mod n" y' n]) huffman@30488: finally show ?thesis unfolding modeq_def . chaieb@26126: qed chaieb@26126: chaieb@26126: lemma cong_mult: assumes xx': "[x = x'] (mod n)" and yy':"[y = y'] (mod n)" chaieb@26126: shows "[x * y = x' * y'] (mod n)" chaieb@26126: proof- huffman@30488: have "(x * y) mod n = (x mod n) * (y mod n) mod n" nipkow@30224: by (simp add: mod_mult_left_eq[of x y n] mod_mult_right_eq[of "x mod n" y n]) huffman@30488: also have "\ = (x' mod n) * (y' mod n) mod n" using xx'[unfolded modeq_def] yy'[unfolded modeq_def] by simp chaieb@26126: also have "\ = (x' * y') mod n" nipkow@30224: by (simp add: mod_mult_left_eq[of x' y' n] mod_mult_right_eq[of "x' mod n" y' n]) huffman@30488: finally show ?thesis unfolding modeq_def . chaieb@26126: qed chaieb@26126: chaieb@26126: lemma cong_exp: "[x = y] (mod n) \ [x^k = y^k] (mod n)" chaieb@26126: by (induct k, auto simp add: cong_refl cong_mult) chaieb@26126: lemma cong_sub: assumes xx': "[x = x'] (mod n)" and yy': "[y = y'] (mod n)" chaieb@26126: and yx: "y <= x" and yx': "y' <= x'" chaieb@26126: shows "[x - y = x' - y'] (mod n)" chaieb@26126: proof- huffman@30488: { fix x a x' a' y b y' b' chaieb@26126: have "(x::nat) + a = x' + a' \ y + b = y' + b' \ y <= x \ y' <= x' chaieb@26126: \ (x - y) + (a + b') = (x' - y') + (a' + b)" by arith} chaieb@26126: note th = this huffman@30488: from xx' yy' obtain q1 q2 q1' q2' where q12: "x + n*q1 = x'+n*q2" chaieb@26126: and q12': "y + n*q1' = y'+n*q2'" unfolding nat_mod by blast+ chaieb@26126: from th[OF q12 q12' yx yx'] huffman@30488: have "(x - y) + n*(q1 + q2') = (x' - y') + n*(q2 + q1')" webertj@49962: by (simp add: distrib_left) chaieb@26126: thus ?thesis unfolding nat_mod by blast chaieb@26126: qed chaieb@26126: huffman@30488: lemma cong_mult_lcancel_eq: assumes an: "coprime a n" chaieb@26126: shows "[a * x = a * y] (mod n) \ [x = y] (mod n)" (is "?lhs \ ?rhs") chaieb@26126: proof chaieb@26126: assume H: "?rhs" from cong_mult[OF cong_refl[of a n] H] show ?lhs . chaieb@26126: next haftmann@57512: assume H: "?lhs" hence H': "[x*a = y*a] (mod n)" by (simp add: mult.commute) chaieb@26126: from cong_mult_rcancel[OF an H'] show ?rhs . chaieb@26126: qed chaieb@26126: huffman@30488: lemma cong_mult_rcancel_eq: assumes an: "coprime a n" chaieb@26126: shows "[x * a = y * a] (mod n) \ [x = y] (mod n)" haftmann@57512: using cong_mult_lcancel_eq[OF an, of x y] by (simp add: mult.commute) chaieb@26126: huffman@30488: lemma cong_add_lcancel_eq: "[a + x = a + y] (mod n) \ [x = y] (mod n)" chaieb@26126: by (simp add: nat_mod) chaieb@26126: chaieb@26126: lemma cong_add_rcancel_eq: "[x + a = y + a] (mod n) \ [x = y] (mod n)" chaieb@26126: by (simp add: nat_mod) chaieb@26126: huffman@30488: lemma cong_add_rcancel: "[x + a = y + a] (mod n) \ [x = y] (mod n)" chaieb@26126: by (simp add: nat_mod) chaieb@26126: chaieb@26126: lemma cong_add_lcancel: "[a + x = a + y] (mod n) \ [x = y] (mod n)" chaieb@26126: by (simp add: nat_mod) chaieb@26126: huffman@30488: lemma cong_add_lcancel_eq_0: "[a + x = a] (mod n) \ [x = 0] (mod n)" chaieb@26126: by (simp add: nat_mod) chaieb@26126: chaieb@26126: lemma cong_add_rcancel_eq_0: "[x + a = a] (mod n) \ [x = 0] (mod n)" chaieb@26126: by (simp add: nat_mod) chaieb@26126: chaieb@26126: lemma cong_imp_eq: assumes xn: "x < n" and yn: "y < n" and xy: "[x = y] (mod n)" chaieb@26126: shows "x = y" huffman@30488: using xy[unfolded modeq_def mod_less[OF xn] mod_less[OF yn]] . chaieb@26126: chaieb@26126: lemma cong_divides_modulus: "[x = y] (mod m) \ n dvd m ==> [x = y] (mod n)" chaieb@26126: apply (auto simp add: nat_mod dvd_def) chaieb@26126: apply (rule_tac x="k*q1" in exI) chaieb@26126: apply (rule_tac x="k*q2" in exI) chaieb@26126: by simp huffman@30488: chaieb@26126: lemma cong_0_divides: "[x = 0] (mod n) \ n dvd x" by simp chaieb@26126: chaieb@26126: lemma cong_1_divides:"[x = 1] (mod n) ==> n dvd x - 1" chaieb@26126: apply (cases "x\1", simp_all) chaieb@26126: using cong_sub_cases[of x 1 n] by auto chaieb@26126: chaieb@26126: lemma cong_divides: "[x = y] (mod n) \ n dvd x \ n dvd y" chaieb@26126: apply (auto simp add: nat_mod dvd_def) chaieb@26126: apply (rule_tac x="k + q1 - q2" in exI, simp add: add_mult_distrib2 diff_mult_distrib2) chaieb@26126: apply (rule_tac x="k + q2 - q1" in exI, simp add: add_mult_distrib2 diff_mult_distrib2) chaieb@26126: done chaieb@26126: huffman@30488: lemma cong_coprime: assumes xy: "[x = y] (mod n)" chaieb@26126: shows "coprime n x \ coprime n y" chaieb@26126: proof- chaieb@26126: {assume "n=0" hence ?thesis using xy by simp} chaieb@26126: moreover chaieb@26126: {assume nz: "n \ 0" huffman@30488: have "coprime n x \ coprime (x mod n) n" chaieb@26126: by (simp add: coprime_mod[OF nz, of x] coprime_commute[of n x]) chaieb@26126: also have "\ \ coprime (y mod n) n" using xy[unfolded modeq_def] by simp chaieb@26126: also have "\ \ coprime y n" by (simp add: coprime_mod[OF nz, of y]) chaieb@26126: finally have ?thesis by (simp add: coprime_commute) } chaieb@26126: ultimately show ?thesis by blast chaieb@26126: qed chaieb@26126: chaieb@26126: lemma cong_mod: "~(n = 0) \ [a mod n = a] (mod n)" by (simp add: modeq_def) chaieb@26126: huffman@30488: lemma mod_mult_cong: "~(a = 0) \ ~(b = 0) chaieb@26126: \ [x mod (a * b) = y] (mod a) \ [x = y] (mod a)" chaieb@26126: by (simp add: modeq_def mod_mult2_eq mod_add_left_eq) chaieb@26126: chaieb@26126: lemma cong_mod_mult: "[x = y] (mod n) \ m dvd n \ [x = y] (mod m)" chaieb@26126: apply (auto simp add: nat_mod dvd_def) chaieb@26126: apply (rule_tac x="k*q1" in exI) chaieb@26126: apply (rule_tac x="k*q2" in exI, simp) chaieb@26126: done chaieb@26126: chaieb@26126: (* Some things when we know more about the order. *) chaieb@26126: chaieb@26126: lemma cong_le: "y <= x \ [x = y] (mod n) \ (\q. x = q * n + y)" chaieb@26126: using nat_mod_lemma[of x y n] chaieb@26126: apply auto chaieb@26126: apply (simp add: nat_mod) chaieb@26126: apply (rule_tac x="q" in exI) chaieb@26126: apply (rule_tac x="q + q" in exI) nipkow@29667: by (auto simp: algebra_simps) chaieb@26126: chaieb@26126: lemma cong_to_1: "[a = 1] (mod n) \ a = 0 \ n = 1 \ (\m. a = 1 + m * n)" chaieb@26126: proof- huffman@30488: {assume "n = 0 \ n = 1\ a = 0 \ a = 1" hence ?thesis chaieb@26126: apply (cases "n=0", simp_all add: cong_commute) chaieb@26126: apply (cases "n=1", simp_all add: cong_commute modeq_def) huffman@30488: apply arith wenzelm@41541: apply (cases "a=1") wenzelm@41541: apply (simp_all add: modeq_def cong_commute) wenzelm@41541: done } chaieb@26126: moreover chaieb@26126: {assume n: "n\0" "n\1" and a:"a\0" "a \ 1" hence a': "a \ 1" by simp chaieb@26126: hence ?thesis using cong_le[OF a', of n] by auto } chaieb@26126: ultimately show ?thesis by auto chaieb@26126: qed chaieb@26126: chaieb@26126: (* Some basic theorems about solving congruences. *) chaieb@26126: chaieb@26126: chaieb@26126: lemma cong_solve: assumes an: "coprime a n" shows "\x. [a * x = b] (mod n)" chaieb@26126: proof- chaieb@26126: {assume "a=0" hence ?thesis using an by (simp add: modeq_def)} chaieb@26126: moreover chaieb@26126: {assume az: "a\0" huffman@30488: from bezout_add_strong[OF az, of n] chaieb@26126: obtain d x y where dxy: "d dvd a" "d dvd n" "a*x = n*y + d" by blast chaieb@26126: from an[unfolded coprime, rule_format, of d] dxy(1,2) have d1: "d = 1" by blast chaieb@26126: hence "a*x*b = (n*y + 1)*b" using dxy(3) by simp chaieb@26126: hence "a*(x*b) = n*(y*b) + b" by algebra chaieb@26126: hence "a*(x*b) mod n = (n*(y*b) + b) mod n" by simp chaieb@26126: hence "a*(x*b) mod n = b mod n" by (simp add: mod_add_left_eq) chaieb@26126: hence "[a*(x*b) = b] (mod n)" unfolding modeq_def . chaieb@26126: hence ?thesis by blast} chaieb@26126: ultimately show ?thesis by blast chaieb@26126: qed chaieb@26126: chaieb@26126: lemma cong_solve_unique: assumes an: "coprime a n" and nz: "n \ 0" chaieb@26126: shows "\!x. x < n \ [a * x = b] (mod n)" chaieb@26126: proof- chaieb@26126: let ?P = "\x. x < n \ [a * x = b] (mod n)" chaieb@26126: from cong_solve[OF an] obtain x where x: "[a*x = b] (mod n)" by blast chaieb@26126: let ?x = "x mod n" chaieb@26126: from x have th: "[a * ?x = b] (mod n)" nipkow@30224: by (simp add: modeq_def mod_mult_right_eq[of a x n]) chaieb@26126: from mod_less_divisor[ of n x] nz th have Px: "?P ?x" by simp chaieb@26126: {fix y assume Py: "y < n" "[a * y = b] (mod n)" chaieb@26126: from Py(2) th have "[a * y = a*?x] (mod n)" by (simp add: modeq_def) chaieb@26126: hence "[y = ?x] (mod n)" by (simp add: cong_mult_lcancel_eq[OF an]) chaieb@26126: with mod_less[OF Py(1)] mod_less_divisor[ of n x] nz chaieb@26126: have "y = ?x" by (simp add: modeq_def)} chaieb@26126: with Px show ?thesis by blast chaieb@26126: qed chaieb@26126: chaieb@26126: lemma cong_solve_unique_nontrivial: chaieb@26126: assumes p: "prime p" and pa: "coprime p a" and x0: "0 < x" and xp: "x < p" chaieb@26126: shows "\!y. 0 < y \ y < p \ [x * y = a] (mod p)" chaieb@26126: proof- chaieb@26126: from p have p1: "p > 1" using prime_ge_2[OF p] by arith chaieb@26126: hence p01: "p \ 0" "p \ 1" by arith+ chaieb@26126: from pa have ap: "coprime a p" by (simp add: coprime_commute) chaieb@26126: from prime_coprime[OF p, of x] dvd_imp_le[of p x] x0 xp have px:"coprime x p" chaieb@26126: by (auto simp add: coprime_commute) huffman@30488: from cong_solve_unique[OF px p01(1)] chaieb@26126: obtain y where y: "y < p" "[x * y = a] (mod p)" "\z. z < p \ [x * z = a] (mod p) \ z = y" by blast chaieb@26126: {assume y0: "y = 0" chaieb@26126: with y(2) have th: "p dvd a" by (simp add: cong_commute[of 0 a p]) chaieb@26126: with p coprime_prime[OF pa, of p] have False by simp} huffman@30488: with y show ?thesis unfolding Ex1_def using neq0_conv by blast chaieb@26126: qed chaieb@26126: lemma cong_unique_inverse_prime: chaieb@26126: assumes p: "prime p" and x0: "0 < x" and xp: "x < p" chaieb@26126: shows "\!y. 0 < y \ y < p \ [x * y = 1] (mod p)" chaieb@26126: using cong_solve_unique_nontrivial[OF p coprime_1[of p] x0 xp] . chaieb@26126: chaieb@26126: (* Forms of the Chinese remainder theorem. *) chaieb@26126: huffman@30488: lemma cong_chinese: huffman@30488: assumes ab: "coprime a b" and xya: "[x = y] (mod a)" chaieb@26126: and xyb: "[x = y] (mod b)" chaieb@26126: shows "[x = y] (mod a*b)" chaieb@26126: using ab xya xyb huffman@30488: by (simp add: cong_sub_cases[of x y a] cong_sub_cases[of x y b] huffman@30488: cong_sub_cases[of x y "a*b"]) chaieb@26126: (cases "x \ y", simp_all add: divides_mul[of a _ b]) chaieb@26126: chaieb@26126: lemma chinese_remainder_unique: chaieb@26126: assumes ab: "coprime a b" and az: "a \ 0" and bz: "b\0" chaieb@26126: shows "\!x. x < a * b \ [x = m] (mod a) \ [x = n] (mod b)" chaieb@26126: proof- chaieb@26126: from az bz have abpos: "a*b > 0" by simp huffman@30488: from chinese_remainder[OF ab az bz] obtain x q1 q2 where chaieb@26126: xq12: "x = m + q1 * a" "x = n + q2 * b" by blast huffman@30488: let ?w = "x mod (a*b)" chaieb@26126: have wab: "?w < a*b" by (simp add: mod_less_divisor[OF abpos]) chaieb@26126: from xq12(1) have "?w mod a = ((m + q1 * a) mod (a*b)) mod a" by simp haftmann@54221: also have "\ = m mod a" by (simp add: mod_mult2_eq) chaieb@26126: finally have th1: "[?w = m] (mod a)" by (simp add: modeq_def) chaieb@26126: from xq12(2) have "?w mod b = ((n + q2 * b) mod (a*b)) mod b" by simp haftmann@57512: also have "\ = ((n + q2 * b) mod (b*a)) mod b" by (simp add: mult.commute) haftmann@54221: also have "\ = n mod b" by (simp add: mod_mult2_eq) chaieb@26126: finally have th2: "[?w = n] (mod b)" by (simp add: modeq_def) chaieb@26126: {fix y assume H: "y < a*b" "[y = m] (mod a)" "[y = n] (mod b)" chaieb@26126: with th1 th2 have H': "[y = ?w] (mod a)" "[y = ?w] (mod b)" chaieb@26126: by (simp_all add: modeq_def) huffman@30488: from cong_chinese[OF ab H'] mod_less[OF H(1)] mod_less[OF wab] chaieb@26126: have "y = ?w" by (simp add: modeq_def)} chaieb@26126: with th1 th2 wab show ?thesis by blast chaieb@26126: qed chaieb@26126: chaieb@26126: lemma chinese_remainder_coprime_unique: huffman@30488: assumes ab: "coprime a b" and az: "a \ 0" and bz: "b \ 0" chaieb@26126: and ma: "coprime m a" and nb: "coprime n b" chaieb@26126: shows "\!x. coprime x (a * b) \ x < a * b \ [x = m] (mod a) \ [x = n] (mod b)" chaieb@26126: proof- chaieb@26126: let ?P = "\x. x < a * b \ [x = m] (mod a) \ [x = n] (mod b)" chaieb@26126: from chinese_remainder_unique[OF ab az bz] huffman@30488: obtain x where x: "x < a * b" "[x = m] (mod a)" "[x = n] (mod b)" chaieb@26126: "\y. ?P y \ y = x" by blast chaieb@26126: from ma nb cong_coprime[OF x(2)] cong_coprime[OF x(3)] chaieb@26126: have "coprime x a" "coprime x b" by (simp_all add: coprime_commute) chaieb@26126: with coprime_mul[of x a b] have "coprime x (a*b)" by simp chaieb@26126: with x show ?thesis by blast chaieb@26126: qed chaieb@26126: chaieb@26126: (* Euler totient function. *) chaieb@26126: chaieb@26126: definition phi_def: "\ n = card { m. 0 < m \ m <= n \ coprime m n }" nipkow@31197: chaieb@26126: lemma phi_0[simp]: "\ 0 = 0" nipkow@31197: unfolding phi_def by auto chaieb@26126: chaieb@26126: lemma phi_finite[simp]: "finite ({ m. 0 < m \ m <= n \ coprime m n })" chaieb@26126: proof- chaieb@26126: have "{ m. 0 < m \ m <= n \ coprime m n } \ {0..n}" by auto chaieb@26126: thus ?thesis by (auto intro: finite_subset) chaieb@26126: qed chaieb@26126: chaieb@26126: declare coprime_1[presburger] chaieb@26126: lemma phi_1[simp]: "\ 1 = 1" chaieb@26126: proof- huffman@30488: {fix m chaieb@26126: have "0 < m \ m <= 1 \ coprime m 1 \ m = 1" by presburger } chaieb@26126: thus ?thesis by (simp add: phi_def) chaieb@26126: qed chaieb@26126: chaieb@26126: lemma [simp]: "\ (Suc 0) = Suc 0" using phi_1 by simp chaieb@26126: chaieb@26126: lemma phi_alt: "\(n) = card { m. coprime m n \ m < n}" chaieb@26126: proof- chaieb@26126: {assume "n=0 \ n=1" hence ?thesis by (cases "n=0", simp_all)} chaieb@26126: moreover chaieb@26126: {assume n: "n\0" "n\1" chaieb@26126: {fix m chaieb@26126: from n have "0 < m \ m <= n \ coprime m n \ coprime m n \ m < n" wenzelm@32960: apply (cases "m = 0", simp_all) wenzelm@32960: apply (cases "m = 1", simp_all) wenzelm@32960: apply (cases "m = n", auto) wenzelm@32960: done } chaieb@26126: hence ?thesis unfolding phi_def by simp} chaieb@26126: ultimately show ?thesis by auto chaieb@26126: qed chaieb@26126: chaieb@26126: lemma phi_finite_lemma[simp]: "finite {m. coprime m n \ m < n}" (is "finite ?S") chaieb@26126: by (rule finite_subset[of "?S" "{0..n}"], auto) chaieb@26126: chaieb@26126: lemma phi_another: assumes n: "n\1" chaieb@26126: shows "\ n = card {m. 0 < m \ m < n \ coprime m n }" chaieb@26126: proof- huffman@30488: {fix m chaieb@26126: from n have "0 < m \ m < n \ coprime m n \ coprime m n \ m < n" chaieb@26126: by (cases "m=0", auto)} chaieb@26126: thus ?thesis unfolding phi_alt by auto chaieb@26126: qed chaieb@26126: chaieb@26126: lemma phi_limit: "\ n \ n" chaieb@26126: proof- chaieb@26126: have "{ m. coprime m n \ m < n} \ {0 .. m < n}"] chaieb@26126: show ?thesis unfolding phi_alt by auto chaieb@26126: qed chaieb@26126: chaieb@26126: lemma stupid[simp]: "{m. (0::nat) < m \ m < n} = {1..1" chaieb@26126: shows "\(n) \ n - 1" chaieb@26126: proof- chaieb@26126: show ?thesis huffman@30488: unfolding phi_another[OF n] finite_number_segment[of n, symmetric] chaieb@26126: by (rule card_mono[of "{m. 0 < m \ m < n}" "{m. 0 < m \ m < n \ coprime m n}"], auto) chaieb@26126: qed chaieb@26126: chaieb@26126: lemma phi_lowerbound_1_strong: assumes n: "n \ 1" chaieb@26126: shows "\(n) \ 1" chaieb@26126: proof- chaieb@26126: let ?S = "{ m. 0 < m \ m <= n \ coprime m n }" huffman@30488: from card_0_eq[of ?S] n have "\ n \ 0" unfolding phi_alt chaieb@26126: apply auto chaieb@26126: apply (cases "n=1", simp_all) chaieb@26126: apply (rule exI[where x=1], simp) chaieb@26126: done chaieb@26126: thus ?thesis by arith chaieb@26126: qed chaieb@26126: chaieb@26126: lemma phi_lowerbound_1: "2 <= n ==> 1 <= \(n)" chaieb@26126: using phi_lowerbound_1_strong[of n] by auto chaieb@26126: chaieb@26126: lemma phi_lowerbound_2: assumes n: "3 <= n" shows "2 <= \ (n)" chaieb@26126: proof- chaieb@26126: let ?S = "{ m. 0 < m \ m <= n \ coprime m n }" huffman@30488: have inS: "{1, n - 1} \ ?S" using n coprime_plus1[of "n - 1"] chaieb@26126: by (auto simp add: coprime_commute) chaieb@26126: from n have c2: "card {1, n - 1} = 2" by (auto simp add: card_insert_if) huffman@30488: from card_mono[of ?S "{1, n - 1}", simplified inS c2] show ?thesis chaieb@26126: unfolding phi_def by auto chaieb@26126: qed chaieb@26126: chaieb@26126: lemma phi_prime: "\ n = n - 1 \ n\0 \ n\1 \ prime n" chaieb@26126: proof- chaieb@26126: {assume "n=0 \ n=1" hence ?thesis by (cases "n=1", simp_all)} chaieb@26126: moreover chaieb@26126: {assume n: "n\0" "n\1" chaieb@26126: let ?S = "{m. 0 < m \ m < n}" chaieb@26126: have fS: "finite ?S" by simp chaieb@26126: let ?S' = "{m. 0 < m \ m < n \ coprime m n}" chaieb@26126: have fS':"finite ?S'" apply (rule finite_subset[of ?S' ?S]) by auto chaieb@26126: {assume H: "\ n = n - 1 \ n\0 \ n\1" huffman@30488: hence ceq: "card ?S' = card ?S" chaieb@26126: using n finite_number_segment[of n] phi_another[OF n(2)] by simp chaieb@26126: {fix m assume m: "0 < m" "m < n" "\ coprime m n" wenzelm@32960: hence mS': "m \ ?S'" by auto wenzelm@32960: have "insert m ?S' \ ?S" using m by auto wenzelm@32960: from m have "card (insert m ?S') \ card ?S" wenzelm@32960: by - (rule card_mono[of ?S "insert m ?S'"], auto) wenzelm@32960: hence False wenzelm@32960: unfolding card_insert_disjoint[of "?S'" m, OF fS' mS'] ceq wenzelm@32960: by simp } chaieb@26126: hence "\m. 0 m < n \ coprime m n" by blast chaieb@26126: hence "prime n" unfolding prime using n by (simp add: coprime_commute)} chaieb@26126: moreover chaieb@26126: {assume H: "prime n" huffman@30488: hence "?S = ?S'" unfolding prime using n wenzelm@32960: by (auto simp add: coprime_commute) chaieb@26126: hence "card ?S = card ?S'" by simp chaieb@26126: hence "\ n = n - 1" unfolding phi_another[OF n(2)] by simp} chaieb@26126: ultimately have ?thesis using n by blast} chaieb@26126: ultimately show ?thesis by (cases "n=0") blast+ chaieb@26126: qed chaieb@26126: chaieb@26126: (* Multiplicativity property. *) chaieb@26126: chaieb@26126: lemma phi_multiplicative: assumes ab: "coprime a b" chaieb@26126: shows "\ (a * b) = \ a * \ b" chaieb@26126: proof- huffman@30488: {assume "a = 0 \ b = 0 \ a = 1 \ b = 1" chaieb@26126: hence ?thesis chaieb@26126: by (cases "a=0", simp, cases "b=0", simp, cases"a=1", simp_all) } chaieb@26126: moreover chaieb@26126: {assume a: "a\0" "a\1" and b: "b\0" "b\1" chaieb@26126: hence ab0: "a*b \ 0" by simp chaieb@26126: let ?S = "\k. {m. coprime m k \ m < k}" chaieb@26126: let ?f = "\x. (x mod a, x mod b)" chaieb@26126: have eq: "?f ` (?S (a*b)) = (?S a \ ?S b)" chaieb@26126: proof- chaieb@26126: {fix x assume x:"x \ ?S (a*b)" wenzelm@32960: hence x': "coprime x (a*b)" "x < a*b" by simp_all wenzelm@32960: hence xab: "coprime x a" "coprime x b" by (simp_all add: coprime_mul_eq) wenzelm@32960: from mod_less_divisor a b have xab':"x mod a < a" "x mod b < b" by auto wenzelm@32960: from xab xab' have "?f x \ (?S a \ ?S b)" wenzelm@32960: by (simp add: coprime_mod[OF a(1)] coprime_mod[OF b(1)])} chaieb@26126: moreover chaieb@26126: {fix x y assume x: "x \ ?S a" and y: "y \ ?S b" wenzelm@32960: hence x': "coprime x a" "x < a" and y': "coprime y b" "y < b" by simp_all wenzelm@32960: from chinese_remainder_coprime_unique[OF ab a(1) b(1) x'(1) y'(1)] wenzelm@32960: obtain z where z: "coprime z (a * b)" "z < a * b" "[z = x] (mod a)" wenzelm@32960: "[z = y] (mod b)" by blast wenzelm@32960: hence "(x,y) \ ?f ` (?S (a*b))" wenzelm@32960: using y'(2) mod_less_divisor[of b y] x'(2) mod_less_divisor[of a x] wenzelm@32960: by (auto simp add: image_iff modeq_def)} chaieb@26126: ultimately show ?thesis by auto chaieb@26126: qed chaieb@26126: have finj: "inj_on ?f (?S (a*b))" chaieb@26126: unfolding inj_on_def chaieb@26126: proof(clarify) huffman@30488: fix x y assume H: "coprime x (a * b)" "x < a * b" "coprime y (a * b)" wenzelm@32960: "y < a * b" "x mod a = y mod a" "x mod b = y mod b" huffman@30488: hence cp: "coprime x a" "coprime x b" "coprime y a" "coprime y b" wenzelm@32960: by (simp_all add: coprime_mul_eq) chaieb@26126: from chinese_remainder_coprime_unique[OF ab a(1) b(1) cp(3,4)] H chaieb@26126: show "x = y" unfolding modeq_def by blast chaieb@26126: qed chaieb@26126: from card_image[OF finj, unfolded eq] have ?thesis chaieb@26126: unfolding phi_alt by simp } chaieb@26126: ultimately show ?thesis by auto chaieb@26126: qed chaieb@26126: chaieb@26126: (* Fermat's Little theorem / Fermat-Euler theorem. *) chaieb@26126: chaieb@26126: chaieb@26126: lemma nproduct_mod: chaieb@26126: assumes fS: "finite S" and n0: "n \ 0" chaieb@26126: shows "[setprod (\m. a(m) mod n) S = setprod a S] (mod n)" chaieb@26126: proof- chaieb@26126: have th1:"[1 = 1] (mod n)" by (simp add: modeq_def) chaieb@26126: from cong_mult chaieb@26126: have th3:"\x1 y1 x2 y2. chaieb@26126: [x1 = x2] (mod n) \ [y1 = y2] (mod n) \ [x1 * y1 = x2 * y2] (mod n)" chaieb@26126: by blast chaieb@26126: have th4:"\x\S. [a x mod n = a x] (mod n)" by (simp add: modeq_def) haftmann@51489: from setprod.related [where h="(\m. a(m) mod n)" and g=a, OF th1 th3 fS, OF th4] show ?thesis by (simp add: fS) chaieb@26126: qed chaieb@26126: chaieb@26126: lemma nproduct_cmul: chaieb@26126: assumes fS:"finite S" haftmann@31021: shows "setprod (\m. (c::'a::{comm_monoid_mult})* a(m)) S = c ^ (card S) * setprod a S" haftmann@57418: unfolding setprod.distrib setprod_constant[OF fS, of c] .. chaieb@26126: chaieb@26126: lemma coprime_nproduct: chaieb@26126: assumes fS: "finite S" and Sn: "\x\S. coprime n (a x)" chaieb@26126: shows "coprime n (setprod a S)" haftmann@51489: using fS by (rule finite_subset_induct) haftmann@27368: (insert Sn, auto simp add: coprime_mul) chaieb@26126: chaieb@26126: lemma fermat_little: assumes an: "coprime a n" chaieb@26126: shows "[a ^ (\ n) = 1] (mod n)" chaieb@26126: proof- chaieb@26126: {assume "n=0" hence ?thesis by simp} chaieb@26126: moreover chaieb@26126: {assume "n=1" hence ?thesis by (simp add: modeq_def)} chaieb@26126: moreover chaieb@26126: {assume nz: "n \ 0" and n1: "n \ 1" chaieb@26126: let ?S = "{m. coprime m n \ m < n}" chaieb@26126: let ?P = "\ ?S" chaieb@26126: have fS: "finite ?S" by simp chaieb@26126: have cardfS: "\ n = card ?S" unfolding phi_alt .. chaieb@26126: {fix m assume m: "m \ ?S" chaieb@26126: hence "coprime m n" by simp huffman@30488: with coprime_mul[of n a m] an have "coprime (a*m) n" wenzelm@32960: by (simp add: coprime_commute)} chaieb@26126: hence Sn: "\m\ ?S. coprime (a*m) n " by blast chaieb@26126: from coprime_nproduct[OF fS, of n "\m. m"] have nP:"coprime ?P n" chaieb@26126: by (simp add: coprime_commute) chaieb@26126: have Paphi: "[?P*a^ (\ n) = ?P*1] (mod n)" chaieb@26126: proof- hoelzl@57129: let ?h = "\m. (a * m) mod n" hoelzl@57129: hoelzl@57129: have eq0: "(\i\?S. ?h i) = (\i\?S. i)" hoelzl@57129: proof (rule setprod.reindex_bij_betw) hoelzl@57129: have "inj_on (\i. ?h i) ?S" hoelzl@57129: proof (rule inj_onI) hoelzl@57129: fix x y assume "?h x = ?h y" hoelzl@57129: then have "[a * x = a * y] (mod n)" hoelzl@57129: by (simp add: modeq_def) hoelzl@57129: moreover assume "x \ ?S" "y \ ?S" hoelzl@57129: ultimately show "x = y" hoelzl@57129: by (auto intro: cong_imp_eq cong_mult_lcancel an) hoelzl@57129: qed hoelzl@57129: moreover have "?h ` ?S = ?S" hoelzl@57129: proof safe hoelzl@57129: fix y assume "coprime y n" then show "coprime (?h y) n" hoelzl@57129: by (metis an nz coprime_commute coprime_mod coprime_mul_eq) hoelzl@57129: next hoelzl@57129: fix y assume y: "coprime y n" "y < n" hoelzl@57129: from cong_solve_unique[OF an nz] obtain x where x: "x < n" "[a * x = y] (mod n)" hoelzl@57129: by blast hoelzl@57129: then show "y \ ?h ` ?S" hoelzl@57129: using cong_coprime[OF x(2)] coprime_mul_eq[of n a x] an y x hoelzl@57129: by (intro image_eqI[of _ _ x]) (auto simp add: coprime_commute modeq_def) hoelzl@57129: qed (insert nz, simp) hoelzl@57129: ultimately show "bij_betw ?h ?S ?S" hoelzl@57129: by (simp add: bij_betw_def) chaieb@26126: qed chaieb@26126: from nproduct_mod[OF fS nz, of "op * a"] hoelzl@57129: have "[(\i\?S. a * i) = (\i\?S. ?h i)] (mod n)" wenzelm@32960: by (simp add: cong_commute) hoelzl@57129: also have "[(\i\?S. ?h i) = ?P] (mod n)" hoelzl@57129: using eq0 fS an by (simp add: setprod_def modeq_def) chaieb@26126: finally show "[?P*a^ (\ n) = ?P*1] (mod n)" haftmann@57512: unfolding cardfS mult.commute[of ?P "a^ (card ?S)"] wenzelm@32960: nproduct_cmul[OF fS, symmetric] mult_1_right by simp chaieb@26126: qed chaieb@26126: from cong_mult_lcancel[OF nP Paphi] have ?thesis . } chaieb@26126: ultimately show ?thesis by blast chaieb@26126: qed chaieb@26126: chaieb@26126: lemma fermat_little_prime: assumes p: "prime p" and ap: "coprime a p" chaieb@26126: shows "[a^ (p - 1) = 1] (mod p)" chaieb@26126: using fermat_little[OF ap] p[unfolded phi_prime[symmetric]] chaieb@26126: by simp chaieb@26126: chaieb@26126: chaieb@26126: (* Lucas's theorem. *) chaieb@26126: chaieb@26126: lemma lucas_coprime_lemma: chaieb@26126: assumes m: "m\0" and am: "[a^m = 1] (mod n)" chaieb@26126: shows "coprime a n" chaieb@26126: proof- chaieb@26126: {assume "n=1" hence ?thesis by simp} chaieb@26126: moreover chaieb@26126: {assume "n = 0" hence ?thesis using am m exp_eq_1[of a m] by simp} chaieb@26126: moreover chaieb@26126: {assume n: "n\0" "n\1" chaieb@26126: from m obtain m' where m': "m = Suc m'" by (cases m, blast+) chaieb@26126: {fix d chaieb@26126: assume d: "d dvd a" "d dvd n" huffman@30488: from n have n1: "1 < n" by arith chaieb@26126: from am mod_less[OF n1] have am1: "a^m mod n = 1" unfolding modeq_def by simp chaieb@26126: from dvd_mult2[OF d(1), of "a^m'"] have dam:"d dvd a^m" by (simp add: m') chaieb@26126: from dvd_mod_iff[OF d(2), of "a^m"] dam am1 chaieb@26126: have "d = 1" by simp } chaieb@26126: hence ?thesis unfolding coprime by auto chaieb@26126: } huffman@30488: ultimately show ?thesis by blast chaieb@26126: qed chaieb@26126: chaieb@26126: lemma lucas_weak: huffman@30488: assumes n: "n \ 2" and an:"[a^(n - 1) = 1] (mod n)" chaieb@26126: and nm: "\m. 0 m < n - 1 \ \ [a^m = 1] (mod n)" chaieb@26126: shows "prime n" chaieb@26126: proof- chaieb@26126: from n have n1: "n \ 1" "n\0" "n - 1 \ 0" "n - 1 > 0" "n - 1 < n" by arith+ chaieb@26126: from lucas_coprime_lemma[OF n1(3) an] have can: "coprime a n" . chaieb@26126: from fermat_little[OF can] have afn: "[a ^ \ n = 1] (mod n)" . chaieb@26126: {assume "\ n \ n - 1" chaieb@26126: with phi_limit_strong[OF n1(1)] phi_lowerbound_1[OF n] chaieb@26126: have c:"\ n > 0 \ \ n < n - 1" by arith chaieb@26126: from nm[rule_format, OF c] afn have False ..} chaieb@26126: hence "\ n = n - 1" by blast chaieb@26126: with phi_prime[of n] n1(1,2) show ?thesis by simp chaieb@26126: qed chaieb@26126: huffman@30488: lemma nat_exists_least_iff: "(\(n::nat). P n) \ (\n. P n \ (\m < n. \ P m))" chaieb@26126: (is "?lhs \ ?rhs") chaieb@26126: proof chaieb@26126: assume ?rhs thus ?lhs by blast chaieb@26126: next chaieb@26126: assume H: ?lhs then obtain n where n: "P n" by blast chaieb@26126: let ?x = "Least P" chaieb@26126: {fix m assume m: "m < ?x" chaieb@26126: from not_less_Least[OF m] have "\ P m" .} chaieb@26126: with LeastI_ex[OF H] show ?rhs by blast chaieb@26126: qed chaieb@26126: huffman@30488: lemma nat_exists_least_iff': "(\(n::nat). P n) \ (P (Least P) \ (\m < (Least P). \ P m))" chaieb@26126: (is "?lhs \ ?rhs") chaieb@26126: proof- chaieb@26126: {assume ?rhs hence ?lhs by blast} huffman@30488: moreover chaieb@26126: { assume H: ?lhs then obtain n where n: "P n" by blast chaieb@26126: let ?x = "Least P" chaieb@26126: {fix m assume m: "m < ?x" chaieb@26126: from not_less_Least[OF m] have "\ P m" .} chaieb@26126: with LeastI_ex[OF H] have ?rhs by blast} chaieb@26126: ultimately show ?thesis by blast chaieb@26126: qed huffman@30488: chaieb@26126: lemma power_mod: "((x::nat) mod m)^n mod m = x^n mod m" chaieb@26126: proof(induct n) chaieb@26126: case 0 thus ?case by simp chaieb@26126: next huffman@30488: case (Suc n) huffman@30488: have "(x mod m)^(Suc n) mod m = ((x mod m) * (((x mod m) ^ n) mod m)) mod m" nipkow@30224: by (simp add: mod_mult_right_eq[symmetric]) chaieb@26126: also have "\ = ((x mod m) * (x^n mod m)) mod m" using Suc.hyps by simp chaieb@26126: also have "\ = x^(Suc n) mod m" nipkow@30224: by (simp add: mod_mult_left_eq[symmetric] mod_mult_right_eq[symmetric]) chaieb@26126: finally show ?case . chaieb@26126: qed chaieb@26126: chaieb@26126: lemma lucas: huffman@30488: assumes n2: "n \ 2" and an1: "[a^(n - 1) = 1] (mod n)" chaieb@26126: and pn: "\p. prime p \ p dvd n - 1 \ \ [a^((n - 1) div p) = 1] (mod n)" chaieb@26126: shows "prime n" chaieb@26126: proof- chaieb@26126: from n2 have n01: "n\0" "n\1" "n - 1 \ 0" by arith+ chaieb@26126: from mod_less_divisor[of n 1] n01 have onen: "1 mod n = 1" by simp huffman@30488: from lucas_coprime_lemma[OF n01(3) an1] cong_coprime[OF an1] chaieb@26126: have an: "coprime a n" "coprime (a^(n - 1)) n" by (simp_all add: coprime_commute) chaieb@26126: {assume H0: "\m. 0 < m \ m < n - 1 \ [a ^ m = 1] (mod n)" (is "EX m. ?P m") huffman@30488: from H0[unfolded nat_exists_least_iff[of ?P]] obtain m where chaieb@26126: m: "0 < m" "m < n - 1" "[a ^ m = 1] (mod n)" "\k ?P k" by blast huffman@30488: {assume nm1: "(n - 1) mod m > 0" huffman@30488: from mod_less_divisor[OF m(1)] have th0:"(n - 1) mod m < m" by blast chaieb@26126: let ?y = "a^ ((n - 1) div m * m)" chaieb@26126: note mdeq = mod_div_equality[of "(n - 1)" m] huffman@30488: from coprime_exp[OF an(1)[unfolded coprime_commute[of a n]], wenzelm@32960: of "(n - 1) div m * m"] huffman@30488: have yn: "coprime ?y n" by (simp add: coprime_commute) huffman@30488: have "?y mod n = (a^m)^((n - 1) div m) mod n" wenzelm@32960: by (simp add: algebra_simps power_mult) huffman@30488: also have "\ = (a^m mod n)^((n - 1) div m) mod n" wenzelm@32960: using power_mod[of "a^m" n "(n - 1) div m"] by simp huffman@30488: also have "\ = 1" using m(3)[unfolded modeq_def onen] onen wenzelm@32960: by (simp add: power_Suc0) huffman@30488: finally have th3: "?y mod n = 1" . huffman@30488: have th2: "[?y * a ^ ((n - 1) mod m) = ?y* 1] (mod n)" wenzelm@32960: using an1[unfolded modeq_def onen] onen wenzelm@32960: mod_div_equality[of "(n - 1)" m, symmetric] wenzelm@32960: by (simp add:power_add[symmetric] modeq_def th3 del: One_nat_def) chaieb@26126: from cong_mult_lcancel[of ?y n "a^((n - 1) mod m)" 1, OF yn th2] huffman@30488: have th1: "[a ^ ((n - 1) mod m) = 1] (mod n)" . huffman@30488: from m(4)[rule_format, OF th0] nm1 wenzelm@32960: less_trans[OF mod_less_divisor[OF m(1), of "n - 1"] m(2)] th1 chaieb@26126: have False by blast } chaieb@26126: hence "(n - 1) mod m = 0" by auto chaieb@26126: then have mn: "m dvd n - 1" by presburger chaieb@26126: then obtain r where r: "n - 1 = m*r" unfolding dvd_def by blast chaieb@26126: from n01 r m(2) have r01: "r\0" "r\1" by - (rule ccontr, simp)+ chaieb@26126: from prime_factor[OF r01(2)] obtain p where p: "prime p" "p dvd r" by blast chaieb@26126: hence th: "prime p \ p dvd n - 1" unfolding r by (auto intro: dvd_mult) chaieb@26126: have "(a ^ ((n - 1) div p)) mod n = (a^(m*r div p)) mod n" using r chaieb@26126: by (simp add: power_mult) chaieb@26126: also have "\ = (a^(m*(r div p))) mod n" using div_mult1_eq[of m r p] p(2)[unfolded dvd_eq_mod_eq_0] by simp chaieb@26126: also have "\ = ((a^m)^(r div p)) mod n" by (simp add: power_mult) chaieb@26126: also have "\ = ((a^m mod n)^(r div p)) mod n" using power_mod[of "a^m" "n" "r div p" ] .. chaieb@26158: also have "\ = 1" using m(3) onen by (simp add: modeq_def power_Suc0) huffman@30488: finally have "[(a ^ ((n - 1) div p))= 1] (mod n)" chaieb@26126: using onen by (simp add: modeq_def) chaieb@26126: with pn[rule_format, OF th] have False by blast} chaieb@26126: hence th: "\m. 0 < m \ m < n - 1 \ \ [a ^ m = 1] (mod n)" by blast chaieb@26126: from lucas_weak[OF n2 an1 th] show ?thesis . chaieb@26126: qed chaieb@26126: chaieb@26126: (* Definition of the order of a number mod n (0 in non-coprime case). *) chaieb@26126: chaieb@26126: definition "ord n a = (if coprime n a then Least (\d. d > 0 \ [a ^d = 1] (mod n)) else 0)" chaieb@26126: chaieb@26126: (* This has the expected properties. *) chaieb@26126: chaieb@26126: lemma coprime_ord: huffman@30488: assumes na: "coprime n a" chaieb@26126: shows "ord n a > 0 \ [a ^(ord n a) = 1] (mod n) \ (\m. 0 < m \ m < ord n a \ \ [a^ m = 1] (mod n))" chaieb@26126: proof- chaieb@26126: let ?P = "\d. 0 < d \ [a ^ d = 1] (mod n)" chaieb@26126: from euclid[of a] obtain p where p: "prime p" "a < p" by blast chaieb@26126: from na have o: "ord n a = Least ?P" by (simp add: ord_def) chaieb@26126: {assume "n=0 \ n=1" with na have "\m>0. ?P m" apply auto apply (rule exI[where x=1]) by (simp add: modeq_def)} chaieb@26126: moreover huffman@30488: {assume "n\0 \ n\1" hence n2:"n \ 2" by arith chaieb@26126: from na have na': "coprime a n" by (simp add: coprime_commute) chaieb@26126: from phi_lowerbound_1[OF n2] fermat_little[OF na'] chaieb@26126: have ex: "\m>0. ?P m" by - (rule exI[where x="\ n"], auto) } chaieb@26126: ultimately have ex: "\m>0. ?P m" by blast huffman@30488: from nat_exists_least_iff'[of ?P] ex na show ?thesis chaieb@26126: unfolding o[symmetric] by auto chaieb@26126: qed chaieb@26126: (* With the special value 0 for non-coprime case, it's more convenient. *) chaieb@26126: lemma ord_works: chaieb@26126: "[a ^ (ord n a) = 1] (mod n) \ (\m. 0 < m \ m < ord n a \ ~[a^ m = 1] (mod n))" chaieb@26126: apply (cases "coprime n a") chaieb@26126: using coprime_ord[of n a] chaieb@26126: by (blast, simp add: ord_def modeq_def) chaieb@26126: huffman@30488: lemma ord: "[a^(ord n a) = 1] (mod n)" using ord_works by blast huffman@30488: lemma ord_minimal: "0 < m \ m < ord n a \ ~[a^m = 1] (mod n)" chaieb@26126: using ord_works by blast chaieb@26126: lemma ord_eq_0: "ord n a = 0 \ ~coprime n a" wenzelm@41541: by (cases "coprime n a", simp add: coprime_ord, simp add: ord_def) chaieb@26126: chaieb@26126: lemma ord_divides: chaieb@26126: "[a ^ d = 1] (mod n) \ ord n a dvd d" (is "?lhs \ ?rhs") chaieb@26126: proof chaieb@26126: assume rh: ?rhs chaieb@26126: then obtain k where "d = ord n a * k" unfolding dvd_def by blast chaieb@26126: hence "[a ^ d = (a ^ (ord n a) mod n)^k] (mod n)" chaieb@26126: by (simp add : modeq_def power_mult power_mod) huffman@30488: also have "[(a ^ (ord n a) mod n)^k = 1] (mod n)" huffman@30488: using ord[of a n, unfolded modeq_def] chaieb@26158: by (simp add: modeq_def power_mod power_Suc0) chaieb@26126: finally show ?lhs . huffman@30488: next chaieb@26126: assume lh: ?lhs chaieb@26126: { assume H: "\ coprime n a" chaieb@26126: hence o: "ord n a = 0" by (simp add: ord_def) chaieb@26126: {assume d: "d=0" with o H have ?rhs by (simp add: modeq_def)} chaieb@26126: moreover chaieb@26126: {assume d0: "d\0" then obtain d' where d': "d = Suc d'" by (cases d, auto) huffman@30488: from H[unfolded coprime] huffman@30488: obtain p where p: "p dvd n" "p dvd a" "p \ 1" by auto huffman@30488: from lh[unfolded nat_mod] chaieb@26126: obtain q1 q2 where q12:"a ^ d + n * q1 = 1 + n * q2" by blast chaieb@26126: hence "a ^ d + n * q1 - n * q2 = 1" by simp nipkow@31952: with dvd_diff_nat [OF dvd_add [OF divides_rexp[OF p(2), of d'] dvd_mult2[OF p(1), of q1]] dvd_mult2[OF p(1), of q2]] d' have "p dvd 1" by simp chaieb@26126: with p(3) have False by simp chaieb@26126: hence ?rhs ..} chaieb@26126: ultimately have ?rhs by blast} chaieb@26126: moreover chaieb@26126: {assume H: "coprime n a" chaieb@26126: let ?o = "ord n a" chaieb@26126: let ?q = "d div ord n a" chaieb@26126: let ?r = "d mod ord n a" huffman@30488: from cong_exp[OF ord[of a n], of ?q] chaieb@26158: have eqo: "[(a^?o)^?q = 1] (mod n)" by (simp add: modeq_def power_Suc0) chaieb@26126: from H have onz: "?o \ 0" by (simp add: ord_eq_0) chaieb@26126: hence op: "?o > 0" by simp chaieb@26126: from mod_div_equality[of d "ord n a"] lh haftmann@57512: have "[a^(?o*?q + ?r) = 1] (mod n)" by (simp add: modeq_def mult.commute) huffman@30488: hence "[(a^?o)^?q * (a^?r) = 1] (mod n)" chaieb@26126: by (simp add: modeq_def power_mult[symmetric] power_add[symmetric]) chaieb@26126: hence th: "[a^?r = 1] (mod n)" nipkow@30224: using eqo mod_mult_left_eq[of "(a^?o)^?q" "a^?r" n] chaieb@26126: apply (simp add: modeq_def del: One_nat_def) nipkow@30224: by (simp add: mod_mult_left_eq[symmetric]) chaieb@26126: {assume r: "?r = 0" hence ?rhs by (simp add: dvd_eq_mod_eq_0)} chaieb@26126: moreover huffman@30488: {assume r: "?r \ 0" chaieb@26126: with mod_less_divisor[OF op, of d] have r0o:"?r >0 \ ?r < ?o" by simp huffman@30488: from conjunct2[OF ord_works[of a n], rule_format, OF r0o] th chaieb@26126: have ?rhs by blast} chaieb@26126: ultimately have ?rhs by blast} chaieb@26126: ultimately show ?rhs by blast chaieb@26126: qed chaieb@26126: chaieb@26126: lemma order_divides_phi: "coprime n a \ ord n a dvd \ n" chaieb@26126: using ord_divides fermat_little coprime_commute by simp huffman@30488: lemma order_divides_expdiff: chaieb@26126: assumes na: "coprime n a" chaieb@26126: shows "[a^d = a^e] (mod n) \ [d = e] (mod (ord n a))" chaieb@26126: proof- huffman@30488: {fix n a d e chaieb@26126: assume na: "coprime n a" and ed: "(e::nat) \ d" chaieb@26126: hence "\c. d = e + c" by arith chaieb@26126: then obtain c where c: "d = e + c" by arith chaieb@26126: from na have an: "coprime a n" by (simp add: coprime_commute) huffman@30488: from coprime_exp[OF na, of e] chaieb@26126: have aen: "coprime (a^e) n" by (simp add: coprime_commute) huffman@30488: from coprime_exp[OF na, of c] chaieb@26126: have acn: "coprime (a^c) n" by (simp add: coprime_commute) chaieb@26126: have "[a^d = a^e] (mod n) \ [a^(e + c) = a^(e + 0)] (mod n)" chaieb@26126: using c by simp chaieb@26126: also have "\ \ [a^e* a^c = a^e *a^0] (mod n)" by (simp add: power_add) chaieb@26126: also have "\ \ [a ^ c = 1] (mod n)" chaieb@26126: using cong_mult_lcancel_eq[OF aen, of "a^c" "a^0"] by simp chaieb@26126: also have "\ \ ord n a dvd c" by (simp only: ord_divides) chaieb@26126: also have "\ \ [e + c = e + 0] (mod ord n a)" chaieb@26126: using cong_add_lcancel_eq[of e c 0 "ord n a", simplified cong_0_divides] chaieb@26126: by simp chaieb@26126: finally have "[a^d = a^e] (mod n) \ [d = e] (mod (ord n a))" chaieb@26126: using c by simp } chaieb@26126: note th = this chaieb@26126: have "e \ d \ d \ e" by arith chaieb@26126: moreover chaieb@26126: {assume ed: "e \ d" from th[OF na ed] have ?thesis .} chaieb@26126: moreover chaieb@26126: {assume de: "d \ e" chaieb@26126: from th[OF na de] have ?thesis by (simp add: cong_commute) } chaieb@26126: ultimately show ?thesis by blast chaieb@26126: qed chaieb@26126: chaieb@26126: (* Another trivial primality characterization. *) chaieb@26126: chaieb@26126: lemma prime_prime_factor: chaieb@26126: "prime n \ n \ 1\ (\p. prime p \ p dvd n \ p = n)" chaieb@26126: proof- chaieb@26126: {assume n: "n=0 \ n=1" hence ?thesis using prime_0 two_is_prime by auto} chaieb@26126: moreover chaieb@26126: {assume n: "n\0" "n\1" chaieb@26126: {assume pn: "prime n" huffman@30488: chaieb@26126: from pn[unfolded prime_def] have "\p. prime p \ p dvd n \ p = n" wenzelm@32960: using n wenzelm@32960: apply (cases "n = 0 \ n=1",simp) wenzelm@32960: by (clarsimp, erule_tac x="p" in allE, auto)} chaieb@26126: moreover chaieb@26126: {assume H: "\p. prime p \ p dvd n \ p = n" chaieb@26126: from n have n1: "n > 1" by arith chaieb@26126: {fix m assume m: "m dvd n" "m\1" wenzelm@32960: from prime_factor[OF m(2)] obtain p where wenzelm@32960: p: "prime p" "p dvd m" by blast wenzelm@32960: from dvd_trans[OF p(2) m(1)] p(1) H have "p = n" by blast wenzelm@32960: with p(2) have "n dvd m" by simp nipkow@33657: hence "m=n" using dvd_antisym[OF m(1)] by simp } chaieb@26126: with n1 have "prime n" unfolding prime_def by auto } huffman@30488: ultimately have ?thesis using n by blast} huffman@30488: ultimately show ?thesis by auto chaieb@26126: qed chaieb@26126: chaieb@26126: lemma prime_divisor_sqrt: wenzelm@53077: "prime n \ n \ 1 \ (\d. d dvd n \ d\<^sup>2 \ n \ d = 1)" chaieb@26126: proof- huffman@30488: {assume "n=0 \ n=1" hence ?thesis using prime_0 prime_1 chaieb@26126: by (auto simp add: nat_power_eq_0_iff)} chaieb@26126: moreover chaieb@26126: {assume n: "n\0" "n\1" chaieb@26126: hence np: "n > 1" by arith wenzelm@53077: {fix d assume d: "d dvd n" "d\<^sup>2 \ n" and H: "\m. m dvd n \ m=1 \ m=n" chaieb@26126: from H d have d1n: "d = 1 \ d=n" by blast chaieb@26126: {assume dn: "d=n" wenzelm@53077: have "n\<^sup>2 > n*1" using n by (simp add: power2_eq_square) wenzelm@32960: with dn d(2) have "d=1" by simp} chaieb@26126: with d1n have "d = 1" by blast } chaieb@26126: moreover wenzelm@53077: {fix d assume d: "d dvd n" and H: "\d'. d' dvd n \ d'\<^sup>2 \ n \ d' = 1" chaieb@26126: from d n have "d \ 0" apply - apply (rule ccontr) by simp chaieb@26126: hence dp: "d > 0" by simp chaieb@26126: from d[unfolded dvd_def] obtain e where e: "n= d*e" by blast chaieb@26126: from n dp e have ep:"e > 0" by simp wenzelm@53077: have "d\<^sup>2 \ n \ e\<^sup>2 \ n" using dp ep wenzelm@32960: by (auto simp add: e power2_eq_square mult_le_cancel_left) chaieb@26126: moreover wenzelm@53077: {assume h: "d\<^sup>2 \ n" wenzelm@32960: from H[rule_format, of d] h d have "d = 1" by blast} chaieb@26126: moreover wenzelm@53077: {assume h: "e\<^sup>2 \ n" haftmann@57512: from e have "e dvd n" unfolding dvd_def by (simp add: mult.commute) wenzelm@32960: with H[rule_format, of e] h have "e=1" by simp wenzelm@32960: with e have "d = n" by simp} chaieb@26126: ultimately have "d=1 \ d=n" by blast} chaieb@26126: ultimately have ?thesis unfolding prime_def using np n(2) by blast} chaieb@26126: ultimately show ?thesis by auto chaieb@26126: qed chaieb@26126: lemma prime_prime_factor_sqrt: wenzelm@53077: "prime n \ n \ 0 \ n \ 1 \ \ (\p. prime p \ p dvd n \ p\<^sup>2 \ n)" chaieb@26126: (is "?lhs \?rhs") chaieb@26126: proof- chaieb@26126: {assume "n=0 \ n=1" hence ?thesis using prime_0 prime_1 by auto} chaieb@26126: moreover chaieb@26126: {assume n: "n\0" "n\1" chaieb@26126: {assume H: ?lhs huffman@30488: from H[unfolded prime_divisor_sqrt] n wenzelm@41541: have ?rhs wenzelm@41541: apply clarsimp wenzelm@41541: apply (erule_tac x="p" in allE) wenzelm@41541: apply simp wenzelm@41541: done chaieb@26126: } chaieb@26126: moreover chaieb@26126: {assume H: ?rhs wenzelm@53077: {fix d assume d: "d dvd n" "d\<^sup>2 \ n" "d\1" wenzelm@32960: from prime_factor[OF d(3)] wenzelm@32960: obtain p where p: "prime p" "p dvd d" by blast wenzelm@32960: from n have np: "n > 0" by arith wenzelm@32960: from d(1) n have "d \ 0" by - (rule ccontr, auto) wenzelm@32960: hence dp: "d > 0" by arith wenzelm@32960: from mult_mono[OF dvd_imp_le[OF p(2) dp] dvd_imp_le[OF p(2) dp]] d(2) wenzelm@53077: have "p\<^sup>2 \ n" unfolding power2_eq_square by arith wenzelm@32960: with H n p(1) dvd_trans[OF p(2) d(1)] have False by blast} chaieb@26126: with n prime_divisor_sqrt have ?lhs by auto} chaieb@26126: ultimately have ?thesis by blast } chaieb@26126: ultimately show ?thesis by (cases "n=0 \ n=1", auto) chaieb@26126: qed chaieb@26126: (* Pocklington theorem. *) chaieb@26126: chaieb@26126: lemma pocklington_lemma: chaieb@26126: assumes n: "n \ 2" and nqr: "n - 1 = q*r" and an: "[a^ (n - 1) = 1] (mod n)" chaieb@26126: and aq:"\p. prime p \ p dvd q \ coprime (a^ ((n - 1) div p) - 1) n" chaieb@26126: and pp: "prime p" and pn: "p dvd n" chaieb@26126: shows "[p = 1] (mod q)" chaieb@26126: proof- chaieb@26126: from pp prime_0 prime_1 have p01: "p \ 0" "p \ 1" by - (rule ccontr, simp)+ huffman@30488: from cong_1_divides[OF an, unfolded nqr, unfolded dvd_def] chaieb@26126: obtain k where k: "a ^ (q * r) - 1 = n*k" by blast chaieb@26126: from pn[unfolded dvd_def] obtain l where l: "n = p*l" by blast chaieb@26126: {assume a0: "a = 0" chaieb@26126: hence "a^ (n - 1) = 0" using n by (simp add: power_0_left) chaieb@26126: with n an mod_less[of 1 n] have False by (simp add: power_0_left modeq_def)} chaieb@26126: hence a0: "a\0" .. wenzelm@41541: from n nqr have aqr0: "a ^ (q * r) \ 0" using a0 by simp chaieb@26126: hence "(a ^ (q * r) - 1) + 1 = a ^ (q * r)" by simp chaieb@26126: with k l have "a ^ (q * r) = p*l*k + 1" by simp haftmann@57514: hence "a ^ (r * q) + p * 0 = 1 + p * (l*k)" by (simp add: ac_simps) chaieb@26126: hence odq: "ord p (a^r) dvd q" chaieb@26126: unfolding ord_divides[symmetric] power_mult[symmetric] nat_mod by blast chaieb@26126: from odq[unfolded dvd_def] obtain d where d: "q = ord p (a^r) * d" by blast huffman@30488: {assume d1: "d \ 1" chaieb@26126: from prime_factor[OF d1] obtain P where P: "prime P" "P dvd d" by blast chaieb@26126: from d dvd_mult[OF P(2), of "ord p (a^r)"] have Pq: "P dvd q" by simp chaieb@26126: from aq P(1) Pq have caP:"coprime (a^ ((n - 1) div P) - 1) n" by blast chaieb@26126: from Pq obtain s where s: "q = P*s" unfolding dvd_def by blast chaieb@26126: have P0: "P \ 0" using P(1) prime_0 by - (rule ccontr, simp) chaieb@26126: from P(2) obtain t where t: "d = P*t" unfolding dvd_def by blast chaieb@26126: from d s t P0 have s': "ord p (a^r) * t = s" by algebra chaieb@26126: have "ord p (a^r) * t*r = r * ord p (a^r) * t" by algebra chaieb@26126: hence exps: "a^(ord p (a^r) * t*r) = ((a ^ r) ^ ord p (a^r)) ^ t" chaieb@26126: by (simp only: power_mult) huffman@30488: have "[((a ^ r) ^ ord p (a^r)) ^ t= 1^t] (mod p)" chaieb@26126: by (rule cong_exp, rule ord) huffman@30488: then have th: "[((a ^ r) ^ ord p (a^r)) ^ t= 1] (mod p)" chaieb@26158: by (simp add: power_Suc0) chaieb@26126: from cong_1_divides[OF th] exps have pd0: "p dvd a^(ord p (a^r) * t*r) - 1" by simp chaieb@26126: from nqr s s' have "(n - 1) div P = ord p (a^r) * t*r" using P0 by simp chaieb@26126: with caP have "coprime (a^(ord p (a^r) * t*r) - 1) n" by simp chaieb@26126: with p01 pn pd0 have False unfolding coprime by auto} huffman@30488: hence d1: "d = 1" by blast huffman@30488: hence o: "ord p (a^r) = q" using d by simp chaieb@26126: from pp phi_prime[of p] have phip: " \ p = p - 1" by simp chaieb@26126: {fix d assume d: "d dvd p" "d dvd a" "d \ 1" chaieb@26126: from pp[unfolded prime_def] d have dp: "d = p" by blast chaieb@26126: from n have n12:"Suc (n - 2) = n - 1" by arith chaieb@26126: with divides_rexp[OF d(2)[unfolded dp], of "n - 2"] chaieb@26126: have th0: "p dvd a ^ (n - 1)" by simp chaieb@26126: from n have n0: "n \ 0" by simp huffman@30488: from d(2) an n12[symmetric] have a0: "a \ 0" chaieb@26126: by - (rule ccontr, simp add: modeq_def) wenzelm@41541: have th1: "a^ (n - 1) \ 0" using n d(2) dp a0 by auto huffman@30488: from coprime_minus1[OF th1, unfolded coprime] chaieb@26126: dvd_trans[OF pn cong_1_divides[OF an]] th0 d(3) dp chaieb@26126: have False by auto} huffman@30488: hence cpa: "coprime p a" using coprime by auto huffman@30488: from coprime_exp[OF cpa, of r] coprime_commute chaieb@26126: have arp: "coprime (a^r) p" by blast chaieb@26126: from fermat_little[OF arp, simplified ord_divides] o phip chaieb@26126: have "q dvd (p - 1)" by simp chaieb@26126: then obtain d where d:"p - 1 = q * d" unfolding dvd_def by blast chaieb@26126: from prime_0 pp have p0:"p \ 0" by - (rule ccontr, auto) chaieb@26126: from p0 d have "p + q * 0 = 1 + q * d" by simp chaieb@26126: with nat_mod[of p 1 q, symmetric] chaieb@26126: show ?thesis by blast chaieb@26126: qed chaieb@26126: chaieb@26126: lemma pocklington: wenzelm@53077: assumes n: "n \ 2" and nqr: "n - 1 = q*r" and sqr: "n \ q\<^sup>2" chaieb@26126: and an: "[a^ (n - 1) = 1] (mod n)" chaieb@26126: and aq:"\p. prime p \ p dvd q \ coprime (a^ ((n - 1) div p) - 1) n" chaieb@26126: shows "prime n" chaieb@26126: unfolding prime_prime_factor_sqrt[of n] chaieb@26126: proof- wenzelm@53015: let ?ths = "n \ 0 \ n \ 1 \ \ (\p. prime p \ p dvd n \ p\<^sup>2 \ n)" chaieb@26126: from n have n01: "n\0" "n\1" by arith+ wenzelm@53077: {fix p assume p: "prime p" "p dvd n" "p\<^sup>2 \ n" chaieb@26126: from p(3) sqr have "p^(Suc 1) \ q^(Suc 1)" by (simp add: power2_eq_square) chaieb@26126: hence pq: "p \ q" unfolding exp_mono_le . chaieb@26126: from pocklington_lemma[OF n nqr an aq p(1,2)] cong_1_divides chaieb@26126: have th: "q dvd p - 1" by blast chaieb@26126: have "p - 1 \ 0"using prime_ge_2[OF p(1)] by arith chaieb@26126: with divides_ge[OF th] pq have False by arith } chaieb@26126: with n01 show ?ths by blast chaieb@26126: qed chaieb@26126: chaieb@26126: (* Variant for application, to separate the exponentiation. *) chaieb@26126: lemma pocklington_alt: wenzelm@53077: assumes n: "n \ 2" and nqr: "n - 1 = q*r" and sqr: "n \ q\<^sup>2" chaieb@26126: and an: "[a^ (n - 1) = 1] (mod n)" chaieb@26126: and aq:"\p. prime p \ p dvd q \ (\b. [a^((n - 1) div p) = b] (mod n) \ coprime (b - 1) n)" chaieb@26126: shows "prime n" chaieb@26126: proof- chaieb@26126: {fix p assume p: "prime p" "p dvd q" huffman@30488: from aq[rule_format] p obtain b where chaieb@26126: b: "[a^((n - 1) div p) = b] (mod n)" "coprime (b - 1) n" by blast chaieb@26126: {assume a0: "a=0" chaieb@26126: from n an have "[0 = 1] (mod n)" unfolding a0 power_0_left by auto chaieb@26126: hence False using n by (simp add: modeq_def dvd_eq_mod_eq_0[symmetric])} chaieb@26126: hence a0: "a\ 0" .. chaieb@26126: hence a1: "a \ 1" by arith chaieb@26126: from one_le_power[OF a1] have ath: "1 \ a ^ ((n - 1) div p)" . chaieb@26126: {assume b0: "b = 0" huffman@30488: from p(2) nqr have "(n - 1) mod p = 0" wenzelm@32960: apply (simp only: dvd_eq_mod_eq_0[symmetric]) by (rule dvd_mult2, simp) huffman@30488: with mod_div_equality[of "n - 1" p] huffman@30488: have "(n - 1) div p * p= n - 1" by auto chaieb@26126: hence eq: "(a^((n - 1) div p))^p = a^(n - 1)" wenzelm@32960: by (simp only: power_mult[symmetric]) chaieb@26126: from prime_ge_2[OF p(1)] have pS: "Suc (p - 1) = p" by arith chaieb@26126: from b(1) have d: "n dvd a^((n - 1) div p)" unfolding b0 cong_0_divides . chaieb@26126: from divides_rexp[OF d, of "p - 1"] pS eq cong_divides[OF an] n chaieb@26126: have False by simp} huffman@30488: then have b0: "b \ 0" .. huffman@30488: hence b1: "b \ 1" by arith chaieb@26126: from cong_coprime[OF cong_sub[OF b(1) cong_refl[of 1] ath b1]] b(2) nqr chaieb@26126: have "coprime (a ^ ((n - 1) div p) - 1) n" by (simp add: coprime_commute)} huffman@30488: hence th: "\p. prime p \ p dvd q \ coprime (a ^ ((n - 1) div p) - 1) n " chaieb@26126: by blast chaieb@26126: from pocklington[OF n nqr sqr an th] show ?thesis . chaieb@26126: qed chaieb@26126: chaieb@26126: (* Prime factorizations. *) chaieb@26126: chaieb@26126: definition "primefact ps n = (foldr op * ps 1 = n \ (\p\ set ps. prime p))" chaieb@26126: chaieb@26126: lemma primefact: assumes n: "n \ 0" chaieb@26126: shows "\ps. primefact ps n" chaieb@26126: using n chaieb@26126: proof(induct n rule: nat_less_induct) chaieb@26126: fix n assume H: "\m 0 \ (\ps. primefact ps m)" and n: "n\0" chaieb@26126: let ?ths = "\ps. primefact ps n" huffman@30488: {assume "n = 1" chaieb@26126: hence "primefact [] n" by (simp add: primefact_def) chaieb@26126: hence ?ths by blast } chaieb@26126: moreover chaieb@26126: {assume n1: "n \ 1" chaieb@26126: with n have n2: "n \ 2" by arith chaieb@26126: from prime_factor[OF n1] obtain p where p: "prime p" "p dvd n" by blast chaieb@26126: from p(2) obtain m where m: "n = p*m" unfolding dvd_def by blast chaieb@26126: from n m have m0: "m > 0" "m\0" by auto chaieb@26126: from prime_ge_2[OF p(1)] have "1 < p" by arith chaieb@26126: with m0 m have mn: "m < n" by auto chaieb@26126: from H[rule_format, OF mn m0(2)] obtain ps where ps: "primefact ps m" .. chaieb@26126: from ps m p(1) have "primefact (p#ps) n" by (simp add: primefact_def) chaieb@26126: hence ?ths by blast} chaieb@26126: ultimately show ?ths by blast chaieb@26126: qed chaieb@26126: huffman@30488: lemma primefact_contains: chaieb@26126: assumes pf: "primefact ps n" and p: "prime p" and pn: "p dvd n" chaieb@26126: shows "p \ set ps" chaieb@26126: using pf p pn chaieb@26126: proof(induct ps arbitrary: p n) chaieb@26126: case Nil thus ?case by (auto simp add: primefact_def) chaieb@26126: next chaieb@26126: case (Cons q qs p n) huffman@30488: from Cons.prems[unfolded primefact_def] chaieb@26126: have q: "prime q" "q * foldr op * qs 1 = n" "\p \set qs. prime p" and p: "prime p" "p dvd q * foldr op * qs 1" by simp_all chaieb@26126: {assume "p dvd q" chaieb@26126: with p(1) q(1) have "p = q" unfolding prime_def by auto chaieb@26126: hence ?case by simp} chaieb@26126: moreover chaieb@26126: { assume h: "p dvd foldr op * qs 1" huffman@30488: from q(3) have pqs: "primefact qs (foldr op * qs 1)" chaieb@26126: by (simp add: primefact_def) chaieb@26126: from Cons.hyps[OF pqs p(1) h] have ?case by simp} chaieb@26126: ultimately show ?case using prime_divprod[OF p] by blast chaieb@26126: qed chaieb@26126: haftmann@37602: lemma primefact_variant: "primefact ps n \ foldr op * ps 1 = n \ list_all prime ps" haftmann@37602: by (auto simp add: primefact_def list_all_iff) chaieb@26126: chaieb@26126: (* Variant of Lucas theorem. *) chaieb@26126: chaieb@26126: lemma lucas_primefact: huffman@30488: assumes n: "n \ 2" and an: "[a^(n - 1) = 1] (mod n)" huffman@30488: and psn: "foldr op * ps 1 = n - 1" chaieb@26126: and psp: "list_all (\p. prime p \ \ [a^((n - 1) div p) = 1] (mod n)) ps" chaieb@26126: shows "prime n" chaieb@26126: proof- chaieb@26126: {fix p assume p: "prime p" "p dvd n - 1" "[a ^ ((n - 1) div p) = 1] (mod n)" huffman@30488: from psn psp have psn1: "primefact ps (n - 1)" chaieb@26126: by (auto simp add: list_all_iff primefact_variant) chaieb@26126: from p(3) primefact_contains[OF psn1 p(1,2)] psp chaieb@26126: have False by (induct ps, auto)} chaieb@26126: with lucas[OF n an] show ?thesis by blast chaieb@26126: qed chaieb@26126: chaieb@26126: (* Variant of Pocklington theorem. *) chaieb@26126: chaieb@26126: lemma mod_le: assumes n: "n \ (0::nat)" shows "m mod n \ m" chaieb@26126: proof- chaieb@26126: from mod_div_equality[of m n] huffman@30488: have "\x. x + m mod n = m" by blast chaieb@26126: then show ?thesis by auto chaieb@26126: qed huffman@30488: chaieb@26126: chaieb@26126: lemma pocklington_primefact: wenzelm@53077: assumes n: "n \ 2" and qrn: "q*r = n - 1" and nq2: "n \ q\<^sup>2" huffman@30488: and arnb: "(a^r) mod n = b" and psq: "foldr op * ps 1 = q" chaieb@26126: and bqn: "(b^q) mod n = 1" chaieb@26126: and psp: "list_all (\p. prime p \ coprime ((b^(q div p)) mod n - 1) n) ps" chaieb@26126: shows "prime n" chaieb@26126: proof- chaieb@26126: from bqn psp qrn chaieb@26126: have bqn: "a ^ (n - 1) mod n = 1" huffman@30488: and psp: "list_all (\p. prime p \ coprime (a^(r *(q div p)) mod n - 1) n) ps" unfolding arnb[symmetric] power_mod nipkow@29667: by (simp_all add: power_mult[symmetric] algebra_simps) chaieb@26126: from n have n0: "n > 0" by arith chaieb@26126: from mod_div_equality[of "a^(n - 1)" n] chaieb@26126: mod_less_divisor[OF n0, of "a^(n - 1)"] huffman@30488: have an1: "[a ^ (n - 1) = 1] (mod n)" chaieb@26126: unfolding nat_mod bqn chaieb@26126: apply - chaieb@26126: apply (rule exI[where x="0"]) chaieb@26126: apply (rule exI[where x="a^(n - 1) div n"]) nipkow@29667: by (simp add: algebra_simps) chaieb@26126: {fix p assume p: "prime p" "p dvd q" chaieb@26126: from psp psq have pfpsq: "primefact ps q" chaieb@26126: by (auto simp add: primefact_variant list_all_iff) huffman@30488: from psp primefact_contains[OF pfpsq p] chaieb@26126: have p': "coprime (a ^ (r * (q div p)) mod n - 1) n" chaieb@26126: by (simp add: list_all_iff) chaieb@26126: from prime_ge_2[OF p(1)] have p01: "p \ 0" "p \ 1" "p =Suc(p - 1)" by arith+ huffman@30488: from div_mult1_eq[of r q p] p(2) chaieb@26126: have eq1: "r* (q div p) = (n - 1) div p" haftmann@57512: unfolding qrn[symmetric] dvd_eq_mod_eq_0 by (simp add: mult.commute) chaieb@26126: have ath: "\a (b::nat). a <= b \ a \ 0 ==> 1 <= a \ 1 <= b" by arith chaieb@26126: from n0 have n00: "n \ 0" by arith chaieb@26126: from mod_le[OF n00] chaieb@26126: have th10: "a ^ ((n - 1) div p) mod n \ a ^ ((n - 1) div p)" . chaieb@26126: {assume "a ^ ((n - 1) div p) mod n = 0" chaieb@26126: then obtain s where s: "a ^ ((n - 1) div p) = n*s" wenzelm@32960: unfolding mod_eq_0_iff by blast chaieb@26126: hence eq0: "(a^((n - 1) div p))^p = (n*s)^p" by simp chaieb@26126: from qrn[symmetric] have qn1: "q dvd n - 1" unfolding dvd_def by auto chaieb@26126: from dvd_trans[OF p(2) qn1] div_mod_equality'[of "n - 1" p] huffman@30488: have npp: "(n - 1) div p * p = n - 1" by (simp add: dvd_eq_mod_eq_0) chaieb@26126: with eq0 have "a^ (n - 1) = (n*s)^p" wenzelm@32960: by (simp add: power_mult[symmetric]) chaieb@26126: hence "1 = (n*s)^(Suc (p - 1)) mod n" using bqn p01 by simp haftmann@57512: also have "\ = 0" by (simp add: mult.assoc) chaieb@26126: finally have False by simp } huffman@30488: then have th11: "a ^ ((n - 1) div p) mod n \ 0" by auto huffman@30488: have th1: "[a ^ ((n - 1) div p) mod n = a ^ ((n - 1) div p)] (mod n)" huffman@30488: unfolding modeq_def by simp chaieb@26126: from cong_sub[OF th1 cong_refl[of 1]] ath[OF th10 th11] chaieb@26126: have th: "[a ^ ((n - 1) div p) mod n - 1 = a ^ ((n - 1) div p) - 1] (mod n)" huffman@30488: by blast huffman@30488: from cong_coprime[OF th] p'[unfolded eq1] chaieb@26126: have "coprime (a ^ ((n - 1) div p) - 1) n" by (simp add: coprime_commute) } chaieb@26126: with pocklington[OF n qrn[symmetric] nq2 an1] huffman@30488: show ?thesis by blast chaieb@26126: qed chaieb@26126: chaieb@26126: end