| author | paulson <lp15@cam.ac.uk> | 
| Fri, 28 Feb 2025 13:50:18 +0000 | |
| changeset 82218 | cbf9f856d3e0 | 
| parent 80401 | 31bf95336f16 | 
| permissions | -rw-r--r-- | 
| 65435 | 1  | 
(* Title: HOL/Number_Theory/Gauss.thy  | 
| 65413 | 2  | 
Authors: Jeremy Avigad, David Gray, and Adam Kramer  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
3  | 
|
| 65413 | 4  | 
Ported by lcp but unfinished.  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
5  | 
*)  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
6  | 
|
| 60526 | 7  | 
section \<open>Gauss' Lemma\<close>  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
8  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
9  | 
theory Gauss  | 
| 65413 | 10  | 
imports Euler_Criterion  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
11  | 
begin  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
12  | 
|
| 65413 | 13  | 
lemma cong_prime_prod_zero_nat:  | 
14  | 
"[a * b = 0] (mod p) \<Longrightarrow> prime p \<Longrightarrow> [a = 0] (mod p) \<or> [b = 0] (mod p)"  | 
|
15  | 
for a :: nat  | 
|
| 
64631
 
7705926ee595
removed dangerous simp rule: prime computations can be excessively long
 
haftmann 
parents: 
64282 
diff
changeset
 | 
16  | 
by (auto simp add: cong_altdef_nat prime_dvd_mult_iff)  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
17  | 
|
| 65413 | 18  | 
lemma cong_prime_prod_zero_int:  | 
19  | 
"[a * b = 0] (mod p) \<Longrightarrow> prime p \<Longrightarrow> [a = 0] (mod p) \<or> [b = 0] (mod p)"  | 
|
20  | 
for a :: int  | 
|
| 
68707
 
5b269897df9d
de-applying and removal of obsolete aliases
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
21  | 
by (simp add: cong_0_iff prime_dvd_mult_iff)  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
22  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
23  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
24  | 
locale GAUSS =  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
25  | 
fixes p :: "nat"  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
26  | 
fixes a :: "int"  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
27  | 
assumes p_prime: "prime p"  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
28  | 
assumes p_ge_2: "2 < p"  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
29  | 
assumes p_a_relprime: "[a \<noteq> 0](mod p)"  | 
| 65413 | 30  | 
assumes a_nonzero: "0 < a"  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
31  | 
begin  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
32  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
33  | 
definition "A = {0::int <.. ((int p - 1) div 2)}"
 | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
34  | 
definition "B = (\<lambda>x. x * a) ` A"  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
35  | 
definition "C = (\<lambda>x. x mod p) ` B"  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
36  | 
definition "D = C \<inter> {.. (int p - 1) div 2}"
 | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
37  | 
definition "E = C \<inter> {(int p - 1) div 2 <..}"
 | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
