src/HOL/Number_Theory/Gauss.thy
author wenzelm
Fri, 07 Apr 2017 21:17:18 +0200
changeset 65435 378175f44328
parent 65413 cb7f9d7d35e6
child 66817 0b12755ccbb2
permissions -rw-r--r--
tuned headers;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
65435
378175f44328 tuned headers;
wenzelm
parents: 65413
diff changeset
     1
(*  Title:      HOL/Number_Theory/Gauss.thy
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
     2
    Authors:    Jeremy Avigad, David Gray, and Adam Kramer
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     3
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
     4
Ported by lcp but unfinished.
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     5
*)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     6
60526
fad653acf58f isabelle update_cartouches;
wenzelm
parents: 59559
diff changeset
     7
section \<open>Gauss' Lemma\<close>
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     8
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     9
theory Gauss
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    10
  imports Euler_Criterion
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    11
begin
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    12
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    13
lemma cong_prime_prod_zero_nat:
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    14
  "[a * b = 0] (mod p) \<Longrightarrow> prime p \<Longrightarrow> [a = 0] (mod p) \<or> [b = 0] (mod p)"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    15
  for a :: nat
64631
7705926ee595 removed dangerous simp rule: prime computations can be excessively long
haftmann
parents: 64282
diff changeset
    16
  by (auto simp add: cong_altdef_nat prime_dvd_mult_iff)
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    17
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    18
lemma cong_prime_prod_zero_int:
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    19
  "[a * b = 0] (mod p) \<Longrightarrow> prime p \<Longrightarrow> [a = 0] (mod p) \<or> [b = 0] (mod p)"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    20
  for a :: int
64631
7705926ee595 removed dangerous simp rule: prime computations can be excessively long
haftmann
parents: 64282
diff changeset
    21
  by (auto simp add: cong_altdef_int prime_dvd_mult_iff)
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    22
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    23
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    24
locale GAUSS =
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    25
  fixes p :: "nat"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    26
  fixes a :: "int"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    27
  assumes p_prime: "prime p"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    28
  assumes p_ge_2: "2 < p"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    29
  assumes p_a_relprime: "[a \<noteq> 0](mod p)"
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    30
  assumes a_nonzero: "0 < a"
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    31
begin
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    32
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    33
definition "A = {0::int <.. ((int p - 1) div 2)}"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    34
definition "B = (\<lambda>x. x * a) ` A"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    35
definition "C = (\<lambda>x. x mod p) ` B"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    36
definition "D = C \<inter> {.. (int p - 1) div 2}"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    37
definition "E = C \<inter> {(int p - 1) div 2 <..}"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    38
definition "F = (\<lambda>x. (int p - x)) ` E"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    39
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    40
60526
fad653acf58f isabelle update_cartouches;
wenzelm
parents: 59559
diff changeset
    41
subsection \<open>Basic properties of p\<close>
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    42
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    43
lemma odd_p: "odd p"
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    44
  by (metis p_prime p_ge_2 prime_odd_nat)
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    45
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    46
lemma p_minus_one_l: "(int p - 1) div 2 < p"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    47
proof -
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    48
  have "(p - 1) div 2 \<le> (p - 1) div 1"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    49
    by (metis div_by_1 div_le_dividend)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    50
  also have "\<dots> = p - 1" by simp
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    51
  finally show ?thesis
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    52
    using p_ge_2 by arith
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    53
qed
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    54
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    55
lemma p_eq2: "int p = (2 * ((int p - 1) div 2)) + 1"
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    56
  using odd_p p_ge_2 nonzero_mult_div_cancel_left [of 2 "p - 1"] by simp
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    57
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    58
lemma p_odd_int: obtains z :: int where "int p = 2 * z + 1" "0 < z"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    59
proof
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    60
  let ?z = "(int p - 1) div 2"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    61
  show "int p = 2 * ?z + 1" by (rule p_eq2)
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    62
  show "0 < ?z"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    63
    using p_ge_2 by linarith
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    64
qed
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    65
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    66
60526
fad653acf58f isabelle update_cartouches;
wenzelm
parents: 59559
diff changeset
    67
