author | urbanc |
Tue, 13 Dec 2005 18:11:21 +0100 | |
changeset 18396 | b3e7da94b51f |
parent 18369 | 694ea14ab4f2 |
child 20217 | 25b068a99d2b |
permissions | -rw-r--r-- |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
1 |
(* Title: HOL/Quadratic_Reciprocity/Gauss.thy |
14981 | 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 {* Gauss' Lemma *} |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
7 |
|
18369 | 8 |
theory Gauss imports Euler begin |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
9 |
|
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
10 |
locale GAUSS = |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
11 |
fixes p :: "int" |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
12 |
fixes a :: "int" |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
13 |
fixes A :: "int set" |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
14 |
fixes B :: "int set" |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
15 |
fixes C :: "int set" |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
16 |
fixes D :: "int set" |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
17 |
fixes E :: "int set" |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
18 |
fixes F :: "int set" |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
19 |
|
16663 | 20 |
assumes p_prime: "zprime p" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
21 |
assumes p_g_2: "2 < p" |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
22 |
assumes p_a_relprime: "~[a = 0](mod p)" |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
23 |
assumes a_nonzero: "0 < a" |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
24 |
|
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
25 |
defines A_def: "A == {(x::int). 0 < x & x \<le> ((p - 1) div 2)}" |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
26 |
defines B_def: "B == (%x. x * a) ` A" |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
27 |
defines C_def: "C == (StandardRes p) ` B" |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
28 |
defines D_def: "D == C \<inter> {x. x \<le> ((p - 1) div 2)}" |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
29 |
defines E_def: "E == C \<inter> {x. ((p - 1) div 2) < x}" |
18369 | 30 |
defines F_def: "F == (%x. (p - x)) ` E" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
31 |
|
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
32 |
subsection {* Basic properties of p *} |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
33 |
|
18369 | 34 |
lemma (in GAUSS) p_odd: "p \<in> zOdd" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
35 |
by (auto simp add: p_prime p_g_2 zprime_zOdd_eq_grt_2) |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
36 |
|
18369 | 37 |
lemma (in GAUSS) p_g_0: "0 < p" |
38 |
using p_g_2 by auto |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
39 |
|
18369 | 40 |
lemma (in GAUSS) int_nat: "int (nat ((p - 1) div 2)) = (p - 1) div 2" |
41 |
using insert p_g_2 by (auto simp add: pos_imp_zdiv_nonneg_iff) |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
42 |
|
18369 | 43 |
lemma (in GAUSS) p_minus_one_l: "(p - 1) div 2 < p" |
44 |
proof - |
|
45 |
have "(p - 1) div 2 \<le> (p - 1) div 1" |
|
46 |
by (rule zdiv_mono2) (auto simp add: p_g_0) |
|
47 |
also have "\<dots> = p - 1" by simp |
|
48 |
finally show ?thesis by simp |
|
49 |
qed |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
50 |
|
18369 | 51 |
lemma (in GAUSS) p_eq: "p = (2 * (p - 1) div 2) + 1" |
52 |
using zdiv_zmult_self2 [of 2 "p - 1"] by auto |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
53 |
|
18369 | 54 |
lemma zodd_imp_zdiv_eq: "x \<in> zOdd ==> 2 * (x - 1) div 2 = 2 * ((x - 1) div 2)" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
55 |
apply (frule odd_minus_one_even) |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
56 |
apply (simp add: zEven_def) |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
57 |
apply (subgoal_tac "2 \<noteq> 0") |
18369 | 58 |
apply (frule_tac b = "2 :: int" and a = "x - 1" in zdiv_zmult_self2) |
59 |
apply (auto simp add: even_div_2_prop2) |
|
60 |
done |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
61 |
|
18369 | 62 |
lemma (in GAUSS) p_eq2: "p = (2 * ((p - 1) div 2)) + 1" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
63 |
apply (insert p_eq p_prime p_g_2 zprime_zOdd_eq_grt_2 [of p], auto) |
18369 | 64 |
apply (frule zodd_imp_zdiv_eq, auto) |
65 |
done |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
66 |
|
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
67 |
subsection {* Basic Properties of the Gauss Sets *} |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
68 |
|
18369 | 69 |
lemma (in GAUSS) finite_A: "finite (A)" |
70 |
apply (auto simp add: A_def) |
|
71 |
apply (subgoal_tac "{x. 0 < x & x \<le> (p - 1) div 2} \<subseteq> {x. 0 \<le> x & x < 1 + (p - 1) div 2}") |
|
72 |
apply (auto simp add: bdd_int_set_l_finite finite_subset) |
|
73 |
done |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
74 |
|
18369 | 75 |
lemma (in GAUSS) finite_B: "finite (B)" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
76 |
by (auto simp add: B_def finite_A finite_imageI) |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
77 |
|
18369 | 78 |
lemma (in GAUSS) finite_C: "finite (C)" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
79 |
by (auto simp add: C_def finite_B finite_imageI) |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
80 |
|
18369 | 81 |
lemma (in GAUSS) finite_D: "finite (D)" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
82 |
by (auto simp add: D_def finite_Int finite_C) |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
83 |
|
18369 | 84 |
lemma (in GAUSS) finite_E: "finite (E)" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
85 |
by (auto simp add: E_def finite_Int finite_C) |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
86 |
|
18369 | 87 |
lemma (in GAUSS) finite_F: "finite (F)" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
88 |
by (auto simp add: F_def finite_E finite_imageI) |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
89 |
|
18369 | 90 |
lemma (in GAUSS) C_eq: "C = D \<union> E" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
91 |
by (auto simp add: C_def D_def E_def) |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
92 |
|
18369 | 93 |
lemma (in GAUSS) A_card_eq: "card A = nat ((p - 1) div 2)" |
94 |
apply (auto simp add: A_def) |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
95 |
apply (insert int_nat) |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
96 |
apply (erule subst) |
18369 | 97 |
apply (auto simp add: card_bdd_int_set_l_le) |
98 |
done |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
99 |
|
18369 | 100 |
lemma (in GAUSS) inj_on_xa_A: "inj_on (%x. x * a) A" |
101 |
using a_nonzero by (simp add: A_def inj_on_def) |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
102 |
|
18369 | 103 |
lemma (in GAUSS) A_res: "ResSet p A" |
104 |
apply (auto simp add: A_def ResSet_def) |
|
105 |
apply (rule_tac m = p in zcong_less_eq) |
|
106 |
apply (insert p_g_2, auto) |
|
107 |
apply (subgoal_tac [1-2] "(p - 1) div 2 < p") |
|
108 |
apply (auto, auto simp add: p_minus_one_l) |
|
109 |
done |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
110 |
|
18369 | 111 |
lemma (in GAUSS) B_res: "ResSet p B" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
112 |
apply (insert p_g_2 p_a_relprime p_minus_one_l) |
18369 | 113 |
apply (auto simp add: B_def) |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
114 |
apply (rule ResSet_image) |
18369 | 115 |
apply (auto simp add: A_res) |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
116 |
apply (auto simp add: A_def) |
18369 | 117 |
proof - |
118 |
fix x fix y |
|
119 |
assume a: "[x * a = y * a] (mod p)" |
|
120 |
assume b: "0 < x" |
|
121 |
assume c: "x \<le> (p - 1) div 2" |
|
122 |
assume d: "0 < y" |
|
123 |
assume e: "y \<le> (p - 1) div 2" |
|
124 |
from a p_a_relprime p_prime a_nonzero zcong_cancel [of p a x y] |
|
125 |
have "[x = y](mod p)" |
|
126 |
by (simp add: zprime_imp_zrelprime zcong_def p_g_0 order_le_less) |
|
127 |
with zcong_less_eq [of x y p] p_minus_one_l |
|
128 |
order_le_less_trans [of x "(p - 1) div 2" p] |
|
129 |
order_le_less_trans [of y "(p - 1) div 2" p] show "x = y" |
|
130 |
by (simp add: prems p_minus_one_l p_g_0) |
|
131 |
qed |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
132 |
|
18369 | 133 |
lemma (in GAUSS) SR_B_inj: "inj_on (StandardRes p) B" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
134 |
apply (auto simp add: B_def StandardRes_def inj_on_def A_def prems) |
18369 | 135 |
proof - |
136 |
fix x fix y |
|
137 |
assume a: "x * a mod p = y * a mod p" |
|
138 |
assume b: "0 < x" |
|
139 |
assume c: "x \<le> (p - 1) div 2" |
|
140 |
assume d: "0 < y" |
|
141 |
assume e: "y \<le> (p - 1) div 2" |
|
142 |
assume f: "x \<noteq> y" |
|
143 |
from a have "[x * a = y * a](mod p)" |
|
144 |
by (simp add: zcong_zmod_eq p_g_0) |
|
145 |
with p_a_relprime p_prime a_nonzero zcong_cancel [of p a x y] |
|
146 |
have "[x = y](mod p)" |
|
147 |
by (simp add: zprime_imp_zrelprime zcong_def p_g_0 order_le_less) |
|
148 |
with zcong_less_eq [of x y p] p_minus_one_l |
|
149 |
order_le_less_trans [of x "(p - 1) div 2" p] |
|
150 |
order_le_less_trans [of y "(p - 1) div 2" p] have "x = y" |
|
151 |
by (simp add: prems p_minus_one_l p_g_0) |
|
152 |
then have False |
|
153 |
by (simp add: f) |
|
154 |
then show "a = 0" |
|
155 |
by simp |
|
156 |
qed |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
157 |
|
18369 | 158 |
lemma (in GAUSS) inj_on_pminusx_E: "inj_on (%x. p - x) E" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
159 |
apply (auto simp add: E_def C_def B_def A_def) |
18369 | 160 |
apply (rule_tac g = "%x. -1 * (x - p)" in inj_on_inverseI) |
161 |
apply auto |
|
162 |
done |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
163 |
|
18369 | 164 |
lemma (in GAUSS) A_ncong_p: "x \<in> A ==> ~[x = 0](mod p)" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
165 |
apply (auto simp add: A_def) |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
166 |
apply (frule_tac m = p in zcong_not_zero) |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
167 |
apply (insert p_minus_one_l) |
18369 | 168 |
apply auto |
169 |
done |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
170 |
|
18369 | 171 |
lemma (in GAUSS) A_greater_zero: "x \<in> A ==> 0 < x" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
172 |
by (auto simp add: A_def) |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
173 |
|
18369 | 174 |
lemma (in GAUSS) B_ncong_p: "x \<in> B ==> ~[x = 0](mod p)" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
175 |
apply (auto simp add: B_def) |
18369 | 176 |
apply (frule A_ncong_p) |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
177 |
apply (insert p_a_relprime p_prime a_nonzero) |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
178 |
apply (frule_tac a = x and b = a in zcong_zprime_prod_zero_contra) |
18369 | 179 |
apply (auto simp add: A_greater_zero) |
180 |
done |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
181 |
|
18369 | 182 |
lemma (in GAUSS) B_greater_zero: "x \<in> B ==> 0 < x" |
183 |
using a_nonzero by (auto simp add: B_def mult_pos_pos A_greater_zero) |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
184 |
|
18369 | 185 |
lemma (in GAUSS) C_ncong_p: "x \<in> C ==> ~[x = 0](mod p)" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
186 |
apply (auto simp add: C_def) |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
187 |
apply (frule B_ncong_p) |
18369 | 188 |
apply (subgoal_tac "[x = StandardRes p x](mod p)") |
189 |
defer apply (simp add: StandardRes_prop1) |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
190 |
apply (frule_tac a = x and b = "StandardRes p x" and c = 0 in zcong_trans) |
18369 | 191 |
apply auto |
192 |
done |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
193 |
|
18369 | 194 |
lemma (in GAUSS) C_greater_zero: "y \<in> C ==> 0 < y" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
195 |
apply (auto simp add: C_def) |
18369 | 196 |
proof - |
197 |
fix x |
|
198 |
assume a: "x \<in> B" |
|
199 |
from p_g_0 have "0 \<le> StandardRes p x" |
|
200 |
by (simp add: StandardRes_lbound) |
|
201 |
moreover have "~[x = 0] (mod p)" |
|
202 |
by (simp add: a B_ncong_p) |
|
203 |
then have "StandardRes p x \<noteq> 0" |
|
204 |
by (simp add: StandardRes_prop3) |
|
205 |
ultimately show "0 < StandardRes p x" |
|
206 |
by (simp add: order_le_less) |
|
207 |
qed |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
208 |
|
18369 | 209 |
lemma (in GAUSS) D_ncong_p: "x \<in> D ==> ~[x = 0](mod p)" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
210 |
by (auto simp add: D_def C_ncong_p) |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
211 |
|
18369 | 212 |
lemma (in GAUSS) E_ncong_p: "x \<in> E ==> ~[x = 0](mod p)" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
213 |
by (auto simp add: E_def C_ncong_p) |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
214 |
|
18369 | 215 |
lemma (in GAUSS) F_ncong_p: "x \<in> F ==> ~[x = 0](mod p)" |
216 |
apply (auto simp add: F_def) |
|
217 |
proof - |
|
218 |
fix x assume a: "x \<in> E" assume b: "[p - x = 0] (mod p)" |
|
219 |
from E_ncong_p have "~[x = 0] (mod p)" |
|
220 |
by (simp add: a) |
|
221 |
moreover from a have "0 < x" |
|
222 |
by (simp add: a E_def C_greater_zero) |
|
223 |
moreover from a have "x < p" |
|
224 |
by (auto simp add: E_def C_def p_g_0 StandardRes_ubound) |
|
225 |
ultimately have "~[p - x = 0] (mod p)" |
|
226 |
by (simp add: zcong_not_zero) |
|
227 |
from this show False by (simp add: b) |
|
228 |
qed |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
229 |
|
18369 | 230 |
lemma (in GAUSS) F_subset: "F \<subseteq> {x. 0 < x & x \<le> ((p - 1) div 2)}" |
231 |
apply (auto simp add: F_def E_def) |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
232 |
apply (insert p_g_0) |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
233 |
apply (frule_tac x = xa in StandardRes_ubound) |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
234 |
apply (frule_tac x = x in StandardRes_ubound) |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
235 |
apply (subgoal_tac "xa = StandardRes p xa") |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
236 |
apply (auto simp add: C_def StandardRes_prop2 StandardRes_prop1) |
18369 | 237 |
proof - |
238 |
from zodd_imp_zdiv_eq p_prime p_g_2 zprime_zOdd_eq_grt_2 have |
|
239 |
"2 * (p - 1) div 2 = 2 * ((p - 1) div 2)" |
|
240 |
by simp |
|
241 |
with p_eq2 show " !!x. [| (p - 1) div 2 < StandardRes p x; x \<in> B |] |
|
242 |
==> p - StandardRes p x \<le> (p - 1) div 2" |
|
243 |
by simp |
|
244 |
qed |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
245 |
|
18369 | 246 |
lemma (in GAUSS) D_subset: "D \<subseteq> {x. 0 < x & x \<le> ((p - 1) div 2)}" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
247 |
by (auto simp add: D_def C_greater_zero) |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
248 |
|
18369 | 249 |
lemma (in GAUSS) F_eq: "F = {x. \<exists>y \<in> A. ( x = p - (StandardRes p (y*a)) & (p - 1) div 2 < StandardRes p (y*a))}" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
250 |
by (auto simp add: F_def E_def D_def C_def B_def A_def) |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
251 |
|
18369 | 252 |
lemma (in GAUSS) D_eq: "D = {x. \<exists>y \<in> A. ( x = StandardRes p (y*a) & StandardRes p (y*a) \<le> (p - 1) div 2)}" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
253 |
by (auto simp add: D_def C_def B_def A_def) |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
254 |
|
18369 | 255 |
lemma (in GAUSS) D_leq: "x \<in> D ==> x \<le> (p - 1) div 2" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
256 |
by (auto simp add: D_eq) |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
257 |
|
18369 | 258 |
lemma (in GAUSS) F_ge: "x \<in> F ==> x \<le> (p - 1) div 2" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
259 |
apply (auto simp add: F_eq A_def) |
18369 | 260 |
proof - |
261 |
fix y |
|
262 |
assume "(p - 1) div 2 < StandardRes p (y * a)" |
|
263 |
then have "p - StandardRes p (y * a) < p - ((p - 1) div 2)" |
|
264 |
by arith |
|
265 |
also from p_eq2 have "... = 2 * ((p - 1) div 2) + 1 - ((p - 1) div 2)" |
|
266 |
by auto |
|
267 |
also have "2 * ((p - 1) div 2) + 1 - (p - 1) div 2 = (p - 1) div 2 + 1" |
|
268 |
by arith |
|
269 |
finally show "p - StandardRes p (y * a) \<le> (p - 1) div 2" |
|
270 |
using zless_add1_eq [of "p - StandardRes p (y * a)" "(p - 1) div 2"] by auto |
|
271 |
qed |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
272 |
|
18369 | 273 |
lemma (in GAUSS) all_A_relprime: "\<forall>x \<in> A. zgcd(x, p) = 1" |
274 |
using p_prime p_minus_one_l by (auto simp add: A_def zless_zprime_imp_zrelprime) |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
275 |
|
18369 | 276 |
lemma (in GAUSS) A_prod_relprime: "zgcd((setprod id A),p) = 1" |
277 |
using all_A_relprime finite_A by (simp add: all_relprime_prod_relprime) |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
278 |
|
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
279 |
subsection {* Relationships Between Gauss Sets *} |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
280 |
|
18369 | 281 |
lemma (in GAUSS) B_card_eq_A: "card B = card A" |
282 |
using finite_A by (simp add: finite_A B_def inj_on_xa_A card_image) |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
283 |
|
18369 | 284 |
lemma (in GAUSS) B_card_eq: "card B = nat ((p - 1) div 2)" |
285 |
by (simp add: B_card_eq_A A_card_eq) |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
286 |
|
18369 | 287 |
lemma (in GAUSS) F_card_eq_E: "card F = card E" |
288 |
using finite_E by (simp add: F_def inj_on_pminusx_E card_image) |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
289 |
|
18369 | 290 |
lemma (in GAUSS) C_card_eq_B: "card C = card B" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
291 |
apply (insert finite_B) |
18369 | 292 |
apply (subgoal_tac "inj_on (StandardRes p) B") |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
293 |
apply (simp add: B_def C_def card_image) |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
294 |
apply (rule StandardRes_inj_on_ResSet) |
18369 | 295 |
apply (simp add: B_res) |
296 |
done |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
297 |
|
18369 | 298 |
lemma (in GAUSS) D_E_disj: "D \<inter> E = {}" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
299 |
by (auto simp add: D_def E_def) |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
300 |
|
18369 | 301 |
lemma (in GAUSS) C_card_eq_D_plus_E: "card C = card D + card E" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
302 |
by (auto simp add: C_eq card_Un_disjoint D_E_disj finite_D finite_E) |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
303 |
|
18369 | 304 |
lemma (in GAUSS) C_prod_eq_D_times_E: "setprod id E * setprod id D = setprod id C" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
305 |
apply (insert D_E_disj finite_D finite_E C_eq) |
15392 | 306 |
apply (frule setprod_Un_disjoint [of D E id]) |
18369 | 307 |
apply auto |
308 |
done |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
309 |
|
18369 | 310 |
lemma (in GAUSS) C_B_zcong_prod: "[setprod id C = setprod id B] (mod p)" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
311 |
apply (auto simp add: C_def) |
18369 | 312 |
apply (insert finite_B SR_B_inj) |
313 |
apply (frule_tac f1 = "StandardRes p" in setprod_reindex_id [symmetric], auto) |
|
15392 | 314 |
apply (rule setprod_same_function_zcong) |
18369 | 315 |
apply (auto simp add: StandardRes_prop1 zcong_sym p_g_0) |
316 |
done |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
317 |
|
18369 | 318 |
lemma (in GAUSS) F_Un_D_subset: "(F \<union> D) \<subseteq> A" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
319 |
apply (rule Un_least) |
18369 | 320 |
apply (auto simp add: A_def F_subset D_subset) |
321 |
done |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
322 |
|
18369 | 323 |
lemma two_eq: "2 * (x::int) = x + x" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
324 |
by arith |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
325 |
|
18369 | 326 |
lemma (in GAUSS) F_D_disj: "(F \<inter> D) = {}" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
327 |
apply (simp add: F_eq D_eq) |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
328 |
apply (auto simp add: F_eq D_eq) |
18369 | 329 |
proof - |
330 |
fix y fix ya |
|
331 |
assume "p - StandardRes p (y * a) = StandardRes p (ya * a)" |
|
332 |
then have "p = StandardRes p (y * a) + StandardRes p (ya * a)" |
|
333 |
by arith |
|
334 |
moreover have "p dvd p" |
|
335 |
by auto |
|
336 |
ultimately have "p dvd (StandardRes p (y * a) + StandardRes p (ya * a))" |
|
337 |
by auto |
|
338 |
then have a: "[StandardRes p (y * a) + StandardRes p (ya * a) = 0] (mod p)" |
|
339 |
by (auto simp add: zcong_def) |
|
340 |
have "[y * a = StandardRes p (y * a)] (mod p)" |
|
341 |
by (simp only: zcong_sym StandardRes_prop1) |
|
342 |
moreover have "[ya * a = StandardRes p (ya * a)] (mod p)" |
|
343 |
by (simp only: zcong_sym StandardRes_prop1) |
|
344 |
ultimately have "[y * a + ya * a = |
|
345 |
StandardRes p (y * a) + StandardRes p (ya * a)] (mod p)" |
|
346 |
by (rule zcong_zadd) |
|
347 |
with a have "[y * a + ya * a = 0] (mod p)" |
|
348 |
apply (elim zcong_trans) |
|
349 |
by (simp only: zcong_refl) |
|
350 |
also have "y * a + ya * a = a * (y + ya)" |
|
351 |
by (simp add: zadd_zmult_distrib2 zmult_commute) |
|
352 |
finally have "[a * (y + ya) = 0] (mod p)" . |
|
353 |
with p_prime a_nonzero zcong_zprime_prod_zero [of p a "y + ya"] |
|
354 |
p_a_relprime |
|
355 |
have a: "[y + ya = 0] (mod p)" |
|
356 |
by auto |
|
357 |
assume b: "y \<in> A" and c: "ya: A" |
|
358 |
with A_def have "0 < y + ya" |
|
359 |
by auto |
|
360 |
moreover from b c A_def have "y + ya \<le> (p - 1) div 2 + (p - 1) div 2" |
|
361 |
by auto |
|
362 |
moreover from b c p_eq2 A_def have "y + ya < p" |
|
363 |
by auto |
|
364 |
ultimately show False |
|
365 |
apply simp |
|
366 |
apply (frule_tac m = p in zcong_not_zero) |
|
367 |
apply (auto simp add: a) |
|
368 |
done |
|
369 |
qed |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
370 |
|
18369 | 371 |
lemma (in GAUSS) F_Un_D_card: "card (F \<union> D) = nat ((p - 1) div 2)" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
372 |
apply (insert F_D_disj finite_F finite_D) |
18369 | 373 |
proof - |
374 |
have "card (F \<union> D) = card E + card D" |
|
375 |
by (auto simp add: finite_F finite_D F_D_disj |
|
376 |
card_Un_disjoint F_card_eq_E) |
|
377 |
then have "card (F \<union> D) = card C" |
|
378 |
by (simp add: C_card_eq_D_plus_E) |
|
379 |
from this show "card (F \<union> D) = nat ((p - 1) div 2)" |
|
380 |
by (simp add: C_card_eq_B B_card_eq) |
|
381 |
qed |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
382 |
|
18369 | 383 |
lemma (in GAUSS) F_Un_D_eq_A: "F \<union> D = A" |
384 |
using finite_A F_Un_D_subset A_card_eq F_Un_D_card by (auto simp add: card_seteq) |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
385 |
|
18369 | 386 |
lemma (in GAUSS) prod_D_F_eq_prod_A: |
387 |
"(setprod id D) * (setprod id F) = setprod id A" |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
388 |
apply (insert F_D_disj finite_D finite_F) |
15392 | 389 |
apply (frule setprod_Un_disjoint [of F D id]) |
18369 | 390 |
apply (auto simp add: F_Un_D_eq_A) |
391 |
done |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
392 |
|
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
393 |
lemma (in GAUSS) prod_F_zcong: |
18369 | 394 |
"[setprod id F = ((-1) ^ (card E)) * (setprod id E)] (mod p)" |
395 |
proof - |
|
396 |
have "setprod id F = setprod id (op - p ` E)" |
|
397 |
by (auto simp add: F_def) |
|
398 |
then have "setprod id F = setprod (op - p) E" |
|
399 |
apply simp |
|
400 |
apply (insert finite_E inj_on_pminusx_E) |
|
401 |
apply (frule_tac f = "op - p" in setprod_reindex_id, auto) |
|
402 |
done |
|
403 |
then have one: |
|
404 |
"[setprod id F = setprod (StandardRes p o (op - p)) E] (mod p)" |
|
405 |
apply simp |
|
406 |
apply (insert p_g_0 finite_E) |
|
407 |
by (auto simp add: StandardRes_prod) |
|
408 |
moreover have a: "\<forall>x \<in> E. [p - x = 0 - x] (mod p)" |
|
409 |
apply clarify |
|
410 |
apply (insert zcong_id [of p]) |
|
411 |
apply (rule_tac a = p and m = p and c = x and d = x in zcong_zdiff, auto) |
|
412 |
done |
|
413 |
moreover have b: "\<forall>x \<in> E. [StandardRes p (p - x) = p - x](mod p)" |
|
414 |
apply clarify |
|
415 |
apply (simp add: StandardRes_prop1 zcong_sym) |
|
416 |
done |
|
417 |
moreover have "\<forall>x \<in> E. [StandardRes p (p - x) = - x](mod p)" |
|
418 |
apply clarify |
|
419 |
apply (insert a b) |
|
420 |
apply (rule_tac b = "p - x" in zcong_trans, auto) |
|
421 |
done |
|
422 |
ultimately have c: |
|
423 |
"[setprod (StandardRes p o (op - p)) E = setprod (uminus) E](mod p)" |
|
424 |
apply simp |
|
425 |
apply (insert finite_E p_g_0) |
|
426 |
apply (rule setprod_same_function_zcong |
|
427 |
[of E "StandardRes p o (op - p)" uminus p], auto) |
|
428 |
done |
|
429 |
then have two: "[setprod id F = setprod (uminus) E](mod p)" |
|
430 |
apply (insert one c) |
|
431 |
apply (rule zcong_trans [of "setprod id F" |
|
15392 | 432 |
"setprod (StandardRes p o op - p) E" p |
18369 | 433 |
"setprod uminus E"], auto) |
434 |
done |
|
435 |
also have "setprod uminus E = (setprod id E) * (-1)^(card E)" |
|
436 |
using finite_E by (induct set: Finites) auto |
|
437 |
then have "setprod uminus E = (-1) ^ (card E) * (setprod id E)" |
|
438 |
by (simp add: zmult_commute) |
|
439 |
with two show ?thesis |
|
440 |
by simp |
|
15392 | 441 |
qed |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
442 |
|
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
443 |
subsection {* Gauss' Lemma *} |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
444 |
|
15392 | 445 |
lemma (in GAUSS) aux: "setprod id A * -1 ^ card E * a ^ card A * -1 ^ card E = setprod id A * a ^ card A" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
446 |
by (auto simp add: finite_E neg_one_special) |
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
447 |
|
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
448 |
theorem (in GAUSS) pre_gauss_lemma: |
18369 | 449 |
"[a ^ nat((p - 1) div 2) = (-1) ^ (card E)] (mod p)" |
450 |
proof - |
|
451 |
have "[setprod id A = setprod id F * setprod id D](mod p)" |
|
452 |
by (auto simp add: prod_D_F_eq_prod_A zmult_commute) |
|
453 |
then have "[setprod id A = ((-1)^(card E) * setprod id E) * |
|
454 |
setprod id D] (mod p)" |
|
455 |
apply (rule zcong_trans) |
|
456 |
apply (auto simp add: prod_F_zcong zcong_scalar) |
|
457 |
done |
|
458 |
then have "[setprod id A = ((-1)^(card E) * setprod id C)] (mod p)" |
|
459 |
apply (rule zcong_trans) |
|
460 |
apply (insert C_prod_eq_D_times_E, erule subst) |
|
461 |
apply (subst zmult_assoc, auto) |
|
462 |
done |
|
463 |
then have "[setprod id A = ((-1)^(card E) * setprod id B)] (mod p)" |
|
464 |
apply (rule zcong_trans) |
|
465 |
apply (simp add: C_B_zcong_prod zcong_scalar2) |
|
466 |
done |
|
467 |
then have "[setprod id A = ((-1)^(card E) * |
|
468 |
(setprod id ((%x. x * a) ` A)))] (mod p)" |
|
469 |
by (simp add: B_def) |
|
470 |
then have "[setprod id A = ((-1)^(card E) * (setprod (%x. x * a) A))] |
|
471 |
(mod p)" |
|
472 |
by (simp add:finite_A inj_on_xa_A setprod_reindex_id[symmetric]) |
|
473 |
moreover have "setprod (%x. x * a) A = |
|
474 |
setprod (%x. a) A * setprod id A" |
|
475 |
using finite_A by (induct set: Finites) auto |
|
476 |
ultimately have "[setprod id A = ((-1)^(card E) * (setprod (%x. a) A * |
|
477 |
setprod id A))] (mod p)" |
|
478 |
by simp |
|
479 |
then have "[setprod id A = ((-1)^(card E) * a^(card A) * |
|
480 |
setprod id A)](mod p)" |
|
481 |
apply (rule zcong_trans) |
|
482 |
apply (simp add: zcong_scalar2 zcong_scalar finite_A setprod_constant zmult_assoc) |
|
483 |
done |
|
484 |
then have a: "[setprod id A * (-1)^(card E) = |
|
485 |
((-1)^(card E) * a^(card A) * setprod id A * (-1)^(card E))](mod p)" |
|
486 |
by (rule zcong_scalar) |
|
487 |
then have "[setprod id A * (-1)^(card E) = setprod id A * |
|
488 |
(-1)^(card E) * a^(card A) * (-1)^(card E)](mod p)" |
|
489 |
apply (rule zcong_trans) |
|
490 |
apply (simp add: a mult_commute mult_left_commute) |
|
491 |
done |
|
492 |
then have "[setprod id A * (-1)^(card E) = setprod id A * |
|
493 |
a^(card A)](mod p)" |
|
494 |
apply (rule zcong_trans) |
|
495 |
apply (simp add: aux) |
|
496 |
done |
|
497 |
with this zcong_cancel2 [of p "setprod id A" "-1 ^ card E" "a ^ card A"] |
|
498 |
p_g_0 A_prod_relprime have "[-1 ^ card E = a ^ card A](mod p)" |
|
499 |
by (simp add: order_less_imp_le) |
|
500 |
from this show ?thesis |
|
501 |
by (simp add: A_card_eq zcong_sym) |
|
15392 | 502 |
qed |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
503 |
|
15392 | 504 |
theorem (in GAUSS) gauss_lemma: "(Legendre a p) = (-1) ^ (card E)" |
505 |
proof - |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
506 |
from Euler_Criterion p_prime p_g_2 have |
18369 | 507 |
"[(Legendre a p) = a^(nat (((p) - 1) div 2))] (mod p)" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
508 |
by auto |
15392 | 509 |
moreover note pre_gauss_lemma |
510 |
ultimately have "[(Legendre a p) = (-1) ^ (card E)] (mod p)" |
|
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
511 |
by (rule zcong_trans) |
15392 | 512 |
moreover from p_a_relprime have "(Legendre a p) = 1 | (Legendre a p) = (-1)" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
513 |
by (auto simp add: Legendre_def) |
15392 | 514 |
moreover have "(-1::int) ^ (card E) = 1 | (-1::int) ^ (card E) = -1" |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
515 |
by (rule neg_one_power) |
15392 | 516 |
ultimately show ?thesis |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
517 |
by (auto simp add: p_g_2 one_not_neg_one_mod_m zcong_sym) |
15392 | 518 |
qed |
13871
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
paulson
parents:
diff
changeset
|
519 |
|
16775
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents:
16733
diff
changeset
|
520 |
end |