38  | 
definition "F = (\<lambda>x. (int p - x)) ` E"  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
39  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
40  | 
|
| 60526 | 41  | 
subsection \<open>Basic properties of p\<close>  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
42  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
43  | 
lemma odd_p: "odd p"  | 
| 65413 | 44  | 
by (metis p_prime p_ge_2 prime_odd_nat)  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
45  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
46  | 
lemma p_minus_one_l: "(int p - 1) div 2 < p"  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
47  | 
proof -  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
48  | 
have "(p - 1) div 2 \<le> (p - 1) div 1"  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
49  | 
by (metis div_by_1 div_le_dividend)  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
50  | 
also have "\<dots> = p - 1" by simp  | 
| 65413 | 51  | 
finally show ?thesis  | 
52  | 
using p_ge_2 by arith  | 
|
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
53  | 
qed  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
54  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
55  | 
lemma p_eq2: "int p = (2 * ((int p - 1) div 2)) + 1"  | 
| 65413 | 56  | 
using odd_p p_ge_2 nonzero_mult_div_cancel_left [of 2 "p - 1"] by simp  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
57  | 
|
| 65413 | 58  | 
lemma p_odd_int: obtains z :: int where "int p = 2 * z + 1" "0 < z"  | 
59  | 
proof  | 
|
60  | 
let ?z = "(int p - 1) div 2"  | 
|
61  | 
show "int p = 2 * ?z + 1" by (rule p_eq2)  | 
|
62  | 
show "0 < ?z"  | 
|
63  | 
using p_ge_2 by linarith  | 
|
64  | 
qed  | 
|
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
65  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
66  | 
|
| 60526 | 67  | 
subsection \<open>Basic Properties of the Gauss Sets\<close>  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
68  | 
|
| 65413 | 69  | 
lemma finite_A: "finite A"  | 
70  | 
by (auto simp add: A_def)  | 
|
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
71  | 
|
| 65413 | 72  | 
lemma finite_B: "finite B"  | 
73  | 
by (auto simp add: B_def finite_A)  | 
|
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
74  | 
|
| 65413 | 75  | 
lemma finite_C: "finite C"  | 
76  | 
by (auto simp add: C_def finite_B)  | 
|
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
77  | 
|
| 65413 | 78  | 
lemma finite_D: "finite D"  | 
79  | 
by (auto simp add: D_def finite_C)  | 
|
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
80  | 
|
| 65413 | 81  | 
lemma finite_E: "finite E"  | 
82  | 
by (auto simp add: E_def finite_C)  | 
|
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
83  | 
|
| 65413 | 84  | 
lemma finite_F: "finite F"  | 
85  | 
by (auto simp add: F_def finite_E)  | 
|
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
86  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
87  | 
lemma C_eq: "C = D \<union> E"  | 
| 65413 | 88  | 
by (auto simp add: C_def D_def E_def)  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
89  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
90  | 
lemma A_card_eq: "card A = nat ((int p - 1) div 2)"  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
91  | 
by (auto simp add: A_def)  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
92  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
93  | 
lemma inj_on_xa_A: "inj_on (\<lambda>x. x * a) A"  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
94  | 
using a_nonzero by (simp add: A_def inj_on_def)  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
95  | 
|
| 65413 | 96  | 
definition ResSet :: "int \<Rightarrow> int set \<Rightarrow> bool"  | 
97  | 
where "ResSet m X \<longleftrightarrow> (\<forall>y1 y2. y1 \<in> X \<and> y2 \<in> X \<and> [y1 = y2] (mod m) \<longrightarrow> y1 = y2)"  | 
|
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
98  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
99  | 
lemma ResSet_image:  | 
| 65413 | 100  | 
"0 < m \<Longrightarrow> ResSet m A \<Longrightarrow> \<forall>x \<in> A. \<forall>y \<in> A. ([f x = f y](mod m) \<longrightarrow> x = y) \<Longrightarrow> ResSet m (f ` A)"  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
101  | 
by (auto simp add: ResSet_def)  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
102  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
103  | 
lemma A_res: "ResSet p A"  | 
| 65413 | 104  | 
using p_ge_2 by (auto simp add: A_def ResSet_def intro!: cong_less_imp_eq_int)  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
105  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
106  | 
lemma B_res: "ResSet p B"  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
107  | 
proof -  | 
| 65413 | 108  | 
have *: "x = y"  | 
109  | 
if a: "[x * a = y * a] (mod p)"  | 
|
110  | 
and b: "0 < x"  | 
|
111  | 
and c: "x \<le> (int p - 1) div 2"  | 
|
112  | 
and d: "0 < y"  | 
|
113  | 
and e: "y \<le> (int p - 1) div 2"  | 
|
114  | 
for x y  | 
|
115  | 
proof -  | 
|
116  | 
from p_a_relprime have "\<not> p dvd a"  | 
|
| 
68707
 
5b269897df9d
de-applying and removal of obsolete aliases
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
117  | 
by (simp add: cong_0_iff)  | 
| 67051 | 118  | 
with p_prime prime_imp_coprime [of _ "nat \<bar>a\<bar>"]  | 
119  | 
have "coprime a (int p)"  | 
|
| 67118 | 120  | 
by (simp_all add: ac_simps)  | 
| 
68707
 
5b269897df9d
de-applying and removal of obsolete aliases
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
121  | 
with a cong_mult_rcancel [of a "int p" x y] have "[x = y] (mod p)"  | 
| 65413 | 122  | 
by simp  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
123  | 
with cong_less_imp_eq_int [of x y p] p_minus_one_l  | 
| 65413 | 124  | 
order_le_less_trans [of x "(int p - 1) div 2" p]  | 
125  | 
order_le_less_trans [of y "(int p - 1) div 2" p]  | 
|
126  | 
show ?thesis  | 
|
| 62348 | 127  | 
by (metis b c cong_less_imp_eq_int d e zero_less_imp_eq_int of_nat_0_le_iff)  | 
| 65413 | 128  | 
qed  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
129  | 
show ?thesis  | 
| 
68707
 