subsection \<open>Basic Properties of the Gauss Sets\<close>
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    68
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    69
lemma finite_A: "finite A"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    70
  by (auto simp add: A_def)
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    71
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    72
lemma finite_B: "finite B"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    73
  by (auto simp add: B_def finite_A)
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    74
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    75
lemma finite_C: "finite C"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    76
  by (auto simp add: C_def finite_B)
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    77
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    78
lemma finite_D: "finite D"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    79
  by (auto simp add: D_def finite_C)
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    80
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    81
lemma finite_E: "finite E"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    82
  by (auto simp add: E_def finite_C)
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    83
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    84
lemma finite_F: "finite F"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    85
  by (auto simp add: F_def finite_E)
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    86
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    87
lemma C_eq: "C = D \<union> E"
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    88
  by (auto simp add: C_def D_def E_def)
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    89
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    90
lemma A_card_eq: "card A = nat ((int p - 1) div 2)"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    91
  by (auto simp add: A_def)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    92
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    93
lemma inj_on_xa_A: "inj_on (\<lambda>x. x * a) A"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    94
  using a_nonzero by (simp add: A_def inj_on_def)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    95
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    96
definition ResSet :: "int \<Rightarrow> int set \<Rightarrow> bool"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
    97
  where "ResSet m X \<longleftrightarrow> (\<forall>y1 y2. y1 \<in> X \<and> y2 \<in> X \<and> [y1 = y2] (mod m) \<longrightarrow> y1 = y2)"
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    98
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    99
lemma ResSet_image:
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   100
  "0 < m \<Longrightarrow> ResSet m A \<Longrightarrow> \<forall>x \<in> A. \<forall>y \<in> A. ([f x = f y](mod m) \<longrightarrow> x = y) \<Longrightarrow> ResSet m (f ` A)"
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   101
  by (auto simp add: ResSet_def)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   102
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   103
lemma A_res: "ResSet p A"
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   104
  using p_ge_2 by (auto simp add: A_def ResSet_def intro!: cong_less_imp_eq_int)
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   105
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   106
lemma B_res: "ResSet p B"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   107
proof -
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   108
  have *: "x = y"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   109
    if a: "[x * a = y * a] (mod p)"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   110
    and b: "0 < x"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   111
    and c: "x \<le> (int p - 1) div 2"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   112
    and d: "0 < y"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   113
    and e: "y \<le> (int p - 1) div 2"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   114
    for x y
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   115
  proof -
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   116
    from p_a_relprime have "\<not> p dvd a"
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62429
diff changeset
   117
      by (simp add: cong_altdef_int)
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   118
    with p_prime have "coprime a (int p)"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   119
      by (subst gcd.commute, intro prime_imp_coprime) auto
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   120
    with a cong_mult_rcancel_int [of a "int p" x y] have "[x = y] (mod p)"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   121
      by simp
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   122
    with cong_less_imp_eq_int [of x y p] p_minus_one_l
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   123
      order_le_less_trans [of x "(int p - 1) div 2" p]
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   124
      order_le_less_trans [of y "(int p - 1) div 2" p]
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   125
    show ?thesis
