| author | huffman | 
| Wed, 24 Dec 2008 08:16:45 -0800 | |
| changeset 29166 | c23b2d108612 | 
| parent 26510 | a329af578d69 | 
| child 30042 | 31039ee583fa | 
| 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)" | |
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25760diff
changeset | 80 | by (metis number_of_is_id power2_eq_square succ_bin_simps) | 
| 16974 | 81 | with prems show False | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 82 | by (simp add: QuadRes_def) | 
| 16974 | 83 | qed | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 84 | |
| 16663 | 85 | 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 | 86 | ~(QuadRes p a); ~([j = 0] (mod p)) |] ==> | 
| 16974 | 87 | card (MultInvPair a p j) = 2" | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 88 | apply (auto simp add: MultInvPair_def) | 
| 16974 | 89 | 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 | 90 | apply auto | 
| 26510 | 91 | apply (metis MultInvPair_distinct Pls_def StandardRes_def aux number_of_is_id one_is_num_one) | 
| 20369 | 92 | done | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 93 | |
| 19670 | 94 | |
| 95 | subsection {* Properties of SetS *}
 | |
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 96 | |
| 16974 | 97 | 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 | 98 | 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 | 99 | |
| 16974 | 100 | 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 | 101 | by (auto simp add: SetS_def MultInvPair_def) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 102 | |
| 16663 | 103 | 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 | 104 | ~(QuadRes p a) |] ==> | 
| 16974 | 105 | \<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 | 106 | apply (auto simp add: SetS_def) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 107 | apply (frule StandardRes_SRStar_prop1a) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 108 | apply (rule MultInvPair_card_two, auto) | 
| 19670 | 109 | done | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 110 | |
| 16974 | 111 | lemma Union_SetS_finite: "2 < p ==> finite (Union (SetS a p))" | 
| 15402 | 112 | 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 | 113 | |
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 114 | lemma card_setsum_aux: "[| finite S; \<forall>X \<in> S. finite (X::int set); | 
| 16974 | 115 | \<forall>X \<in> S. card X = n |] ==> setsum card S = setsum (%x. n) S" | 
| 22274 | 116 | by (induct set: finite) auto | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 117 | |
| 16663 | 118 | lemma SetS_card: "[| zprime p; 2 < p; ~([a = 0] (mod p)); ~(QuadRes p a) |] ==> | 
| 16974 | 119 | int(card(SetS a p)) = (p - 1) div 2" | 
| 120 | proof - | |
| 121 | assume "zprime p" and "2 < p" and "~([a = 0] (mod p))" and "~(QuadRes p a)" | |
| 122 | then have "(p - 1) = 2 * int(card(SetS a p))" | |
| 123 | proof - | |
| 124 | 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 | 125 | by (auto simp add: prems MultInvPair_prop2 SRStar_card) | 
| 16974 | 126 | 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 | 127 | by (auto simp add: prems SetS_finite SetS_elems_finite | 
| 15402 | 128 | MultInvPair_prop1c [of p a] card_Union_disjoint) | 
| 16974 | 129 | also have "... = int(setsum (%x.2) (SetS a p))" | 
| 19670 | 130 | using prems | 
| 131 | by (auto simp add: SetS_elems_card SetS_finite SetS_elems_finite | |
| 15047 | 132 | card_setsum_aux simp del: setsum_constant) | 
| 16974 | 133 | 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 | 134 | by (auto simp add: prems SetS_finite setsum_const2) | 
| 16974 | 135 | finally show ?thesis . | 
| 136 | qed | |
| 137 | from this show ?thesis | |
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 138 | by auto | 
| 16974 | 139 | qed | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 140 | |
| 16663 | 141 | 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 | 142 | ~(QuadRes p a); x \<in> (SetS a p) |] ==> | 
| 16974 | 143 | [\<Prod>x = a] (mod p)" | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 144 | apply (auto simp add: SetS_def MultInvPair_def) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 145 | apply (frule StandardRes_SRStar_prop1a) | 
| 16974 | 146 | 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 | 147 | apply (auto simp add: StandardRes_prop2 MultInvPair_distinct) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 148 | apply (frule_tac m = p and x = x and y = "(a * MultInv p x)" in | 
| 16974 | 149 | StandardRes_prop4) | 
| 150 | 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 | 151 | 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 | 152 | b = "x * (a * MultInv p x)" and | 
| 16974 | 153 | 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 | 154 | apply (frule_tac p = p and x = x in MultInv_prop2, auto) | 
| 25760 | 155 | apply (metis StandardRes_SRStar_prop3 mult_1_right mult_commute zcong_sym zcong_zmult_prop1) | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 156 | apply (auto simp add: zmult_ac) | 
| 19670 | 157 | done | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 158 | |
| 16974 | 159 | 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 | 160 | by arith | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 161 | |
| 16974 | 162 | 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 | 163 | by auto | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 164 | |
| 18369 | 165 | lemma SRStar_d22set_prop: "2 < p \<Longrightarrow> (SRStar p) = {1} \<union> (d22set (p - 1))"
 | 