5b269897df9d
de-applying and removal of obsolete aliases
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
130  | 
using p_ge_2 p_a_relprime p_minus_one_l  | 
| 
 
5b269897df9d
de-applying and removal of obsolete aliases
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
131  | 
by (metis "*" A_def A_res B_def GAUSS.ResSet_image GAUSS_axioms greaterThanAtMost_iff odd_p odd_pos of_nat_0_less_iff)  | 
| 65413 | 132  | 
qed  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
133  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
134  | 
lemma SR_B_inj: "inj_on (\<lambda>x. x mod p) B"  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
135  | 
proof -  | 
| 65413 | 136  | 
have False  | 
137  | 
if a: "x * a mod p = y * a mod p"  | 
|
138  | 
and b: "0 < x"  | 
|
139  | 
and c: "x \<le> (int p - 1) div 2"  | 
|
140  | 
and d: "0 < y"  | 
|
141  | 
and e: "y \<le> (int p - 1) div 2"  | 
|
142  | 
and f: "x \<noteq> y"  | 
|
143  | 
for x y  | 
|
144  | 
proof -  | 
|
145  | 
from a have a': "[x * a = y * a](mod p)"  | 
|
| 66888 | 146  | 
using cong_def by blast  | 
| 65413 | 147  | 
from p_a_relprime have "\<not>p dvd a"  | 
| 
68707
 
5b269897df9d
de-applying and removal of obsolete aliases
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
148  | 
by (simp add: cong_0_iff)  | 
| 67051 | 149  | 
with p_prime prime_imp_coprime [of _ "nat \<bar>a\<bar>"]  | 
150  | 
have "coprime a (int p)"  | 
|
| 67118 | 151  | 
by (simp_all add: ac_simps)  | 
| 
68707
 
5b269897df9d
de-applying and removal of obsolete aliases
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
152  | 
with a' cong_mult_rcancel [of a "int p" x y]  | 
| 
63534
 
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
 
eberlm <eberlm@in.tum.de> 
parents: 
62429 
diff
changeset
 | 
153  | 
have "[x = y] (mod p)" by simp  | 
| 65413 | 154  | 
with cong_less_imp_eq_int [of x y p] p_minus_one_l  | 
155  | 
order_le_less_trans [of x "(int p - 1) div 2" p]  | 
|
156  | 
order_le_less_trans [of y "(int p - 1) div 2" p]  | 
|
157  | 
have "x = y"  | 
|
158  | 
by (metis b c cong_less_imp_eq_int d e zero_less_imp_eq_int of_nat_0_le_iff)  | 
|
159  | 
then show ?thesis  | 
|
160  | 
by (simp add: f)  | 
|
161  | 
qed  | 
|
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
162  | 
then show ?thesis  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
163  | 
by (auto simp add: B_def inj_on_def A_def) metis  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
164  | 
qed  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
165  | 
|
| 65413 | 166  | 
lemma nonzero_mod_p: "0 < x \<Longrightarrow> x < int p \<Longrightarrow> [x \<noteq> 0](mod p)"  | 
167  | 
for x :: int  | 
|
| 66888 | 168  | 
by (simp add: cong_def)  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
169  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
170  | 
lemma A_ncong_p: "x \<in> A \<Longrightarrow> [x \<noteq> 0](mod p)"  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
171  | 
by (rule nonzero_mod_p) (auto simp add: A_def)  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
172  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
173  | 
lemma A_greater_zero: "x \<in> A \<Longrightarrow> 0 < x"  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
174  | 
by (auto simp add: A_def)  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
175  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
176  | 
lemma B_ncong_p: "x \<in> B \<Longrightarrow> [x \<noteq> 0](mod p)"  | 
| 65413 | 177  | 
by (auto simp: B_def p_prime p_a_relprime A_ncong_p dest: cong_prime_prod_zero_int)  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
178  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
179  | 
lemma B_greater_zero: "x \<in> B \<Longrightarrow> 0 < x"  | 
| 56544 | 180  | 
using a_nonzero by (auto simp add: B_def A_greater_zero)  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
181  | 
|
| 66888 | 182  | 
lemma B_mod_greater_zero:  | 
183  | 
"0 < x mod int p" if "x \<in> B"  | 
|
184  | 
proof -  | 
|
185  | 
from that have "x mod int p \<noteq> 0"  | 
|
186  | 
using B_ncong_p cong_def cong_mult_self_left by blast  | 
|
187  | 
moreover from that have "0 < x"  | 
|
188  | 
by (rule B_greater_zero)  | 
|
189  | 
then have "0 \<le> x mod int p"  | 
|
190  | 
by (auto simp add: mod_int_pos_iff intro: neq_le_trans)  | 
|
191  | 
ultimately show ?thesis  | 
|
192  | 
by simp  | 
|
193  | 
qed  | 
|
194  | 
||
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
195  | 
lemma C_greater_zero: "y \<in> C \<Longrightarrow> 0 < y"  | 
| 66888 | 196  | 
by (auto simp add: C_def B_mod_greater_zero)  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
197  | 
|
| 65413 | 198  | 
lemma F_subset: "F \<subseteq> {x. 0 < x \<and> x \<le> ((int p - 1) div 2)}"
 | 
