src/HOL/NumberTheory/Quadratic_Reciprocity.thy
author blanchet
Wed, 04 Mar 2009 10:45:52 +0100
changeset 30240 5b25fee0362c
parent 23431 25ca91279a9b
permissions -rw-r--r--
Merge.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
20346
b138816322c5 title fixed
webertj
parents: 20217
diff changeset
     1
(*  Title:      HOL/NumberTheory/Quadratic_Reciprocity.thy
14981
e73f8140af78 Merged in license change from Isabelle2004
kleing
parents: 14434
diff changeset
     2
    ID:         $Id$
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
     3
    Authors:    Jeremy Avigad, David Gray, and Adam Kramer
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
     4
*)
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
     5
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
     6
header {* The law of Quadratic reciprocity *}
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
     7
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
     8
theory Quadratic_Reciprocity
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
     9
imports Gauss
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    10
begin
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
    11
19670
2e4a143c73c5 prefer 'definition' over low-level defs;
wenzelm
parents: 18369
diff changeset
    12
text {*
2e4a143c73c5 prefer 'definition' over low-level defs;
wenzelm
parents: 18369
diff changeset
    13
  Lemmas leading up to the proof of theorem 3.3 in Niven and
2e4a143c73c5 prefer 'definition' over low-level defs;
wenzelm
parents: 18369
diff changeset
    14
  Zuckerman's presentation.
2e4a143c73c5 prefer 'definition' over low-level defs;
wenzelm
parents: 18369
diff changeset
    15
*}
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
    16
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
    17
context GAUSS
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
    18
begin
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
    19
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
    20
lemma QRLemma1: "a * setsum id A =
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    21
  p * setsum (%x. ((x * a) div p)) A + setsum id D + setsum id E"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    22
proof -
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
    23
  from finite_A have "a * setsum id A = setsum (%x. a * x) A"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
    24
    by (auto simp add: setsum_const_mult id_def)
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
    25
  also have "setsum (%x. a * x) = setsum (%x. x * a)"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
    26
    by (auto simp add: zmult_commute)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    27
  also have "setsum (%x. x * a) A = setsum id B"
16733
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16663
diff changeset
    28
    by (simp add: B_def setsum_reindex_id[OF inj_on_xa_A])
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    29
  also have "... = setsum (%x. p * (x div p) + StandardRes p x) B"
16733
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16663
diff changeset
    30
    by (auto simp add: StandardRes_def zmod_zdiv_equality)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    31
  also have "... = setsum (%x. p * (x div p)) B + setsum (StandardRes p) B"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
    32
    by (rule setsum_addf)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    33
  also have "setsum (StandardRes p) B = setsum id C"
16733
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16663
diff changeset
    34
    by (auto simp add: C_def setsum_reindex_id[OF SR_B_inj])
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    35
  also from C_eq have "... = setsum id (D \<union> E)"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
    36
    by auto
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    37
  also from finite_D finite_E have "... = setsum id D + setsum id E"
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
    38
    by (rule setsum_Un_disjoint) (auto simp add: D_def E_def)
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
    39
  also have "setsum (%x. p * (x div p)) B =
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    40
      setsum ((%x. p * (x div p)) o (%x. (x * a))) A"
16733
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16663
diff changeset
    41
    by (auto simp add: B_def setsum_reindex inj_on_xa_A)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    42
  also have "... = setsum (%x. p * ((x * a) div p)) A"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
    43
    by (auto simp add: o_def)
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
    44
  also from finite_A have "setsum (%x. p * ((x * a) div p)) A =
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    45
    p * setsum (%x. ((x * a) div p)) A"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
    46
    by (auto simp add: setsum_const_mult)
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
    47
  finally show ?thesis by arith
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    48
qed
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
    49
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
    50
lemma QRLemma2: "setsum id A = p * int (card E) - setsum id E +
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
    51
  setsum id D"
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    52
proof -
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    53
  from F_Un_D_eq_A have "setsum id A = setsum id (D \<union> F)"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
    54
    by (simp add: Un_commute)
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
    55
  also from F_D_disj finite_D finite_F
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
    56
  have "... = setsum id D + setsum id F"
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
    57
    by (auto simp add: Int_commute intro: setsum_Un_disjoint)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    58
  also from F_def have "F = (%x. (p - x)) ` E"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
    59
    by auto
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
    60
  also from finite_E inj_on_pminusx_E have "setsum id ((%x. (p - x)) ` E) =
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    61
      setsum (%x. (p - x)) E"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    62
    by (auto simp add: setsum_reindex)
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    63
  also from finite_E have "setsum (op - p) E = setsum (%x. p) E - setsum id E"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    64
    by (auto simp add: setsum_subtractf id_def)
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    65
  also from finite_E have "setsum (%x. p) E = p * int(card E)"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
    66
    by (intro setsum_const)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    67
  finally show ?thesis
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
    68
    by arith
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    69
qed
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
    70
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
    71