| 166 | apply (induct p rule: d22set.induct) | |
| 167 | apply auto | |
| 16733 
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
 nipkow parents: 
16663diff
changeset | 168 | apply (simp add: SRStar_def d22set.simps) | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 169 | apply (simp add: SRStar_def d22set.simps, clarify) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 170 | apply (frule aux1) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 171 | apply (frule aux2, auto) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 172 | apply (simp_all add: SRStar_def) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 173 | apply (simp add: d22set.simps) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 174 | apply (frule d22set_le) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 175 | apply (frule d22set_g_1, auto) | 
| 18369 | 176 | done | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 177 | |
| 16663 | 178 | lemma Union_SetS_setprod_prop1: "[| zprime p; 2 < p; ~([a = 0] (mod p)); ~(QuadRes p a) |] ==> | 
| 15392 | 179 | [\<Prod>(Union (SetS a p)) = a ^ nat ((p - 1) div 2)] (mod p)" | 
| 180 | proof - | |
| 16663 | 181 | assume "zprime p" and "2 < p" and "~([a = 0] (mod p))" and "~(QuadRes p a)" | 
| 15392 | 182 | then have "[\<Prod>(Union (SetS a p)) = | 
| 183 | 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 | 184 | by (auto simp add: SetS_finite SetS_elems_finite | 
| 15392 | 185 | MultInvPair_prop1c setprod_Union_disjoint) | 
| 186 | also have "[setprod (setprod (%x. x)) (SetS a p) = | |
| 187 | setprod (%x. a) (SetS a p)] (mod p)" | |
| 18369 | 188 | by (rule setprod_same_function_zcong) | 
| 189 | (auto simp add: prems SetS_setprod_prop SetS_finite) | |
| 15392 | 190 | also (zcong_trans) have "[setprod (%x. a) (SetS a p) = | 
| 191 | a^(card (SetS a p))] (mod p)" | |
| 192 | by (auto simp add: prems SetS_finite setprod_constant) | |
| 193 | finally (zcong_trans) show ?thesis | |
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 194 | apply (rule zcong_trans) | 
| 15392 | 195 | apply (subgoal_tac "card(SetS a p) = nat((p - 1) div 2)", auto) | 
| 196 | 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 | 197 | apply (auto simp add: prems SetS_card) | 
| 18369 | 198 | done | 
| 15392 | 199 | qed | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 200 | |
| 16663 | 201 | lemma Union_SetS_setprod_prop2: "[| zprime p; 2 < p; ~([a = 0](mod p)) |] ==> | 
| 16974 | 202 | \<Prod>(Union (SetS a p)) = zfact (p - 1)" | 
| 203 | proof - | |
| 204 | assume "zprime p" and "2 < p" and "~([a = 0](mod p))" | |
| 15392 | 205 | 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 | 206 | by (auto simp add: MultInvPair_prop2) | 
| 15392 | 207 |   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 | 208 | by (auto simp add: prems SRStar_d22set_prop) | 
| 15392 | 209 | also have "... = zfact(p - 1)" | 
| 210 | proof - | |
| 18369 | 211 | have "~(1 \<in> d22set (p - 1)) & finite( d22set (p - 1))" | 
| 25760 | 212 | by (metis d22set_fin d22set_g_1 linorder_neq_iff) | 
| 18369 | 213 |     then have "\<Prod>({1} \<union> (d22set (p - 1))) = \<Prod>(d22set (p - 1))"
 | 