| 
80401
 
31bf95336f16
dropped references to theorems from transitional theory Divides.thy
 
haftmann 
parents: 
76262 
diff
changeset
 | 
199  | 
using p_ge_2 by (auto simp add: F_def E_def C_def intro: p_odd_int)  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
200  | 
|
| 65413 | 201  | 
lemma D_subset: "D \<subseteq> {x. 0 < x \<and> x \<le> ((p - 1) div 2)}"
 | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
202  | 
by (auto simp add: D_def C_greater_zero)  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
203  | 
|
| 65413 | 204  | 
lemma F_eq: "F = {x. \<exists>y \<in> A. (x = p - ((y * a) mod p) \<and> (int p - 1) div 2 < (y * a) mod p)}"
 | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
205  | 
by (auto simp add: F_def E_def D_def C_def B_def A_def)  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
206  | 
|
| 65413 | 207  | 
lemma D_eq: "D = {x. \<exists>y \<in> A. (x = (y * a) mod p \<and> (y * a) mod p \<le> (int p - 1) div 2)}"
 | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
208  | 
by (auto simp add: D_def C_def B_def A_def)  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
209  | 
|
| 65413 | 210  | 
lemma all_A_relprime:  | 
| 67051 | 211  | 
"coprime x p" if "x \<in> A"  | 
212  | 
proof -  | 
|
213  | 
from A_ncong_p [OF that] have "\<not> int p dvd x"  | 
|
| 
68707
 
5b269897df9d
de-applying and removal of obsolete aliases
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
214  | 
by (simp add: cong_0_iff)  | 
| 67051 | 215  | 
with p_prime show ?thesis  | 
216  | 
by (metis (no_types) coprime_commute prime_imp_coprime prime_nat_int_transfer)  | 
|
217  | 
qed  | 
|
218  | 
||
219  | 
lemma A_prod_relprime: "coprime (prod id A) p"  | 
|
220  | 
by (auto intro: prod_coprime_left all_A_relprime)  | 
|
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
221  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
222  | 
|
| 60526 | 223  | 
subsection \<open>Relationships Between Gauss Sets\<close>  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
224  | 
|
| 65413 | 225  | 
lemma StandardRes_inj_on_ResSet: "ResSet m X \<Longrightarrow> inj_on (\<lambda>b. b mod m) X"  | 
| 66888 | 226  | 
by (auto simp add: ResSet_def inj_on_def cong_def)  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
227  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
228  | 
lemma B_card_eq_A: "card B = card A"  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
229  | 
using finite_A by (simp add: finite_A B_def inj_on_xa_A card_image)  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
230  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
231  | 
lemma B_card_eq: "card B = nat ((int p - 1) div 2)"  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
232  | 
by (simp add: B_card_eq_A A_card_eq)  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
233  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
234  | 
lemma F_card_eq_E: "card F = card E"  | 
| 76262 | 235  | 
using finite_E by (simp add: F_def card_image)  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
236  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
237  | 
lemma C_card_eq_B: "card C = card B"  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
238  | 
proof -  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
239  | 
have "inj_on (\<lambda>x. x mod p) B"  | 
| 65413 | 240  | 
by (metis SR_B_inj)  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
241  | 
then show ?thesis  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
242  | 
by (metis C_def card_image)  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
243  | 
qed  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
244  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
245  | 
lemma D_E_disj: "D \<inter> E = {}"
 | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
