| author | wenzelm | 
| Fri, 14 Dec 2007 17:56:08 +0100 | |
| changeset 25624 | 04b67ee73327 | 
| parent 23373 | ead82c82da9e | 
| child 25675 | 2488fc510178 | 
| 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/Euler.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 {* Euler's criterion *}
 | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 7 | |
| 16974 | 8 | theory Euler imports Residues EvenOdd begin | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 9 | |
| 19670 | 10 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20369diff
changeset | 11 | MultInvPair :: "int => int => int => int set" where | 
| 19670 | 12 |   "MultInvPair a p j = {StandardRes p j, StandardRes p (a * (MultInv p j))}"
 | 
| 13 | ||
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20369diff
changeset | 14 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20369diff
changeset | 15 | SetS :: "int => int => int set set" where | 
| 19670 | 16 | "SetS a p = (MultInvPair a p ` SRStar p)" | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 17 | |
| 19670 | 18 | |
| 19 | subsection {* Property for MultInvPair *}
 | |
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 20 | |
| 19670 | 21 | lemma MultInvPair_prop1a: | 
| 22 | "[| zprime p; 2 < p; ~([a = 0](mod p)); | |
| 23 | X \<in> (SetS a p); Y \<in> (SetS a p); | |
| 24 |       ~((X \<inter> Y) = {}) |] ==> X = Y"
 | |
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 25 | apply (auto simp add: SetS_def) | 
| 16974 | 26 | apply (drule StandardRes_SRStar_prop1a)+ defer 1 | 
| 27 | apply (drule StandardRes_SRStar_prop1a)+ | |
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 28 | apply (auto simp add: MultInvPair_def StandardRes_prop2 zcong_sym) | 
| 20369 | 29 | apply (drule notE, rule MultInv_zcong_prop1, auto)[] | 
| 30 | apply (drule notE, rule MultInv_zcong_prop2, auto simp add: zcong_sym)[] | |
| 31 | apply (drule MultInv_zcong_prop2, auto simp add: zcong_sym)[] | |
| 32 | apply (drule MultInv_zcong_prop3, auto simp add: zcong_sym)[] | |
| 33 | apply (drule MultInv_zcong_prop1, auto)[] | |
| 34 | apply (drule MultInv_zcong_prop2, auto simp add: zcong_sym)[] | |
| 35 | apply (drule MultInv_zcong_prop2, auto simp add: zcong_sym)[] | |
| 36 | apply (drule MultInv_zcong_prop3, auto simp add: zcong_sym)[] | |
| 19670 | 37 | done | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 38 | |
| 19670 | 39 | lemma MultInvPair_prop1b: | 
| 40 | "[| zprime p; 2 < p; ~([a = 0](mod p)); | |
| 41 | X \<in> (SetS a p); Y \<in> (SetS a p); | |
| 42 |       X \<noteq> Y |] ==> X \<inter> Y = {}"
 | |
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 43 | apply (rule notnotD) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 44 | apply (rule notI) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 45 | apply (drule MultInvPair_prop1a, auto) | 
| 19670 | 46 | done | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 47 | |
| 16663 | 48 | lemma MultInvPair_prop1c: "[| zprime p; 2 < p; ~([a = 0](mod p)) |] ==> | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 49 |     \<forall>X \<in> SetS a p. \<forall>Y \<in> SetS a p. X \<noteq> Y --> X\<inter>Y = {}"
 | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 50 | by (auto simp add: MultInvPair_prop1b) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 51 | |