lemma QRLemma3: "(a - 1) * setsum id A =
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    72
    p * (setsum (%x. ((x * a) div p)) A - int(card E)) + 2 * setsum id E"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    73
proof -
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    74
  have "(a - 1) * setsum id A = a * setsum id A - setsum id A"
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
    75
    by (auto simp add: zdiff_zmult_distrib)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    76
  also note QRLemma1
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
    77
  also from QRLemma2 have "p * (\<Sum>x \<in> A. x * a div p) + setsum id D +
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
    78
     setsum id E - setsum id A =
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
    79
      p * (\<Sum>x \<in> A. x * a div p) + setsum id D +
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    80
      setsum id E - (p * int (card E) - setsum id E + setsum id D)"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
    81
    by auto
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
    82
  also have "... = p * (\<Sum>x \<in> A. x * a div p) -
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
    83
      p * int (card E) + 2 * setsum id E"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
    84
    by arith
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    85
  finally show ?thesis
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
    86
    by (auto simp only: zdiff_zmult_distrib2)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    87
qed
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
    88
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
    89
lemma QRLemma4: "a \<in> zOdd ==>
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    90
    (setsum (%x. ((x * a) div p)) A \<in> zEven) = (int(card E): zEven)"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    91
proof -
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    92
  assume a_odd: "a \<in> zOdd"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
    93
  from QRLemma3 have a: "p * (setsum (%x. ((x * a) div p)) A - int(card E)) =
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
    94
      (a - 1) * setsum id A - 2 * setsum id E"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
    95
    by arith
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
    96
  from a_odd have "a - 1 \<in> zEven"
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
    97
    by (rule odd_minus_one_even)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
    98
  hence "(a - 1) * setsum id A \<in> zEven"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
    99
    by (rule even_times_either)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   100
  moreover have "2 * setsum id E \<in> zEven"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   101
    by (auto simp add: zEven_def)
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   102
  ultimately have "(a - 1) * setsum id A - 2 * setsum id E \<in> zEven"
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   103
    by (rule even_minus_even)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   104
  with a have "p * (setsum (%x. ((x * a) div p)) A - int(card E)): zEven"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   105
    by simp
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   106
  hence "p \<in> zEven | (setsum (%x. ((x * a) div p)) A - int(card E)): zEven"
14434
5f14c1207499 patch to NumberTheory problems caused by Parity
paulson
parents: 14387
diff changeset
   107
    by (rule EvenOdd.even_product)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   108
  with p_odd have "(setsum (%x. ((x * a) div p)) A - int(card E)): zEven"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   109
    by (auto simp add: odd_iff_not_even)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   110
  thus ?thesis
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   111
    by (auto simp only: even_diff [symmetric])
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   112
qed
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   113
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   114
lemma QRLemma5: "a \<in> zOdd ==>
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   115
   (-1::int)^(card E) = (-1::int)^(nat(setsum (%x. ((x * a) div p)) A))"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   116
proof -
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   117
  assume "a \<in> zOdd"
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 23365
diff changeset
   118
  from QRLemma4 [OF this] have
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 23365
diff changeset
   119
    "(int(card E): zEven) = (setsum (%x. ((x * a) div p)) A \<in> zEven)" ..
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   120
  moreover have "0 \<le> int(card E)"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   121
    by auto
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   122
  moreover have "0 \<le> setsum (%x. ((x * a) div p)) A"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   123
    proof (intro setsum_nonneg)
15537
5538d3244b4d continued eliminating sumr
nipkow
parents: 15402
diff changeset
   124
      show "\<forall>x \<in> A. 0 \<le> x * a div p"
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   125
      proof
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   126
        fix x
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   127
        assume "x \<in> A"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   128
        then have "0 \<le> x"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   129
          by (auto simp add: A_def)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   130
        with a_nonzero have "0 \<le> x * a"
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 13871
diff changeset
   131
          by (auto simp add: zero_le_mult_iff)
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   132
        with p_g_2 show "0 \<le> x * a div p"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   133
          by (auto simp add: pos_imp_zdiv_nonneg_iff)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   134
      qed
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   135
    qed
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   136
  ultimately have "(-1::int)^nat((int (card E))) =
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   137
      (-1)^nat(((\<Sum>x \<in> A. x * a div p)))"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   138
    by (intro neg_one_power_parity, auto)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   139
  also have "nat (int(card E)) = card E"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   140
    by auto
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   141
  finally show ?thesis .
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   142
qed
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   143
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   144
end
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   145
16663
13e9c402308b prime is a predicate now.
nipkow
parents: 15541
diff changeset
   146
lemma MainQRLemma: "[| a \<in> zOdd; 0 < a; ~([a = 0] (mod p)); zprime p; 2 < p;
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   147
  A = {x. 0 < x & x \<le> (p - 1) div 2} |] ==>
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   148
  (Legendre a p) = (-1::int)^(nat(setsum (%x. ((x * a) div p)) A))"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   149
  apply (subst GAUSS.gauss_lemma)
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   150
  apply (auto simp add: GAUSS_def)
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   151
  apply (subst GAUSS.QRLemma5)
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   152
  apply (auto simp add: GAUSS_def)
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   153
  apply (simp add: GAUSS.A_def [OF GAUSS.intro] GAUSS_def)
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   154
  done
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   155
19670
2e4a143c73c5 prefer 'definition' over low-level defs;
wenzelm
parents: 18369
diff changeset
   156