62348
9a5f43dac883 dropped various legacy fact bindings
haftmann
parents: 60688
diff changeset
   126
      by (metis b c cong_less_imp_eq_int d e zero_less_imp_eq_int of_nat_0_le_iff)
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   127
  qed
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   128
  show ?thesis
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   129
    apply (insert p_ge_2 p_a_relprime p_minus_one_l)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   130
    apply (auto simp add: B_def)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   131
    apply (rule ResSet_image)
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   132
      apply (auto simp add: A_res)
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   133
    apply (auto simp add: A_def *)
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   134
    done
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   135
qed
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   136
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   137
lemma SR_B_inj: "inj_on (\<lambda>x. x mod p) B"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   138
proof -
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   139
  have False
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   140
    if a: "x * a mod p = y * a mod p"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   141
    and b: "0 < x"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   142
    and c: "x \<le> (int p - 1) div 2"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   143
    and d: "0 < y"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   144
    and e: "y \<le> (int p - 1) div 2"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   145
    and f: "x \<noteq> y"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   146
    for x y
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   147
  proof -
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   148
    from a have a': "[x * a = y * a](mod p)"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   149
      by (metis cong_int_def)
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   150
    from p_a_relprime have "\<not>p dvd a"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   151
      by (simp add: cong_altdef_int)
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   152
    with p_prime have "coprime a (int p)"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   153
      by (subst gcd.commute, intro prime_imp_coprime) auto
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   154
    with a' cong_mult_rcancel_int [of a "int p" x y]
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62429
diff changeset
   155
    have "[x = y] (mod p)" by simp
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   156
    with cong_less_imp_eq_int [of x y p] p_minus_one_l
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   157
      order_le_less_trans [of x "(int p - 1) div 2" p]
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   158
      order_le_less_trans [of y "(int p - 1) div 2" p]
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   159
    have "x = y"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   160
      by (metis b c cong_less_imp_eq_int d e zero_less_imp_eq_int of_nat_0_le_iff)
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   161
    then show ?thesis
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   162
      by (simp add: f)
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   163
  qed
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   164
  then show ?thesis
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   165
    by (auto simp add: B_def inj_on_def A_def) metis
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   166
qed
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   167
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   168
lemma inj_on_pminusx_E: "inj_on (\<lambda>x. p - x) E"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   169
  apply (auto simp add: E_def C_def B_def A_def)
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   170
  apply (rule inj_on_inverseI [where g = "op - (int p)"])
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   171
  apply auto
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   172
  done
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   173
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   174
lemma nonzero_mod_p: "0 < x \<Longrightarrow> x < int p \<Longrightarrow> [x \<noteq> 0](mod p)"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   175
  for x :: int
59545
12a6088ed195 explicit equivalence for strict order on lattices
haftmann
parents: 58889
diff changeset
   176
  by (simp add: cong_int_def)
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   177
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   178
lemma A_ncong_p: "x \<in> A \<Longrightarrow> [x \<noteq> 0](mod p)"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   179
  by (rule nonzero_mod_p) (auto simp add: A_def)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   180
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   181
lemma A_greater_zero: "x \<in> A \<Longrightarrow> 0 < x"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   182
  by (auto simp add: A_def)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   183
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   184
lemma B_ncong_p: "x \<in> B \<Longrightarrow> [x \<noteq> 0](mod p)"
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   185
  by (auto simp: B_def p_prime p_a_relprime A_ncong_p dest: cong_prime_prod_zero_int)
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   186
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   187
lemma B_greater_zero: "x \<in> B \<Longrightarrow> 0 < x"
56544
b60d5d119489 made mult_pos_pos a simp rule
nipkow
parents: 55730
diff changeset
   188
  using a_nonzero by (auto simp add: B_def A_greater_zero)
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   189
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   190
lemma C_greater_zero: "y \<in> C \<Longrightarrow> 0 < y"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   191
proof (auto simp add: C_def)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   192
  fix x :: int
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   193
  assume x: "x \<in> B"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   194
  moreover from x have "x mod int p \<noteq> 0"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   195
    using B_ncong_p cong_int_def by simp
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   196
  moreover have "int y = 0 \<or> 0 < int y" for y
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   197
    by linarith
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   198
  ultimately show "0 < x mod int p"
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   199
    by (metis (no_types) B_greater_zero Divides.transfer_int_nat_functions(2) zero_less_imp_eq_int)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   200
qed
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   201
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   202
lemma F_subset: "F \<subseteq> {x. 0 < x \<and> x \<le> ((int p - 1) div 2)}"
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   203
  apply (auto simp add: F_def E_def C_def)
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   204
   apply (metis p_ge_2 Divides.pos_mod_bound nat_int zless_nat_conj)
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   205
  apply (auto intro: p_odd_int)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   206
  done
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   207
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   208
lemma D_subset: "D \<subseteq> {x. 0 < x \<and> x \<le> ((p - 1) div 2)}"
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   209
  by (auto simp add: D_def C_greater_zero)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   210
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   211
lemma F_eq: "F = {x. \<exists>y \<in> A. (x = p - ((y * a) mod p) \<and> (int p - 1) div 2 < (y * a) mod p)}"
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   212
  by (auto simp add: F_def E_def D_def C_def B_def A_def)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   213
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   214
lemma D_eq: "D = {x. \<exists>y \<in> A. (x = (y * a) mod p \<and> (y * a) mod p \<le> (int p - 1) div 2)}"
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   215
  by (auto simp add: D_def C_def B_def A_def)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   216
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   217
lemma all_A_relprime:
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   218
  assumes "x \<in> A"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   219
  shows "gcd x p = 1"
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   220
  using p_prime A_ncong_p [OF assms]