246  | 
by (auto simp add: D_def E_def)  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
247  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
248  | 
lemma C_card_eq_D_plus_E: "card C = card D + card E"  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
249  | 
by (auto simp add: C_eq card_Un_disjoint D_E_disj finite_D finite_E)  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
250  | 
|
| 64272 | 251  | 
lemma C_prod_eq_D_times_E: "prod id E * prod id D = prod id C"  | 
252  | 
by (metis C_eq D_E_disj finite_D finite_E inf_commute prod.union_disjoint sup_commute)  | 
|
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
253  | 
|
| 64272 | 254  | 
lemma C_B_zcong_prod: "[prod id C = prod id B] (mod p)"  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
255  | 
apply (auto simp add: C_def)  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
256  | 
apply (insert finite_B SR_B_inj)  | 
| 64272 | 257  | 
apply (drule prod.reindex [of "\<lambda>x. x mod int p" B id])  | 
| 57418 | 258  | 
apply auto  | 
| 66888 | 259  | 
apply (rule cong_prod)  | 
260  | 
apply (auto simp add: cong_def)  | 
|
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
261  | 
done  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
262  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
263  | 
lemma F_Un_D_subset: "(F \<union> D) \<subseteq> A"  | 
| 65413 | 264  | 
by (intro Un_least subset_trans [OF F_subset] subset_trans [OF D_subset]) (auto simp: A_def)  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
265  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
266  | 
lemma F_D_disj: "(F \<inter> D) = {}"
 | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
267  | 
proof (auto simp add: F_eq D_eq)  | 
| 65413 | 268  | 
fix y z :: int  | 
269  | 
assume "p - (y * a) mod p = (z * a) mod p"  | 
|
270  | 
then have "[(y * a) mod p + (z * a) mod p = 0] (mod p)"  | 
|
| 66888 | 271  | 
by (metis add.commute diff_eq_eq dvd_refl cong_def dvd_eq_mod_eq_0 mod_0)  | 
| 65413 | 272  | 
moreover have "[y * a = (y * a) mod p] (mod p)"  | 
| 66888 | 273  | 
by (metis cong_def mod_mod_trivial)  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
274  | 
ultimately have "[a * (y + z) = 0] (mod p)"  | 
| 66888 | 275  | 
by (metis cong_def mod_add_left_eq mod_add_right_eq mult.commute ring_class.ring_distribs(1))  | 
| 65413 | 276  | 
with p_prime a_nonzero p_a_relprime have a: "[y + z = 0] (mod p)"  | 
| 
63534
 
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
 
eberlm <eberlm@in.tum.de> 
parents: 
62429 
diff
changeset
 | 
