src/HOL/Old_Number_Theory/Pocklington.thy
author wenzelm
Fri, 06 Aug 2010 12:37:00 +0200
changeset 38159 e9b4835a54ee
parent 37602 501b0cae5aa8
child 41541 1fa4725c4656
permissions -rw-r--r--
modernized specifications; tuned headers;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
38159
e9b4835a54ee modernized specifications;
wenzelm
parents: 37602
diff changeset
     1
(*  Title:      HOL/Old_Number_Theory/Pocklington.thy
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
     2
    Author:     Amine Chaieb
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
     3
*)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
     4
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
     5
header {* Pocklington's Theorem for Primes *}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
     6
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
     7
theory Pocklington
38159
e9b4835a54ee modernized specifications;
wenzelm
parents: 37602
diff changeset
     8
imports Primes
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
     9
begin
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    10
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    11
definition modeq:: "nat => nat => nat => bool"    ("(1[_ = _] '(mod _'))")
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    12
  where "[a = b] (mod p) == ((a mod p) = (b mod p))"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    13
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    14
definition modneq:: "nat => nat => nat => bool"    ("(1[_ \<noteq> _] '(mod _'))")
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    15
  where "[a \<noteq> b] (mod p) == ((a mod p) \<noteq> (b mod p))"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    16
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    17
lemma modeq_trans:
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    18
  "\<lbrakk> [a = b] (mod p); [b = c] (mod p) \<rbrakk> \<Longrightarrow> [a = c] (mod p)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    19
  by (simp add:modeq_def)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    20
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    21
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    22
lemma nat_mod_lemma: assumes xyn: "[x = y] (mod n)" and xy:"y \<le> x"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    23
  shows "\<exists>q. x = y + n * q"
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27487
diff changeset
    24
using xyn xy unfolding modeq_def using nat_mod_eq_lemma by blast
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    25
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
    26
lemma nat_mod[algebra]: "[x = y] (mod n) \<longleftrightarrow> (\<exists>q1 q2. x + n * q1 = y + n * q2)"
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27487
diff changeset
    27
unfolding modeq_def nat_mod_eq_iff ..
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    28
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    29
(* Lemmas about previously defined terms.                                    *)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    30
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
    31
lemma prime: "prime p \<longleftrightarrow> p \<noteq> 0 \<and> p\<noteq>1 \<and> (\<forall>m. 0 < m \<and> m < p \<longrightarrow> coprime p m)"
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
    32
  (is "?lhs \<longleftrightarrow> ?rhs")
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    33
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    34
  {assume "p=0 \<or> p=1" hence ?thesis using prime_0 prime_1 by (cases "p=0", simp_all)}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    35
  moreover
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    36
  {assume p0: "p\<noteq>0" "p\<noteq>1"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    37
    {assume H: "?lhs"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    38
      {fix m assume m: "m > 0" "m < p"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
    39
        {assume "m=1" hence "coprime p m" by simp}
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
    40
        moreover
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
    41
        {assume "p dvd m" hence "p \<le> m" using dvd_imp_le m by blast with m(2)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
    42
          have "coprime p m" by simp}
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
    43
        ultimately have "coprime p m" using prime_coprime[OF H, of m] by blast}
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    44
      hence ?rhs using p0 by auto}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    45
    moreover
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    46
    { assume H: "\<forall>m. 0 < m \<and> m < p \<longrightarrow> coprime p m"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    47
      from prime_factor[OF p0(2)] obtain q where q: "prime q" "q dvd p" by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    48
      from prime_ge_2[OF q(1)] have q0: "q > 0" by arith
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    49
      from dvd_imp_le[OF q(2)] p0 have qp: "q \<le> p" by arith
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    50
      {assume "q = p" hence ?lhs using q(1) by blast}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    51
      moreover
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    52
      {assume "q\<noteq>p" with qp have qplt: "q < p" by arith
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
    53
        from H[rule_format, of q] qplt q0 have "coprime p q" by arith
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
    54
        with coprime_prime[of p q q] q have False by simp hence ?lhs by blast}
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    55
      ultimately have ?lhs by blast}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    56
    ultimately have ?thesis by blast}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    57
  ultimately show ?thesis  by (cases"p=0 \<or> p=1", auto)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    58
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    59
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    60
lemma finite_number_segment: "card { m. 0 < m \<and> m < n } = n - 1"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    61
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    62
  have "{ m. 0 < m \<and> m < n } = {1..<n}" by auto
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    63
  thus ?thesis by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    64
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    65
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    66
lemma coprime_mod: assumes n: "n \<noteq> 0" shows "coprime (a mod n) n \<longleftrightarrow> coprime a n"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    67
  using n dvd_mod_iff[of _ n a] by (auto simp add: coprime)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    68
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    69
(* Congruences.                                                              *)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    70
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
    71
lemma cong_mod_01[simp,presburger]:
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    72
  "[x = y] (mod 0) \<longleftrightarrow> x = y" "[x = y] (mod 1)" "[x = 0] (mod n) \<longleftrightarrow> n dvd x"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    73
  by (simp_all add: modeq_def, presburger)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    74
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
    75
lemma cong_sub_cases:
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    76
  "[x = y] (mod n) \<longleftrightarrow> (if x <= y then [y - x = 0] (mod n) else [x - y = 0] (mod n))"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    77
apply (auto simp add: nat_mod)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    78
apply (rule_tac x="q2" in exI)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    79
apply (rule_tac x="q1" in exI, simp)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    80
apply (rule_tac x="q2" in exI)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    81
apply (rule_tac x="q1" in exI, simp)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    82
apply (rule_tac x="q1" in exI)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    83
apply (rule_tac x="q2" in exI, simp)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    84
apply (rule_tac x="q1" in exI)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    85
apply (rule_tac x="q2" in exI, simp)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    86
done
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    87
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    88
lemma cong_mult_lcancel: assumes an: "coprime a n" and axy:"[a * x = a * y] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    89
  shows "[x = y] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    90
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    91
  {assume "a = 0" with an axy coprime_0'[of n] have ?thesis by (simp add: modeq_def) }
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    92
  moreover
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    93
  {assume az: "a\<noteq>0"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    94
    {assume xy: "x \<le> y" hence axy': "a*x \<le> a*y" by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    95
      with axy cong_sub_cases[of "a*x" "a*y" n]  have "[a*(y - x) = 0] (mod n)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
    96
        by (simp only: if_True diff_mult_distrib2)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
    97
      hence th: "n dvd a*(y -x)" by simp
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
    98
      from coprime_divprod[OF th] an have "n dvd y - x"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
    99
        by (simp add: coprime_commute)
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   100
      hence ?thesis using xy cong_sub_cases[of x y n] by simp}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   101
    moreover
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   102
    {assume H: "\<not>x \<le> y" hence xy: "y \<le> x"  by arith
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   103
      from H az have axy': "\<not> a*x \<le> a*y" by auto
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   104
      with axy H cong_sub_cases[of "a*x" "a*y" n]  have "[a*(x - y) = 0] (mod n)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   105
        by (simp only: if_False diff_mult_distrib2)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   106
      hence th: "n dvd a*(x - y)" by simp
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   107
      from coprime_divprod[OF th] an have "n dvd x - y"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   108
        by (simp add: coprime_commute)
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   109
      hence ?thesis using xy cong_sub_cases[of x y n] by simp}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   110
    ultimately have ?thesis by blast}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   111
  ultimately show ?thesis by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   112
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   113
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   114
lemma cong_mult_rcancel: assumes an: "coprime a n" and axy:"[x*a = y*a] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   115
  shows "[x = y] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   116
  using cong_mult_lcancel[OF an axy[unfolded mult_commute[of _a]]] .
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   117
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   118
lemma cong_refl: "[x = x] (mod n)" by (simp add: modeq_def)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   119
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   120
lemma eq_imp_cong: "a = b \<Longrightarrow> [a = b] (mod n)" by (simp add: cong_refl)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   121
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   122
lemma cong_commute: "[x = y] (mod n) \<longleftrightarrow> [y = x] (mod n)"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   123
  by (auto simp add: modeq_def)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   124
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   125
lemma cong_trans[trans]: "[x = y] (mod n) \<Longrightarrow> [y = z] (mod n) \<Longrightarrow> [x = z] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   126
  by (simp add: modeq_def)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   127
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   128
lemma cong_add: assumes xx': "[x = x'] (mod n)" and yy':"[y = y'] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   129
  shows "[x + y = x' + y'] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   130
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   131
  have "(x + y) mod n = (x mod n + y mod n) mod n"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   132
    by (simp add: mod_add_left_eq[of x y n] mod_add_right_eq[of "x mod n" y n])
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   133
  also have "\<dots> = (x' mod n + y' mod n) mod n" using xx' yy' modeq_def by simp
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   134
  also have "\<dots> = (x' + y') mod n"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   135
    by (simp add: mod_add_left_eq[of x' y' n] mod_add_right_eq[of "x' mod n" y' n])
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   136
  finally show ?thesis unfolding modeq_def .
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   137
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   138
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   139
lemma cong_mult: assumes xx': "[x = x'] (mod n)" and yy':"[y = y'] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   140
  shows "[x * y = x' * y'] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   141
proof-
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   142
  have "(x * y) mod n = (x mod n) * (y mod n) mod n"
30224
79136ce06bdb removed and renamed redundant lemmas
nipkow
parents: 30042
diff changeset
   143
    by (simp add: mod_mult_left_eq[of x y n] mod_mult_right_eq[of "x mod n" y n])
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   144
  also have "\<dots> = (x' mod n) * (y' mod n) mod n" using xx'[unfolded modeq_def] yy'[unfolded modeq_def] by simp
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   145
  also have "\<dots> = (x' * y') mod n"
30224
79136ce06bdb removed and renamed redundant lemmas
nipkow
parents: 30042
diff changeset
   146
    by (simp add: mod_mult_left_eq[of x' y' n] mod_mult_right_eq[of "x' mod n" y' n])
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   147
  finally show ?thesis unfolding modeq_def .
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   148
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   149
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   150
lemma cong_exp: "[x = y] (mod n) \<Longrightarrow> [x^k = y^k] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   151
  by (induct k, auto simp add: cong_refl cong_mult)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   152
lemma cong_sub: assumes xx': "[x = x'] (mod n)" and yy': "[y = y'] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   153
  and yx: "y <= x" and yx': "y' <= x'"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   154
  shows "[x - y = x' - y'] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   155
proof-
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   156
  { fix x a x' a' y b y' b'
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   157
    have "(x::nat) + a = x' + a' \<Longrightarrow> y + b = y' + b' \<Longrightarrow> y <= x \<Longrightarrow> y' <= x'
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   158
      \<Longrightarrow> (x - y) + (a + b') = (x' - y') + (a' + b)" by arith}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   159
  note th = this
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   160
  from xx' yy' obtain q1 q2 q1' q2' where q12: "x + n*q1 = x'+n*q2"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   161
    and q12': "y + n*q1' = y'+n*q2'" unfolding nat_mod by blast+
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   162
  from th[OF q12 q12' yx yx']
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   163
  have "(x - y) + n*(q1 + q2') = (x' - y') + n*(q2 + q1')"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   164
    by (simp add: right_distrib)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   165
  thus ?thesis unfolding nat_mod by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   166
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   167
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   168
lemma cong_mult_lcancel_eq: assumes an: "coprime a n"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   169
  shows "[a * x = a * y] (mod n) \<longleftrightarrow> [x = y] (mod n)" (is "?lhs \<longleftrightarrow> ?rhs")
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   170
proof
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   171
  assume H: "?rhs" from cong_mult[OF cong_refl[of a n] H] show ?lhs .
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   172
next
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   173
  assume H: "?lhs" hence H': "[x*a = y*a] (mod n)" by (simp add: mult_commute)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   174
  from cong_mult_rcancel[OF an H'] show ?rhs  .
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   175
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   176
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   177
lemma cong_mult_rcancel_eq: assumes an: "coprime a n"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   178
  shows "[x * a = y * a] (mod n) \<longleftrightarrow> [x = y] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   179
using cong_mult_lcancel_eq[OF an, of x y] by (simp add: mult_commute)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   180
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   181
lemma cong_add_lcancel_eq: "[a + x = a + y] (mod n) \<longleftrightarrow> [x = y] (mod n)"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   182
  by (simp add: nat_mod)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   183
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   184
lemma cong_add_rcancel_eq: "[x + a = y + a] (mod n) \<longleftrightarrow> [x = y] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   185
  by (simp add: nat_mod)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   186
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   187
lemma cong_add_rcancel: "[x + a = y + a] (mod n) \<Longrightarrow> [x = y] (mod n)"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   188
  by (simp add: nat_mod)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   189
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   190
lemma cong_add_lcancel: "[a + x = a + y] (mod n) \<Longrightarrow> [x = y] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   191
  by (simp add: nat_mod)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   192
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   193
lemma cong_add_lcancel_eq_0: "[a + x = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   194
  by (simp add: nat_mod)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   195
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   196
lemma cong_add_rcancel_eq_0: "[x + a = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   197
  by (simp add: nat_mod)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   198
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   199
lemma cong_imp_eq: assumes xn: "x < n" and yn: "y < n" and xy: "[x = y] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   200
  shows "x = y"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   201
  using xy[unfolded modeq_def mod_less[OF xn] mod_less[OF yn]] .
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   202
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   203
lemma cong_divides_modulus: "[x = y] (mod m) \<Longrightarrow> n dvd m ==> [x = y] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   204
  apply (auto simp add: nat_mod dvd_def)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   205
  apply (rule_tac x="k*q1" in exI)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   206
  apply (rule_tac x="k*q2" in exI)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   207
  by simp
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   208
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   209
lemma cong_0_divides: "[x = 0] (mod n) \<longleftrightarrow> n dvd x" by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   210
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   211
lemma cong_1_divides:"[x = 1] (mod n) ==> n dvd x - 1"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   212
  apply (cases "x\<le>1", simp_all)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   213
  using cong_sub_cases[of x 1 n] by auto
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   214
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   215
lemma cong_divides: "[x = y] (mod n) \<Longrightarrow> n dvd x \<longleftrightarrow> n dvd y"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   216
apply (auto simp add: nat_mod dvd_def)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   217
apply (rule_tac x="k + q1 - q2" in exI, simp add: add_mult_distrib2 diff_mult_distrib2)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   218
apply (rule_tac x="k + q2 - q1" in exI, simp add: add_mult_distrib2 diff_mult_distrib2)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   219
done
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   220
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   221
lemma cong_coprime: assumes xy: "[x = y] (mod n)"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   222
  shows "coprime n x \<longleftrightarrow> coprime n y"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   223
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   224
  {assume "n=0" hence ?thesis using xy by simp}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   225
  moreover
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   226
  {assume nz: "n \<noteq> 0"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   227
  have "coprime n x \<longleftrightarrow> coprime (x mod n) n"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   228
    by (simp add: coprime_mod[OF nz, of x] coprime_commute[of n x])
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   229
  also have "\<dots> \<longleftrightarrow> coprime (y mod n) n" using xy[unfolded modeq_def] by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   230
  also have "\<dots> \<longleftrightarrow> coprime y n" by (simp add: coprime_mod[OF nz, of y])
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   231
  finally have ?thesis by (simp add: coprime_commute) }
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   232
ultimately show ?thesis by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   233
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   234
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   235
lemma cong_mod: "~(n = 0) \<Longrightarrow> [a mod n = a] (mod n)" by (simp add: modeq_def)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   236
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   237
lemma mod_mult_cong: "~(a = 0) \<Longrightarrow> ~(b = 0)
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   238
  \<Longrightarrow> [x mod (a * b) = y] (mod a) \<longleftrightarrow> [x = y] (mod a)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   239
  by (simp add: modeq_def mod_mult2_eq mod_add_left_eq)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   240
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   241
lemma cong_mod_mult: "[x = y] (mod n) \<Longrightarrow> m dvd n \<Longrightarrow> [x = y] (mod m)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   242
  apply (auto simp add: nat_mod dvd_def)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   243
  apply (rule_tac x="k*q1" in exI)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   244
  apply (rule_tac x="k*q2" in exI, simp)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   245
  done
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   246
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   247
(* Some things when we know more about the order.                            *)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   248
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   249
lemma cong_le: "y <= x \<Longrightarrow> [x = y] (mod n) \<longleftrightarrow> (\<exists>q. x = q * n + y)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   250
  using nat_mod_lemma[of x y n]
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   251
  apply auto
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   252
  apply (simp add: nat_mod)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   253
  apply (rule_tac x="q" in exI)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   254
  apply (rule_tac x="q + q" in exI)
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28854
diff changeset
   255
  by (auto simp: algebra_simps)
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   256
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   257
lemma cong_to_1: "[a = 1] (mod n) \<longleftrightarrow> a = 0 \<and> n = 1 \<or> (\<exists>m. a = 1 + m * n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   258
proof-
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   259
  {assume "n = 0 \<or> n = 1\<or> a = 0 \<or> a = 1" hence ?thesis
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   260
      apply (cases "n=0", simp_all add: cong_commute)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   261
      apply (cases "n=1", simp_all add: cong_commute modeq_def)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   262
      apply arith
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   263
      by (cases "a=1", simp_all add: modeq_def cong_commute)}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   264
  moreover
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   265
  {assume n: "n\<noteq>0" "n\<noteq>1" and a:"a\<noteq>0" "a \<noteq> 1" hence a': "a \<ge> 1" by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   266
    hence ?thesis using cong_le[OF a', of n] by auto }
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   267
  ultimately show ?thesis by auto
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   268
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   269
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   270
(* Some basic theorems about solving congruences.                            *)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   271
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   272
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   273
lemma cong_solve: assumes an: "coprime a n" shows "\<exists>x. [a * x = b] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   274
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   275
  {assume "a=0" hence ?thesis using an by (simp add: modeq_def)}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   276
  moreover
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   277
  {assume az: "a\<noteq>0"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   278
  from bezout_add_strong[OF az, of n]
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   279
  obtain d x y where dxy: "d dvd a" "d dvd n" "a*x = n*y + d" by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   280
  from an[unfolded coprime, rule_format, of d] dxy(1,2) have d1: "d = 1" by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   281
  hence "a*x*b = (n*y + 1)*b" using dxy(3) by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   282
  hence "a*(x*b) = n*(y*b) + b" by algebra
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   283
  hence "a*(x*b) mod n = (n*(y*b) + b) mod n" by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   284
  hence "a*(x*b) mod n = b mod n" by (simp add: mod_add_left_eq)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   285
  hence "[a*(x*b) = b] (mod n)" unfolding modeq_def .
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   286
  hence ?thesis by blast}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   287
ultimately  show ?thesis by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   288
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   289
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   290
lemma cong_solve_unique: assumes an: "coprime a n" and nz: "n \<noteq> 0"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   291
  shows "\<exists>!x. x < n \<and> [a * x = b] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   292
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   293
  let ?P = "\<lambda>x. x < n \<and> [a * x = b] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   294
  from cong_solve[OF an] obtain x where x: "[a*x = b] (mod n)" by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   295
  let ?x = "x mod n"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   296
  from x have th: "[a * ?x = b] (mod n)"
30224
79136ce06bdb removed and renamed redundant lemmas
nipkow
parents: 30042
diff changeset
   297
    by (simp add: modeq_def mod_mult_right_eq[of a x n])
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   298
  from mod_less_divisor[ of n x] nz th have Px: "?P ?x" by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   299
  {fix y assume Py: "y < n" "[a * y = b] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   300
    from Py(2) th have "[a * y = a*?x] (mod n)" by (simp add: modeq_def)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   301
    hence "[y = ?x] (mod n)" by (simp add: cong_mult_lcancel_eq[OF an])
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   302
    with mod_less[OF Py(1)] mod_less_divisor[ of n x] nz
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   303
    have "y = ?x" by (simp add: modeq_def)}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   304
  with Px show ?thesis by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   305
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   306
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   307
lemma cong_solve_unique_nontrivial:
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   308
  assumes p: "prime p" and pa: "coprime p a" and x0: "0 < x" and xp: "x < p"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   309
  shows "\<exists>!y. 0 < y \<and> y < p \<and> [x * y = a] (mod p)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   310
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   311
  from p have p1: "p > 1" using prime_ge_2[OF p] by arith
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   312
  hence p01: "p \<noteq> 0" "p \<noteq> 1" by arith+
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   313
  from pa have ap: "coprime a p" by (simp add: coprime_commute)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   314
  from prime_coprime[OF p, of x] dvd_imp_le[of p x] x0 xp have px:"coprime x p"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   315
    by (auto simp add: coprime_commute)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   316
  from cong_solve_unique[OF px p01(1)]
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   317
  obtain y where y: "y < p" "[x * y = a] (mod p)" "\<forall>z. z < p \<and> [x * z = a] (mod p) \<longrightarrow> z = y" by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   318
  {assume y0: "y = 0"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   319
    with y(2) have th: "p dvd a" by (simp add: cong_commute[of 0 a p])
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   320
    with p coprime_prime[OF pa, of p] have False by simp}
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   321
  with y show ?thesis unfolding Ex1_def using neq0_conv by blast
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   322
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   323
lemma cong_unique_inverse_prime:
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   324
  assumes p: "prime p" and x0: "0 < x" and xp: "x < p"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   325
  shows "\<exists>!y. 0 < y \<and> y < p \<and> [x * y = 1] (mod p)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   326
  using cong_solve_unique_nontrivial[OF p coprime_1[of p] x0 xp] .
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   327
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   328
(* Forms of the Chinese remainder theorem.                                   *)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   329
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   330
lemma cong_chinese:
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   331
  assumes ab: "coprime a b" and  xya: "[x = y] (mod a)"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   332
  and xyb: "[x = y] (mod b)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   333
  shows "[x = y] (mod a*b)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   334
  using ab xya xyb
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   335
  by (simp add: cong_sub_cases[of x y a] cong_sub_cases[of x y b]
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   336
    cong_sub_cases[of x y "a*b"])
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   337
(cases "x \<le> y", simp_all add: divides_mul[of a _ b])
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   338
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   339
lemma chinese_remainder_unique:
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   340
  assumes ab: "coprime a b" and az: "a \<noteq> 0" and bz: "b\<noteq>0"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   341
  shows "\<exists>!x. x < a * b \<and> [x = m] (mod a) \<and> [x = n] (mod b)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   342
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   343
  from az bz have abpos: "a*b > 0" by simp
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   344
  from chinese_remainder[OF ab az bz] obtain x q1 q2 where
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   345
    xq12: "x = m + q1 * a" "x = n + q2 * b" by blast
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   346
  let ?w = "x mod (a*b)"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   347
  have wab: "?w < a*b" by (simp add: mod_less_divisor[OF abpos])
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   348
  from xq12(1) have "?w mod a = ((m + q1 * a) mod (a*b)) mod a" by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   349
  also have "\<dots> = m mod a" apply (simp add: mod_mult2_eq)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   350
    apply (subst mod_add_left_eq)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   351
    by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   352
  finally have th1: "[?w = m] (mod a)" by (simp add: modeq_def)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   353
  from xq12(2) have "?w mod b = ((n + q2 * b) mod (a*b)) mod b" by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   354
  also have "\<dots> = ((n + q2 * b) mod (b*a)) mod b" by (simp add: mult_commute)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   355
  also have "\<dots> = n mod b" apply (simp add: mod_mult2_eq)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   356
    apply (subst mod_add_left_eq)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   357
    by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   358
  finally have th2: "[?w = n] (mod b)" by (simp add: modeq_def)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   359
  {fix y assume H: "y < a*b" "[y = m] (mod a)" "[y = n] (mod b)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   360
    with th1 th2 have H': "[y = ?w] (mod a)" "[y = ?w] (mod b)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   361
      by (simp_all add: modeq_def)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   362
    from cong_chinese[OF ab H'] mod_less[OF H(1)] mod_less[OF wab]
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   363
    have "y = ?w" by (simp add: modeq_def)}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   364
  with th1 th2 wab show ?thesis by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   365
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   366
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   367
lemma chinese_remainder_coprime_unique:
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   368
  assumes ab: "coprime a b" and az: "a \<noteq> 0" and bz: "b \<noteq> 0"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   369
  and ma: "coprime m a" and nb: "coprime n b"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   370
  shows "\<exists>!x. coprime x (a * b) \<and> x < a * b \<and> [x = m] (mod a) \<and> [x = n] (mod b)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   371
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   372
  let ?P = "\<lambda>x. x < a * b \<and> [x = m] (mod a) \<and> [x = n] (mod b)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   373
  from chinese_remainder_unique[OF ab az bz]
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   374
  obtain x where x: "x < a * b" "[x = m] (mod a)" "[x = n] (mod b)"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   375
    "\<forall>y. ?P y \<longrightarrow> y = x" by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   376
  from ma nb cong_coprime[OF x(2)] cong_coprime[OF x(3)]
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   377
  have "coprime x a" "coprime x b" by (simp_all add: coprime_commute)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   378
  with coprime_mul[of x a b] have "coprime x (a*b)" by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   379
  with x show ?thesis by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   380
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   381
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   382
(* Euler totient function.                                                   *)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   383
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   384
definition phi_def: "\<phi> n = card { m. 0 < m \<and> m <= n \<and> coprime m n }"
31197
c1c163ec6c44 fine-tuned elimination of comprehensions involving x=t.
nipkow
parents: 31021
diff changeset
   385
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   386
lemma phi_0[simp]: "\<phi> 0 = 0"
31197
c1c163ec6c44 fine-tuned elimination of comprehensions involving x=t.
nipkow
parents: 31021
diff changeset
   387
  unfolding phi_def by auto
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   388
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   389
lemma phi_finite[simp]: "finite ({ m. 0 < m \<and> m <= n \<and> coprime m n })"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   390
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   391
  have "{ m. 0 < m \<and> m <= n \<and> coprime m n } \<subseteq> {0..n}" by auto
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   392
  thus ?thesis by (auto intro: finite_subset)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   393
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   394
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   395
declare coprime_1[presburger]
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   396
lemma phi_1[simp]: "\<phi> 1 = 1"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   397
proof-
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   398
  {fix m
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   399
    have "0 < m \<and> m <= 1 \<and> coprime m 1 \<longleftrightarrow> m = 1" by presburger }
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   400
  thus ?thesis by (simp add: phi_def)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   401
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   402
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   403
lemma [simp]: "\<phi> (Suc 0) = Suc 0" using phi_1 by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   404
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   405
lemma phi_alt: "\<phi>(n) = card { m. coprime m n \<and> m < n}"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   406
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   407
  {assume "n=0 \<or> n=1" hence ?thesis by (cases "n=0", simp_all)}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   408
  moreover
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   409
  {assume n: "n\<noteq>0" "n\<noteq>1"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   410
    {fix m
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   411
      from n have "0 < m \<and> m <= n \<and> coprime m n \<longleftrightarrow> coprime m n \<and> m < n"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   412
        apply (cases "m = 0", simp_all)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   413
        apply (cases "m = 1", simp_all)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   414
        apply (cases "m = n", auto)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   415
        done }
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   416
    hence ?thesis unfolding phi_def by simp}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   417
  ultimately show ?thesis by auto
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   418
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   419
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   420
lemma phi_finite_lemma[simp]: "finite {m. coprime m n \<and>  m < n}" (is "finite ?S")
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   421
  by (rule finite_subset[of "?S" "{0..n}"], auto)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   422
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   423
lemma phi_another: assumes n: "n\<noteq>1"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   424
  shows "\<phi> n = card {m. 0 < m \<and> m < n \<and> coprime m n }"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   425
proof-
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   426
  {fix m
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   427
    from n have "0 < m \<and> m < n \<and> coprime m n \<longleftrightarrow> coprime m n \<and> m < n"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   428
      by (cases "m=0", auto)}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   429
  thus ?thesis unfolding phi_alt by auto
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   430
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   431
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   432
lemma phi_limit: "\<phi> n \<le> n"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   433
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   434
  have "{ m. coprime m n \<and> m < n} \<subseteq> {0 ..<n}" by auto
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   435
  with card_mono[of "{0 ..<n}" "{ m. coprime m n \<and> m < n}"]
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   436
  show ?thesis unfolding phi_alt by auto
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   437
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   438
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   439
lemma stupid[simp]: "{m. (0::nat) < m \<and> m < n} = {1..<n}"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   440
  by auto
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   441
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   442
lemma phi_limit_strong: assumes n: "n\<noteq>1"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   443
  shows "\<phi>(n) \<le> n - 1"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   444
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   445
  show ?thesis
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   446
    unfolding phi_another[OF n] finite_number_segment[of n, symmetric]
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   447
    by (rule card_mono[of "{m. 0 < m \<and> m < n}" "{m. 0 < m \<and> m < n \<and> coprime m n}"], auto)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   448
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   449
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   450
lemma phi_lowerbound_1_strong: assumes n: "n \<ge> 1"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   451
  shows "\<phi>(n) \<ge> 1"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   452
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   453
  let ?S = "{ m. 0 < m \<and> m <= n \<and> coprime m n }"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   454
  from card_0_eq[of ?S] n have "\<phi> n \<noteq> 0" unfolding phi_alt
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   455
    apply auto
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   456
    apply (cases "n=1", simp_all)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   457
    apply (rule exI[where x=1], simp)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   458
    done
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   459
  thus ?thesis by arith
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   460
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   461
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   462
lemma phi_lowerbound_1: "2 <= n ==> 1 <= \<phi>(n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   463
  using phi_lowerbound_1_strong[of n] by auto
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   464
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   465
lemma phi_lowerbound_2: assumes n: "3 <= n" shows "2 <= \<phi> (n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   466
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   467
  let ?S = "{ m. 0 < m \<and> m <= n \<and> coprime m n }"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   468
  have inS: "{1, n - 1} \<subseteq> ?S" using n coprime_plus1[of "n - 1"]
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   469
    by (auto simp add: coprime_commute)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   470
  from n have c2: "card {1, n - 1} = 2" by (auto simp add: card_insert_if)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   471
  from card_mono[of ?S "{1, n - 1}", simplified inS c2] show ?thesis
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   472
    unfolding phi_def by auto
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   473
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   474
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   475
lemma phi_prime: "\<phi> n = n - 1 \<and> n\<noteq>0 \<and> n\<noteq>1 \<longleftrightarrow> prime n"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   476
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   477
  {assume "n=0 \<or> n=1" hence ?thesis by (cases "n=1", simp_all)}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   478
  moreover
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   479
  {assume n: "n\<noteq>0" "n\<noteq>1"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   480
    let ?S = "{m. 0 < m \<and> m < n}"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   481
    have fS: "finite ?S" by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   482
    let ?S' = "{m. 0 < m \<and> m < n \<and> coprime m n}"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   483
    have fS':"finite ?S'" apply (rule finite_subset[of ?S' ?S]) by auto
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   484
    {assume H: "\<phi> n = n - 1 \<and> n\<noteq>0 \<and> n\<noteq>1"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   485
      hence ceq: "card ?S' = card ?S"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   486
      using n finite_number_segment[of n] phi_another[OF n(2)] by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   487
      {fix m assume m: "0 < m" "m < n" "\<not> coprime m n"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   488
        hence mS': "m \<notin> ?S'" by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   489
        have "insert m ?S' \<le> ?S" using m by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   490
        from m have "card (insert m ?S') \<le> card ?S"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   491
          by - (rule card_mono[of ?S "insert m ?S'"], auto)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   492
        hence False
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   493
          unfolding card_insert_disjoint[of "?S'" m, OF fS' mS'] ceq
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   494
          by simp }
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   495
      hence "\<forall>m. 0 <m \<and> m < n \<longrightarrow> coprime m n" by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   496
      hence "prime n" unfolding prime using n by (simp add: coprime_commute)}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   497
    moreover
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   498
    {assume H: "prime n"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   499
      hence "?S = ?S'" unfolding prime using n
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   500
        by (auto simp add: coprime_commute)
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   501
      hence "card ?S = card ?S'" by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   502
      hence "\<phi> n = n - 1" unfolding phi_another[OF n(2)] by simp}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   503
    ultimately have ?thesis using n by blast}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   504
  ultimately show ?thesis by (cases "n=0") blast+
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   505
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   506
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   507
(* Multiplicativity property.                                                *)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   508
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   509
lemma phi_multiplicative: assumes ab: "coprime a b"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   510
  shows "\<phi> (a * b) = \<phi> a * \<phi> b"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   511
proof-
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   512
  {assume "a = 0 \<or> b = 0 \<or> a = 1 \<or> b = 1"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   513
    hence ?thesis
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   514
      by (cases "a=0", simp, cases "b=0", simp, cases"a=1", simp_all) }
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   515
  moreover
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   516
  {assume a: "a\<noteq>0" "a\<noteq>1" and b: "b\<noteq>0" "b\<noteq>1"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   517
    hence ab0: "a*b \<noteq> 0" by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   518
    let ?S = "\<lambda>k. {m. coprime m k \<and> m < k}"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   519
    let ?f = "\<lambda>x. (x mod a, x mod b)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   520
    have eq: "?f ` (?S (a*b)) = (?S a \<times> ?S b)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   521
    proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   522
      {fix x assume x:"x \<in> ?S (a*b)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   523
        hence x': "coprime x (a*b)" "x < a*b" by simp_all
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   524
        hence xab: "coprime x a" "coprime x b" by (simp_all add: coprime_mul_eq)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   525
        from mod_less_divisor a b have xab':"x mod a < a" "x mod b < b" by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   526
        from xab xab' have "?f x \<in> (?S a \<times> ?S b)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   527
          by (simp add: coprime_mod[OF a(1)] coprime_mod[OF b(1)])}
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   528
      moreover
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   529
      {fix x y assume x: "x \<in> ?S a" and y: "y \<in> ?S b"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   530
        hence x': "coprime x a" "x < a" and y': "coprime y b" "y < b" by simp_all
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   531
        from chinese_remainder_coprime_unique[OF ab a(1) b(1) x'(1) y'(1)]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   532
        obtain z where z: "coprime z (a * b)" "z < a * b" "[z = x] (mod a)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   533
          "[z = y] (mod b)" by blast
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   534
        hence "(x,y) \<in> ?f ` (?S (a*b))"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   535
          using y'(2) mod_less_divisor[of b y] x'(2) mod_less_divisor[of a x]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   536
          by (auto simp add: image_iff modeq_def)}
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   537
      ultimately show ?thesis by auto
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   538
    qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   539
    have finj: "inj_on ?f (?S (a*b))"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   540
      unfolding inj_on_def
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   541
    proof(clarify)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   542
      fix x y assume H: "coprime x (a * b)" "x < a * b" "coprime y (a * b)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   543
        "y < a * b" "x mod a = y mod a" "x mod b = y mod b"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   544
      hence cp: "coprime x a" "coprime x b" "coprime y a" "coprime y b"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   545
        by (simp_all add: coprime_mul_eq)
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   546
      from chinese_remainder_coprime_unique[OF ab a(1) b(1) cp(3,4)] H
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   547
      show "x = y" unfolding modeq_def by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   548
    qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   549
    from card_image[OF finj, unfolded eq] have ?thesis
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   550
      unfolding phi_alt by simp }
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   551
  ultimately show ?thesis by auto
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   552
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   553
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   554
(* Fermat's Little theorem / Fermat-Euler theorem.                           *)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   555
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   556
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   557
lemma nproduct_mod:
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   558
  assumes fS: "finite S" and n0: "n \<noteq> 0"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   559
  shows "[setprod (\<lambda>m. a(m) mod n) S = setprod a S] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   560
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   561
  have th1:"[1 = 1] (mod n)" by (simp add: modeq_def)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   562
  from cong_mult
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   563
  have th3:"\<forall>x1 y1 x2 y2.
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   564
    [x1 = x2] (mod n) \<and> [y1 = y2] (mod n) \<longrightarrow> [x1 * y1 = x2 * y2] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   565
    by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   566
  have th4:"\<forall>x\<in>S. [a x mod n = a x] (mod n)" by (simp add: modeq_def)
28854
nipkow
parents: 28668
diff changeset
   567
  from fold_image_related[where h="(\<lambda>m. a(m) mod n)" and g=a, OF th1 th3 fS, OF th4] show ?thesis unfolding setprod_def by (simp add: fS)
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   568
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   569
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   570
lemma nproduct_cmul:
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   571
  assumes fS:"finite S"
31021
53642251a04f farewell to class recpower
haftmann
parents: 30738
diff changeset
   572
  shows "setprod (\<lambda>m. (c::'a::{comm_monoid_mult})* a(m)) S = c ^ (card S) * setprod a S"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   573
unfolding setprod_timesf setprod_constant[OF fS, of c] ..
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   574
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   575
lemma coprime_nproduct:
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   576
  assumes fS: "finite S" and Sn: "\<forall>x\<in>S. coprime n (a x)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   577
  shows "coprime n (setprod a S)"
27368
9f90ac19e32b established Plain theory and image
haftmann
parents: 26158
diff changeset
   578
  using fS unfolding setprod_def by (rule finite_subset_induct)
9f90ac19e32b established Plain theory and image
haftmann
parents: 26158
diff changeset
   579
    (insert Sn, auto simp add: coprime_mul)
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   580
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   581
lemma fermat_little: assumes an: "coprime a n"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   582
  shows "[a ^ (\<phi> n) = 1] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   583
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   584
  {assume "n=0" hence ?thesis by simp}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   585
  moreover
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   586
  {assume "n=1" hence ?thesis by (simp add: modeq_def)}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   587
  moreover
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   588
  {assume nz: "n \<noteq> 0" and n1: "n \<noteq> 1"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   589
    let ?S = "{m. coprime m n \<and> m < n}"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   590
    let ?P = "\<Prod> ?S"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   591
    have fS: "finite ?S" by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   592
    have cardfS: "\<phi> n = card ?S" unfolding phi_alt ..
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   593
    {fix m assume m: "m \<in> ?S"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   594
      hence "coprime m n" by simp
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   595
      with coprime_mul[of n a m] an have "coprime (a*m) n"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   596
        by (simp add: coprime_commute)}
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   597
    hence Sn: "\<forall>m\<in> ?S. coprime (a*m) n " by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   598
    from coprime_nproduct[OF fS, of n "\<lambda>m. m"] have nP:"coprime ?P n"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   599
      by (simp add: coprime_commute)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   600
    have Paphi: "[?P*a^ (\<phi> n) = ?P*1] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   601
    proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   602
      let ?h = "\<lambda>m. m mod n"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   603
      {fix m assume mS: "m\<in> ?S"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   604
        hence "?h m \<in> ?S" by simp}
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   605
      hence hS: "?h ` ?S = ?S"by (auto simp add: image_iff)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   606
      have "a\<noteq>0" using an n1 nz apply- apply (rule ccontr) by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   607
      hence inj: "inj_on (op * a) ?S" unfolding inj_on_def by simp
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   608
28854
nipkow
parents: 28668
diff changeset
   609
      have eq0: "fold_image op * (?h \<circ> op * a) 1 {m. coprime m n \<and> m < n} =
nipkow
parents: 28668
diff changeset
   610
     fold_image op * (\<lambda>m. m) 1 {m. coprime m n \<and> m < n}"
nipkow
parents: 28668
diff changeset
   611
      proof (rule fold_image_eq_general[where h="?h o (op * a)"])
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   612
        show "finite ?S" using fS .
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   613
      next
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   614
        {fix y assume yS: "y \<in> ?S" hence y: "coprime y n" "y < n" by simp_all
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   615
          from cong_solve_unique[OF an nz, of y]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   616
          obtain x where x:"x < n" "[a * x = y] (mod n)" "\<forall>z. z < n \<and> [a * z = y] (mod n) \<longrightarrow> z=x" by blast
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   617
          from cong_coprime[OF x(2)] y(1)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   618
          have xm: "coprime x n" by (simp add: coprime_mul_eq coprime_commute)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   619
          {fix z assume "z \<in> ?S" "(?h \<circ> op * a) z = y"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   620
            hence z: "coprime z n" "z < n" "(?h \<circ> op * a) z = y" by simp_all
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   621
            from x(3)[rule_format, of z] z(2,3) have "z=x"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   622
              unfolding modeq_def mod_less[OF y(2)] by simp}
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   623
          with xm x(1,2) have "\<exists>!x. x \<in> ?S \<and> (?h \<circ> op * a) x = y"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   624
            unfolding modeq_def mod_less[OF y(2)] by auto }
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   625
        thus "\<forall>y\<in>{m. coprime m n \<and> m < n}.
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   626
       \<exists>!x. x \<in> {m. coprime m n \<and> m < n} \<and> ((\<lambda>m. m mod n) \<circ> op * a) x = y" by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   627
      next
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   628
        {fix x assume xS: "x\<in> ?S"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   629
          hence x: "coprime x n" "x < n" by simp_all
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   630
          with an have "coprime (a*x) n"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   631
            by (simp add: coprime_mul_eq[of n a x] coprime_commute)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   632
          hence "?h (a*x) \<in> ?S" using nz
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   633
            by (simp add: coprime_mod[OF nz] mod_less_divisor)}
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   634
        thus " \<forall>x\<in>{m. coprime m n \<and> m < n}.
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   635
       ((\<lambda>m. m mod n) \<circ> op * a) x \<in> {m. coprime m n \<and> m < n} \<and>
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   636
       ((\<lambda>m. m mod n) \<circ> op * a) x = ((\<lambda>m. m mod n) \<circ> op * a) x" by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   637
      qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   638
      from nproduct_mod[OF fS nz, of "op * a"]
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   639
      have "[(setprod (op *a) ?S) = (setprod (?h o (op * a)) ?S)] (mod n)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   640
        unfolding o_def
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   641
        by (simp add: cong_commute)
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   642
      also have "[setprod (?h o (op * a)) ?S = ?P ] (mod n)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   643
        using eq0 fS an by (simp add: setprod_def modeq_def o_def)
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   644
      finally show "[?P*a^ (\<phi> n) = ?P*1] (mod n)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   645
        unfolding cardfS mult_commute[of ?P "a^ (card ?S)"]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   646
          nproduct_cmul[OF fS, symmetric] mult_1_right by simp
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   647
    qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   648
    from cong_mult_lcancel[OF nP Paphi] have ?thesis . }
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   649
  ultimately show ?thesis by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   650
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   651
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   652
lemma fermat_little_prime: assumes p: "prime p" and ap: "coprime a p"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   653
  shows "[a^ (p - 1) = 1] (mod p)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   654
  using fermat_little[OF ap] p[unfolded phi_prime[symmetric]]
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   655
by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   656
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   657
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   658
(* Lucas's theorem.                                                          *)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   659
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   660
lemma lucas_coprime_lemma:
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   661
  assumes m: "m\<noteq>0" and am: "[a^m = 1] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   662
  shows "coprime a n"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   663
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   664
  {assume "n=1" hence ?thesis by simp}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   665
  moreover
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   666
  {assume "n = 0" hence ?thesis using am m exp_eq_1[of a m] by simp}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   667
  moreover
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   668
  {assume n: "n\<noteq>0" "n\<noteq>1"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   669
    from m obtain m' where m': "m = Suc m'" by (cases m, blast+)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   670
    {fix d
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   671
      assume d: "d dvd a" "d dvd n"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   672
      from n have n1: "1 < n" by arith
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   673
      from am mod_less[OF n1] have am1: "a^m mod n = 1" unfolding modeq_def by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   674
      from dvd_mult2[OF d(1), of "a^m'"] have dam:"d dvd a^m" by (simp add: m')
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   675
      from dvd_mod_iff[OF d(2), of "a^m"] dam am1
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   676
      have "d = 1" by simp }
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   677
    hence ?thesis unfolding coprime by auto
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   678
  }
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   679
  ultimately show ?thesis by blast
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   680
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   681
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   682
lemma lucas_weak:
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   683
  assumes n: "n \<ge> 2" and an:"[a^(n - 1) = 1] (mod n)"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   684
  and nm: "\<forall>m. 0 <m \<and> m < n - 1 \<longrightarrow> \<not> [a^m = 1] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   685
  shows "prime n"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   686
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   687
  from n have n1: "n \<noteq> 1" "n\<noteq>0" "n - 1 \<noteq> 0" "n - 1 > 0" "n - 1 < n" by arith+
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   688
  from lucas_coprime_lemma[OF n1(3) an] have can: "coprime a n" .
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   689
  from fermat_little[OF can] have afn: "[a ^ \<phi> n = 1] (mod n)" .
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   690
  {assume "\<phi> n \<noteq> n - 1"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   691
    with phi_limit_strong[OF n1(1)] phi_lowerbound_1[OF n]
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   692
    have c:"\<phi> n > 0 \<and> \<phi> n < n - 1" by arith
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   693
    from nm[rule_format, OF c] afn have False ..}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   694
  hence "\<phi> n = n - 1" by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   695
  with phi_prime[of n] n1(1,2) show ?thesis by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   696
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   697
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   698
lemma nat_exists_least_iff: "(\<exists>(n::nat). P n) \<longleftrightarrow> (\<exists>n. P n \<and> (\<forall>m < n. \<not> P m))"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   699
  (is "?lhs \<longleftrightarrow> ?rhs")
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   700
proof
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   701
  assume ?rhs thus ?lhs by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   702
next
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   703
  assume H: ?lhs then obtain n where n: "P n" by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   704
  let ?x = "Least P"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   705
  {fix m assume m: "m < ?x"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   706
    from not_less_Least[OF m] have "\<not> P m" .}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   707
  with LeastI_ex[OF H] show ?rhs by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   708
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   709
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   710
lemma nat_exists_least_iff': "(\<exists>(n::nat). P n) \<longleftrightarrow> (P (Least P) \<and> (\<forall>m < (Least P). \<not> P m))"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   711
  (is "?lhs \<longleftrightarrow> ?rhs")
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   712
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   713
  {assume ?rhs hence ?lhs by blast}
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   714
  moreover
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   715
  { assume H: ?lhs then obtain n where n: "P n" by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   716
    let ?x = "Least P"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   717
    {fix m assume m: "m < ?x"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   718
      from not_less_Least[OF m] have "\<not> P m" .}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   719
    with LeastI_ex[OF H] have ?rhs by blast}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   720
  ultimately show ?thesis by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   721
qed
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   722
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   723
lemma power_mod: "((x::nat) mod m)^n mod m = x^n mod m"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   724
proof(induct n)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   725
  case 0 thus ?case by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   726
next
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   727
  case (Suc n)
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   728
  have "(x mod m)^(Suc n) mod m = ((x mod m) * (((x mod m) ^ n) mod m)) mod m"
30224
79136ce06bdb removed and renamed redundant lemmas
nipkow
parents: 30042
diff changeset
   729
    by (simp add: mod_mult_right_eq[symmetric])
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   730
  also have "\<dots> = ((x mod m) * (x^n mod m)) mod m" using Suc.hyps by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   731
  also have "\<dots> = x^(Suc n) mod m"
30224
79136ce06bdb removed and renamed redundant lemmas
nipkow
parents: 30042
diff changeset
   732
    by (simp add: mod_mult_left_eq[symmetric] mod_mult_right_eq[symmetric])
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   733
  finally show ?case .
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   734
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   735
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   736
lemma lucas:
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   737
  assumes n2: "n \<ge> 2" and an1: "[a^(n - 1) = 1] (mod n)"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   738
  and pn: "\<forall>p. prime p \<and> p dvd n - 1 \<longrightarrow> \<not> [a^((n - 1) div p) = 1] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   739
  shows "prime n"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   740
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   741
  from n2 have n01: "n\<noteq>0" "n\<noteq>1" "n - 1 \<noteq> 0" by arith+
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   742
  from mod_less_divisor[of n 1] n01 have onen: "1 mod n = 1" by simp
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   743
  from lucas_coprime_lemma[OF n01(3) an1] cong_coprime[OF an1]
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   744
  have an: "coprime a n" "coprime (a^(n - 1)) n" by (simp_all add: coprime_commute)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   745
  {assume H0: "\<exists>m. 0 < m \<and> m < n - 1 \<and> [a ^ m = 1] (mod n)" (is "EX m. ?P m")
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   746
    from H0[unfolded nat_exists_least_iff[of ?P]] obtain m where
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   747
      m: "0 < m" "m < n - 1" "[a ^ m = 1] (mod n)" "\<forall>k <m. \<not>?P k" by blast
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   748
    {assume nm1: "(n - 1) mod m > 0"
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   749
      from mod_less_divisor[OF m(1)] have th0:"(n - 1) mod m < m" by blast
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   750
      let ?y = "a^ ((n - 1) div m * m)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   751
      note mdeq = mod_div_equality[of "(n - 1)" m]
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   752
      from coprime_exp[OF an(1)[unfolded coprime_commute[of a n]],
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   753
        of "(n - 1) div m * m"]
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   754
      have yn: "coprime ?y n" by (simp add: coprime_commute)
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   755
      have "?y mod n = (a^m)^((n - 1) div m) mod n"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   756
        by (simp add: algebra_simps power_mult)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   757
      also have "\<dots> = (a^m mod n)^((n - 1) div m) mod n"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   758
        using power_mod[of "a^m" n "(n - 1) div m"] by simp
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   759
      also have "\<dots> = 1" using m(3)[unfolded modeq_def onen] onen
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   760
        by (simp add: power_Suc0)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   761
      finally have th3: "?y mod n = 1"  .
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   762
      have th2: "[?y * a ^ ((n - 1) mod m) = ?y* 1] (mod n)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   763
        using an1[unfolded modeq_def onen] onen
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   764
          mod_div_equality[of "(n - 1)" m, symmetric]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   765
        by (simp add:power_add[symmetric] modeq_def th3 del: One_nat_def)
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   766
      from cong_mult_lcancel[of ?y n "a^((n - 1) mod m)" 1, OF yn th2]
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   767
      have th1: "[a ^ ((n - 1) mod m) = 1] (mod n)"  .
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   768
      from m(4)[rule_format, OF th0] nm1
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   769
        less_trans[OF mod_less_divisor[OF m(1), of "n - 1"] m(2)] th1
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   770
      have False by blast }
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   771
    hence "(n - 1) mod m = 0" by auto
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   772
    then have mn: "m dvd n - 1" by presburger
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   773
    then obtain r where r: "n - 1 = m*r" unfolding dvd_def by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   774
    from n01 r m(2) have r01: "r\<noteq>0" "r\<noteq>1" by - (rule ccontr, simp)+
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   775
    from prime_factor[OF r01(2)] obtain p where p: "prime p" "p dvd r" by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   776
    hence th: "prime p \<and> p dvd n - 1" unfolding r by (auto intro: dvd_mult)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   777
    have "(a ^ ((n - 1) div p)) mod n = (a^(m*r div p)) mod n" using r
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   778
      by (simp add: power_mult)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   779
    also have "\<dots> = (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
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   780
    also have "\<dots> = ((a^m)^(r div p)) mod n" by (simp add: power_mult)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   781
    also have "\<dots> = ((a^m mod n)^(r div p)) mod n" using power_mod[of "a^m" "n" "r div p" ] ..
26158
9dc286ee452b Fixed proofs
chaieb
parents: 26126
diff changeset
   782
    also have "\<dots> = 1" using m(3) onen by (simp add: modeq_def power_Suc0)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   783
    finally have "[(a ^ ((n - 1) div p))= 1] (mod n)"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   784
      using onen by (simp add: modeq_def)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   785
    with pn[rule_format, OF th] have False by blast}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   786
  hence th: "\<forall>m. 0 < m \<and> m < n - 1 \<longrightarrow> \<not> [a ^ m = 1] (mod n)" by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   787
  from lucas_weak[OF n2 an1 th] show ?thesis .
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   788
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   789
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   790
(* Definition of the order of a number mod n (0 in non-coprime case).        *)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   791
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   792
definition "ord n a = (if coprime n a then Least (\<lambda>d. d > 0 \<and> [a ^d = 1] (mod n)) else 0)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   793
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   794
(* This has the expected properties.                                         *)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   795
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   796
lemma coprime_ord:
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   797
  assumes na: "coprime n a"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   798
  shows "ord n a > 0 \<and> [a ^(ord n a) = 1] (mod n) \<and> (\<forall>m. 0 < m \<and> m < ord n a \<longrightarrow> \<not> [a^ m = 1] (mod n))"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   799
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   800
  let ?P = "\<lambda>d. 0 < d \<and> [a ^ d = 1] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   801
  from euclid[of a] obtain p where p: "prime p" "a < p" by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   802
  from na have o: "ord n a = Least ?P" by (simp add: ord_def)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   803
  {assume "n=0 \<or> n=1" with na have "\<exists>m>0. ?P m" apply auto apply (rule exI[where x=1]) by (simp  add: modeq_def)}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   804
  moreover
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   805
  {assume "n\<noteq>0 \<and> n\<noteq>1" hence n2:"n \<ge> 2" by arith
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   806
    from na have na': "coprime a n" by (simp add: coprime_commute)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   807
    from phi_lowerbound_1[OF n2] fermat_little[OF na']
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   808
    have ex: "\<exists>m>0. ?P m" by - (rule exI[where x="\<phi> n"], auto) }
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   809
  ultimately have ex: "\<exists>m>0. ?P m" by blast
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   810
  from nat_exists_least_iff'[of ?P] ex na show ?thesis
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   811
    unfolding o[symmetric] by auto
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   812
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   813
(* With the special value 0 for non-coprime case, it's more convenient.      *)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   814
lemma ord_works:
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   815
 "[a ^ (ord n a) = 1] (mod n) \<and> (\<forall>m. 0 < m \<and> m < ord n a \<longrightarrow> ~[a^ m = 1] (mod n))"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   816
apply (cases "coprime n a")
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   817
using coprime_ord[of n a]
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   818
by (blast, simp add: ord_def modeq_def)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   819
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   820
lemma ord: "[a^(ord n a) = 1] (mod n)" using ord_works by blast
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   821
lemma ord_minimal: "0 < m \<Longrightarrow> m < ord n a \<Longrightarrow> ~[a^m = 1] (mod n)"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   822
  using ord_works by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   823
lemma ord_eq_0: "ord n a = 0 \<longleftrightarrow> ~coprime n a"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   824
by (cases "coprime n a", simp add: neq0_conv coprime_ord, simp add: neq0_conv ord_def)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   825
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   826
lemma ord_divides:
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   827
 "[a ^ d = 1] (mod n) \<longleftrightarrow> ord n a dvd d" (is "?lhs \<longleftrightarrow> ?rhs")
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   828
proof
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   829
  assume rh: ?rhs
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   830
  then obtain k where "d = ord n a * k" unfolding dvd_def by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   831
  hence "[a ^ d = (a ^ (ord n a) mod n)^k] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   832
    by (simp add : modeq_def power_mult power_mod)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   833
  also have "[(a ^ (ord n a) mod n)^k = 1] (mod n)"
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   834
    using ord[of a n, unfolded modeq_def]
26158
9dc286ee452b Fixed proofs
chaieb
parents: 26126
diff changeset
   835
    by (simp add: modeq_def power_mod power_Suc0)
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   836
  finally  show ?lhs .
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   837
next
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   838
  assume lh: ?lhs
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   839
  { assume H: "\<not> coprime n a"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   840
    hence o: "ord n a = 0" by (simp add: ord_def)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   841
    {assume d: "d=0" with o H have ?rhs by (simp add: modeq_def)}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   842
    moreover
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   843
    {assume d0: "d\<noteq>0" then obtain d' where d': "d = Suc d'" by (cases d, auto)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   844
      from H[unfolded coprime]
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   845
      obtain p where p: "p dvd n" "p dvd a" "p \<noteq> 1" by auto
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   846
      from lh[unfolded nat_mod]
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   847
      obtain q1 q2 where q12:"a ^ d + n * q1 = 1 + n * q2" by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   848
      hence "a ^ d + n * q1 - n * q2 = 1" by simp
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31197
diff changeset
   849
      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
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   850
      with p(3) have False by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   851
      hence ?rhs ..}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   852
    ultimately have ?rhs by blast}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   853
  moreover
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   854
  {assume H: "coprime n a"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   855
    let ?o = "ord n a"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   856
    let ?q = "d div ord n a"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   857
    let ?r = "d mod ord n a"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   858
    from cong_exp[OF ord[of a n], of ?q]
26158
9dc286ee452b Fixed proofs
chaieb
parents: 26126
diff changeset
   859
    have eqo: "[(a^?o)^?q = 1] (mod n)"  by (simp add: modeq_def power_Suc0)
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   860
    from H have onz: "?o \<noteq> 0" by (simp add: ord_eq_0)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   861
    hence op: "?o > 0" by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   862
    from mod_div_equality[of d "ord n a"] lh
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   863
    have "[a^(?o*?q + ?r) = 1] (mod n)" by (simp add: modeq_def mult_commute)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   864
    hence "[(a^?o)^?q * (a^?r) = 1] (mod n)"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   865
      by (simp add: modeq_def power_mult[symmetric] power_add[symmetric])
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   866
    hence th: "[a^?r = 1] (mod n)"
30224
79136ce06bdb removed and renamed redundant lemmas
nipkow
parents: 30042
diff changeset
   867
      using eqo mod_mult_left_eq[of "(a^?o)^?q" "a^?r" n]
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   868
      apply (simp add: modeq_def del: One_nat_def)
30224
79136ce06bdb removed and renamed redundant lemmas
nipkow
parents: 30042
diff changeset
   869
      by (simp add: mod_mult_left_eq[symmetric])
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   870
    {assume r: "?r = 0" hence ?rhs by (simp add: dvd_eq_mod_eq_0)}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   871
    moreover
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   872
    {assume r: "?r \<noteq> 0"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   873
      with mod_less_divisor[OF op, of d] have r0o:"?r >0 \<and> ?r < ?o" by simp
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   874
      from conjunct2[OF ord_works[of a n], rule_format, OF r0o] th
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   875
      have ?rhs by blast}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   876
    ultimately have ?rhs by blast}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   877
  ultimately  show ?rhs by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   878
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   879
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   880
lemma order_divides_phi: "coprime n a \<Longrightarrow> ord n a dvd \<phi> n"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   881
using ord_divides fermat_little coprime_commute by simp
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   882
lemma order_divides_expdiff:
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   883
  assumes na: "coprime n a"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   884
  shows "[a^d = a^e] (mod n) \<longleftrightarrow> [d = e] (mod (ord n a))"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   885
proof-
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   886
  {fix n a d e
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   887
    assume na: "coprime n a" and ed: "(e::nat) \<le> d"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   888
    hence "\<exists>c. d = e + c" by arith
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   889
    then obtain c where c: "d = e + c" by arith
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   890
    from na have an: "coprime a n" by (simp add: coprime_commute)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   891
    from coprime_exp[OF na, of e]
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   892
    have aen: "coprime (a^e) n" by (simp add: coprime_commute)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   893
    from coprime_exp[OF na, of c]
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   894
    have acn: "coprime (a^c) n" by (simp add: coprime_commute)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   895
    have "[a^d = a^e] (mod n) \<longleftrightarrow> [a^(e + c) = a^(e + 0)] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   896
      using c by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   897
    also have "\<dots> \<longleftrightarrow> [a^e* a^c = a^e *a^0] (mod n)" by (simp add: power_add)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   898
    also have  "\<dots> \<longleftrightarrow> [a ^ c = 1] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   899
      using cong_mult_lcancel_eq[OF aen, of "a^c" "a^0"] by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   900
    also  have "\<dots> \<longleftrightarrow> ord n a dvd c" by (simp only: ord_divides)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   901
    also have "\<dots> \<longleftrightarrow> [e + c = e + 0] (mod ord n a)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   902
      using cong_add_lcancel_eq[of e c 0 "ord n a", simplified cong_0_divides]
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   903
      by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   904
    finally have "[a^d = a^e] (mod n) \<longleftrightarrow> [d = e] (mod (ord n a))"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   905
      using c by simp }
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   906
  note th = this
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   907
  have "e \<le> d \<or> d \<le> e" by arith
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   908
  moreover
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   909
  {assume ed: "e \<le> d" from th[OF na ed] have ?thesis .}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   910
  moreover
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   911
  {assume de: "d \<le> e"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   912
    from th[OF na de] have ?thesis by (simp add: cong_commute) }
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   913
  ultimately show ?thesis by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   914
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   915
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   916
(* Another trivial primality characterization.                               *)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   917
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   918
lemma prime_prime_factor:
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   919
  "prime n \<longleftrightarrow> n \<noteq> 1\<and> (\<forall>p. prime p \<and> p dvd n \<longrightarrow> p = n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   920
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   921
  {assume n: "n=0 \<or> n=1" hence ?thesis using prime_0 two_is_prime by auto}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   922
  moreover
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   923
  {assume n: "n\<noteq>0" "n\<noteq>1"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   924
    {assume pn: "prime n"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   925
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   926
      from pn[unfolded prime_def] have "\<forall>p. prime p \<and> p dvd n \<longrightarrow> p = n"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   927
        using n
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   928
        apply (cases "n = 0 \<or> n=1",simp)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   929
        by (clarsimp, erule_tac x="p" in allE, auto)}
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   930
    moreover
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   931
    {assume H: "\<forall>p. prime p \<and> p dvd n \<longrightarrow> p = n"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   932
      from n have n1: "n > 1" by arith
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   933
      {fix m assume m: "m dvd n" "m\<noteq>1"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   934
        from prime_factor[OF m(2)] obtain p where
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   935
          p: "prime p" "p dvd m" by blast
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   936
        from dvd_trans[OF p(2) m(1)] p(1) H have "p = n" by blast
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   937
        with p(2) have "n dvd m"  by simp
33657
a4179bf442d1 renamed lemmas "anti_sym" -> "antisym"
nipkow
parents: 32960
diff changeset
   938
        hence "m=n"  using dvd_antisym[OF m(1)] by simp }
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   939
      with n1 have "prime n"  unfolding prime_def by auto }
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   940
    ultimately have ?thesis using n by blast}
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   941
  ultimately       show ?thesis by auto
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   942
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   943
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   944
lemma prime_divisor_sqrt:
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   945
  "prime n \<longleftrightarrow> n \<noteq> 1 \<and> (\<forall>d. d dvd n \<and> d^2 \<le> n \<longrightarrow> d = 1)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   946
proof-
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   947
  {assume "n=0 \<or> n=1" hence ?thesis using prime_0 prime_1
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   948
    by (auto simp add: nat_power_eq_0_iff)}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   949
  moreover
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   950
  {assume n: "n\<noteq>0" "n\<noteq>1"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   951
    hence np: "n > 1" by arith
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   952
    {fix d assume d: "d dvd n" "d^2 \<le> n" and H: "\<forall>m. m dvd n \<longrightarrow> m=1 \<or> m=n"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   953
      from H d have d1n: "d = 1 \<or> d=n" by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   954
      {assume dn: "d=n"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   955
        have "n^2 > n*1" using n
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   956
          by (simp add: power2_eq_square mult_less_cancel1)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   957
        with dn d(2) have "d=1" by simp}
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   958
      with d1n have "d = 1" by blast  }
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   959
    moreover
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   960
    {fix d assume d: "d dvd n" and H: "\<forall>d'. d' dvd n \<and> d'^2 \<le> n \<longrightarrow> d' = 1"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   961
      from d n have "d \<noteq> 0" apply - apply (rule ccontr) by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   962
      hence dp: "d > 0" by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   963
      from d[unfolded dvd_def] obtain e where e: "n= d*e" by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   964
      from n dp e have ep:"e > 0" by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   965
      have "d^2 \<le> n \<or> e^2 \<le> n" using dp ep
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   966
        by (auto simp add: e power2_eq_square mult_le_cancel_left)
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   967
      moreover
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   968
      {assume h: "d^2 \<le> n"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   969
        from H[rule_format, of d] h d have "d = 1" by blast}
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   970
      moreover
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   971
      {assume h: "e^2 \<le> n"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   972
        from e have "e dvd n" unfolding dvd_def by (simp add: mult_commute)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   973
        with H[rule_format, of e] h have "e=1" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   974
        with e have "d = n" by simp}
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   975
      ultimately have "d=1 \<or> d=n"  by blast}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   976
    ultimately have ?thesis unfolding prime_def using np n(2) by blast}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   977
  ultimately show ?thesis by auto
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   978
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   979
lemma prime_prime_factor_sqrt:
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   980
  "prime n \<longleftrightarrow> n \<noteq> 0 \<and> n \<noteq> 1 \<and> \<not> (\<exists>p. prime p \<and> p dvd n \<and> p^2 \<le> n)"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   981
  (is "?lhs \<longleftrightarrow>?rhs")
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   982
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   983
  {assume "n=0 \<or> n=1" hence ?thesis using prime_0 prime_1 by auto}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   984
  moreover
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   985
  {assume n: "n\<noteq>0" "n\<noteq>1"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   986
    {assume H: ?lhs
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
   987
      from H[unfolded prime_divisor_sqrt] n
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   988
      have ?rhs  apply clarsimp by (erule_tac x="p" in allE, simp add: prime_1)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   989
    }
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   990
    moreover
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   991
    {assume H: ?rhs
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
   992
      {fix d assume d: "d dvd n" "d^2 \<le> n" "d\<noteq>1"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   993
        from prime_factor[OF d(3)]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   994
        obtain p where p: "prime p" "p dvd d" by blast
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   995
        from n have np: "n > 0" by arith
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   996
        from d(1) n have "d \<noteq> 0" by - (rule ccontr, auto)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   997
        hence dp: "d > 0" by arith
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   998
        from mult_mono[OF dvd_imp_le[OF p(2) dp] dvd_imp_le[OF p(2) dp]] d(2)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   999
        have "p^2 \<le> n" unfolding power2_eq_square by arith
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
  1000
        with H n p(1) dvd_trans[OF p(2) d(1)] have False  by blast}
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1001
      with n prime_divisor_sqrt  have ?lhs by auto}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1002
    ultimately have ?thesis by blast }
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1003
  ultimately show ?thesis by (cases "n=0 \<or> n=1", auto)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1004
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1005
(* Pocklington theorem. *)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1006
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1007
lemma pocklington_lemma:
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1008
  assumes n: "n \<ge> 2" and nqr: "n - 1 = q*r" and an: "[a^ (n - 1) = 1] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1009
  and aq:"\<forall>p. prime p \<and> p dvd q \<longrightarrow> coprime (a^ ((n - 1) div p) - 1) n"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1010
  and pp: "prime p" and pn: "p dvd n"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1011
  shows "[p = 1] (mod q)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1012
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1013
  from pp prime_0 prime_1 have p01: "p \<noteq> 0" "p \<noteq> 1" by - (rule ccontr, simp)+
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1014
  from cong_1_divides[OF an, unfolded nqr, unfolded dvd_def]
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1015
  obtain k where k: "a ^ (q * r) - 1 = n*k" by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1016
  from pn[unfolded dvd_def] obtain l where l: "n = p*l" by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1017
  {assume a0: "a = 0"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1018
    hence "a^ (n - 1) = 0" using n by (simp add: power_0_left)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1019
    with n an mod_less[of 1 n]  have False by (simp add: power_0_left modeq_def)}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1020
  hence a0: "a\<noteq>0" ..
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1021
  from n nqr have aqr0: "a ^ (q * r) \<noteq> 0" using a0 by (simp add: neq0_conv)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1022
  hence "(a ^ (q * r) - 1) + 1  = a ^ (q * r)" by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1023
  with k l have "a ^ (q * r) = p*l*k + 1" by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1024
  hence "a ^ (r * q) + p * 0 = 1 + p * (l*k)" by (simp add: mult_ac)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1025
  hence odq: "ord p (a^r) dvd q"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1026
    unfolding ord_divides[symmetric] power_mult[symmetric] nat_mod  by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1027
  from odq[unfolded dvd_def] obtain d where d: "q = ord p (a^r) * d" by blast
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1028
  {assume d1: "d \<noteq> 1"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1029
    from prime_factor[OF d1] obtain P where P: "prime P" "P dvd d" by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1030
    from d dvd_mult[OF P(2), of "ord p (a^r)"] have Pq: "P dvd q" by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1031
    from aq P(1) Pq have caP:"coprime (a^ ((n - 1) div P) - 1) n" by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1032
    from Pq obtain s where s: "q = P*s" unfolding dvd_def by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1033
    have P0: "P \<noteq> 0" using P(1) prime_0 by - (rule ccontr, simp)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1034
    from P(2) obtain t where t: "d = P*t" unfolding dvd_def by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1035
    from d s t P0  have s': "ord p (a^r) * t = s" by algebra
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1036
    have "ord p (a^r) * t*r = r * ord p (a^r) * t" by algebra
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1037
    hence exps: "a^(ord p (a^r) * t*r) = ((a ^ r) ^ ord p (a^r)) ^ t"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1038
      by (simp only: power_mult)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1039
    have "[((a ^ r) ^ ord p (a^r)) ^ t= 1^t] (mod p)"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1040
      by (rule cong_exp, rule ord)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1041
    then have th: "[((a ^ r) ^ ord p (a^r)) ^ t= 1] (mod p)"
26158
9dc286ee452b Fixed proofs
chaieb
parents: 26126
diff changeset
  1042
      by (simp add: power_Suc0)
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1043
    from cong_1_divides[OF th] exps have pd0: "p dvd a^(ord p (a^r) * t*r) - 1" by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1044
    from nqr s s' have "(n - 1) div P = ord p (a^r) * t*r" using P0 by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1045
    with caP have "coprime (a^(ord p (a^r) * t*r) - 1) n" by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1046
    with p01 pn pd0 have False unfolding coprime by auto}
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1047
  hence d1: "d = 1" by blast
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1048
  hence o: "ord p (a^r) = q" using d by simp
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1049
  from pp phi_prime[of p] have phip: " \<phi> p = p - 1" by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1050
  {fix d assume d: "d dvd p" "d dvd a" "d \<noteq> 1"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1051
    from pp[unfolded prime_def] d have dp: "d = p" by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1052
    from n have n12:"Suc (n - 2) = n - 1" by arith
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1053
    with divides_rexp[OF d(2)[unfolded dp], of "n - 2"]
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1054
    have th0: "p dvd a ^ (n - 1)" by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1055
    from n have n0: "n \<noteq> 0" by simp
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1056
    from d(2) an n12[symmetric] have a0: "a \<noteq> 0"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1057
      by - (rule ccontr, simp add: modeq_def)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1058
    have th1: "a^ (n - 1) \<noteq> 0" using n d(2) dp a0 by (auto simp add: neq0_conv)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1059
    from coprime_minus1[OF th1, unfolded coprime]
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1060
      dvd_trans[OF pn cong_1_divides[OF an]] th0 d(3) dp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1061
    have False by auto}
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1062
  hence cpa: "coprime p a" using coprime by auto
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1063
  from coprime_exp[OF cpa, of r] coprime_commute
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1064
  have arp: "coprime (a^r) p" by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1065
  from fermat_little[OF arp, simplified ord_divides] o phip
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1066
  have "q dvd (p - 1)" by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1067
  then obtain d where d:"p - 1 = q * d" unfolding dvd_def by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1068
  from prime_0 pp have p0:"p \<noteq> 0" by -  (rule ccontr, auto)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1069
  from p0 d have "p + q * 0 = 1 + q * d" by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1070
  with nat_mod[of p 1 q, symmetric]
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1071
  show ?thesis by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1072
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1073
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1074
lemma pocklington:
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1075
  assumes n: "n \<ge> 2" and nqr: "n - 1 = q*r" and sqr: "n \<le> q^2"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1076
  and an: "[a^ (n - 1) = 1] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1077
  and aq:"\<forall>p. prime p \<and> p dvd q \<longrightarrow> coprime (a^ ((n - 1) div p) - 1) n"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1078
  shows "prime n"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1079
unfolding prime_prime_factor_sqrt[of n]
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1080
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1081
  let ?ths = "n \<noteq> 0 \<and> n \<noteq> 1 \<and> \<not> (\<exists>p. prime p \<and> p dvd n \<and> p\<twosuperior> \<le> n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1082
  from n have n01: "n\<noteq>0" "n\<noteq>1" by arith+
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1083
  {fix p assume p: "prime p" "p dvd n" "p^2 \<le> n"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1084
    from p(3) sqr have "p^(Suc 1) \<le> q^(Suc 1)" by (simp add: power2_eq_square)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1085
    hence pq: "p \<le> q" unfolding exp_mono_le .
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1086
    from pocklington_lemma[OF n nqr an aq p(1,2)]  cong_1_divides
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1087
    have th: "q dvd p - 1" by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1088
    have "p - 1 \<noteq> 0"using prime_ge_2[OF p(1)] by arith
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1089
    with divides_ge[OF th] pq have False by arith }
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1090
  with n01 show ?ths by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1091
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1092
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1093
(* Variant for application, to separate the exponentiation.                  *)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1094
lemma pocklington_alt:
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1095
  assumes n: "n \<ge> 2" and nqr: "n - 1 = q*r" and sqr: "n \<le> q^2"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1096
  and an: "[a^ (n - 1) = 1] (mod n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1097
  and aq:"\<forall>p. prime p \<and> p dvd q \<longrightarrow> (\<exists>b. [a^((n - 1) div p) = b] (mod n) \<and> coprime (b - 1) n)"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1098
  shows "prime n"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1099
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1100
  {fix p assume p: "prime p" "p dvd q"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1101
    from aq[rule_format] p obtain b where
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1102
      b: "[a^((n - 1) div p) = b] (mod n)" "coprime (b - 1) n" by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1103
    {assume a0: "a=0"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1104
      from n an have "[0 = 1] (mod n)" unfolding a0 power_0_left by auto
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1105
      hence False using n by (simp add: modeq_def dvd_eq_mod_eq_0[symmetric])}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1106
    hence a0: "a\<noteq> 0" ..
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1107
    hence a1: "a \<ge> 1" by arith
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1108
    from one_le_power[OF a1] have ath: "1 \<le> a ^ ((n - 1) div p)" .
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1109
    {assume b0: "b = 0"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1110
      from p(2) nqr have "(n - 1) mod p = 0"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
  1111
        apply (simp only: dvd_eq_mod_eq_0[symmetric]) by (rule dvd_mult2, simp)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1112
      with mod_div_equality[of "n - 1" p]
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1113
      have "(n - 1) div p * p= n - 1" by auto
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1114
      hence eq: "(a^((n - 1) div p))^p = a^(n - 1)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
  1115
        by (simp only: power_mult[symmetric])
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1116
      from prime_ge_2[OF p(1)] have pS: "Suc (p - 1) = p" by arith
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1117
      from b(1) have d: "n dvd a^((n - 1) div p)" unfolding b0 cong_0_divides .
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1118
      from divides_rexp[OF d, of "p - 1"] pS eq cong_divides[OF an] n
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1119
      have False by simp}
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1120
    then have b0: "b \<noteq> 0" ..
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1121
    hence b1: "b \<ge> 1" by arith
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1122
    from cong_coprime[OF cong_sub[OF b(1) cong_refl[of 1] ath b1]] b(2) nqr
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1123
    have "coprime (a ^ ((n - 1) div p) - 1) n" by (simp add: coprime_commute)}
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1124
  hence th: "\<forall>p. prime p \<and> p dvd q \<longrightarrow> coprime (a ^ ((n - 1) div p) - 1) n "
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1125
    by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1126
  from pocklington[OF n nqr sqr an th] show ?thesis .
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1127
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1128
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1129
(* Prime factorizations.                                                     *)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1130
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1131
definition "primefact ps n = (foldr op * ps  1 = n \<and> (\<forall>p\<in> set ps. prime p))"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1132
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1133
lemma primefact: assumes n: "n \<noteq> 0"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1134
  shows "\<exists>ps. primefact ps n"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1135
using n
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1136
proof(induct n rule: nat_less_induct)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1137
  fix n assume H: "\<forall>m<n. m \<noteq> 0 \<longrightarrow> (\<exists>ps. primefact ps m)" and n: "n\<noteq>0"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1138
  let ?ths = "\<exists>ps. primefact ps n"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1139
  {assume "n = 1"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1140
    hence "primefact [] n" by (simp add: primefact_def)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1141
    hence ?ths by blast }
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1142
  moreover
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1143
  {assume n1: "n \<noteq> 1"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1144
    with n have n2: "n \<ge> 2" by arith
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1145
    from prime_factor[OF n1] obtain p where p: "prime p" "p dvd n" by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1146
    from p(2) obtain m where m: "n = p*m" unfolding dvd_def by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1147
    from n m have m0: "m > 0" "m\<noteq>0" by auto
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1148
    from prime_ge_2[OF p(1)] have "1 < p" by arith
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1149
    with m0 m have mn: "m < n" by auto
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1150
    from H[rule_format, OF mn m0(2)] obtain ps where ps: "primefact ps m" ..
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1151
    from ps m p(1) have "primefact (p#ps) n" by (simp add: primefact_def)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1152
    hence ?ths by blast}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1153
  ultimately show ?ths by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1154
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1155
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1156
lemma primefact_contains:
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1157
  assumes pf: "primefact ps n" and p: "prime p" and pn: "p dvd n"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1158
  shows "p \<in> set ps"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1159
  using pf p pn
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1160
proof(induct ps arbitrary: p n)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1161
  case Nil thus ?case by (auto simp add: primefact_def)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1162
next
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1163
  case (Cons q qs p n)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1164
  from Cons.prems[unfolded primefact_def]
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1165
  have q: "prime q" "q * foldr op * qs 1 = n" "\<forall>p \<in>set qs. prime p"  and p: "prime p" "p dvd q * foldr op * qs 1" by simp_all
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1166
  {assume "p dvd q"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1167
    with p(1) q(1) have "p = q" unfolding prime_def by auto
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1168
    hence ?case by simp}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1169
  moreover
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1170
  { assume h: "p dvd foldr op * qs 1"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1171
    from q(3) have pqs: "primefact qs (foldr op * qs 1)"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1172
      by (simp add: primefact_def)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1173
    from Cons.hyps[OF pqs p(1) h] have ?case by simp}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1174
  ultimately show ?case using prime_divprod[OF p] by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1175
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1176
37602
501b0cae5aa8 tuned whitespace
haftmann
parents: 33657
diff changeset
  1177
lemma primefact_variant: "primefact ps n \<longleftrightarrow> foldr op * ps 1 = n \<and> list_all prime ps"
501b0cae5aa8 tuned whitespace
haftmann
parents: 33657
diff changeset
  1178
  by (auto simp add: primefact_def list_all_iff)
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1179
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1180
(* Variant of Lucas theorem.                                                 *)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1181
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1182
lemma lucas_primefact:
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1183
  assumes n: "n \<ge> 2" and an: "[a^(n - 1) = 1] (mod n)"
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1184
  and psn: "foldr op * ps 1 = n - 1"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1185
  and psp: "list_all (\<lambda>p. prime p \<and> \<not> [a^((n - 1) div p) = 1] (mod n)) ps"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1186
  shows "prime n"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1187
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1188
  {fix p assume p: "prime p" "p dvd n - 1" "[a ^ ((n - 1) div p) = 1] (mod n)"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1189
    from psn psp have psn1: "primefact ps (n - 1)"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1190
      by (auto simp add: list_all_iff primefact_variant)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1191
    from p(3) primefact_contains[OF psn1 p(1,2)] psp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1192
    have False by (induct ps, auto)}
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1193
  with lucas[OF n an] show ?thesis by blast
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1194
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1195
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1196
(* Variant of Pocklington theorem.                                           *)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1197
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1198
lemma mod_le: assumes n: "n \<noteq> (0::nat)" shows "m mod n \<le> m"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1199
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1200
    from mod_div_equality[of m n]
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1201
    have "\<exists>x. x + m mod n = m" by blast
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1202
    then show ?thesis by auto
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1203
qed
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1204
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1205
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1206
lemma pocklington_primefact:
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1207
  assumes n: "n \<ge> 2" and qrn: "q*r = n - 1" and nq2: "n \<le> q^2"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1208
  and arnb: "(a^r) mod n = b" and psq: "foldr op * ps 1 = q"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1209
  and bqn: "(b^q) mod n = 1"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1210
  and psp: "list_all (\<lambda>p. prime p \<and> coprime ((b^(q div p)) mod n - 1) n) ps"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1211
  shows "prime n"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1212
proof-
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1213
  from bqn psp qrn
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1214
  have bqn: "a ^ (n - 1) mod n = 1"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1215
    and psp: "list_all (\<lambda>p. prime p \<and> coprime (a^(r *(q div p)) mod n - 1) n) ps"  unfolding arnb[symmetric] power_mod
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28854
diff changeset
  1216
    by (simp_all add: power_mult[symmetric] algebra_simps)
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1217
  from n  have n0: "n > 0" by arith
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1218
  from mod_div_equality[of "a^(n - 1)" n]
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1219
    mod_less_divisor[OF n0, of "a^(n - 1)"]
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1220
  have an1: "[a ^ (n - 1) = 1] (mod n)"
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1221
    unfolding nat_mod bqn
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1222
    apply -
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1223
    apply (rule exI[where x="0"])
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1224
    apply (rule exI[where x="a^(n - 1) div n"])
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28854
diff changeset
  1225
    by (simp add: algebra_simps)
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1226
  {fix p assume p: "prime p" "p dvd q"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1227
    from psp psq have pfpsq: "primefact ps q"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1228
      by (auto simp add: primefact_variant list_all_iff)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1229
    from psp primefact_contains[OF pfpsq p]
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1230
    have p': "coprime (a ^ (r * (q div p)) mod n - 1) n"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1231
      by (simp add: list_all_iff)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1232
    from prime_ge_2[OF p(1)] have p01: "p \<noteq> 0" "p \<noteq> 1" "p =Suc(p - 1)" by arith+
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1233
    from div_mult1_eq[of r q p] p(2)
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1234
    have eq1: "r* (q div p) = (n - 1) div p"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1235
      unfolding qrn[symmetric] dvd_eq_mod_eq_0 by (simp add: mult_commute)
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1236
    have ath: "\<And>a (b::nat). a <= b \<Longrightarrow> a \<noteq> 0 ==> 1 <= a \<and> 1 <= b" by arith
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1237
    from n0 have n00: "n \<noteq> 0" by arith
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1238
    from mod_le[OF n00]
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1239
    have th10: "a ^ ((n - 1) div p) mod n \<le> a ^ ((n - 1) div p)" .
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1240
    {assume "a ^ ((n - 1) div p) mod n = 0"
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1241
      then obtain s where s: "a ^ ((n - 1) div p) = n*s"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
  1242
        unfolding mod_eq_0_iff by blast
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1243
      hence eq0: "(a^((n - 1) div p))^p = (n*s)^p" by simp
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1244
      from qrn[symmetric] have qn1: "q dvd n - 1" unfolding dvd_def by auto
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1245
      from dvd_trans[OF p(2) qn1] div_mod_equality'[of "n - 1" p]
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1246
      have npp: "(n - 1) div p * p = n - 1" by (simp add: dvd_eq_mod_eq_0)
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1247
      with eq0 have "a^ (n - 1) = (n*s)^p"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
  1248
        by (simp add: power_mult[symmetric])
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1249
      hence "1 = (n*s)^(Suc (p - 1)) mod n" using bqn p01 by simp
28668
e79e196039a1 fixed and reactivated HOL/Library/Pocklington.thy -- by Mark Hillebrand;
wenzelm
parents: 27668
diff changeset
  1250
      also have "\<dots> = 0" by (simp add: mult_assoc)
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1251
      finally have False by simp }
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1252
      then have th11: "a ^ ((n - 1) div p) mod n \<noteq> 0" by auto
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1253
    have th1: "[a ^ ((n - 1) div p) mod n = a ^ ((n - 1) div p)] (mod n)"
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1254
      unfolding modeq_def by simp
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1255
    from cong_sub[OF th1 cong_refl[of 1]]  ath[OF th10 th11]
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1256
    have th: "[a ^ ((n - 1) div p) mod n - 1 = a ^ ((n - 1) div p) - 1] (mod n)"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1257
      by blast
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1258
    from cong_coprime[OF th] p'[unfolded eq1]
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1259
    have "coprime (a ^ ((n - 1) div p) - 1) n" by (simp add: coprime_commute) }
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1260
  with pocklington[OF n qrn[symmetric] nq2 an1]
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30264
diff changeset
  1261
  show ?thesis by blast
26126
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1262
qed
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1263
f536ac0f92ca Pocklington's Primality criterion
chaieb
parents:
diff changeset
  1264
end