src/HOL/Number_Theory/Euler_Criterion.thy
author haftmann
Sat, 02 Dec 2017 16:50:53 +0000
changeset 67118 ccab07d1196c
parent 67091 1393c2340eec
child 68707 5b269897df9d
permissions -rw-r--r--
more simplification rules
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
64282
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
     1
(* Author: Jaime Mendizabal Roche *)
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
     2
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
     3
theory Euler_Criterion
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
     4
imports Residues
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
     5
begin
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
     6
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
     7
context
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
     8
  fixes p :: "nat"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
     9
  fixes a :: "int"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    10
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    11
  assumes p_prime: "prime p"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    12
  assumes p_ge_2: "2 < p"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    13
  assumes p_a_relprime: "[a \<noteq> 0](mod p)"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    14
begin
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    15
67083
6b2c0681ef28 new simp rule
haftmann
parents: 67051
diff changeset
    16
private lemma odd_p: "odd p"
6b2c0681ef28 new simp rule
haftmann
parents: 67051
diff changeset
    17
  using p_ge_2 p_prime prime_odd_nat by blast
6b2c0681ef28 new simp rule
haftmann
parents: 67051
diff changeset
    18
6b2c0681ef28 new simp rule
haftmann
parents: 67051
diff changeset
    19
private lemma p_minus_1_int:
67118
ccab07d1196c more simplification rules
haftmann
parents: 67091
diff changeset
    20
  "int (p - 1) = int p - 1"
ccab07d1196c more simplification rules
haftmann
parents: 67091
diff changeset
    21
  by (metis of_nat_1 of_nat_diff p_prime prime_ge_1_nat)
ccab07d1196c more simplification rules
haftmann
parents: 67091
diff changeset
    22
  
67083
6b2c0681ef28 new simp rule
haftmann
parents: 67051
diff changeset
    23
private lemma p_not_eq_Suc_0 [simp]:
6b2c0681ef28 new simp rule
haftmann
parents: 67051
diff changeset
    24
  "p \<noteq> Suc 0"
6b2c0681ef28 new simp rule
haftmann
parents: 67051
diff changeset
    25
  using p_ge_2 by simp
6b2c0681ef28 new simp rule
haftmann
parents: 67051
diff changeset
    26
6b2c0681ef28 new simp rule
haftmann
parents: 67051
diff changeset
    27
private lemma one_mod_int_p_eq [simp]:
6b2c0681ef28 new simp rule
haftmann
parents: 67051
diff changeset
    28
  "1 mod int p = 1"
6b2c0681ef28 new simp rule
haftmann
parents: 67051
diff changeset
    29
proof -
6b2c0681ef28 new simp rule
haftmann
parents: 67051
diff changeset
    30
  from p_ge_2 have "int 1 mod int p = int 1"
6b2c0681ef28 new simp rule
haftmann
parents: 67051
diff changeset
    31
    by simp
6b2c0681ef28 new simp rule
haftmann
parents: 67051
diff changeset
    32
  then show ?thesis
6b2c0681ef28 new simp rule
haftmann
parents: 67051
diff changeset
    33
    by simp
6b2c0681ef28 new simp rule
haftmann
parents: 67051
diff changeset
    34
qed
64282
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    35
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 64282
diff changeset
    36
private lemma E_1:
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 64282
diff changeset
    37
  assumes "QuadRes (int p) a"
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 64282
diff changeset
    38
  shows "[a ^ ((p - 1) div 2) = 1] (mod int p)"
64282
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    39
proof -
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 64282
diff changeset
    40
  from assms obtain b where b: "[b ^ 2 = a] (mod int p)"
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 64282
diff changeset
    41
    unfolding QuadRes_def by blast
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 64282
diff changeset
    42
  then have "[a ^ ((p - 1) div 2) = b ^ (2 * ((p - 1) div 2))] (mod int p)"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 65465
