| author | chaieb | 
| Fri, 27 Mar 2009 17:35:21 +0000 | |
| changeset 30748 | fe67d729a61c | 
| parent 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/Residues.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 {* Residue Sets *}
 | 
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
7  | 
|
| 18369 | 8  | 
theory Residues imports Int2 begin  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
9  | 
|
| 19670 | 10  | 
text {*
 | 
11  | 
\medskip Define the residue of a set, the standard residue,  | 
|
12  | 
quadratic residues, and prove some basic properties. *}  | 
|
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
13  | 
|
| 19670 | 14  | 
definition  | 
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19670 
diff
changeset
 | 
15  | 
ResSet :: "int => int set => bool" where  | 
| 19670 | 16  | 
"ResSet m X = (\<forall>y1 y2. (y1 \<in> X & y2 \<in> X & [y1 = y2] (mod m) --> y1 = y2))"  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
17  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19670 
diff
changeset
 | 
18  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19670 
diff
changeset
 | 
19  | 
StandardRes :: "int => int => int" where  | 
| 19670 | 20  | 
"StandardRes m x = x mod m"  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
21  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19670 
diff
changeset
 | 
22  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19670 
diff
changeset
 | 
23  | 
QuadRes :: "int => int => bool" where  | 
| 19670 | 24  | 
"QuadRes m x = (\<exists>y. ([(y ^ 2) = x] (mod m)))"  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
25  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19670 
diff
changeset
 | 
26  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19670 
diff
changeset
 | 
27  | 
Legendre :: "int => int => int" where  | 
| 19670 | 28  | 
"Legendre a p = (if ([a = 0] (mod p)) then 0  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
29  | 
else if (QuadRes p a) then 1  | 
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
30  | 
else -1)"  | 
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
31  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19670 
diff
changeset
 | 
32  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19670 
diff
changeset
 | 
33  | 
SR :: "int => int set" where  | 
| 19670 | 34  | 
  "SR p = {x. (0 \<le> x) & (x < p)}"
 | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
35  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19670 
diff
changeset
 | 
36  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19670 
diff
changeset
 | 
37  | 
SRStar :: "int => int set" where  | 
| 19670 | 38  | 
  "SRStar p = {x. (0 < x) & (x < p)}"
 | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
39  | 
|
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
40  | 
|
| 19670 | 41  | 
subsection {* Some useful properties of StandardRes *}
 | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