2e4a143c73c5 prefer 'definition' over low-level defs;
wenzelm
parents: 18369
diff changeset
   157
subsection {* Stuff about S, S1 and S2 *}
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   158
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   159
locale QRTEMP =
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   160
  fixes p     :: "int"
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   161
  fixes q     :: "int"
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   162
16663
13e9c402308b prime is a predicate now.
nipkow
parents: 15541
diff changeset
   163
  assumes p_prime: "zprime p"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   164
  assumes p_g_2: "2 < p"
16663
13e9c402308b prime is a predicate now.
nipkow
parents: 15541
diff changeset
   165
  assumes q_prime: "zprime q"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   166
  assumes q_g_2: "2 < q"
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   167
  assumes p_neq_q:      "p \<noteq> q"
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   168
begin
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   169
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   170
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21288
diff changeset
   171
  P_set :: "int set" where
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   172
  "P_set = {x. 0 < x & x \<le> ((p - 1) div 2) }"
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   173
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21288
diff changeset
   174
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21288
diff changeset
   175
  Q_set :: "int set" where
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   176
  "Q_set = {x. 0 < x & x \<le> ((q - 1) div 2) }"
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   177
  
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21288
diff changeset
   178
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21288
diff changeset
   179
  S :: "(int * int) set" where
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   180
  "S = P_set <*> Q_set"
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   181
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21288
diff changeset
   182
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21288
diff changeset
   183
  S1 :: "(int * int) set" where
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   184
  "S1 = { (x, y). (x, y):S & ((p * y) < (q * x)) }"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   185
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21288
diff changeset
   186
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21288
diff changeset
   187
  S2 :: "(int * int) set" where
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   188
  "S2 = { (x, y). (x, y):S & ((q * x) < (p * y)) }"
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   189
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21288
diff changeset
   190
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21288
diff changeset
   191
  f1 :: "int => (int * int) set" where
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   192
  "f1 j = { (j1, y). (j1, y):S & j1 = j & (y \<le> (q * j) div p) }"
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   193
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21288
diff changeset
   194
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21288
diff changeset
   195
  f2 :: "int => (int * int) set" where
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   196
  "f2 j = { (x, j1). (x, j1):S & j1 = j & (x \<le> (p * j) div q) }"
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   197
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   198
lemma p_fact: "0 < (p - 1) div 2"
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   199
proof -
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   200
  from p_g_2 have "2 \<le> p - 1" by arith
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   201
  then have "2 div 2 \<le> (p - 1) div 2" by (rule zdiv_mono1, auto)
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   202
  then show ?thesis by auto
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   203
qed
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   204
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   205
lemma q_fact: "0 < (q - 1) div 2"
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   206
proof -
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   207
  from q_g_2 have "2 \<le> q - 1" by arith
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   208
  then have "2 div 2 \<le> (q - 1) div 2" by (rule zdiv_mono1, auto)
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   209
  then show ?thesis by auto
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   210
qed
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   211
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   212
lemma pb_neq_qa: "[|1 \<le> b; b \<le> (q - 1) div 2 |] ==>
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   213
    (p * b \<noteq> q * a)"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   214