277  | 
by (auto dest!: cong_prime_prod_zero_int)  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
278  | 
assume b: "y \<in> A" and c: "z \<in> A"  | 
| 65413 | 279  | 
then have "0 < y + z"  | 
280  | 
by (auto simp: A_def)  | 
|
281  | 
moreover from b c p_eq2 have "y + z < p"  | 
|
282  | 
by (auto simp: A_def)  | 
|
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
283  | 
ultimately show False  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
284  | 
by (metis a nonzero_mod_p)  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
285  | 
qed  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
286  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
287  | 
lemma F_Un_D_card: "card (F \<union> D) = nat ((p - 1) div 2)"  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
288  | 
proof -  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
289  | 
have "card (F \<union> D) = card E + card D"  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
290  | 
by (auto simp add: finite_F finite_D F_D_disj card_Un_disjoint F_card_eq_E)  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
291  | 
then have "card (F \<union> D) = card C"  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
292  | 
by (simp add: C_card_eq_D_plus_E)  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
293  | 
then show "card (F \<union> D) = nat ((p - 1) div 2)"  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
294  | 
by (simp add: C_card_eq_B B_card_eq)  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
295  | 
qed  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
296  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
297  | 
lemma F_Un_D_eq_A: "F \<union> D = A"  | 
| 65413 | 298  | 
using finite_A F_Un_D_subset A_card_eq F_Un_D_card by (auto simp add: card_seteq)  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
299  | 
|
| 65413 | 300  | 
lemma prod_D_F_eq_prod_A: "prod id D * prod id F = prod id A"  | 
| 64272 | 301  | 
by (metis F_D_disj F_Un_D_eq_A Int_commute Un_commute finite_D finite_F prod.union_disjoint)  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
302  | 
|
| 65413 | 303  | 
lemma prod_F_zcong: "[prod id F = ((-1) ^ (card E)) * prod id E] (mod p)"  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
304  | 
proof -  | 
| 67399 | 305  | 
have FE: "prod id F = prod ((-) p) E"  | 
| 76262 | 306  | 
using finite_E prod.reindex[OF inj_on_diff_left] by (auto simp add: F_def)  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
307  | 
then have "\<forall>x \<in> E. [(p-x) mod p = - x](mod p)"  | 
| 66888 | 308  | 
by (metis cong_def minus_mod_self1 mod_mod_trivial)  | 
| 67399 | 309  | 
then have "[prod ((\<lambda>x. x mod p) \<circ> ((-) p)) E = prod (uminus) E](mod p)"  | 
310  | 
using finite_E p_ge_2 cong_prod [of E "(\<lambda>x. x mod p) \<circ> ((-) p)" uminus p]  | 
|
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
311  | 
by auto  | 
| 64272 | 312  | 
then have two: "[prod id F = prod (uminus) E](mod p)"  | 
| 66888 | 313  | 
by (metis FE cong_cong_mod_int cong_refl cong_prod minus_mod_self1)  | 
| 65413 | 314  | 
have "prod uminus E = (-1) ^ card E * prod id E"  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
315  | 
using finite_E by (induct set: finite) auto  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
316  | 
with two show ?thesis  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
317  | 
by simp  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
318  | 
qed  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
319  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
320  | 
|
| 60526 | 321  | 
subsection \<open>Gauss' Lemma\<close>  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
322  | 
|
| 64272 | 323  | 
lemma aux: "prod id A * (- 1) ^ card E * a ^ card A * (- 1) ^ card E = prod id A * a ^ card A"  | 
| 66888 | 324  | 
by auto  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
325  | 
|
| 65413 | 326  | 
theorem pre_gauss_lemma: "[a ^ nat((int p - 1) div 2) = (-1) ^ (card E)] (mod p)"  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
327  | 
proof -  | 
| 64272 | 328  | 
have "[prod id A = prod id F * prod id D](mod p)"  | 
| 69654 | 329  | 
by (auto simp: prod_D_F_eq_prod_A mult.commute cong del: prod.cong_simp)  | 
| 64272 | 330  | 
then have "[prod id A = ((-1)^(card E) * prod id E) * prod id D] (mod p)"  | 
| 66888 | 331  | 
by (rule cong_trans) (metis cong_scalar_right prod_F_zcong)  | 
| 64272 | 332  | 
then have "[prod id A = ((-1)^(card E) * prod id C)] (mod p)"  | 
| 66888 | 333  | 
using finite_D finite_E by (auto simp add: ac_simps C_prod_eq_D_times_E C_eq D_E_disj prod.union_disjoint)  | 
| 64272 | 334  | 
then have "[prod id A = ((-1)^(card E) * prod id B)] (mod p)"  | 
| 66888 | 335  | 
by (rule cong_trans) (metis C_B_zcong_prod cong_scalar_left)  | 
| 65413 | 336  | 
then have "[prod id A = ((-1)^(card E) * prod id ((\<lambda>x. x * a) ` A))] (mod p)"  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
337  | 
by (simp add: B_def)  | 
| 65413 | 338  | 
then have "[prod id A = ((-1)^(card E) * prod (\<lambda>x. x * a) A)] (mod p)"  | 
| 64272 | 339  | 
by (simp add: inj_on_xa_A prod.reindex)  | 
| 65413 | 340  | 
moreover have "prod (\<lambda>x. x * a) A = prod (\<lambda>x. a) A * prod id A"  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
341  | 
using finite_A by (induct set: finite) auto  | 
| 65413 | 342  | 
ultimately have "[prod id A = ((-1)^(card E) * (prod (\<lambda>x. a) A * prod id A))] (mod p)"  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
343  | 
by simp  | 
| 65413 | 344  | 
then have "[prod id A = ((-1)^(card E) * a^(card A) * prod id A)](mod p)"  | 
| 66888 | 345  | 
by (rule cong_trans)  | 
| 76262 | 346  | 
(simp add: cong_scalar_left cong_scalar_right finite_A ac_simps)  | 
| 64272 | 347  | 
then have a: "[prod id A * (-1)^(card E) =  | 
348  | 
((-1)^(card E) * a^(card A) * prod id A * (-1)^(card E))](mod p)"  | 
|
| 66888 | 349  | 
by (rule cong_scalar_right)  | 
| 64272 | 350  | 
then have "[prod id A * (-1)^(card E) = prod id A *  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
351  | 
(-1)^(card E) * a^(card A) * (-1)^(card E)](mod p)"  | 
| 66888 | 352  | 
by (rule cong_trans) (simp add: a ac_simps)  | 
| 64272 | 353  | 
then have "[prod id A * (-1)^(card E) = prod id A * a^(card A)](mod p)"  | 
| 69654 | 354  | 
by (rule cong_trans) (simp add: aux cong del: prod.cong_simp)  | 
| 
58410
 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 