diff changeset
    43
    by (simp add: cong_pow cong_sym power_mult)
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 64282
diff changeset
    44
  then have "[a ^ ((p - 1) div 2) = b ^ (p - 1)] (mod int p)"
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 64282
diff changeset
    45
    using odd_p by force
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 64282
diff changeset
    46
  moreover have "[b ^ (p - 1) = 1] (mod int p)"
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 64282
diff changeset
    47
  proof -
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 64282
diff changeset
    48
    have "[nat \<bar>b\<bar> ^ (p - 1) = 1] (mod p)"
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 64282
diff changeset
    49
    using p_prime proof (rule fermat_theorem)
67118
ccab07d1196c more simplification rules
haftmann
parents: 67091
diff changeset
    50
      from b p_a_relprime show "\<not> p dvd nat \<bar>b\<bar>"
ccab07d1196c more simplification rules
haftmann
parents: 67091
diff changeset
    51
        by (auto simp add: cong_altdef_int power2_eq_square)
ccab07d1196c more simplification rules
haftmann
parents: 67091
diff changeset
    52
          (metis cong_altdef_int cong_dvd_iff dvd_mult2) 
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 64282
diff changeset
    53
    qed
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 64282
diff changeset
    54
    then have "nat \<bar>b\<bar> ^ (p - 1) mod p = 1 mod p"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 65465
diff changeset
    55
      by (simp add: cong_def)
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 64282
diff changeset
    56
    then have "int (nat \<bar>b\<bar> ^ (p - 1) mod p) = int (1 mod p)"
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 64282
diff changeset
    57
      by simp
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 64282
diff changeset
    58
    moreover from odd_p have "\<bar>b\<bar> ^ (p - Suc 0) = b ^ (p - Suc 0)"
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 64282
diff changeset
    59
      by (simp add: power_even_abs)
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 64282
diff changeset
    60
    ultimately show ?thesis
67083
6b2c0681ef28 new simp rule
haftmann
parents: 67051
diff changeset
    61
      by (auto simp add: cong_def zmod_int)
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 64282
diff changeset
    62
  qed
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 64282
diff changeset
    63
  ultimately show ?thesis
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 65465
diff changeset
    64
    by (auto intro: cong_trans)
64282
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    65
qed
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    66
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    67
private definition S1 :: "int set" where "S1 = {0 <.. int p - 1}"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    68
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    69
private definition P :: "int \<Rightarrow> int \<Rightarrow> bool" where
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    70
  "P x y \<longleftrightarrow> [x * y = a] (mod p) \<and> y \<in> S1"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    71
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    72
private definition f_1 :: "int \<Rightarrow> int" where
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    73
  "f_1 x = (THE y. P x y)"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    74
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    75
private definition f :: "int \<Rightarrow> int set" where
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    76
  "f x = {x, f_1 x}"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    77
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    78
private definition S2 :: "int set set" where "S2 = f ` S1"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    79
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    80
private lemma P_lemma: assumes "x \<in> S1"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    81
  shows "\<exists>! y. P x y"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    82
proof -
67091
1393c2340eec more symbols;
wenzelm
parents: 67083
diff changeset
    83
  have "\<not> p dvd x" using assms zdvd_not_zless S1_def by auto
64282
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    84
  hence co_xp: "coprime x p" using p_prime prime_imp_coprime_int[of p x] 
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
    85
    by (simp add: ac_simps)
64282
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    86
  then obtain y' where y': "[x * y' = 1] (mod p)" using cong_solve_coprime_int by blast
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    87
  moreover define y where "y = y' * a mod p"
67083
6b2c0681ef28 new simp rule
haftmann
parents: 67051
diff changeset
    88
  ultimately have "[x * y = a] (mod p)"
6b2c0681ef28 new simp rule
haftmann
parents: 67051
diff changeset
    89
    using mod_mult_right_eq [of x "y' * a" p]
6b2c0681ef28 new simp rule
haftmann
parents: 67051
diff changeset
    90
    cong_scalar_right [of "x * y'" 1 "int p" a]
6b2c0681ef28 new simp rule
haftmann
parents: 67051
diff changeset
    91
    by (auto simp add: cong_def ac_simps)