42  | 
|
| 18369 | 43  | 
lemma StandardRes_prop1: "[x = StandardRes m x] (mod m)"  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
44  | 
by (auto simp add: StandardRes_def zcong_zmod)  | 
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
45  | 
|
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
46  | 
lemma StandardRes_prop2: "0 < m ==> (StandardRes m x1 = StandardRes m x2)  | 
| 18369 | 47  | 
= ([x1 = x2] (mod m))"  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
48  | 
by (auto simp add: StandardRes_def zcong_zmod_eq)  | 
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
49  | 
|
| 18369 | 50  | 
lemma StandardRes_prop3: "(~[x = 0] (mod p)) = (~(StandardRes p x = 0))"  | 
| 30042 | 51  | 
by (auto simp add: StandardRes_def zcong_def dvd_eq_mod_eq_0)  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
52  | 
|
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
53  | 
lemma StandardRes_prop4: "2 < m  | 
| 18369 | 54  | 
==> [StandardRes m x * StandardRes m y = (x * y)] (mod m)"  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
55  | 
by (auto simp add: StandardRes_def zcong_zmod_eq  | 
| 29948 | 56  | 
mod_mult_eq [of x y m])  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
57  | 
|
| 18369 | 58  | 
lemma StandardRes_lbound: "0 < p ==> 0 \<le> StandardRes p x"  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
59  | 
by (auto simp add: StandardRes_def pos_mod_sign)  | 
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
60  | 
|
| 18369 | 61  | 
lemma StandardRes_ubound: "0 < p ==> StandardRes p x < p"  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
62  | 
by (auto simp add: StandardRes_def pos_mod_bound)  | 
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
63  | 
|
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
64  | 
lemma StandardRes_eq_zcong:  | 
| 18369 | 65  | 
"(StandardRes m x = 0) = ([x = 0](mod m))"  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
66  | 
by (auto simp add: StandardRes_def zcong_eq_zdvd_prop dvd_def)  | 
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
67  | 
|
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
68  | 
|
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
69  | 
subsection {* Relations between StandardRes, SRStar, and SR *}
 | 
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
70  | 
|
| 18369 | 71  | 
lemma SRStar_SR_prop: "x \<in> SRStar p ==> x \<in> SR p"  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
72  | 
by (auto simp add: SRStar_def SR_def)  | 
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
73  | 
|
| 18369 | 74  | 
lemma StandardRes_SR_prop: "x \<in> SR p ==> StandardRes p x = x"  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
75  | 
by (auto simp add: SR_def StandardRes_def mod_pos_pos_trivial)  | 
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
76  | 
|
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
77  | 
lemma StandardRes_SRStar_prop1: "2 < p ==> (StandardRes p x \<in> SRStar p)  | 
| 18369 | 78  | 
= (~[x = 0] (mod p))"  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
79  | 
apply (auto simp add: StandardRes_prop3 StandardRes_def  | 
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
80  | 
SRStar_def pos_mod_bound)  | 
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
81  | 
apply (subgoal_tac "0 < p")  | 
| 18369 | 82  | 
apply (drule_tac a = x in pos_mod_sign, arith, simp)  | 
83  | 
done  | 
|
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
84  | 
|
| 18369 | 85  | 
lemma StandardRes_SRStar_prop1a: "x \<in> SRStar p ==> ~([x = 0] (mod p))"  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
86  | 
by (auto simp add: SRStar_def zcong_def zdvd_not_zless)  | 
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
87  | 
|
| 16663 | 88  | 
lemma StandardRes_SRStar_prop2: "[| 2 < p; zprime p; x \<in> SRStar p |]  | 
| 18369 | 89  | 
==> StandardRes p (MultInv p x) \<in> SRStar p"  | 
90  | 
apply (frule_tac x = "(MultInv p x)" in StandardRes_SRStar_prop1, simp)  | 
|
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
91  | 
apply (rule MultInv_prop3)  | 
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
92  | 
apply (auto simp add: SRStar_def zcong_def zdvd_not_zless)  | 
| 18369 | 93  | 
done  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
94  | 
|
| 18369 | 95  | 
lemma StandardRes_SRStar_prop3: "x \<in> SRStar p ==> StandardRes p x = x"  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
96  | 
by (auto simp add: SRStar_SR_prop StandardRes_SR_prop)  | 
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
97  | 
|
| 16663 | 98  | 
lemma StandardRes_SRStar_prop4: "[| zprime p; 2 < p; x \<in> SRStar p |]  | 
| 18369 | 99  | 
==> StandardRes p x \<in> SRStar p"  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
100  | 
by (frule StandardRes_SRStar_prop3, auto)  | 
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
101  | 
|
| 16663 | 102  | 
lemma SRStar_mult_prop1: "[| zprime p; 2 < p; x \<in> SRStar p; y \<in> SRStar p|]  | 
| 18369 | 103  | 
==> (StandardRes p (x * y)):SRStar p"  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
104  | 
apply (frule_tac x = x in StandardRes_SRStar_prop4, auto)  | 
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
105  | 
apply (frule_tac x = y in StandardRes_SRStar_prop4, auto)  | 
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
106  | 
apply (auto simp add: StandardRes_SRStar_prop1 zcong_zmult_prop3)  | 
| 18369 | 107  | 
done  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
108  | 
|
| 16663 | 109  | 
lemma SRStar_mult_prop2: "[| 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
 | 