proof
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   215
  assume "p * b = q * a" and "1 \<le> b" and "b \<le> (q - 1) div 2"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   216
  then have "q dvd (p * b)" by (auto simp add: dvd_def)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   217
  with q_prime p_g_2 have "q dvd p | q dvd b"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   218
    by (auto simp add: zprime_zdvd_zmult)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   219
  moreover have "~ (q dvd p)"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   220
  proof
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   221
    assume "q dvd p"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   222
    with p_prime have "q = 1 | q = p"
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   223
      apply (auto simp add: zprime_def QRTEMP_def)
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   224
      apply (drule_tac x = q and R = False in allE)
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   225
      apply (simp add: QRTEMP_def)
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   226
      apply (subgoal_tac "0 \<le> q", simp add: QRTEMP_def)
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   227
      apply (insert prems)
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   228
      apply (auto simp add: QRTEMP_def)
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   229
      done
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   230
    with q_g_2 p_neq_q show False by auto
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   231
  qed
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   232
  ultimately have "q dvd b" by auto
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   233
  then have "q \<le> b"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   234
  proof -
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   235
    assume "q dvd b"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   236
    moreover from prems have "0 < b" by auto
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   237
    ultimately show ?thesis using zdvd_bounds [of q b] by auto
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   238
  qed
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   239
  with prems have "q \<le> (q - 1) div 2" by auto
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   240
  then have "2 * q \<le> 2 * ((q - 1) div 2)" by arith
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   241
  then have "2 * q \<le> q - 1"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   242
  proof -
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   243
    assume "2 * q \<le> 2 * ((q - 1) div 2)"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   244
    with prems have "q \<in> zOdd" by (auto simp add: QRTEMP_def zprime_zOdd_eq_grt_2)
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   245
    with odd_minus_one_even have "(q - 1):zEven" by auto
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   246
    with even_div_2_prop2 have "(q - 1) = 2 * ((q - 1) div 2)" by auto
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   247
    with prems show ?thesis by auto
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   248
  qed
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   249
  then have p1: "q \<le> -1" by arith
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   250
  with q_g_2 show False by auto
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   251
qed
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   252
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   253
lemma P_set_finite: "finite (P_set)"
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   254
  using p_fact by (auto simp add: P_set_def bdd_int_set_l_le_finite)
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   255
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   256
lemma Q_set_finite: "finite (Q_set)"
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   257
  using q_fact by (auto simp add: Q_set_def bdd_int_set_l_le_finite)
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   258
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   259
lemma S_finite: "finite S"
15402
97204f3b4705 REorganized Finite_Set
nipkow
parents: 15392
diff changeset
   260
  by (auto simp add: S_def  P_set_finite Q_set_finite finite_cartesian_product)
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   261
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   262
lemma S1_finite: "finite S1"
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   263
proof -
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   264
  have "finite S" by (auto simp add: S_finite)
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   265
  moreover have "S1 \<subseteq> S" by (auto simp add: S1_def S_def)
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   266
  ultimately show ?thesis by (auto simp add: finite_subset)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   267
qed
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   268
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   269
lemma S2_finite: "finite S2"
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   270
proof -
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   271
  have "finite S" by (auto simp add: S_finite)
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   272
  moreover have "S2 \<subseteq> S" by (auto simp add: S2_def S_def)
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   273
  ultimately show ?thesis by (auto simp add: finite_subset)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   274
qed
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   275
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   276
lemma P_set_card: "(p - 1) div 2 = int (card (P_set))"
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   277
  using p_fact by (auto simp add: P_set_def card_bdd_int_set_l_le)
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   278
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   279
lemma Q_set_card: "(q - 1) div 2 = int (card (Q_set))"
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   280
  using q_fact by (auto simp add: Q_set_def card_bdd_int_set_l_le)
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   281
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   282
lemma S_card: "((p - 1) div 2) * ((q - 1) div 2) = int (card(S))"
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   283
  using P_set_card Q_set_card P_set_finite Q_set_finite
23431
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23373
diff changeset
   284
  by (auto simp add: S_def zmult_int setsum_constant)
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   285
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   286
lemma S1_Int_S2_prop: "S1 \<inter> S2 = {}"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   287
  by (auto simp add: S1_def S2_def)
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   288
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   289
lemma S1_Union_S2_prop: "S = S1 \<union> S2"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   290
  apply (auto simp add: S_def P_set_def Q_set_def S1_def S2_def)
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   291
proof -
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   292
  fix a and b
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   293
  assume "~ q * a < p * b" and b1: "0 < b" and b2: "b \<le> (q - 1) div 2"
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   294
  with zless_linear have "(p * b < q * a) | (p * b = q * a)" by auto
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   295
  moreover from pb_neq_qa b1 b2 have "(p * b \<noteq> q * a)" by auto
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   296
  ultimately show "p * b < q * a" by auto
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   297
qed
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   298
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   299
lemma card_sum_S1_S2: "((p - 1) div 2) * ((q - 1) div 2) =
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   300
    int(card(S1)) + int(card(S2))"
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   301
proof -
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   302
  have "((p - 1) div 2) * ((q - 1) div 2) = int (card(S))"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   303
    by (auto simp add: S_card)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   304
  also have "... = int( card(S1) + card(S2))"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   305
    apply (insert S1_finite S2_finite S1_Int_S2_prop S1_Union_S2_prop)
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   306
    apply (drule card_Un_disjoint, auto)
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   307
    done
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   308
  also have "... = int(card(S1)) + int(card(S2))" by auto
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   309
  finally show ?thesis .
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   310
qed
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   311
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   312
lemma aux1a: "[| 0 < a; a \<le> (p - 1) div 2;
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   313
                             0 < b; b \<le> (q - 1) div 2 |] ==>
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   314
                          (p * b < q * a) = (b \<le> q * a div p)"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   315
proof -
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   316
  assume "0 < a" and "a \<le> (p - 1) div 2" and "0 < b" and "b \<le> (q - 1) div 2"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   317
  have "p * b < q * a ==> b \<le> q * a div p"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   318
  proof -
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   319
    assume "p * b < q * a"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   320
    then have "p * b \<le> q * a" by auto
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   321
    then have "(p * b) div p \<le> (q * a) div p"
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   322
      by (rule zdiv_mono1) (insert p_g_2, auto)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   323
    then show "b \<le> (q * a) div p"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   324
      apply (subgoal_tac "p \<noteq> 0")