| 16663 | 52 | lemma MultInvPair_prop2: "[| zprime p; 2 < p; ~([a = 0](mod p)) |] ==> | 
| 16974 | 53 | Union ( SetS a p) = SRStar p" | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 54 | apply (auto simp add: SetS_def MultInvPair_def StandardRes_SRStar_prop4 | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 55 | SRStar_mult_prop2) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 56 | apply (frule StandardRes_SRStar_prop3) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 57 | apply (rule bexI, auto) | 
| 19670 | 58 | done | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 59 | |
| 16663 | 60 | lemma MultInvPair_distinct: "[| zprime p; 2 < p; ~([a = 0] (mod p)); | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 61 | ~([j = 0] (mod p)); | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 62 | ~(QuadRes p a) |] ==> | 
| 16974 | 63 | ~([j = a * MultInv p j] (mod p))" | 
| 20369 | 64 | proof | 
| 16663 | 65 | assume "zprime p" and "2 < p" and "~([a = 0] (mod p))" and | 
| 16974 | 66 | "~([j = 0] (mod p))" and "~(QuadRes p a)" | 
| 67 | assume "[j = a * MultInv p j] (mod p)" | |
| 68 | then have "[j * j = (a * MultInv p j) * j] (mod p)" | |
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 69 | by (auto simp add: zcong_scalar) | 
| 16974 | 70 | then have a:"[j * j = a * (MultInv p j * j)] (mod p)" | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 71 | by (auto simp add: zmult_ac) | 
| 16974 | 72 | have "[j * j = a] (mod p)" | 
| 73 | proof - | |
| 74 | from prems have b: "[MultInv p j * j = 1] (mod p)" | |
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 75 | by (simp add: MultInv_prop2a) | 
| 16974 | 76 | from b a show ?thesis | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 77 | by (auto simp add: zcong_zmult_prop2) | 
| 16974 | 78 | qed | 
| 79 | then have "[j^2 = a] (mod p)" | |
| 80 | apply(subgoal_tac "2 = Suc(Suc(0))") | |
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 81 | apply (erule ssubst) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 82 | apply (auto simp only: power_Suc power_0) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 83 | by auto | 
| 16974 | 84 | with prems show False | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 85 | by (simp add: QuadRes_def) | 
| 16974 | 86 | qed | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 87 | |
| 16663 | 88 | lemma MultInvPair_card_two: "[| zprime p; 2 < p; ~([a = 0] (mod p)); | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 89 | ~(QuadRes p a); ~([j = 0] (mod p)) |] ==> | 
| 16974 | 90 | card (MultInvPair a p j) = 2" | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 91 | apply (auto simp add: MultInvPair_def) | 
| 16974 | 92 | apply (subgoal_tac "~ (StandardRes p j = StandardRes p (a * MultInv p j))") | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 93 | apply auto | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 94 | apply (simp only: StandardRes_prop2) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 95 | apply (drule MultInvPair_distinct) | 
| 20369 | 96 | apply auto back | 
| 97 | done | |
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 98 | |
| 19670 | 99 | |
| 100 | subsection {* Properties of SetS *}
 | |
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 101 | |
| 16974 | 102 | lemma SetS_finite: "2 < p ==> finite (SetS a p)" | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 103 | by (auto simp add: SetS_def SRStar_finite [of p] finite_imageI) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 104 | |
| 16974 | 105 | lemma SetS_elems_finite: "\<forall>X \<in> SetS a p. finite X" | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 106 | by (auto simp add: SetS_def MultInvPair_def) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 107 | |
| 16663 | 108 | lemma SetS_elems_card: "[| zprime p; 2 < p; ~([a = 0] (mod p)); | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 109 | ~(QuadRes p a) |] ==> | 
| 16974 | 110 | \<forall>X \<in> SetS a p. card X = 2" | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 111 | apply (auto simp add: SetS_def) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 112 | apply (frule StandardRes_SRStar_prop1a) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 113 | apply (rule MultInvPair_card_two, auto) | 
| 19670 | 114 | done | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 115 | |
| 16974 | 116 | lemma Union_SetS_finite: "2 < p ==> finite (Union (SetS a p))" | 
| 15402 | 117 | by (auto simp add: SetS_finite SetS_elems_finite finite_Union) | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 118 | |
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 119 | lemma card_setsum_aux: "[| finite S; \<forall>X \<in> S. finite (X::int set); | 
| 16974 | 120 | \<forall>X \<in> S. card X = n |] ==> setsum card S = setsum (%x. n) S" | 
| 22274 | 121 | by (induct set: finite) auto | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 122 | |
| 16663 | 123 | lemma SetS_card: "[| zprime p; 2 < p; ~([a = 0] (mod p)); ~(QuadRes p a) |] ==> | 
| 16974 | 124 | int(card(SetS a p)) = (p - 1) div 2" | 
| 125 | proof - | |
| 126 | assume "zprime p" and "2 < p" and "~([a = 0] (mod p))" and "~(QuadRes p a)" | |
| 127 | then have "(p - 1) = 2 * int(card(SetS a p))" | |
| 128 | proof - | |
| 129 | have "p - 1 = int(card(Union (SetS a p)))" | |
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 130 | by (auto simp add: prems MultInvPair_prop2 SRStar_card) | 
| 16974 | 131 | also have "... = int (setsum card (SetS a p))" | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 132 | by (auto simp add: prems SetS_finite SetS_elems_finite | 
| 15402 | 133 | MultInvPair_prop1c [of p a] card_Union_disjoint) | 
| 16974 | 134 | also have "... = int(setsum (%x.2) (SetS a p))" | 
| 19670 | 135 | using prems | 
| 136 | by (auto simp add: SetS_elems_card SetS_finite SetS_elems_finite | |
| 15047 | 137 | card_setsum_aux simp del: setsum_constant) | 
| 16974 | 138 | also have "... = 2 * int(card( SetS a p))" | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 139 | by (auto simp add: prems SetS_finite setsum_const2) | 
| 16974 | 140 | finally show ?thesis . | 
| 141 | qed | |
| 142 | from this show ?thesis | |
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 143 | by auto | 
| 16974 | 144 | qed | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 145 | |
| 16663 | 146 | lemma SetS_setprod_prop: "[| zprime p; 2 < p; ~([a = 0] (mod p)); | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 147 | ~(QuadRes p a); x \<in> (SetS a p) |] ==> | 
| 16974 | 148 | [\<Prod>x = a] (mod p)" | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 149 | apply (auto simp add: SetS_def MultInvPair_def) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 150 | apply (frule StandardRes_SRStar_prop1a) | 
| 16974 | 151 | apply (subgoal_tac "StandardRes p x \<noteq> StandardRes p (a * MultInv p x)") | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 152 | apply (auto simp add: StandardRes_prop2 MultInvPair_distinct) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 153 | apply (frule_tac m = p and x = x and y = "(a * MultInv p x)" in | 
| 16974 | 154 | StandardRes_prop4) | 
| 155 | apply (subgoal_tac "[x * (a * MultInv p x) = a * (x * MultInv p x)] (mod p)") | |
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 156 | apply (drule_tac a = "StandardRes p x * StandardRes p (a * MultInv p x)" and | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 157 | b = "x * (a * MultInv p x)" and | 
| 16974 | 158 | c = "a * (x * MultInv p x)" in zcong_trans, force) | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 159 | apply (frule_tac p = p and x = x in MultInv_prop2, auto) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 160 | apply (drule_tac a = "x * MultInv p x" and b = 1 in zcong_zmult_prop2) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 161 | apply (auto simp add: zmult_ac) | 
| 19670 | 162 | done | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 163 | |
| 16974 | 164 | lemma aux1: "[| 0 < x; (x::int) < a; x \<noteq> (a - 1) |] ==> x < a - 1" | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 165 | by arith | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 166 | |
| 16974 | 167 | lemma aux2: "[| (a::int) < c; b < c |] ==> (a \<le> b | b \<le> a)" | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 168 | by auto | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 169 | |
| 18369 | 170 | lemma SRStar_d22set_prop: "2 < p \<Longrightarrow> (SRStar p) = {1} \<union> (d22set (p - 1))"
 | 