63633
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63566
diff changeset
   221
  by (auto simp: cong_altdef_int gcd.commute[of _ "int p"] intro!: prime_imp_coprime)
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   222
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64240
diff changeset
   223
lemma A_prod_relprime: "gcd (prod id A) p = 1"
f76b6dda2e56 setprod -> prod
nipkow
parents: 64240
diff changeset
   224
  by (metis id_def all_A_relprime prod_coprime)
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   225
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   226
60526
fad653acf58f isabelle update_cartouches;
wenzelm
parents: 59559
diff changeset
   227
subsection \<open>Relationships Between Gauss Sets\<close>
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   228
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   229
lemma StandardRes_inj_on_ResSet: "ResSet m X \<Longrightarrow> inj_on (\<lambda>b. b mod m) X"
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   230
  by (auto simp add: ResSet_def inj_on_def cong_int_def)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   231
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   232
lemma B_card_eq_A: "card B = card A"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   233
  using finite_A by (simp add: finite_A B_def inj_on_xa_A card_image)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   234
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   235
lemma B_card_eq: "card B = nat ((int p - 1) div 2)"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   236
  by (simp add: B_card_eq_A A_card_eq)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   237
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   238
lemma F_card_eq_E: "card F = card E"
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   239
  using finite_E by (simp add: F_def inj_on_pminusx_E card_image)
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   240
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   241
lemma C_card_eq_B: "card C = card B"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   242
proof -
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   243
  have "inj_on (\<lambda>x. x mod p) B"
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   244
    by (metis SR_B_inj)
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   245
  then show ?thesis
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   246
    by (metis C_def card_image)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   247
qed
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   248
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   249
lemma D_E_disj: "D \<inter> E = {}"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   250
  by (auto simp add: D_def E_def)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   251
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   252
lemma C_card_eq_D_plus_E: "card C = card D + card E"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   253
  by (auto simp add: C_eq card_Un_disjoint D_E_disj finite_D finite_E)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   254
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64240
diff changeset
   255
lemma C_prod_eq_D_times_E: "prod id E * prod id D = prod id C"
f76b6dda2e56 setprod -> prod
nipkow
parents: 64240
diff changeset
   256
  by (metis C_eq D_E_disj finite_D finite_E inf_commute prod.union_disjoint sup_commute)
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   257
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64240
diff changeset
   258
lemma C_B_zcong_prod: "[prod id C = prod id B] (mod p)"
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   259
  apply (auto simp add: C_def)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   260
  apply (insert finite_B SR_B_inj)
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64240
diff changeset
   261
  apply (drule prod.reindex [of "\<lambda>x. x mod int p" B id])
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56544
diff changeset
   262
  apply auto
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64240
diff changeset
   263
  apply (rule cong_prod_int)
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   264
  apply (auto simp add: cong_int_def)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   265
  done
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   266
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   267
lemma F_Un_D_subset: "(F \<union> D) \<subseteq> A"
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   268
  by (intro Un_least subset_trans [OF F_subset] subset_trans [OF D_subset]) (auto simp: A_def)
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   269
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   270
lemma F_D_disj: "(F \<inter> D) = {}"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   271
proof (auto simp add: F_eq D_eq)
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   272
  fix y z :: int
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   273
  assume "p - (y * a) mod p = (z * a) mod p"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   274
  then have "[(y * a) mod p + (z * a) mod p = 0] (mod p)"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
   275
    by (metis add.commute diff_eq_eq dvd_refl cong_int_def dvd_eq_mod_eq_0 mod_0)
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   276
  moreover have "[y * a = (y * a) mod p] (mod p)"
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   277
    by (metis cong_int_def mod_mod_trivial)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   278
  ultimately have "[a * (y + z) = 0] (mod p)"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
   279
    by (metis cong_int_def mod_add_left_eq mod_add_right_eq mult.commute ring_class.ring_distribs(1))
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   280
  with p_prime a_nonzero p_a_relprime have a: "[y + z = 0] (mod p)"
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62429
diff changeset
   281
    by (auto dest!: cong_prime_prod_zero_int)
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   282
  assume b: "y \<in> A" and c: "z \<in> A"
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   283
  then have "0 < y + z"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   284
    by (auto simp: A_def)
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   285
  moreover from b c p_eq2 have "y + z < p"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   286
    by (auto simp: A_def)
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   287
  ultimately show False
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   288
    by (metis a nonzero_mod_p)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   289