64282
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    92
  moreover have "y \<in> {0 .. int p - 1}" unfolding y_def using p_ge_2 by auto
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    93
  hence "y \<in> S1" using calculation cong_altdef_int p_a_relprime S1_def by auto
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    94
  ultimately have "P x y" unfolding P_def by blast
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    95
  moreover {
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    96
    fix y1 y2
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    97
    assume "P x y1" "P x y2"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
    98
    moreover hence "[y1 = y2] (mod p)" unfolding P_def
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 65465
diff changeset
    99
      using co_xp cong_mult_lcancel_int[of x p y1 y2] cong_sym cong_trans by blast
64282
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   100
    ultimately have "y1 = y2" unfolding P_def S1_def using cong_less_imp_eq_int by auto
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   101
  }
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   102
  ultimately show ?thesis by blast
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   103
qed
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   104
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   105
private lemma f_1_lemma_1: assumes "x \<in> S1"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   106
  shows "P x (f_1 x)" using assms P_lemma theI'[of "P x"] f_1_def by presburger
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   107
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   108
private lemma f_1_lemma_2: assumes "x \<in> S1"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   109
  shows "f_1 (f_1 x) = x"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   110
  using assms f_1_lemma_1[of x] f_1_def P_lemma[of "f_1 x"] P_def by (auto simp: mult.commute)
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   111
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   112
private lemma f_lemma_1: assumes "x \<in> S1"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   113
  shows "f x = f (f_1 x)" using f_def f_1_lemma_2[of x] assms by auto
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   114
67091
1393c2340eec more symbols;
wenzelm
parents: 67083
diff changeset
   115
private lemma l1: assumes "\<not> QuadRes p a" "x \<in> S1"
64282
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   116
  shows "x \<noteq> f_1 x"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   117
  using f_1_lemma_1[of x] assms unfolding P_def QuadRes_def power2_eq_square by fastforce
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   118
67091
1393c2340eec more symbols;
wenzelm
parents: 67083
diff changeset
   119
private lemma l2: assumes "\<not> QuadRes p a" "x \<in> S1"
64282
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   120
  shows "[\<Prod> (f x) = a] (mod p)"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   121
  using assms l1 f_1_lemma_1 P_def f_def by auto
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   122
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   123
private lemma l3: assumes "x \<in> S2"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   124
  shows "finite x" using assms f_def S2_def by auto
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   125
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   126
private lemma l4: "S1 = \<Union> S2" using f_1_lemma_1 P_def f_def S2_def by auto
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   127
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   128
private lemma l5: assumes "x \<in> S2" "y \<in> S2" "x \<noteq> y"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   129
  shows "x \<inter> y = {}"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   130