| 171 | apply (induct p rule: d22set.induct) | |
| 172 | apply auto | |
| 16733 
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
 nipkow parents: 
16663diff
changeset | 173 | apply (simp add: SRStar_def d22set.simps) | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 174 | apply (simp add: SRStar_def d22set.simps, clarify) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 175 | apply (frule aux1) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 176 | apply (frule aux2, auto) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 177 | apply (simp_all add: SRStar_def) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 178 | apply (simp add: d22set.simps) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 179 | apply (frule d22set_le) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 180 | apply (frule d22set_g_1, auto) | 
| 18369 | 181 | done | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 182 | |
| 16663 | 183 | lemma Union_SetS_setprod_prop1: "[| zprime p; 2 < p; ~([a = 0] (mod p)); ~(QuadRes p a) |] ==> | 
| 15392 | 184 | [\<Prod>(Union (SetS a p)) = a ^ nat ((p - 1) div 2)] (mod p)" | 
| 185 | proof - | |
| 16663 | 186 | assume "zprime p" and "2 < p" and "~([a = 0] (mod p))" and "~(QuadRes p a)" | 
| 15392 | 187 | then have "[\<Prod>(Union (SetS a p)) = | 
| 188 | setprod (setprod (%x. x)) (SetS a p)] (mod p)" | |
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 189 | by (auto simp add: SetS_finite SetS_elems_finite | 
| 15392 | 190 | MultInvPair_prop1c setprod_Union_disjoint) | 
| 191 | also have "[setprod (setprod (%x. x)) (SetS a p) = | |
| 192 | setprod (%x. a) (SetS a p)] (mod p)" | |
| 18369 | 193 | by (rule setprod_same_function_zcong) | 
| 194 | (auto simp add: prems SetS_setprod_prop SetS_finite) | |
| 15392 | 195 | also (zcong_trans) have "[setprod (%x. a) (SetS a p) = | 
| 196 | a^(card (SetS a p))] (mod p)" | |
| 197 | by (auto simp add: prems SetS_finite setprod_constant) | |
| 198 | finally (zcong_trans) show ?thesis | |
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 199 | apply (rule zcong_trans) | 
| 15392 | 200 | apply (subgoal_tac "card(SetS a p) = nat((p - 1) div 2)", auto) | 
| 201 | apply (subgoal_tac "nat(int(card(SetS a p))) = nat((p - 1) div 2)", force) | |
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 202 | apply (auto simp add: prems SetS_card) | 
| 18369 | 203 | done | 
| 15392 | 204 | qed | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 205 | |
| 16663 | 206 | lemma Union_SetS_setprod_prop2: "[| zprime p; 2 < p; ~([a = 0](mod p)) |] ==> | 
| 16974 | 207 | \<Prod>(Union (SetS a p)) = zfact (p - 1)" | 
| 208 | proof - | |
| 209 | assume "zprime p" and "2 < p" and "~([a = 0](mod p))" | |
| 15392 | 210 | then have "\<Prod>(Union (SetS a p)) = \<Prod>(SRStar p)" | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 211 | by (auto simp add: MultInvPair_prop2) | 
| 15392 | 212 |   also have "... = \<Prod>({1} \<union> (d22set (p - 1)))"
 | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 213 | by (auto simp add: prems SRStar_d22set_prop) | 