qed
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   290
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   291
lemma F_Un_D_card: "card (F \<union> D) = nat ((p - 1) div 2)"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   292
proof -
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   293
  have "card (F \<union> D) = card E + card D"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   294
    by (auto simp add: finite_F finite_D F_D_disj card_Un_disjoint F_card_eq_E)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   295
  then have "card (F \<union> D) = card C"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   296
    by (simp add: C_card_eq_D_plus_E)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   297
  then show "card (F \<union> D) = nat ((p - 1) div 2)"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   298
    by (simp add: C_card_eq_B B_card_eq)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   299
qed
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   300
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   301
lemma F_Un_D_eq_A: "F \<union> D = A"
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   302
  using finite_A F_Un_D_subset A_card_eq F_Un_D_card by (auto simp add: card_seteq)
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   303
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   304
lemma prod_D_F_eq_prod_A: "prod id D * prod id F = prod id A"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64240
diff changeset
   305
  by (metis F_D_disj F_Un_D_eq_A Int_commute Un_commute finite_D finite_F prod.union_disjoint)
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   306
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   307
lemma prod_F_zcong: "[prod id F = ((-1) ^ (card E)) * prod id E] (mod p)"
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   308
proof -
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64240
diff changeset
   309
  have FE: "prod id F = prod (op - p) E"
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   310
    apply (auto simp add: F_def)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   311
    apply (insert finite_E inj_on_pminusx_E)
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   312
    apply (drule prod.reindex)
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   313
    apply auto
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   314
    done
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   315
  then have "\<forall>x \<in> E. [(p-x) mod p = - x](mod p)"
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   316
    by (metis cong_int_def minus_mod_self1 mod_mod_trivial)
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64240
diff changeset
   317
  then have "[prod ((\<lambda>x. x mod p) o (op - p)) E = prod (uminus) E](mod p)"
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   318
    using finite_E p_ge_2 cong_prod_int [of E "(\<lambda>x. x mod p) o (op - p)" uminus p]
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   319
    by auto
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64240
diff changeset
   320
  then have two: "[prod id F = prod (uminus) E](mod p)"
f76b6dda2e56 setprod -> prod
nipkow
parents: 64240
diff changeset
   321
    by (metis FE cong_cong_mod_int cong_refl_int cong_prod_int minus_mod_self1)
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   322
  have "prod uminus E = (-1) ^ card E * prod id E"
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   323
    using finite_E by (induct set: finite) auto
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   324
  with two show ?thesis
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   325
    by simp
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   326
qed
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   327
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   328
60526
fad653acf58f isabelle update_cartouches;
wenzelm
parents: 59559
diff changeset
   329
subsection \<open>Gauss' Lemma\<close>
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   330
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64240
diff changeset
   331
lemma aux: "prod id A * (- 1) ^ card E * a ^ card A * (- 1) ^ card E = prod id A * a ^ card A"
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   332
  by (metis (no_types) minus_minus mult.commute mult.left_commute power_minus power_one)
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   333
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   334
theorem pre_gauss_lemma: "[a ^ nat((int p - 1) div 2) = (-1) ^ (card E)] (mod p)"
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   335
proof -
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64240
diff changeset
   336
  have "[prod id A = prod id F * prod id D](mod p)"
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   337
    by (auto simp: prod_D_F_eq_prod_A mult.commute cong del: prod.strong_cong)
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64240
diff changeset
   338
  then have "[prod id A = ((-1)^(card E) * prod id E) * prod id D] (mod p)"
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   339
    by (rule cong_trans_int) (metis cong_scalar_int prod_F_zcong)
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64240
diff changeset
   340
  then have "[prod id A = ((-1)^(card E) * prod id C)] (mod p)"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
   341
    by (metis C_prod_eq_D_times_E mult.commute mult.left_commute)
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64240
diff changeset
   342
  then have "[prod id A = ((-1)^(card E) * prod id B)] (mod p)"
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   343
    by (rule cong_trans_int) (metis C_B_zcong_prod cong_scalar2_int)
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   344
  then have "[prod id A = ((-1)^(card E) * prod id ((\<lambda>x. x * a) ` A))] (mod p)"
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   345
    by (simp add: B_def)
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   346
  then have "[prod id A = ((-1)^(card E) * prod (\<lambda>x. x * a) A)] (mod p)"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64240