haftmann 
parents: 
58288 
diff
changeset
 | 
355  | 
with A_prod_relprime have "[(- 1) ^ card E = a ^ card A](mod p)"  | 
| 
68707
 
5b269897df9d
de-applying and removal of obsolete aliases
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
356  | 
by (metis cong_mult_lcancel)  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
357  | 
then show ?thesis  | 
| 66888 | 358  | 
by (simp add: A_card_eq cong_sym)  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
359  | 
qed  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
360  | 
|
| 65413 | 361  | 
theorem gauss_lemma: "Legendre a p = (-1) ^ (card E)"  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
362  | 
proof -  | 
| 65413 | 363  | 
from euler_criterion p_prime p_ge_2 have "[Legendre a p = a^(nat (((p) - 1) div 2))] (mod p)"  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
364  | 
by auto  | 
| 65413 | 365  | 
moreover have "int ((p - 1) div 2) = (int p - 1) div 2"  | 
366  | 
using p_eq2 by linarith  | 
|
367  | 
then have "[a ^ nat (int ((p - 1) div 2)) = a ^ nat ((int p - 1) div 2)] (mod int p)"  | 
|
368  | 
by force  | 
|
369  | 
ultimately have "[Legendre a p = (-1) ^ (card E)] (mod p)"  | 
|
| 66888 | 370  | 
using pre_gauss_lemma cong_trans by blast  | 
| 65413 | 371  | 
moreover from p_a_relprime have "Legendre a p = 1 \<or> Legendre a p = -1"  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
372  | 
by (auto simp add: Legendre_def)  | 
| 65413 | 373  | 
moreover have "(-1::int) ^ (card E) = 1 \<or> (-1::int) ^ (card E) = -1"  | 
| 
64282
 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 
eberlm <eberlm@in.tum.de> 
parents: 
64272 
diff
changeset
 | 
374  | 
using neg_one_even_power neg_one_odd_power by blast  | 
| 
 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 
eberlm <eberlm@in.tum.de> 
parents: 
64272 
diff
changeset
 | 
375  | 
moreover have "[1 \<noteq> - 1] (mod int p)"  | 
| 
68707
 
5b269897df9d
de-applying and removal of obsolete aliases
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
376  | 
using cong_iff_dvd_diff [where 'a=int] nonzero_mod_p[of 2] p_odd_int  | 
| 
 
5b269897df9d
de-applying and removal of obsolete aliases
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
377  | 
by fastforce  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
378  | 
ultimately show ?thesis  | 
| 66888 | 379  | 
by (auto simp add: cong_sym)  | 
| 
55730
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
380  | 
qed  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
381  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
382  | 
end  | 
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
383  | 
|
| 
 
97ff9276e12d
Gauss.thy ported from Old_Number_Theory (unfinished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
384  | 
end  |