author | wenzelm |
Wed, 12 Mar 2025 11:39:00 +0100 | |
changeset 82265 | 4b875a4c83b0 |
parent 80612 | e65eed943bee |
permissions | -rw-r--r-- |
41959 | 1 |
(* Title: HOL/Number_Theory/Residues.thy |
31719 | 2 |
Author: Jeremy Avigad |
3 |
||
41541 | 4 |
An algebraic treatment of residue rings, and resulting proofs of |
41959 | 5 |
Euler's theorem and Wilson's theorem. |
31719 | 6 |
*) |
7 |
||
60526 | 8 |
section \<open>Residue rings\<close> |
31719 | 9 |
|
10 |
theory Residues |
|
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
65066
diff
changeset
|
11 |
imports |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
65066
diff
changeset
|
12 |
Cong |
66453
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
66305
diff
changeset
|
13 |
"HOL-Algebra.Multiplicative_Group" |
65465
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
14 |
Totient |
31719 | 15 |
begin |
16 |
||
80084
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
17 |
lemma (in ring_1) CHAR_dvd_CARD: "CHAR('a) dvd card (UNIV :: 'a set)" |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
18 |
proof (cases "card (UNIV :: 'a set) = 0") |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
19 |
case False |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
20 |
hence [intro]: "CHAR('a) > 0" |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
21 |
by (simp add: card_eq_0_iff finite_imp_CHAR_pos) |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
22 |
define G where "G = \<lparr> carrier = (UNIV :: 'a set), monoid.mult = (+), one = (0 :: 'a) \<rparr>" |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
23 |
define H where "H = (of_nat ` {..<CHAR('a)} :: 'a set)" |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
24 |
interpret group G |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
25 |
proof (rule groupI) |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
26 |
fix x assume x: "x \<in> carrier G" |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
27 |
show "\<exists>y\<in>carrier G. y \<otimes>\<^bsub>G\<^esub> x = \<one>\<^bsub>G\<^esub>" |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
28 |
by (intro bexI[of _ "-x"]) (auto simp: G_def) |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
29 |
qed (auto simp: G_def add_ac) |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
30 |
|
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
31 |
interpret subgroup H G |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
32 |
proof |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
33 |
show "\<one>\<^bsub>G\<^esub> \<in> H" |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
34 |
using False unfolding G_def H_def by force |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
35 |
next |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
36 |
fix x y :: 'a |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
37 |
assume "x \<in> H" "y \<in> H" |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
38 |
then obtain x' y' where [simp]: "x = of_nat x'" "y = of_nat y'" |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
39 |
by (auto simp: H_def) |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
40 |
have "x + y = of_nat ((x' + y') mod CHAR('a))" |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
41 |
by (auto simp flip: of_nat_add simp: of_nat_eq_iff_cong_CHAR) |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
42 |
moreover have "(x' + y') mod CHAR('a) < CHAR('a)" |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
43 |
using H_def \<open>y \<in> H\<close> by fastforce |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
44 |
ultimately show "x \<otimes>\<^bsub>G\<^esub> y \<in> H" |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
45 |
by (auto simp: H_def G_def intro!: imageI) |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
46 |
next |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
47 |
fix x :: 'a |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
48 |
assume x: "x \<in> H" |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
49 |
then obtain x' where [simp]: "x = of_nat x'" and x': "x' < CHAR('a)" |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
50 |
by (auto simp: H_def) |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
51 |
have "CHAR('a) dvd x' + (CHAR('a) - x') mod CHAR('a)" |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
52 |
using mod_eq_0_iff_dvd mod_if x' by fastforce |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
53 |
hence "x + of_nat ((CHAR('a) - x') mod CHAR('a)) = 0" |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
54 |
by (auto simp flip: of_nat_add simp: of_nat_eq_0_iff_char_dvd) |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
55 |
moreover from this have "inv\<^bsub>G\<^esub> x = of_nat ((CHAR('a) - x') mod CHAR('a))" |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
56 |
by (intro inv_equality) (auto simp: G_def add_ac) |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
57 |
moreover have "of_nat ((CHAR('a) - x') mod CHAR('a)) \<in> H" |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
58 |
unfolding H_def using \<open>CHAR('a) > 0\<close> by (intro imageI) auto |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
59 |
ultimately show "inv\<^bsub>G\<^esub> x \<in> H" by force |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
60 |
qed (auto simp: G_def H_def) |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
61 |
|
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
62 |
have "card H dvd card (rcosets\<^bsub>G\<^esub> H) * card H" |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
63 |
by simp |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
64 |
also have "card (rcosets\<^bsub>G\<^esub> H) * card H = Coset.order G" |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
65 |
proof (rule lagrange_finite) |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
66 |
show "finite (carrier G)" |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
67 |
using False card_ge_0_finite by (auto simp: G_def) |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
68 |
qed (fact is_subgroup) |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
69 |
finally have "card H dvd card (UNIV :: 'a set)" |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
70 |
by (simp add: Coset.order_def G_def) |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
71 |
also have "card H = card {..<CHAR('a)}" |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
72 |
unfolding H_def by (intro card_image inj_onI) (auto simp: of_nat_eq_iff_cong_CHAR cong_def) |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
73 |
finally show "CHAR('a) dvd card (UNIV :: 'a set)" |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
74 |
by simp |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
75 |
qed auto |
173548e4d5d0
moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
78522
diff
changeset
|
76 |
|
66305 | 77 |
definition QuadRes :: "int \<Rightarrow> int \<Rightarrow> bool" |
78 |
where "QuadRes p a = (\<exists>y. ([y^2 = a] (mod p)))" |
|
64282
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
64272
diff
changeset
|
79 |
|
66305 | 80 |
definition Legendre :: "int \<Rightarrow> int \<Rightarrow> int" |
81 |
where "Legendre a p = |
|
82 |
(if ([a = 0] (mod p)) then 0 |
|
83 |
else if QuadRes p a then 1 |
|
84 |
else -1)" |
|
85 |
||
64282
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
eberlm <eberlm@in.tum.de>
parents:
64272
diff
changeset
|
86 |
|
60527 | 87 |
subsection \<open>A locale for residue rings\<close> |
31719 | 88 |
|
60527 | 89 |
definition residue_ring :: "int \<Rightarrow> int ring" |
66305 | 90 |
where |
91 |
"residue_ring m = |
|
92 |
\<lparr>carrier = {0..m - 1}, |
|
93 |
monoid.mult = \<lambda>x y. (x * y) mod m, |
|
94 |
one = 1, |
|
95 |
zero = 0, |
|
96 |
add = \<lambda>x y. (x + y) mod m\<rparr>" |
|
31719 | 97 |
|
98 |
locale residues = |
|
99 |
fixes m :: int and R (structure) |
|
100 |
assumes m_gt_one: "m > 1" |
|
75963
884dbbc8e1b3
avoid duplicate fact error on global_interpretation of residues
haftmann
parents:
71252
diff
changeset
|
101 |
defines R_m_def: "R \<equiv> residue_ring m" |
44872 | 102 |
begin |
31719 | 103 |
|
104 |
lemma abelian_group: "abelian_group R" |
|
65066 | 105 |
proof - |
106 |
have "\<exists>y\<in>{0..m - 1}. (x + y) mod m = 0" if "0 \<le> x" "x < m" for x |
|
107 |
proof (cases "x = 0") |
|
108 |
case True |
|
109 |
with m_gt_one show ?thesis by simp |
|
110 |
next |
|
111 |
case False |
|
112 |
then have "(x + (m - x)) mod m = 0" |
|
113 |
by simp |
|
114 |
with m_gt_one that show ?thesis |
|
115 |
by (metis False atLeastAtMost_iff diff_ge_0_iff_ge diff_left_mono int_one_le_iff_zero_less less_le) |
|
116 |
qed |
|
117 |
with m_gt_one show ?thesis |
|
75963
884dbbc8e1b3
avoid duplicate fact error on global_interpretation of residues
haftmann
parents:
71252
diff
changeset
|
118 |
by (fastforce simp add: R_m_def residue_ring_def mod_add_right_eq ac_simps intro!: abelian_groupI) |
65899 | 119 |
qed |
31719 | 120 |
|
121 |
lemma comm_monoid: "comm_monoid R" |
|
78522
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
122 |
proof - |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
123 |
have "\<And>x y z. \<lbrakk>x \<in> carrier R; y \<in> carrier R; z \<in> carrier R\<rbrakk> \<Longrightarrow> x \<otimes> y \<otimes> z = x \<otimes> (y \<otimes> z)" |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
124 |
"\<And>x y. \<lbrakk>x \<in> carrier R; y \<in> carrier R\<rbrakk> \<Longrightarrow> x \<otimes> y = y \<otimes> x" |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
125 |
unfolding R_m_def residue_ring_def |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
126 |
by (simp_all add: algebra_simps mod_mult_right_eq) |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
127 |
then show ?thesis |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
128 |
unfolding R_m_def residue_ring_def |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
129 |
by unfold_locales (use m_gt_one in simp_all) |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
130 |
qed |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
131 |
|
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
132 |
interpretation comm_monoid R |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
133 |
using comm_monoid by blast |
31719 | 134 |
|
135 |
lemma cring: "cring R" |
|
65066 | 136 |
apply (intro cringI abelian_group comm_monoid) |
75963
884dbbc8e1b3
avoid duplicate fact error on global_interpretation of residues
haftmann
parents:
71252
diff
changeset
|
137 |
unfolding R_m_def residue_ring_def |
65066 | 138 |
apply (auto simp add: comm_semiring_class.distrib mod_add_eq mod_mult_left_eq) |
41541 | 139 |
done |
31719 | 140 |
|
141 |
end |
|
142 |
||
143 |
sublocale residues < cring |
|
144 |
by (rule cring) |
|
145 |
||
146 |
||
41541 | 147 |
context residues |
148 |
begin |
|
31719 | 149 |
|
60527 | 150 |
text \<open> |
151 |
These lemmas translate back and forth between internal and |
|
152 |
external concepts. |
|
153 |
\<close> |
|
31719 | 154 |
|
155 |
lemma res_carrier_eq: "carrier R = {0..m - 1}" |
|
75963
884dbbc8e1b3
avoid duplicate fact error on global_interpretation of residues
haftmann
parents:
71252
diff
changeset
|
156 |
by (auto simp: R_m_def residue_ring_def) |
31719 | 157 |
|
158 |
lemma res_add_eq: "x \<oplus> y = (x + y) mod m" |
|
75963
884dbbc8e1b3
avoid duplicate fact error on global_interpretation of residues
haftmann
parents:
71252
diff
changeset
|
159 |
by (auto simp: R_m_def residue_ring_def) |
31719 | 160 |
|
161 |
lemma res_mult_eq: "x \<otimes> y = (x * y) mod m" |
|
75963
884dbbc8e1b3
avoid duplicate fact error on global_interpretation of residues
haftmann
parents:
71252
diff
changeset
|
162 |
by (auto simp: R_m_def residue_ring_def) |
31719 | 163 |
|
164 |
lemma res_zero_eq: "\<zero> = 0" |
|
75963
884dbbc8e1b3
avoid duplicate fact error on global_interpretation of residues
haftmann
parents:
71252
diff
changeset
|
165 |
by (auto simp: R_m_def residue_ring_def) |
31719 | 166 |
|
167 |
lemma res_one_eq: "\<one> = 1" |
|
75963
884dbbc8e1b3
avoid duplicate fact error on global_interpretation of residues
haftmann
parents:
71252
diff
changeset
|
168 |
by (auto simp: R_m_def residue_ring_def units_of_def) |
31719 | 169 |
|
78522
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
170 |
lemma res_units_eq: "Units R = {x. 0 < x \<and> x < m \<and> coprime x m}" (is "_ = ?rhs") |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
171 |
proof |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
172 |
show "Units R \<subseteq> ?rhs" |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
173 |
using zero_less_mult_iff invertible_coprime |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
174 |
by (fastforce simp: Units_def R_m_def residue_ring_def) |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
175 |
next |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
176 |
show "?rhs \<subseteq> Units R" |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
177 |
unfolding Units_def R_m_def residue_ring_def |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
178 |
by (force simp add: cong_def coprime_iff_invertible'_int mult.commute) |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
179 |
qed |
31719 | 180 |
|
181 |
lemma res_neg_eq: "\<ominus> x = (- x) mod m" |
|
78522
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
182 |
proof - |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
183 |
have "\<ominus> x = (THE y. 0 \<le> y \<and> y < m \<and> (x + y) mod m = 0 \<and> (y + x) mod m = 0)" |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
184 |
by (simp add: R_m_def a_inv_def m_inv_def residue_ring_def) |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
185 |
also have "\<dots> = (- x) mod m" |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
186 |
proof - |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
187 |
have "\<And>y. 0 \<le> y \<and> y < m \<and> (x + y) mod m = 0 \<and> (y + x) mod m = 0 \<Longrightarrow> |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
188 |
y = - x mod m" |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
189 |
by (metis minus_add_cancel mod_add_eq plus_int_code(1) zmod_trivial_iff) |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
190 |
then show ?thesis |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
191 |
by (intro the_equality) (use m_gt_one in \<open>simp add: add.commute mod_add_right_eq\<close>) |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
192 |
qed |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
193 |
finally show ?thesis . |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
194 |
qed |
31719 | 195 |
|
44872 | 196 |
lemma finite [iff]: "finite (carrier R)" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
65066
diff
changeset
|
197 |
by (simp add: res_carrier_eq) |
31719 | 198 |
|
44872 | 199 |
lemma finite_Units [iff]: "finite (Units R)" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
65066
diff
changeset
|
200 |
by (simp add: finite_ring_finite_units) |
31719 | 201 |
|
60527 | 202 |
text \<open> |
63167 | 203 |
The function \<open>a \<mapsto> a mod m\<close> maps the integers to the |
60527 | 204 |
residue classes. The following lemmas show that this mapping |
205 |
respects addition and multiplication on the integers. |
|
206 |
\<close> |
|
31719 | 207 |
|
60527 | 208 |
lemma mod_in_carrier [iff]: "a mod m \<in> carrier R" |
209 |
unfolding res_carrier_eq |
|
210 |
using insert m_gt_one by auto |
|
31719 | 211 |
|
212 |
lemma add_cong: "(x mod m) \<oplus> (y mod m) = (x + y) mod m" |
|
75963
884dbbc8e1b3
avoid duplicate fact error on global_interpretation of residues
haftmann
parents:
71252
diff
changeset
|
213 |
by (auto simp: R_m_def residue_ring_def mod_simps) |
31719 | 214 |
|
215 |
lemma mult_cong: "(x mod m) \<otimes> (y mod m) = (x * y) mod m" |
|
75963
884dbbc8e1b3
avoid duplicate fact error on global_interpretation of residues
haftmann
parents:
71252
diff
changeset
|
216 |
by (auto simp: R_m_def residue_ring_def mod_simps) |
31719 | 217 |
|
218 |
lemma zero_cong: "\<zero> = 0" |
|
75963
884dbbc8e1b3
avoid duplicate fact error on global_interpretation of residues
haftmann
parents:
71252
diff
changeset
|
219 |
by (auto simp: R_m_def residue_ring_def) |
31719 | 220 |
|
221 |
lemma one_cong: "\<one> = 1 mod m" |
|
75963
884dbbc8e1b3
avoid duplicate fact error on global_interpretation of residues
haftmann
parents:
71252
diff
changeset
|
222 |
using m_gt_one by (auto simp: R_m_def residue_ring_def) |
31719 | 223 |
|
60527 | 224 |
(* FIXME revise algebra library to use 1? *) |
67341
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents:
67091
diff
changeset
|
225 |
lemma pow_cong: "(x mod m) [^] n = x^n mod m" |
65066 | 226 |
using m_gt_one |
78522
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
227 |
proof (induct n) |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
228 |
case 0 |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
229 |
then show ?case |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
230 |
by (simp add: one_cong) |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
231 |
next |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
232 |
case (Suc n) |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
233 |
then show ?case |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
234 |
by (simp add: mult_cong power_commutes) |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
235 |
qed |
31719 | 236 |
|
237 |
lemma neg_cong: "\<ominus> (x mod m) = (- x) mod m" |
|
55352 | 238 |
by (metis mod_minus_eq res_neg_eq) |
31719 | 239 |
|
60528 | 240 |
lemma (in residues) prod_cong: "finite A \<Longrightarrow> (\<Otimes>i\<in>A. (f i) mod m) = (\<Prod>i\<in>A. f i) mod m" |
55352 | 241 |
by (induct set: finite) (auto simp: one_cong mult_cong) |
31719 | 242 |
|
60528 | 243 |
lemma (in residues) sum_cong: "finite A \<Longrightarrow> (\<Oplus>i\<in>A. (f i) mod m) = (\<Sum>i\<in>A. f i) mod m" |
55352 | 244 |
by (induct set: finite) (auto simp: zero_cong add_cong) |
31719 | 245 |
|
60688
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents:
60528
diff
changeset
|
246 |
lemma mod_in_res_units [simp]: |
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents:
60528
diff
changeset
|
247 |
assumes "1 < m" and "coprime a m" |
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents:
60528
diff
changeset
|
248 |
shows "a mod m \<in> Units R" |
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents:
60528
diff
changeset
|
249 |
proof (cases "a mod m = 0") |
66305 | 250 |
case True |
251 |
with assms show ?thesis |
|
60688
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents:
60528
diff
changeset
|
252 |
by (auto simp add: res_units_eq gcd_red_int [symmetric]) |
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents:
60528
diff
changeset
|
253 |
next |
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents:
60528
diff
changeset
|
254 |
case False |
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents:
60528
diff
changeset
|
255 |
from assms have "0 < m" by simp |
66305 | 256 |
then have "0 \<le> a mod m" by (rule pos_mod_sign [of m a]) |
60688
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents:
60528
diff
changeset
|
257 |
with False have "0 < a mod m" by simp |
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents:
60528
diff
changeset
|
258 |
with assms show ?thesis |
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents:
60528
diff
changeset
|
259 |
by (auto simp add: res_units_eq gcd_red_int [symmetric] ac_simps) |
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents:
60528
diff
changeset
|
260 |
qed |
31719 | 261 |
|
60528 | 262 |
lemma res_eq_to_cong: "(a mod m) = (b mod m) \<longleftrightarrow> [a = b] (mod m)" |
66888 | 263 |
by (auto simp: cong_def) |
31719 | 264 |
|
265 |
||
60528 | 266 |
text \<open>Simplifying with these will translate a ring equation in R to a congruence.\<close> |
66305 | 267 |
lemmas res_to_cong_simps = |
268 |
add_cong mult_cong pow_cong one_cong |
|
269 |
prod_cong sum_cong neg_cong res_eq_to_cong |
|
31719 | 270 |
|
60527 | 271 |
text \<open>Other useful facts about the residue ring.\<close> |
31719 | 272 |
lemma one_eq_neg_one: "\<one> = \<ominus> \<one> \<Longrightarrow> m = 2" |
78522
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
273 |
using one_cong res_neg_eq res_one_eq zmod_zminus1_eq_if by fastforce |
31719 | 274 |
|
275 |
end |
|
276 |
||
277 |
||
60527 | 278 |
subsection \<open>Prime residues\<close> |
31719 | 279 |
|
280 |
locale residues_prime = |
|
63534
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
63417
diff
changeset
|
281 |
fixes p :: nat and R (structure) |
31719 | 282 |
assumes p_prime [intro]: "prime p" |
63534
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
63417
diff
changeset
|
283 |
defines "R \<equiv> residue_ring (int p)" |
31719 | 284 |
|
285 |
sublocale residues_prime < residues p |
|
80612
e65eed943bee
A lot of new material from the Ramsey development, including a couple of new simprules.
paulson <lp15@cam.ac.uk>
parents:
80084
diff
changeset
|
286 |
proof |
e65eed943bee
A lot of new material from the Ramsey development, including a couple of new simprules.
paulson <lp15@cam.ac.uk>
parents:
80084
diff
changeset
|
287 |
show "1 < int p" |
e65eed943bee
A lot of new material from the Ramsey development, including a couple of new simprules.
paulson <lp15@cam.ac.uk>
parents:
80084
diff
changeset
|
288 |
using prime_gt_1_nat by auto |
e65eed943bee
A lot of new material from the Ramsey development, including a couple of new simprules.
paulson <lp15@cam.ac.uk>
parents:
80084
diff
changeset
|
289 |
qed |
31719 | 290 |
|
44872 | 291 |
context residues_prime |
292 |
begin |
|
31719 | 293 |
|
67051 | 294 |
lemma p_coprime_left: |
295 |
"coprime p a \<longleftrightarrow> \<not> p dvd a" |
|
296 |
using p_prime by (auto intro: prime_imp_coprime dest: coprime_common_divisor) |
|
297 |
||
298 |
lemma p_coprime_right: |
|
299 |
"coprime a p \<longleftrightarrow> \<not> p dvd a" |
|
300 |
using p_coprime_left [of a] by (simp add: ac_simps) |
|
301 |
||
302 |
lemma p_coprime_left_int: |
|
303 |
"coprime (int p) a \<longleftrightarrow> \<not> int p dvd a" |
|
304 |
using p_prime by (auto intro: prime_imp_coprime dest: coprime_common_divisor) |
|
305 |
||
306 |
lemma p_coprime_right_int: |
|
307 |
"coprime a (int p) \<longleftrightarrow> \<not> int p dvd a" |
|
78522
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
308 |
using coprime_commute p_coprime_left_int by blast |
67051 | 309 |
|
31719 | 310 |
lemma is_field: "field R" |
65066 | 311 |
proof - |
66837 | 312 |
have "0 < x \<Longrightarrow> x < int p \<Longrightarrow> coprime (int p) x" for x |
313 |
by (rule prime_imp_coprime) (auto simp add: zdvd_not_zless) |
|
65066 | 314 |
then show ?thesis |
66837 | 315 |
by (intro cring.field_intro2 cring) |
316 |
(auto simp add: res_carrier_eq res_one_eq res_zero_eq res_units_eq ac_simps) |
|
65066 | 317 |
qed |
31719 | 318 |
|
319 |
lemma res_prime_units_eq: "Units R = {1..p - 1}" |
|
78522
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
320 |
by (auto simp add: res_units_eq p_coprime_right_int zdvd_not_zless) |
31719 | 321 |
|
322 |
end |
|
323 |
||
324 |
sublocale residues_prime < field |
|
325 |
by (rule is_field) |
|
326 |
||
327 |
||
60527 | 328 |
section \<open>Test cases: Euler's theorem and Wilson's theorem\<close> |
31719 | 329 |
|
60527 | 330 |
subsection \<open>Euler's theorem\<close> |
31719 | 331 |
|
67051 | 332 |
lemma (in residues) totatives_eq: |
333 |
"totatives (nat m) = nat ` Units R" |
|
55261
ad3604df6bc6
new lemmas involving phi from Lehmer AFP entry
paulson <lp15@cam.ac.uk>
parents:
55242
diff
changeset
|
334 |
proof - |
67051 | 335 |
from m_gt_one have "\<bar>m\<bar> > 1" |
336 |
by simp |
|
337 |
then have "totatives (nat \<bar>m\<bar>) = nat ` abs ` Units R" |
|
338 |
by (auto simp add: totatives_def res_units_eq image_iff le_less) |
|
339 |
(use m_gt_one zless_nat_eq_int_zless in force) |
|
340 |
moreover have "\<bar>m\<bar> = m" "abs ` Units R = Units R" |
|
341 |
using m_gt_one by (auto simp add: res_units_eq image_iff) |
|
342 |
ultimately show ?thesis |
|
343 |
by simp |
|
344 |
qed |
|
345 |
||
346 |
lemma (in residues) totient_eq: |
|
347 |
"totient (nat m) = card (Units R)" |
|
348 |
proof - |
|
65465
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
349 |
have *: "inj_on nat (Units R)" |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
350 |
by (rule inj_onI) (auto simp add: res_units_eq) |
67051 | 351 |
then show ?thesis |
352 |
by (simp add: totient_def totatives_eq card_image) |
|
55261
ad3604df6bc6
new lemmas involving phi from Lehmer AFP entry
paulson <lp15@cam.ac.uk>
parents:
55242
diff
changeset
|
353 |
qed |
ad3604df6bc6
new lemmas involving phi from Lehmer AFP entry
paulson <lp15@cam.ac.uk>
parents:
55242
diff
changeset
|
354 |
|
75963
884dbbc8e1b3
avoid duplicate fact error on global_interpretation of residues
haftmann
parents:
71252
diff
changeset
|
355 |
lemma (in residues_prime) prime_totient_eq: "totient p = p - 1" |
78522
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
356 |
using p_prime totient_prime by blast |
31719 | 357 |
|
65465
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
358 |
lemma (in residues) euler_theorem: |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
359 |
assumes "coprime a m" |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
360 |
shows "[a ^ totient (nat m) = 1] (mod m)" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
65066
diff
changeset
|
361 |
proof - |
65465
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
362 |
have "a ^ totient (nat m) mod m = 1 mod m" |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
363 |
by (metis assms finite_Units m_gt_one mod_in_res_units one_cong totient_eq pow_cong units_power_order_eq_one) |
65066 | 364 |
then show ?thesis |
365 |
using res_eq_to_cong by blast |
|
31719 | 366 |
qed |
367 |
||
368 |
lemma euler_theorem: |
|
65465
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
369 |
fixes a m :: nat |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
370 |
assumes "coprime a m" |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
371 |
shows "[a ^ totient m = 1] (mod m)" |
67091 | 372 |
proof (cases "m = 0 \<or> m = 1") |
60527 | 373 |
case True |
44872 | 374 |
then show ?thesis by auto |
31719 | 375 |
next |
60527 | 376 |
case False |
41541 | 377 |
with assms show ?thesis |
66954 | 378 |
using residues.euler_theorem [of "int m" "int a"] cong_int_iff |
379 |
by (auto simp add: residues_def gcd_int_def) fastforce |
|
31719 | 380 |
qed |
381 |
||
382 |
lemma fermat_theorem: |
|
65465
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
383 |
fixes p a :: nat |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
384 |
assumes "prime p" and "\<not> p dvd a" |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
385 |
shows "[a ^ (p - 1) = 1] (mod p)" |
31719 | 386 |
proof - |
65465
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
387 |
from assms prime_imp_coprime [of p a] have "coprime a p" |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
388 |
by (auto simp add: ac_simps) |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
389 |
then have "[a ^ totient p = 1] (mod p)" |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
390 |
by (rule euler_theorem) |
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
391 |
also have "totient p = p - 1" |
65726
f5d64d094efe
More material on totient function
eberlm <eberlm@in.tum.de>
parents:
65465
diff
changeset
|
392 |
by (rule totient_prime) (rule assms) |
65465
067210a08a22
more fundamental euler's totient function on nat rather than int;
haftmann
parents:
65416
diff
changeset
|
393 |
finally show ?thesis . |
31719 | 394 |
qed |
395 |
||
396 |
||
60526 | 397 |
subsection \<open>Wilson's theorem\<close> |
31719 | 398 |
|
60527 | 399 |
lemma (in field) inv_pair_lemma: "x \<in> Units R \<Longrightarrow> y \<in> Units R \<Longrightarrow> |
400 |
{x, inv x} \<noteq> {y, inv y} \<Longrightarrow> {x, inv x} \<inter> {y, inv y} = {}" |
|
78522
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
401 |
by auto |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
402 |
|
31719 | 403 |
|
404 |
lemma (in residues_prime) wilson_theorem1: |
|
405 |
assumes a: "p > 2" |
|
59730
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents:
59667
diff
changeset
|
406 |
shows "[fact (p - 1) = (-1::int)] (mod p)" |
31719 | 407 |
proof - |
60527 | 408 |
let ?Inverse_Pairs = "{{x, inv x}| x. x \<in> Units R - {\<one>, \<ominus> \<one>}}" |
409 |
have UR: "Units R = {\<one>, \<ominus> \<one>} \<union> \<Union>?Inverse_Pairs" |
|
31719 | 410 |
by auto |
78522
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
411 |
have 11: "\<one> \<noteq> \<ominus> \<one>" |
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
412 |
using a one_eq_neg_one by force |
60527 | 413 |
have "(\<Otimes>i\<in>Units R. i) = (\<Otimes>i\<in>{\<one>, \<ominus> \<one>}. i) \<otimes> (\<Otimes>i\<in>\<Union>?Inverse_Pairs. i)" |
31732 | 414 |
apply (subst UR) |
31719 | 415 |
apply (subst finprod_Un_disjoint) |
78522
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
416 |
using inv_one inv_eq_neg_one_eq apply (auto intro!: funcsetI)+ |
31719 | 417 |
done |
60527 | 418 |
also have "(\<Otimes>i\<in>{\<one>, \<ominus> \<one>}. i) = \<ominus> \<one>" |
78522
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
419 |
by (simp add: 11) |
60527 | 420 |
also have "(\<Otimes>i\<in>(\<Union>?Inverse_Pairs). i) = (\<Otimes>A\<in>?Inverse_Pairs. (\<Otimes>y\<in>A. y))" |
78522
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
421 |
by (rule finprod_Union_disjoint) (auto simp: pairwise_def disjnt_def dest!: inv_eq_imp_eq) |
31719 | 422 |
also have "\<dots> = \<one>" |
68447
0beb927eed89
Adjusting Number_Theory for new Algebra
paulson <lp15@cam.ac.uk>
parents:
67341
diff
changeset
|
423 |
apply (rule finprod_one_eqI) |
78522
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
424 |
apply clarsimp |
60527 | 425 |
apply (subst finprod_insert) |
66305 | 426 |
apply auto |
55352 | 427 |
apply (metis inv_eq_self) |
31719 | 428 |
done |
60527 | 429 |
finally have "(\<Otimes>i\<in>Units R. i) = \<ominus> \<one>" |
31719 | 430 |
by simp |
60527 | 431 |
also have "(\<Otimes>i\<in>Units R. i) = (\<Otimes>i\<in>Units R. i mod p)" |
65066 | 432 |
by (rule finprod_cong') (auto simp: res_units_eq) |
60527 | 433 |
also have "\<dots> = (\<Prod>i\<in>Units R. i) mod p" |
65066 | 434 |
by (rule prod_cong) auto |
31719 | 435 |
also have "\<dots> = fact (p - 1) mod p" |
65066 | 436 |
using assms |
78522
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
437 |
by (simp add: res_prime_units_eq int_prod zmod_int prod_int_eq fact_prod) |
60527 | 438 |
finally have "fact (p - 1) mod p = \<ominus> \<one>" . |
439 |
then show ?thesis |
|
66888 | 440 |
by (simp add: cong_def res_neg_eq res_one_eq zmod_int) |
31719 | 441 |
qed |
442 |
||
55352 | 443 |
lemma wilson_theorem: |
60527 | 444 |
assumes "prime p" |
445 |
shows "[fact (p - 1) = - 1] (mod p)" |
|
55352 | 446 |
proof (cases "p = 2") |
59667
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
447 |
case True |
55352 | 448 |
then show ?thesis |
66888 | 449 |
by (simp add: cong_def fact_prod) |
55352 | 450 |
next |
451 |
case False |
|
452 |
then show ?thesis |
|
453 |
using assms prime_ge_2_nat |
|
454 |
by (metis residues_prime.wilson_theorem1 residues_prime.intro le_eq_less_or_eq) |
|
455 |
qed |
|
31719 | 456 |
|
66304 | 457 |
text \<open> |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
65066
diff
changeset
|
458 |
This result can be transferred to the multiplicative group of |
66305 | 459 |
\<open>\<int>/p\<int>\<close> for \<open>p\<close> prime.\<close> |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
65066
diff
changeset
|
460 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
65066
diff
changeset
|
461 |
lemma mod_nat_int_pow_eq: |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
65066
diff
changeset
|
462 |
fixes n :: nat and p a :: int |
66305 | 463 |
shows "a \<ge> 0 \<Longrightarrow> p \<ge> 0 \<Longrightarrow> (nat a ^ n) mod (nat p) = nat ((a ^ n) mod p)" |
78522
918a9ed06898
some refinements in Algebra and Number_Theory
paulson <lp15@cam.ac.uk>
parents:
75963
diff
changeset
|
464 |
by (simp add: nat_mod_as_int) |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
65066
diff
changeset
|
465 |
|
69785
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
466 |
theorem residue_prime_mult_group_has_gen: |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
65066
diff
changeset
|
467 |
fixes p :: nat |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
65066
diff
changeset
|
468 |
assumes prime_p : "prime p" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
65066
diff
changeset
|
469 |
shows "\<exists>a \<in> {1 .. p - 1}. {1 .. p - 1} = {a^i mod p|i . i \<in> UNIV}" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
65066
diff
changeset
|
470 |
proof - |
66305 | 471 |
have "p \<ge> 2" |
472 |
using prime_gt_1_nat[OF prime_p] by simp |
|
473 |
interpret R: residues_prime p "residue_ring p" |
|
474 |
by (simp add: residues_prime_def prime_p) |
|
475 |
have car: "carrier (residue_ring (int p)) - {\<zero>\<^bsub>residue_ring (int p)\<^esub>} = {1 .. int p - 1}" |
|
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
65066
diff
changeset
|
476 |
by (auto simp add: R.zero_cong R.res_carrier_eq) |
66305 | 477 |
|
67341
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents:
67091
diff
changeset
|
478 |
have "x [^]\<^bsub>residue_ring (int p)\<^esub> i = x ^ i mod (int p)" |
66305 | 479 |
if "x \<in> {1 .. int p - 1}" for x and i :: nat |
480 |
using that R.pow_cong[of x i] by auto |
|
481 |
moreover |
|
482 |
obtain a where a: "a \<in> {1 .. int p - 1}" |
|
67341
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents:
67091
diff
changeset
|
483 |
and a_gen: "{1 .. int p - 1} = {a[^]\<^bsub>residue_ring (int p)\<^esub>i|i::nat . i \<in> UNIV}" |
66305 | 484 |
using field.finite_field_mult_group_has_gen[OF R.is_field] |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
65066
diff
changeset
|
485 |
by (auto simp add: car[symmetric] carrier_mult_of) |
66305 | 486 |
moreover |
487 |
have "nat ` {1 .. int p - 1} = {1 .. p - 1}" (is "?L = ?R") |
|
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
65066
diff
changeset
|
488 |
proof |
66305 | 489 |
have "n \<in> ?R" if "n \<in> ?L" for n |
490 |
using that \<open>p\<ge>2\<close> by force |
|
491 |
then show "?L \<subseteq> ?R" by blast |
|
492 |
have "n \<in> ?L" if "n \<in> ?R" for n |
|
66837 | 493 |
using that \<open>p\<ge>2\<close> by (auto intro: rev_image_eqI [of "int n"]) |
66305 | 494 |
then show "?R \<subseteq> ?L" by blast |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
65066
diff
changeset
|
495 |
qed |
66305 | 496 |
moreover |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
65066
diff
changeset
|
497 |
have "nat ` {a^i mod (int p) | i::nat. i \<in> UNIV} = {nat a^i mod p | i . i \<in> UNIV}" (is "?L = ?R") |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
65066
diff
changeset
|
498 |
proof |
66305 | 499 |
have "x \<in> ?R" if "x \<in> ?L" for x |
500 |
proof - |
|
501 |
from that obtain i where i: "x = nat (a^i mod (int p))" |
|
502 |
by blast |
|
503 |
then have "x = nat a ^ i mod p" |
|
504 |
using mod_nat_int_pow_eq[of a "int p" i] a \<open>p\<ge>2\<close> by auto |
|
505 |
with i show ?thesis by blast |
|
506 |
qed |
|
507 |
then show "?L \<subseteq> ?R" by blast |
|
508 |
have "x \<in> ?L" if "x \<in> ?R" for x |
|
509 |
proof - |
|
510 |
from that obtain i where i: "x = nat a^i mod p" |
|
511 |
by blast |
|
512 |
with mod_nat_int_pow_eq[of a "int p" i] a \<open>p\<ge>2\<close> show ?thesis |
|
513 |
by auto |
|
514 |
qed |
|
515 |
then show "?R \<subseteq> ?L" by blast |
|
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
65066
diff
changeset
|
516 |
qed |
66305 | 517 |
ultimately have "{1 .. p - 1} = {nat a^i mod p | i. i \<in> UNIV}" |
518 |
by presburger |
|
519 |
moreover from a have "nat a \<in> {1 .. p - 1}" by force |
|
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
65066
diff
changeset
|
520 |
ultimately show ?thesis .. |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
65066
diff
changeset
|
521 |
qed |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
65066
diff
changeset
|
522 |
|
69785
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
523 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
524 |
subsection \<open>Upper bound for the number of $n$-th roots\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
525 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
526 |
lemma roots_mod_prime_bound: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
527 |
fixes n c p :: nat |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
528 |
assumes "prime p" "n > 0" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
529 |
defines "A \<equiv> {x\<in>{..<p}. [x ^ n = c] (mod p)}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
530 |
shows "card A \<le> n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
531 |
proof - |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
532 |
define R where "R = residue_ring (int p)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
533 |
from assms(1) interpret residues_prime p R |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
534 |
by unfold_locales (simp_all add: R_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
535 |
interpret R: UP_domain R "UP R" by (unfold_locales) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
536 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
537 |
let ?f = "UnivPoly.monom (UP R) \<one>\<^bsub>R\<^esub> n \<ominus>\<^bsub>(UP R)\<^esub> UnivPoly.monom (UP R) (int (c mod p)) 0" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
538 |
have in_carrier: "int (c mod p) \<in> carrier R" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
539 |
using prime_gt_1_nat[OF assms(1)] by (simp add: R_def residue_ring_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
540 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
541 |
have "deg R ?f = n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
542 |
using assms in_carrier by (simp add: R.deg_minus_eq) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
543 |
hence f_not_zero: "?f \<noteq> \<zero>\<^bsub>UP R\<^esub>" using assms by (auto simp add : R.deg_nzero_nzero) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
544 |
have roots_bound: "finite {a \<in> carrier R. UnivPoly.eval R R id a ?f = \<zero>\<^bsub>R\<^esub>} \<and> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
545 |
card {a \<in> carrier R. UnivPoly.eval R R id a ?f = \<zero>\<^bsub>R\<^esub>} \<le> deg R ?f" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
546 |
using finite in_carrier by (intro R.roots_bound[OF _ f_not_zero]) simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
547 |
have subs: "{x \<in> carrier R. x [^]\<^bsub>R\<^esub> n = int (c mod p)} \<subseteq> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
548 |
{a \<in> carrier R. UnivPoly.eval R R id a ?f = \<zero>\<^bsub>R\<^esub>}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
549 |
using in_carrier by (auto simp: R.evalRR_simps) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
550 |
then have "card {x \<in> carrier R. x [^]\<^bsub>R\<^esub> n = int (c mod p)} \<le> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
551 |
card {a \<in> carrier R. UnivPoly.eval R R id a ?f = \<zero>\<^bsub>R\<^esub>}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
552 |
using finite by (intro card_mono) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
553 |
also have "\<dots> \<le> n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
554 |
using \<open>deg R ?f = n\<close> roots_bound by linarith |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
555 |
also { |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
556 |
fix x assume "x \<in> carrier R" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
557 |
hence "x [^]\<^bsub>R\<^esub> n = (x ^ n) mod (int p)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
558 |
by (subst pow_cong [symmetric]) (auto simp: R_def residue_ring_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
559 |
} |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
560 |
hence "{x \<in> carrier R. x [^]\<^bsub>R\<^esub> n = int (c mod p)} = {x \<in> carrier R. [x ^ n = int c] (mod p)}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
561 |
by (fastforce simp: cong_def zmod_int) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
562 |
also have "bij_betw int A {x \<in> carrier R. [x ^ n = int c] (mod p)}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
563 |
by (rule bij_betwI[of int _ _ nat]) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
564 |
(use cong_int_iff in \<open>force simp: R_def residue_ring_def A_def\<close>)+ |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
565 |
from bij_betw_same_card[OF this] have "card {x \<in> carrier R. [x ^ n = int c] (mod p)} = card A" .. |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
566 |
finally show ?thesis . |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
567 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
568 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
68458
diff
changeset
|
569 |
|
31719 | 570 |
end |