30240
blanchet
parents: 23431
diff changeset
   325
      apply (frule div_mult_self1_is_id, force)
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   326
      apply (insert p_g_2, auto)
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   327
      done
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   328
  qed
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   329
  moreover have "b \<le> q * a div p ==> p * b < q * a"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   330
  proof -
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   331
    assume "b \<le> q * a div p"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   332
    then have "p * b \<le> p * ((q * a) div p)"
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   333
      using p_g_2 by (auto simp add: mult_le_cancel_left)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   334
    also have "... \<le> q * a"
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   335
      by (rule zdiv_leq_prop) (insert p_g_2, auto)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   336
    finally have "p * b \<le> q * a" .
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   337
    then have "p * b < q * a | p * b = q * a"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   338
      by (simp only: order_le_imp_less_or_eq)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   339
    moreover have "p * b \<noteq> q * a"
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   340
      by (rule  pb_neq_qa) (insert prems, auto)
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   341
    ultimately show ?thesis by auto
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   342
  qed
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   343
  ultimately show ?thesis ..
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   344
qed
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   345
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   346
lemma aux1b: "[| 0 < a; a \<le> (p - 1) div 2;
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   347
                             0 < b; b \<le> (q - 1) div 2 |] ==>
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   348
                          (q * a < p * b) = (a \<le> p * b div q)"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   349
proof -
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   350
  assume "0 < a" and "a \<le> (p - 1) div 2" and "0 < b" and "b \<le> (q - 1) div 2"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   351
  have "q * a < p * b ==> a \<le> p * b div q"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   352
  proof -
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   353
    assume "q * a < p * b"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   354
    then have "q * a \<le> p * b" by auto
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   355
    then have "(q * a) div q \<le> (p * b) div q"
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   356
      by (rule zdiv_mono1) (insert q_g_2, auto)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   357
    then show "a \<le> (p * b) div q"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   358
      apply (subgoal_tac "q \<noteq> 0")
30240
blanchet
parents: 23431
diff changeset
   359
      apply (frule div_mult_self1_is_id, force)
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   360
      apply (insert q_g_2, auto)
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   361
      done
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   362
  qed
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   363
  moreover have "a \<le> p * b div q ==> q * a < p * b"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   364
  proof -
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   365
    assume "a \<le> p * b div q"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   366
    then have "q * a \<le> q * ((p * b) div q)"
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   367
      using q_g_2 by (auto simp add: mult_le_cancel_left)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   368
    also have "... \<le> p * b"
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   369
      by (rule zdiv_leq_prop) (insert q_g_2, auto)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   370
    finally have "q * a \<le> p * b" .
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   371
    then have "q * a < p * b | q * a = p * b"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   372
      by (simp only: order_le_imp_less_or_eq)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   373
    moreover have "p * b \<noteq> q * a"
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   374
      by (rule  pb_neq_qa) (insert prems, auto)
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   375
    ultimately show ?thesis by auto
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   376
  qed
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   377
  ultimately show ?thesis ..
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   378
qed
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   379
21288
wenzelm
parents: 21233
diff changeset
   380
lemma (in -) aux2: "[| zprime p; zprime q; 2 < p; 2 < q |] ==>
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   381
             (q * ((p - 1) div 2)) div p \<le> (q - 1) div 2"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   382