| 15392 | 214 | also have "... = zfact(p - 1)" | 
| 215 | proof - | |
| 18369 | 216 | have "~(1 \<in> d22set (p - 1)) & finite( d22set (p - 1))" | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 217 | apply (insert prems, auto) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 218 | apply (drule d22set_g_1) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 219 | apply (auto simp add: d22set_fin) | 
| 18369 | 220 | done | 
| 221 |     then have "\<Prod>({1} \<union> (d22set (p - 1))) = \<Prod>(d22set (p - 1))"
 | |
| 222 | by auto | |
| 223 | then show ?thesis | |
| 224 | by (auto simp add: d22set_prod_zfact) | |
| 16974 | 225 | qed | 
| 15392 | 226 | finally show ?thesis . | 
| 16974 | 227 | qed | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 228 | |
| 16663 | 229 | lemma zfact_prop: "[| zprime p; 2 < p; ~([a = 0] (mod p)); ~(QuadRes p a) |] ==> | 
| 16974 | 230 | [zfact (p - 1) = 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 | 231 | apply (frule Union_SetS_setprod_prop1) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 232 | apply (auto simp add: Union_SetS_setprod_prop2) | 
| 18369 | 233 | done | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 234 | |
| 19670 | 235 | text {* \medskip Prove the first part of Euler's Criterion: *}
 | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 236 | |
| 16663 | 237 | lemma Euler_part1: "[| 2 < p; zprime p; ~([x = 0](mod p)); | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 238 | ~(QuadRes p x) |] ==> | 
| 16974 | 239 | [x^(nat (((p) - 1) div 2)) = -1](mod p)" | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 240 | apply (frule zfact_prop, auto) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 241 | apply (frule Wilson_Russ) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 242 | apply (auto simp add: zcong_sym) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 243 | apply (rule zcong_trans, auto) | 
| 18369 | 244 | done | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 245 | |
| 19670 | 246 | text {* \medskip Prove another part of Euler Criterion: *}
 | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 247 | |
| 16974 | 248 | lemma aux_1: "0 < p ==> (a::int) ^ nat (p) = a * a ^ (nat (p) - 1)" | 
| 249 | proof - | |
| 250 | assume "0 < p" | |
| 251 | then have "a ^ (nat p) = a ^ (1 + (nat p - 1))" | |
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 252 | by (auto simp add: diff_add_assoc) | 
| 16974 | 253 | also have "... = (a ^ 1) * a ^ (nat(p) - 1)" | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 254 | by (simp only: zpower_zadd_distrib) | 
| 16974 | 255 | also have "... = a * a ^ (nat(p) - 1)" | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 256 | by auto | 
| 16974 | 257 | finally show ?thesis . | 
| 258 | qed | |
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 259 | |
| 16974 | 260 | lemma aux_2: "[| (2::int) < p; p \<in> zOdd |] ==> 0 < ((p - 1) div 2)" | 
| 261 | proof - | |
| 262 | assume "2 < p" and "p \<in> zOdd" | |
| 263 | then have "(p - 1):zEven" | |
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 264 | by (auto simp add: zEven_def zOdd_def) | 
| 16974 | 265 | then have aux_1: "2 * ((p - 1) div 2) = (p - 1)" | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 266 | by (auto simp add: even_div_2_prop2) | 
| 23373 | 267 | with `2 < p` have "1 < (p - 1)" | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 268 | by auto | 
| 16974 | 269 | then have " 1 < (2 * ((p - 1) div 2))" | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 270 | by (auto simp add: aux_1) | 
| 16974 | 271 | then have "0 < (2 * ((p - 1) div 2)) div 2" | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 272 | by auto | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 273 | then show ?thesis by auto | 
| 16974 | 274 | qed | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 275 | |
| 19670 | 276 | lemma Euler_part2: | 
| 277 | "[| 2 < p; zprime p; [a = 0] (mod p) |] ==> [0 = 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 | 278 | apply (frule zprime_zOdd_eq_grt_2) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 279 | apply (frule aux_2, auto) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 280 | apply (frule_tac a = a in aux_1, auto) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 281 | apply (frule zcong_zmult_prop1, auto) | 
| 18369 | 282 | done | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 283 | |
| 19670 | 284 | text {* \medskip Prove the final part of Euler's Criterion: *}
 | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 285 | |
| 16974 | 286 | lemma aux__1: "[| ~([x = 0] (mod p)); [y ^ 2 = x] (mod p)|] ==> ~(p dvd y)" | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 287 | apply (subgoal_tac "[| ~([x = 0] (mod p)); [y ^ 2 = x] (mod p)|] ==> | 
| 16974 | 288 | ~([y ^ 2 = 0] (mod p))") | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 289 | apply (auto simp add: zcong_sym [of "y^2" x p] intro: zcong_trans) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 290 | apply (auto simp add: zcong_eq_zdvd_prop intro: zpower_zdvd_prop1) | 
| 18369 | 291 | done | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 292 | |
| 16974 | 293 | lemma aux__2: "2 * nat((p - 1) div 2) = nat (2 * ((p - 1) div 2))" | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 294 | by (auto simp add: nat_mult_distrib) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 295 | |
| 16663 | 296 | lemma Euler_part3: "[| 2 < p; zprime p; ~([x = 0](mod p)); QuadRes p x |] ==> | 
| 16974 | 297 | [x^(nat (((p) - 1) div 2)) = 1](mod p)" | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 298 | apply (subgoal_tac "p \<in> zOdd") | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 299 | apply (auto simp add: QuadRes_def) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 300 | apply (frule aux__1, auto) | 
| 16974 | 301 | apply (drule_tac z = "nat ((p - 1) div 2)" in zcong_zpower) | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 302 | apply (auto simp add: zpower_zpower) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 303 | apply (rule zcong_trans) | 
| 16974 | 304 | apply (auto simp add: zcong_sym [of "x ^ nat ((p - 1) div 2)"]) | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 305 | apply (simp add: aux__2) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 306 | apply (frule odd_minus_one_even) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 307 | apply (frule even_div_2_prop2) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 308 | apply (auto intro: Little_Fermat simp add: zprime_zOdd_eq_grt_2) | 
| 18369 | 309 | done | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 310 | |
| 19670 | 311 | |
| 312 | text {* \medskip Finally show Euler's Criterion: *}
 | |
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 313 | |
| 16663 | 314 | theorem Euler_Criterion: "[| 2 < p; zprime p |] ==> [(Legendre a p) = | 
| 16974 | 315 | 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 | 316 | apply (auto simp add: Legendre_def Euler_part2) | 
| 20369 | 317 | apply (frule Euler_part3, auto simp add: zcong_sym)[] | 
| 318 | apply (frule Euler_part1, auto simp add: zcong_sym)[] | |
| 18369 | 319 | done | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 320 | |
| 18369 | 321 | end |