| 214 | by auto | |
| 215 | then show ?thesis | |
| 216 | by (auto simp add: d22set_prod_zfact) | |
| 16974 | 217 | qed | 
| 15392 | 218 | finally show ?thesis . | 
| 16974 | 219 | qed | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 220 | |
| 16663 | 221 | lemma zfact_prop: "[| zprime p; 2 < p; ~([a = 0] (mod p)); ~(QuadRes p a) |] ==> | 
| 16974 | 222 | [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 | 223 | apply (frule Union_SetS_setprod_prop1) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 224 | apply (auto simp add: Union_SetS_setprod_prop2) | 
| 18369 | 225 | done | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 226 | |
| 19670 | 227 | 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 | 228 | |
| 16663 | 229 | 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 | 230 | ~(QuadRes p x) |] ==> | 
| 16974 | 231 | [x^(nat (((p) - 1) div 2)) = -1](mod p)" | 
| 25760 | 232 | by (metis Wilson_Russ number_of_is_id zcong_sym zcong_trans zfact_prop) | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 233 | |
| 19670 | 234 | 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 | 235 | |
| 16974 | 236 | lemma aux_1: "0 < p ==> (a::int) ^ nat (p) = a * a ^ (nat (p) - 1)" | 
| 237 | proof - | |
| 238 | assume "0 < p" | |
| 239 | 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 | 240 | by (auto simp add: diff_add_assoc) | 
| 16974 | 241 | 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 | 242 | by (simp only: zpower_zadd_distrib) | 
| 16974 | 243 | also have "... = a * a ^ (nat(p) - 1)" | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 244 | by auto | 
| 16974 | 245 | finally show ?thesis . | 
| 246 | qed | |
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 247 | |
| 16974 | 248 | lemma aux_2: "[| (2::int) < p; p \<in> zOdd |] ==> 0 < ((p - 1) div 2)" | 
| 249 | proof - | |
| 250 | assume "2 < p" and "p \<in> zOdd" | |
| 251 | then have "(p - 1):zEven" | |
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 252 | by (auto simp add: zEven_def zOdd_def) | 
| 16974 | 253 | 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 | 254 | by (auto simp add: even_div_2_prop2) | 
| 23373 | 255 | with `2 < p` have "1 < (p - 1)" | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 256 | by auto | 
| 16974 | 257 | 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 | 258 | by (auto simp add: aux_1) | 
| 16974 | 259 | 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 | 260 | by auto | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 261 | then show ?thesis by auto | 
| 16974 | 262 | qed | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 263 | |
| 19670 | 264 | lemma Euler_part2: | 
| 265 | "[| 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 | 266 | apply (frule zprime_zOdd_eq_grt_2) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 267 | apply (frule aux_2, auto) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 268 | 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 | 269 | apply (frule zcong_zmult_prop1, auto) | 
| 18369 | 270 | done | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 271 | |
| 19670 | 272 | 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 | 273 | |
| 16974 | 274 | lemma aux__1: "[| ~([x = 0] (mod p)); [y ^ 2 = x] (mod p)|] ==> ~(p dvd y)" | 
| 25760 | 275 | by (metis dvdI power2_eq_square zcong_sym zcong_trans zcong_zero_equiv_div zdvd_trans) | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 276 | |
| 16974 | 277 | 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 | 278 | by (auto simp add: nat_mult_distrib) | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 279 | |
| 16663 | 280 | lemma Euler_part3: "[| 2 < p; zprime p; ~([x = 0](mod p)); QuadRes p x |] ==> | 
| 16974 | 281 | [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 | 282 | apply (subgoal_tac "p \<in> zOdd") | 
| 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 283 | apply (auto simp add: QuadRes_def) | 
| 25675 | 284 | prefer 2 | 
| 285 | apply (metis number_of_is_id numeral_1_eq_1 zprime_zOdd_eq_grt_2) | |
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 286 | apply (frule aux__1, auto) | 
| 16974 | 287 | apply (drule_tac z = "nat ((p - 1) div 2)" in zcong_zpower) | 
| 25675 | 288 | apply (auto simp add: zpower_zpower) | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 289 | apply (rule zcong_trans) | 
| 16974 | 290 | apply (auto simp add: zcong_sym [of "x ^ nat ((p - 1) div 2)"]) | 
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25760diff
changeset | 291 | apply (metis Little_Fermat even_div_2_prop2 mult_Bit0 number_of_is_id odd_minus_one_even one_is_num_one zmult_1 aux__2) | 
| 18369 | 292 | done | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 293 | |
| 19670 | 294 | |
| 295 | text {* \medskip Finally show Euler's Criterion: *}
 | |
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 296 | |
| 16663 | 297 | theorem Euler_Criterion: "[| 2 < p; zprime p |] ==> [(Legendre a p) = | 
| 16974 | 298 | 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 | 299 | apply (auto simp add: Legendre_def Euler_part2) | 
| 20369 | 300 | apply (frule Euler_part3, auto simp add: zcong_sym)[] | 
| 301 | apply (frule Euler_part1, auto simp add: zcong_sym)[] | |
| 18369 | 302 | done | 
| 13871 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 paulson parents: diff
changeset | 303 | |
| 18369 | 304 | end |