110  | 
x \<in> SRStar p |]  | 
| 18369 | 111  | 
==> StandardRes p (a * MultInv p x) \<in> SRStar p"  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
112  | 
apply (frule_tac x = x in StandardRes_SRStar_prop2, auto)  | 
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
113  | 
apply (frule_tac x = "MultInv p x" in StandardRes_SRStar_prop1)  | 
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
114  | 
apply (auto simp add: StandardRes_SRStar_prop1 zcong_zmult_prop3)  | 
| 18369 | 115  | 
done  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
116  | 
|
| 18369 | 117  | 
lemma SRStar_card: "2 < p ==> int(card(SRStar p)) = p - 1"  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
118  | 
by (auto simp add: SRStar_def int_card_bdd_int_set_l_l)  | 
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
119  | 
|
| 18369 | 120  | 
lemma SRStar_finite: "2 < p ==> finite( SRStar p)"  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
121  | 
by (auto simp add: SRStar_def bdd_int_set_l_l_finite)  | 
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
122  | 
|
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
123  | 
|
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
124  | 
subsection {* Properties relating ResSets with StandardRes *}
 | 
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
125  | 
|
| 18369 | 126  | 
lemma aux: "x mod m = y mod m ==> [x = y] (mod m)"  | 
127  | 
apply (subgoal_tac "x = y ==> [x = y](mod m)")  | 
|
128  | 
apply (subgoal_tac "[x mod m = y mod m] (mod m) ==> [x = y] (mod m)")  | 
|
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
129  | 
apply (auto simp add: zcong_zmod [of x y m])  | 
| 18369 | 130  | 
done  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
131  | 
|
| 18369 | 132  | 
lemma StandardRes_inj_on_ResSet: "ResSet m X ==> (inj_on (StandardRes m) X)"  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
133  | 
apply (auto simp add: ResSet_def StandardRes_def inj_on_def)  | 
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
134  | 
apply (drule_tac m = m in aux, auto)  | 
| 18369 | 135  | 
done  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
136  | 
|
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
137  | 
lemma StandardRes_Sum: "[| finite X; 0 < m |]  | 
| 18369 | 138  | 
==> [setsum f X = setsum (StandardRes m o f) X](mod m)"  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
139  | 
apply (rule_tac F = X in finite_induct)  | 
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
140  | 
apply (auto intro!: zcong_zadd simp add: StandardRes_prop1)  | 
| 18369 | 141  | 
done  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
142  | 
|
| 18369 | 143  | 
lemma SR_pos: "0 < m ==> (StandardRes m ` X) \<subseteq> {x. 0 \<le> x & x < m}"
 | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
144  | 
by (auto simp add: StandardRes_ubound StandardRes_lbound)  | 
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
145  | 
|
| 18369 | 146  | 
lemma ResSet_finite: "0 < m ==> ResSet m X ==> finite X"  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
147  | 
apply (rule_tac f = "StandardRes m" in finite_imageD)  | 
| 18369 | 148  | 
  apply (rule_tac B = "{x. (0 :: int) \<le> x & x < m}" in finite_subset)
 | 
149  | 
apply (auto simp add: StandardRes_inj_on_ResSet bdd_int_set_l_finite SR_pos)  | 
|
150  | 
done  | 
|
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
151  | 
|
| 18369 | 152  | 
lemma mod_mod_is_mod: "[x = x mod m](mod m)"  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
153  | 
by (auto simp add: zcong_zmod)  | 
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
154  | 
|
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
155  | 
lemma StandardRes_prod: "[| finite X; 0 < m |]  | 
| 18369 | 156  | 
==> [setprod f X = setprod (StandardRes m o f) X] (mod m)"  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
157  | 
apply (rule_tac F = X in finite_induct)  | 
| 18369 | 158  | 
apply (auto intro!: zcong_zmult simp add: StandardRes_prop1)  | 
159  | 
done  | 
|
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
160  | 
|
| 19670 | 161  | 
lemma ResSet_image:  | 
162  | 
"[| 0 < m; ResSet m A; \<forall>x \<in> A. \<forall>y \<in> A. ([f x = f y](mod m) --> x = y) |] ==>  | 
|
163  | 
ResSet m (f ` A)"  | 
|
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
164  | 
by (auto simp add: ResSet_def)  | 
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
165  | 
|
| 19670 | 166  | 
|
167  | 
subsection {* Property for SRStar *}
 | 
|
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
168  | 
|
| 18369 | 169  | 
lemma ResSet_SRStar_prop: "ResSet p (SRStar p)"  | 
| 
13871
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
170  | 
by (auto simp add: SRStar_def ResSet_def zcong_zless_imp_eq)  | 
| 
 
26e5f5e624f6
Gauss's law of quadratic reciprocity by Avigad, Gray and Kramer
 
paulson 
parents:  
diff
changeset
 | 
171  | 
|
| 18369 | 172  | 
end  |