proof -
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   131
  obtain a b where ab: "x = f a" "a \<in> S1" "y = f b" "b \<in> S1" using assms S2_def by auto
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   132
  hence "a \<noteq> b" "a \<noteq> f_1 b" "f_1 a \<noteq> b" using assms(3) f_lemma_1 by blast+
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   133
  moreover hence "f_1 a \<noteq> f_1 b" using f_1_lemma_2[of a] f_1_lemma_2[of b] ab by force
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   134
  ultimately show ?thesis using f_def ab by fastforce
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   135
qed
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   136
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   137
private lemma l6: "prod Prod S2 = \<Prod> S1" 
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   138
  using prod.Union_disjoint[of S2 "\<lambda>x. x"] l3 l4 l5 unfolding comp_def by auto
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   139
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   140
private lemma l7: "fact n = \<Prod> {0 <.. int n}"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   141
proof (induction n)
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   142
case (Suc n)
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   143
  have "int (Suc n) = int n + 1" by simp
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   144
  hence "insert (int (Suc n)) {0<..int n} = {0<..int (Suc n)}" by auto
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   145
  thus ?case using prod.insert[of "{0<..int n}" "int (Suc n)" "\<lambda>x. x"] Suc fact_Suc by auto
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   146
qed simp
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   147
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   148
private lemma l8: "fact (p - 1) = \<Prod> S1" using l7[of "p - 1"] S1_def p_minus_1_int by presburger
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   149
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   150
private lemma l9: "[prod Prod S2 = -1] (mod p)"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   151
  using l6 l8 wilson_theorem[of p] p_prime by presburger
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   152
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   153
private lemma l10: assumes "card S = n" "\<And>x. x \<in> S \<Longrightarrow> [g x = a] (mod p)"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   154
  shows "[prod g S = a ^ n] (mod p)" using assms
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   155
proof (induction n arbitrary: S)
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   156
case 0
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   157
  thus ?case using card_0_eq[of S] prod.empty prod.infinite by fastforce
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   158
next
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   159
case (Suc n)
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   160
  then obtain x where x: "x \<in> S" by force
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   161
  define S' where "S' = S - {x}"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   162
  hence "[prod g S' = a ^ n] (mod int p)"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   163
    using x Suc(1)[of S'] Suc(2) Suc(3) by (simp add: card_ge_0_finite)
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   164
  moreover have "prod g S = g x * prod g S'"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   165
    using x S'_def Suc(2) prod.remove[of S x g] by fastforce
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 65465
diff changeset
   166
  ultimately show ?case using x Suc(3) cong_mult
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 65465
diff changeset
   167
    by simp blast 
64282
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   168
qed
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   169
67091
1393c2340eec more symbols;
wenzelm
parents: 67083
diff changeset
   170
private lemma l11: assumes "\<not> QuadRes p a"
64282
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   171
  shows "card S2 = (p - 1) div 2"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   172
proof -
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   173
  have "sum card S2 = 2 * card S2"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   174
    using sum.cong[of S2 S2 card "\<lambda>x. 2"] l1 f_def S2_def assms by fastforce
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   175
  moreover have "p - 1 = sum card S2"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   176
    using l4 card_UN_disjoint[of S2 "\<lambda>x. x"] l3 l5 S1_def S2_def by auto
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   177
  ultimately show ?thesis by linarith
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   178
qed
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   179
67091
1393c2340eec more symbols;
wenzelm
parents: 67083
diff changeset
   180
private lemma l12: assumes "\<not> QuadRes p a"
64282
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   181
  shows "[prod Prod S2 = a ^ ((p - 1) div 2)] (mod p)"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   182
  using assms l2 l10 l11 unfolding S2_def by blast
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   183
67091
1393c2340eec more symbols;
wenzelm
parents: 67083
diff changeset
   184
private lemma E_2: assumes "\<not> QuadRes p a"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 65465
diff changeset
   185
  shows "[a ^ ((p - 1) div 2) = -1] (mod p)" using l9 l12 cong_trans cong_sym assms by blast
64282
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   186
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   187
lemma euler_criterion_aux: "[(Legendre a p) = a ^ ((p - 1) div 2)] (mod p)"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 65465
diff changeset
   188
  using p_a_relprime by (auto simp add: Legendre_def
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 65465
diff changeset
   189
    intro!: cong_sym [of _ 1] cong_sym [of _ "- 1"]
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 65465
diff changeset
   190
    dest: E_1 E_2)
64282
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   191
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   192
end
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   193
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   194
theorem euler_criterion: assumes "prime p" "2 < p"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   195
  shows "[(Legendre a p) = a ^ ((p - 1) div 2)] (mod p)"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   196
proof (cases "[a = 0] (mod p)")
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 65465
diff changeset
   197
  case True
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 65465
diff changeset
   198
  then have "[a ^ ((p - 1) div 2) = 0 ^ ((p - 1) div 2)] (mod p)"
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 65465
diff changeset
   199
    using cong_pow by blast
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 65465
diff changeset
   200
  moreover have "(0::int) ^ ((p - 1) div 2) = 0"
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 65465
diff changeset
   201
    using zero_power [of "(p - 1) div 2"] assms(2) by simp
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 65465
diff changeset
   202
  ultimately have "[a ^ ((p - 1) div 2) = 0] (mod p)"
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 65465
diff changeset
   203
    using True assms(1) cong_altdef_int prime_dvd_power_int_iff by auto
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 65465
diff changeset
   204
  then show ?thesis unfolding Legendre_def using True cong_sym
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 65465
diff changeset
   205
    by auto
64282
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   206
next
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 65465
diff changeset
   207
  case False
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 65465
diff changeset
   208
  then show ?thesis
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 65465
diff changeset
   209
    using euler_criterion_aux assms by presburger
64282
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   210
qed
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   211
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   212
hide_fact euler_criterion_aux
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   213
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
diff changeset
   214
end