proof-
16663
13e9c402308b prime is a predicate now.
nipkow
parents: 15541
diff changeset
   383
  assume "zprime p" and "zprime q" and "2 < p" and "2 < q"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   384
  (* Set up what's even and odd *)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   385
  then have "p \<in> zOdd & q \<in> zOdd"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   386
    by (auto simp add:  zprime_zOdd_eq_grt_2)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   387
  then have even1: "(p - 1):zEven & (q - 1):zEven"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   388
    by (auto simp add: odd_minus_one_even)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   389
  then have even2: "(2 * p):zEven & ((q - 1) * p):zEven"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   390
    by (auto simp add: zEven_def)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   391
  then have even3: "(((q - 1) * p) + (2 * p)):zEven"
14434
5f14c1207499 patch to NumberTheory problems caused by Parity
paulson
parents: 14387
diff changeset
   392
    by (auto simp: EvenOdd.even_plus_even)
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   393
  (* using these prove it *)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   394
  from prems have "q * (p - 1) < ((q - 1) * p) + (2 * p)"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   395
    by (auto simp add: int_distrib)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   396
  then have "((p - 1) * q) div 2 < (((q - 1) * p) + (2 * p)) div 2"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   397
    apply (rule_tac x = "((p - 1) * q)" in even_div_2_l)
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   398
    by (auto simp add: even3, auto simp add: zmult_ac)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   399
  also have "((p - 1) * q) div 2 = q * ((p - 1) div 2)"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   400
    by (auto simp add: even1 even_prod_div_2)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   401
  also have "(((q - 1) * p) + (2 * p)) div 2 = (((q - 1) div 2) * p) + p"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   402
    by (auto simp add: even1 even2 even_prod_div_2 even_sum_div_2)
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   403
  finally show ?thesis
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   404
    apply (rule_tac x = " q * ((p - 1) div 2)" and
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   405
                    y = "(q - 1) div 2" in div_prop2)
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   406
    using prems by auto
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   407
qed
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   408
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   409
lemma aux3a: "\<forall>j \<in> P_set. int (card (f1 j)) = (q * j) div p"
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   410
proof
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   411
  fix j
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   412
  assume j_fact: "j \<in> P_set"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   413
  have "int (card (f1 j)) = int (card {y. y \<in> Q_set & y \<le> (q * j) div p})"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   414
  proof -
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   415
    have "finite (f1 j)"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   416
    proof -
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   417
      have "(f1 j) \<subseteq> S" by (auto simp add: f1_def)
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   418
      with S_finite show ?thesis by (auto simp add: finite_subset)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   419
    qed
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   420
    moreover have "inj_on (%(x,y). y) (f1 j)"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   421
      by (auto simp add: f1_def inj_on_def)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   422
    ultimately have "card ((%(x,y). y) ` (f1 j)) = card  (f1 j)"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   423
      by (auto simp add: f1_def card_image)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   424
    moreover have "((%(x,y). y) ` (f1 j)) = {y. y \<in> Q_set & y \<le> (q * j) div p}"
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   425
      using prems by (auto simp add: f1_def S_def Q_set_def P_set_def image_def)
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   426
    ultimately show ?thesis by (auto simp add: f1_def)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   427
  qed
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   428
  also have "... = int (card {y. 0 < y & y \<le> (q * j) div p})"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   429
  proof -
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   430
    have "{y. y \<in> Q_set & y \<le> (q * j) div p} =
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   431
        {y. 0 < y & y \<le> (q * j) div p}"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   432
      apply (auto simp add: Q_set_def)
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   433
    proof -
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   434
      fix x
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   435
      assume "0 < x" and "x \<le> q * j div p"
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   436
      with j_fact P_set_def  have "j \<le> (p - 1) div 2" by auto
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   437
      with q_g_2 have "q * j \<le> q * ((p - 1) div 2)"
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   438
        by (auto simp add: mult_le_cancel_left)
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   439
      with p_g_2 have "q * j div p \<le> q * ((p - 1) div 2) div p"
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   440
        by (auto simp add: zdiv_mono1)
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   441
      also from prems P_set_def have "... \<le> (q - 1) div 2"
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   442
        apply simp
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   443
        apply (insert aux2)
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   444
        apply (simp add: QRTEMP_def)
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   445
        done
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   446
      finally show "x \<le> (q - 1) div 2" using prems by auto
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   447
    qed
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   448
    then show ?thesis by auto
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   449
  qed
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   450
  also have "... = (q * j) div p"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   451
  proof -
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   452
    from j_fact P_set_def have "0 \<le> j" by auto
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14353
diff changeset
   453
    with q_g_2 have "q * 0 \<le> q * j" by (auto simp only: mult_left_mono)
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   454
    then have "0 \<le> q * j" by auto
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   455
    then have "0 div p \<le> (q * j) div p"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   456
      apply (rule_tac a = 0 in zdiv_mono1)
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   457
      apply (insert p_g_2, auto)
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   458
      done
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   459
    also have "0 div p = 0" by auto
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   460
    finally show ?thesis by (auto simp add: card_bdd_int_set_l_le)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   461
  qed
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   462
  finally show "int (card (f1 j)) = q * j div p" .
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   463
qed
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   464
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   465
lemma aux3b: "\<forall>j \<in> Q_set. int (card (f2 j)) = (p * j) div q"
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   466
proof
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   467
  fix j
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   468
  assume j_fact: "j \<in> Q_set"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   469
  have "int (card (f2 j)) = int (card {y. y \<in> P_set & y \<le> (p * j) div q})"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   470
  proof -
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   471
    have "finite (f2 j)"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   472
    proof -
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   473
      have "(f2 j) \<subseteq> S" by (auto simp add: f2_def)
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   474
      with S_finite show ?thesis by (auto simp add: finite_subset)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   475
    qed
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   476
    moreover have "inj_on (%(x,y). x) (f2 j)"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   477
      by (auto simp add: f2_def inj_on_def)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   478
    ultimately have "card ((%(x,y). x) ` (f2 j)) = card  (f2 j)"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   479
      by (auto simp add: f2_def card_image)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   480
    moreover have "((%(x,y). x) ` (f2 j)) = {y. y \<in> P_set & y \<le> (p * j) div q}"
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   481
      using prems by (auto simp add: f2_def S_def Q_set_def P_set_def image_def)
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   482
    ultimately show ?thesis by (auto simp add: f2_def)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   483
  qed
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   484
  also have "... = int (card {y. 0 < y & y \<le> (p * j) div q})"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   485
  proof -
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   486
    have "{y. y \<in> P_set & y \<le> (p * j) div q} =
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   487
        {y. 0 < y & y \<le> (p * j) div q}"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   488
      apply (auto simp add: P_set_def)
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   489
    proof -
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   490
      fix x
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   491
      assume "0 < x" and "x \<le> p * j div q"
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   492
      with j_fact Q_set_def  have "j \<le> (q - 1) div 2" by auto
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   493
      with p_g_2 have "p * j \<le> p * ((q - 1) div 2)"
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   494
        by (auto simp add: mult_le_cancel_left)
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   495
      with q_g_2 have "p * j div q \<le> p * ((q - 1) div 2) div q"
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   496
        by (auto simp add: zdiv_mono1)
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   497
      also from prems have "... \<le> (p - 1) div 2"
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   498
        by (auto simp add: aux2 QRTEMP_def)
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   499
      finally show "x \<le> (p - 1) div 2" using prems by auto
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   500
      qed
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   501
    then show ?thesis by auto
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   502
  qed
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   503
  also have "... = (p * j) div q"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   504
  proof -
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   505
    from j_fact Q_set_def have "0 \<le> j" by auto
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14353
diff changeset
   506
    with p_g_2 have "p * 0 \<le> p * j" by (auto simp only: mult_left_mono)
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   507
    then have "0 \<le> p * j" by auto
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   508
    then have "0 div q \<le> (p * j) div q"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   509
      apply (rule_tac a = 0 in zdiv_mono1)
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   510
      apply (insert q_g_2, auto)
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   511
      done
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   512
    also have "0 div q = 0" by auto
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   513
    finally show ?thesis by (auto simp add: card_bdd_int_set_l_le)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   514
  qed
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   515
  finally show "int (card (f2 j)) = p * j div q" .
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   516
qed
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   517
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   518
lemma S1_card: "int (card(S1)) = setsum (%j. (q * j) div p) P_set"
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   519
proof -
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   520
  have "\<forall>x \<in> P_set. finite (f1 x)"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   521
  proof
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   522
    fix x
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   523
    have "f1 x \<subseteq> S" by (auto simp add: f1_def)
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   524
    with S_finite show "finite (f1 x)" by (auto simp add: finite_subset)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   525
  qed
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   526
  moreover have "(\<forall>x \<in> P_set. \<forall>y \<in> P_set. x \<noteq> y --> (f1 x) \<inter> (f1 y) = {})"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   527
    by (auto simp add: f1_def)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   528
  moreover note P_set_finite
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   529
  ultimately have "int(card (UNION P_set f1)) =
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   530
      setsum (%x. int(card (f1 x))) P_set"
15402
97204f3b4705 REorganized Finite_Set
nipkow
parents: 15392
diff changeset
   531
    by(simp add:card_UN_disjoint int_setsum o_def)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   532
  moreover have "S1 = UNION P_set f1"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   533
    by (auto simp add: f1_def S_def S1_def S2_def P_set_def Q_set_def aux1a)
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   534
  ultimately have "int(card (S1)) = setsum (%j. int(card (f1 j))) P_set"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   535
    by auto
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   536
  also have "... = setsum (%j. q * j div p) P_set"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   537
    using aux3a by(fastsimp intro: setsum_cong)
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   538
  finally show ?thesis .
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   539
qed
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   540
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   541
lemma S2_card: "int (card(S2)) = setsum (%j. (p * j) div q) Q_set"
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   542
proof -
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   543
  have "\<forall>x \<in> Q_set. finite (f2 x)"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   544
  proof
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   545
    fix x
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   546
    have "f2 x \<subseteq> S" by (auto simp add: f2_def)
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   547
    with S_finite show "finite (f2 x)" by (auto simp add: finite_subset)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   548
  qed
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   549
  moreover have "(\<forall>x \<in> Q_set. \<forall>y \<in> Q_set. x \<noteq> y -->
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   550
      (f2 x) \<inter> (f2 y) = {})"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   551
    by (auto simp add: f2_def)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   552
  moreover note Q_set_finite
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   553
  ultimately have "int(card (UNION Q_set f2)) =
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   554
      setsum (%x. int(card (f2 x))) Q_set"