diff changeset
   347
    by (simp add: inj_on_xa_A prod.reindex)
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   348
  moreover have "prod (\<lambda>x. x * a) A = prod (\<lambda>x. a) A * prod id A"
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   349
    using finite_A by (induct set: finite) auto
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   350
  ultimately have "[prod id A = ((-1)^(card E) * (prod (\<lambda>x. a) A * prod id A))] (mod p)"
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   351
    by simp
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   352
  then have "[prod id A = ((-1)^(card E) * a^(card A) * prod id A)](mod p)"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   353
    by (rule cong_trans_int)
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   354
      (simp add: cong_scalar2_int cong_scalar_int finite_A prod_constant mult.assoc)
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64240
diff changeset
   355
  then have a: "[prod id A * (-1)^(card E) =
f76b6dda2e56 setprod -> prod
nipkow
parents: 64240
diff changeset
   356
      ((-1)^(card E) * a^(card A) * prod id A * (-1)^(card E))](mod p)"
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   357
    by (rule cong_scalar_int)
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64240
diff changeset
   358
  then have "[prod id A * (-1)^(card E) = prod id A *
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   359
      (-1)^(card E) * a^(card A) * (-1)^(card E)](mod p)"
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   360
    by (rule cong_trans_int) (simp add: a mult.commute mult.left_commute)
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64240
diff changeset
   361
  then have "[prod id A * (-1)^(card E) = prod id A * a^(card A)](mod p)"
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   362
    by (rule cong_trans_int) (simp add: aux cong del: prod.strong_cong)
58410
6d46ad54a2ab explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
haftmann
parents: 58288
diff changeset
   363
  with A_prod_relprime have "[(- 1) ^ card E = a ^ card A](mod p)"
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   364
    by (metis cong_mult_lcancel_int)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   365
  then show ?thesis
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   366
    by (simp add: A_card_eq cong_sym_int)
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   367
qed
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   368
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   369
theorem gauss_lemma: "Legendre a p = (-1) ^ (card E)"
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   370
proof -
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   371
  from euler_criterion p_prime p_ge_2 have "[Legendre a p = a^(nat (((p) - 1) div 2))] (mod p)"
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   372
    by auto
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   373
  moreover have "int ((p - 1) div 2) = (int p - 1) div 2"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   374
    using p_eq2 by linarith
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   375
  then have "[a ^ nat (int ((p - 1) div 2)) = a ^ nat ((int p - 1) div 2)] (mod int p)"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   376
    by force
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   377
  ultimately have "[Legendre a p = (-1) ^ (card E)] (mod p)"
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   378
    using pre_gauss_lemma cong_trans_int by blast
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   379
  moreover from p_a_relprime have "Legendre a p = 1 \<or> Legendre a p = -1"
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   380
    by (auto simp add: Legendre_def)
65413
cb7f9d7d35e6 misc tuning and modernization;
wenzelm
parents: 64631
diff changeset
   381
  moreover have "(-1::int) ^ (card E) = 1 \<or> (-1::int) ^ (card E) = -1"
64282
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
   382
    using neg_one_even_power neg_one_odd_power by blast
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
   383
  moreover have "[1 \<noteq> - 1] (mod int p)"
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
   384
    using cong_altdef_int nonzero_mod_p[of 2] p_odd_int by fastforce
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   385
  ultimately show ?thesis
64282
261d42f0bfac Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
   386
    by (auto simp add: cong_sym_int)
55730
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   387
qed
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   388
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   389
end
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   390
97ff9276e12d Gauss.thy ported from Old_Number_Theory (unfinished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   391
end