15402
97204f3b4705 REorganized Finite_Set
nipkow
parents: 15392
diff changeset
   555
    by(simp add:card_UN_disjoint int_setsum o_def)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   556
  moreover have "S2 = UNION Q_set f2"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   557
    by (auto simp add: f2_def S_def S1_def S2_def P_set_def Q_set_def aux1b)
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   558
  ultimately have "int(card (S2)) = setsum (%j. int(card (f2 j))) Q_set"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   559
    by auto
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   560
  also have "... = setsum (%j. p * j div q) Q_set"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   561
    using aux3b by(fastsimp intro: setsum_cong)
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   562
  finally show ?thesis .
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   563
qed
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   564
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   565
lemma S1_carda: "int (card(S1)) =
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   566
    setsum (%j. (j * q) div p) P_set"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   567
  by (auto simp add: S1_card zmult_ac)
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   568
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   569
lemma S2_carda: "int (card(S2)) =
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   570
    setsum (%j. (j * p) div q) Q_set"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   571
  by (auto simp add: S2_card zmult_ac)
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   572
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   573
lemma pq_sum_prop: "(setsum (%j. (j * p) div q) Q_set) +
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   574
    (setsum (%j. (j * q) div p) P_set) = ((p - 1) div 2) * ((q - 1) div 2)"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   575
proof -
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   576
  have "(setsum (%j. (j * p) div q) Q_set) +
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   577
      (setsum (%j. (j * q) div p) P_set) = int (card S2) + int (card S1)"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   578
    by (auto simp add: S1_carda S2_carda)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   579
  also have "... = int (card S1) + int (card S2)"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   580
    by auto
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   581
  also have "... = ((p - 1) div 2) * ((q - 1) div 2)"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   582
    by (auto simp add: card_sum_S1_S2)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   583
  finally show ?thesis .
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   584
qed
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   585
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   586
21288
wenzelm
parents: 21233
diff changeset
   587
lemma (in -) pq_prime_neq: "[| zprime p; zprime q; p \<noteq> q |] ==> (~[p = 0] (mod q))"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   588
  apply (auto simp add: zcong_eq_zdvd_prop zprime_def)
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   589
  apply (drule_tac x = q in allE)
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   590
  apply (drule_tac x = p in allE)
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   591
  apply auto
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   592
  done
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   593
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   594
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   595
lemma QR_short: "(Legendre p q) * (Legendre q p) =
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   596
    (-1::int)^nat(((p - 1) div 2)*((q - 1) div 2))"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   597
proof -
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   598
  from prems have "~([p = 0] (mod q))"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   599
    by (auto simp add: pq_prime_neq QRTEMP_def)
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   600
  with prems Q_set_def have a1: "(Legendre p q) = (-1::int) ^
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   601
      nat(setsum (%x. ((x * p) div q)) Q_set)"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   602
    apply (rule_tac p = q in  MainQRLemma)
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   603
    apply (auto simp add: zprime_zOdd_eq_grt_2 QRTEMP_def)
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   604
    done
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   605
  from prems have "~([q = 0] (mod p))"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   606
    apply (rule_tac p = q and q = p in pq_prime_neq)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   607
    apply (simp add: QRTEMP_def)+
16733
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16663
diff changeset
   608
    done
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   609
  with prems P_set_def have a2: "(Legendre q p) =
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   610
      (-1::int) ^ nat(setsum (%x. ((x * q) div p)) P_set)"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   611
    apply (rule_tac p = p in  MainQRLemma)
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   612
    apply (auto simp add: zprime_zOdd_eq_grt_2 QRTEMP_def)
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   613
    done
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   614
  from a1 a2 have "(Legendre p q) * (Legendre q p) =
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   615
      (-1::int) ^ nat(setsum (%x. ((x * p) div q)) Q_set) *
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   616
        (-1::int) ^ nat(setsum (%x. ((x * q) div p)) P_set)"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   617
    by auto
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   618
  also have "... = (-1::int) ^ (nat(setsum (%x. ((x * p) div q)) Q_set) +
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   619
                   nat(setsum (%x. ((x * q) div p)) P_set))"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   620
    by (auto simp add: zpower_zadd_distrib)
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   621
  also have "nat(setsum (%x. ((x * p) div q)) Q_set) +
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   622
      nat(setsum (%x. ((x * q) div p)) P_set) =
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   623
        nat((setsum (%x. ((x * p) div q)) Q_set) +
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   624
          (setsum (%x. ((x * q) div p)) P_set))"
20898
113c9516a2d7 attribute symmetric: zero_var_indexes;
wenzelm
parents: 20432
diff changeset
   625
    apply (rule_tac z = "setsum (%x. ((x * p) div q)) Q_set" in
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   626
      nat_add_distrib [symmetric])
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   627
    apply (auto simp add: S1_carda [symmetric] S2_carda [symmetric])
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   628
    done
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   629
  also have "... = nat(((p - 1) div 2) * ((q - 1) div 2))"
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   630
    by (auto simp add: pq_sum_prop)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   631
  finally show ?thesis .
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   632
qed
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   633
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   634
end
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 20898
diff changeset
   635
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   636
theorem Quadratic_Reciprocity:
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   637
     "[| p \<in> zOdd; zprime p; q \<in> zOdd; zprime q;
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   638
         p \<noteq> q |]
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   639
      ==> (Legendre p q) * (Legendre q p) =
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 14981
diff changeset
   640
          (-1::int)^nat(((p - 1) div 2)*((q - 1) div 2))"
18369
694ea14ab4f2 tuned sources and proofs
wenzelm
parents: 16733
diff changeset
   641
  by (auto simp add: QRTEMP.QR_short zprime_zOdd_eq_grt_2 [symmetric]
13871
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   642
                     QRTEMP_def)
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   643
26e5f5e624f6 Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff changeset
   644
end