author | wenzelm |
Mon, 12 Apr 2021 14:14:47 +0200 | |
changeset 73563 | 55b66a45bc94 |
parent 69785 | 9e326f6f8a24 |
child 74362 | 0135a0c77b64 |
permissions | -rw-r--r-- |
69785
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1 |
(* |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
2 |
File: HOL/Number_Theory/Residue_Primitive_Roots.thy |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
3 |
Author: Manuel Eberl |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
4 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
5 |
Primitive roots in residue rings: Definition and existence criteria |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
6 |
*) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
7 |
section \<open>Primitive roots in residue rings and Carmichael's function\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
8 |
theory Residue_Primitive_Roots |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
9 |
imports Pocklington |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
10 |
begin |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
11 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
12 |
text \<open> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
13 |
This theory develops the notions of primitive roots (generators) in residue rings. It also |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
14 |
provides a definition and all the basic properties of Carmichael's function $\lambda(n)$, which |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
15 |
is strongly related to this. The proofs mostly follow Apostol's presentation |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
16 |
\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
17 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
18 |
subsection \<open>Primitive roots in residue rings\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
19 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
20 |
text \<open> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
21 |
A primitive root of a residue ring modulo \<open>n\<close> is an element \<open>g\<close> that \<^emph>\<open>generates\<close> the |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
22 |
ring, i.\,e.\ such that for each \<open>x\<close> coprime to \<open>n\<close> there exists an \<open>i\<close> such that $x = g^i$. |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
23 |
A simpler definition is that \<open>g\<close> must have the same order as the cardinality of the |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
24 |
multiplicative group, which is $\varphi(n)$. |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
25 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
26 |
Note that for convenience, this definition does \<^emph>\<open>not\<close> demand \<open>g < n\<close>. |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
27 |
\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
28 |
inductive residue_primroot :: "nat \<Rightarrow> nat \<Rightarrow> bool" where |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
29 |
"n > 0 \<Longrightarrow> coprime n g \<Longrightarrow> ord n g = totient n \<Longrightarrow> residue_primroot n g" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
30 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
31 |
lemma residue_primroot_def [code]: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
32 |
"residue_primroot n x \<longleftrightarrow> n > 0 \<and> coprime n x \<and> ord n x = totient n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
33 |
by (simp add: residue_primroot.simps) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
34 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
35 |
lemma not_residue_primroot_0 [simp]: "~residue_primroot 0 x" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
36 |
by (auto simp: residue_primroot_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
37 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
38 |
lemma residue_primroot_mod [simp]: "residue_primroot n (x mod n) = residue_primroot n x" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
39 |
by (cases "n = 0") (simp_all add: residue_primroot_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
40 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
41 |
lemma residue_primroot_cong: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
42 |
assumes "[x = x'] (mod n)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
43 |
shows "residue_primroot n x = residue_primroot n x'" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
44 |
proof - |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
45 |
have "residue_primroot n x = residue_primroot n (x mod n)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
46 |
by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
47 |
also have "x mod n = x' mod n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
48 |
using assms by (simp add: cong_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
49 |
also have "residue_primroot n (x' mod n) = residue_primroot n x'" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
50 |
by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
51 |
finally show ?thesis . |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
52 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
53 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
54 |
lemma not_residue_primroot_0_right [simp]: "residue_primroot n 0 \<longleftrightarrow> n = 1" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
55 |
by (auto simp: residue_primroot_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
56 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
57 |
lemma residue_primroot_1_iff: "residue_primroot n (Suc 0) \<longleftrightarrow> n \<in> {1, 2}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
58 |
proof |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
59 |
assume *: "residue_primroot n (Suc 0)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
60 |
with totient_gt_1[of n] have "n \<le> 2" by (cases "n \<le> 2") (auto simp: residue_primroot_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
61 |
hence "n \<in> {0, 1, 2}" by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
62 |
thus "n \<in> {1, 2}" using * by (auto simp: residue_primroot_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
63 |
qed (auto simp: residue_primroot_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
64 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
65 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
66 |
subsection \<open>Primitive roots modulo a prime\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
67 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
68 |
text \<open> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
69 |
For prime \<open>p\<close>, we now analyse the number of elements in the ring $\mathbb{Z}/p\mathbb{Z}$ |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
70 |
whose order is precisely \<open>d\<close> for each \<open>d\<close>. |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
71 |
\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
72 |
context |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
73 |
fixes n :: nat and \<psi> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
74 |
assumes n: "n > 1" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
75 |
defines "\<psi> \<equiv> (\<lambda>d. card {x\<in>totatives n. ord n x = d})" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
76 |
begin |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
77 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
78 |
lemma elements_with_ord_restrict_totatives: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
79 |
"d > 0 \<Longrightarrow> {x\<in>{..<n}. ord n x = d} = {x\<in>totatives n. ord n x = d}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
80 |
using n by (auto simp: totatives_def coprime_commute intro!: Nat.gr0I le_neq_trans) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
81 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
82 |
lemma prime_elements_with_ord: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
83 |
assumes "\<psi> d \<noteq> 0" and "prime n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
84 |
and a: "a \<in> totatives n" "ord n a = d" "a \<noteq> 1" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
85 |
shows "inj_on (\<lambda>k. a ^ k mod n) {..<d}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
86 |
and "{x\<in>{..<n}. [x ^ d = 1] (mod n)} = (\<lambda>k. a ^ k mod n) ` {..<d}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
87 |
and "bij_betw (\<lambda>k. a ^ k mod n) (totatives d) {x\<in>{..<n}. ord n x = d}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
88 |
proof - |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
89 |
show inj: "inj_on (\<lambda>k. a ^ k mod n) {..<d}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
90 |
using inj_power_mod[of n a] a by (auto simp: totatives_def coprime_commute) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
91 |
from a have "d > 0" by (auto simp: totatives_def coprime_commute) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
92 |
moreover have "d \<noteq> 1" using a n |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
93 |
by (auto simp: ord_eq_Suc_0_iff totatives_less cong_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
94 |
ultimately have d: "d > 1" by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
95 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
96 |
have *: "(\<lambda>k. a ^ k mod n) ` {..<d} = {x\<in>{..<n}. [x ^ d = 1] (mod n)}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
97 |
proof (rule card_seteq) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
98 |
have "card {x\<in>{..<n}. [x ^ d = 1] (mod n)} \<le> d" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
99 |
using assms a by (intro roots_mod_prime_bound) (auto simp: totatives_def coprime_commute) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
100 |
also have "\<dots> = card ((\<lambda>k. a ^ k mod n) ` {..<d})" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
101 |
using inj by (subst card_image) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
102 |
finally show "card {x \<in> {..<n}. [x ^ d = 1] (mod n)} \<le> \<dots>" . |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
103 |
next |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
104 |
show "(\<lambda>k. a ^ k mod n) ` {..<d} \<subseteq> {x \<in> {..<n}. [x ^ d = 1] (mod n)}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
105 |
proof safe |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
106 |
fix k assume "k < d" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
107 |
have "[(a ^ d) ^ k = 1 ^ k] (mod n)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
108 |
by (intro cong_pow) (use a in \<open>auto simp: ord_divides'\<close>) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
109 |
thus "[(a ^ k mod n) ^ d = 1] (mod n)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
110 |
by (simp add: power_mult [symmetric] cong_def power_mod mult.commute) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
111 |
qed (use \<open>prime n\<close> in \<open>auto dest: prime_gt_1_nat\<close>) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
112 |
qed auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
113 |
thus "{x\<in>{..<n}. [x ^ d = 1] (mod n)} = (\<lambda>k. a ^ k mod n) ` {..<d}" .. |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
114 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
115 |
show "bij_betw (\<lambda>k. a ^ k mod n) (totatives d) {x\<in>{..<n}. ord n x = d}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
116 |
unfolding bij_betw_def |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
117 |
proof (intro conjI inj_on_subset[OF inj] equalityI subsetI) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
118 |
fix b assume "b \<in> (\<lambda>k. a ^ k mod n) ` totatives d" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
119 |
then obtain k where "b = a ^ k mod n" "k \<in> totatives d" by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
120 |
thus "b \<in> {b \<in> {..<n}. ord n b = d}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
121 |
using n a by (simp add: ord_power totatives_def coprime_commute) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
122 |
next |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
123 |
fix b assume "b \<in> {x \<in> {..<n}. ord n x = d}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
124 |
hence b: "ord n b = d" "b < n" by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
125 |
with d have "coprime n b" using ord_eq_0[of n b] by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
126 |
from b have "b \<in> {x\<in>{..<n}. [x ^ d = 1] (mod n)}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
127 |
by (auto simp: ord_divides') |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
128 |
with * obtain k where k: "k < d" "b = a ^ k mod n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
129 |
by blast |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
130 |
with b(2) n a d have "d div gcd k d = ord n b" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
131 |
using \<open>coprime n b\<close> by (auto simp: ord_power) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
132 |
also have "ord n b = d" by (simp add: b) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
133 |
finally have "coprime k d" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
134 |
unfolding coprime_iff_gcd_eq_1 using d a by (subst (asm) div_eq_dividend_iff) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
135 |
with k b d have "k \<in> totatives d" by (auto simp: totatives_def intro!: Nat.gr0I) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
136 |
with k show "b \<in> (\<lambda>k. a ^ k mod n) ` totatives d" by blast |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
137 |
qed (use d n in \<open>auto simp: totatives_less\<close>) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
138 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
139 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
140 |
lemma prime_card_elements_with_ord: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
141 |
assumes "\<psi> d \<noteq> 0" and "prime n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
142 |
shows "\<psi> d = totient d" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
143 |
proof (cases "d = 1") |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
144 |
case True |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
145 |
have "\<psi> 1 = 1" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
146 |
using elements_with_ord_1[of n] n by (simp add: \<psi>_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
147 |
thus ?thesis using True by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
148 |
next |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
149 |
case False |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
150 |
from assms obtain a where a: "a \<in> totatives n" "ord n a = d" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
151 |
by (auto simp: \<psi>_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
152 |
from a have "d > 0" by (auto intro!: Nat.gr0I simp: ord_eq_0 totatives_def coprime_commute) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
153 |
from a and False have "a \<noteq> 1" by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
154 |
from bij_betw_same_card[OF prime_elements_with_ord(3)[OF assms a this]] show "\<psi> d = totient d" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
155 |
using elements_with_ord_restrict_totatives[of d] False a \<open>d > 0\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
156 |
by (simp add: \<psi>_def totient_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
157 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
158 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
159 |
lemma prime_sum_card_elements_with_ord_eq_totient: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
160 |
"(\<Sum>d | d dvd totient n. \<psi> d) = totient n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
161 |
proof - |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
162 |
have "totient n = card (totatives n)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
163 |
by (simp add: totient_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
164 |
also have "totatives n = (\<Union>d\<in>{d. d dvd totient n}. {x\<in>totatives n. ord n x = d})" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
165 |
by (force simp: order_divides_totient totatives_def coprime_commute) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
166 |
also have "card \<dots> = (\<Sum>d | d dvd totient n. \<psi> d)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
167 |
unfolding \<psi>_def using n by (subst card_UN_disjoint) (auto intro!: finite_divisors_nat) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
168 |
finally show ?thesis .. |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
169 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
170 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
171 |
text \<open> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
172 |
We can now show that the number of elements of order \<open>d\<close> is $\varphi(d)$ if $d\mid p - 1$ |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
173 |
and 0 otherwise. |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
174 |
\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
175 |
theorem prime_card_elements_with_ord_eq_totient: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
176 |
assumes "prime n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
177 |
shows "\<psi> d = (if d dvd n - 1 then totient d else 0)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
178 |
proof (cases "d dvd totient n") |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
179 |
case False |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
180 |
thus ?thesis using order_divides_totient[of n] assms |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
181 |
by (auto simp: \<psi>_def totient_prime totatives_def coprime_commute[of n]) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
182 |
next |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
183 |
case True |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
184 |
have "\<psi> d = totient d" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
185 |
proof (rule ccontr) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
186 |
assume neq: "\<psi> d \<noteq> totient d" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
187 |
have le: "\<psi> d \<le> totient d" if "d dvd totient n" for d |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
188 |
using prime_card_elements_with_ord[of d] assms by (cases "\<psi> d = 0") auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
189 |
from neq and le[of d] and True have less: "\<psi> d < totient d" by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
190 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
191 |
have "totient n = (\<Sum>d | d dvd totient n. \<psi> d)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
192 |
using prime_sum_card_elements_with_ord_eq_totient .. |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
193 |
also have "\<dots> < (\<Sum>d | d dvd totient n. totient d)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
194 |
by (rule sum_strict_mono_ex1) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
195 |
(use n le less assms True in \<open>auto intro!: finite_divisors_nat\<close>) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
196 |
also have "\<dots> = totient n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
197 |
using totient_divisor_sum . |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
198 |
finally show False by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
199 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
200 |
with True show ?thesis using assms by (simp add: totient_prime) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
201 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
202 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
203 |
text \<open> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
204 |
As a corollary, we get that the number of primitive roots modulo a prime \<open>p\<close> is |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
205 |
$\varphi(p - 1)$. Since this number is positive, we also get that there \<^emph>\<open>is\<close> at least |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
206 |
one primitive root modulo \<open>p\<close>. |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
207 |
\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
208 |
lemma |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
209 |
assumes "prime n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
210 |
shows prime_card_primitive_roots: "card {x\<in>totatives n. ord n x = n - 1} = totient (n - 1)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
211 |
"card {x\<in>{..<n}. ord n x = n - 1} = totient (n - 1)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
212 |
and prime_primitive_root_exists: "\<exists>x. residue_primroot n x" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
213 |
proof - |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
214 |
show *: "card {x\<in>totatives n. ord n x = n - 1} = totient (n - 1)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
215 |
using prime_card_elements_with_ord_eq_totient[of "n - 1"] assms |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
216 |
by (auto simp: totient_prime \<psi>_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
217 |
thus "card {x\<in>{..<n}. ord n x = n - 1} = totient (n - 1)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
218 |
using assms n elements_with_ord_restrict_totatives[of "n - 1"] by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
219 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
220 |
note * |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
221 |
also have "totient (n - 1) > 0" using n by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
222 |
finally show "\<exists>x. residue_primroot n x" using assms prime_gt_1_nat[of n] |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
223 |
by (subst (asm) card_gt_0_iff) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
224 |
(auto simp: residue_primroot_def totient_prime totatives_def coprime_commute) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
225 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
226 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
227 |
end |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
228 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
229 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
230 |
subsection \<open>Primitive roots modulo powers of an odd prime\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
231 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
232 |
text \<open> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
233 |
Any primitive root \<open>g\<close> modulo an odd prime \<open>p\<close> is also a primitive root modulo $p^k$ for all |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
234 |
$k > 0$ if $[g^{p - 1} \neq 1]\ (\text{mod}\ p^2)$. To show this, we first need the |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
235 |
following lemma. |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
236 |
\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
237 |
lemma residue_primroot_power_prime_power_neq_1: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
238 |
assumes "k \<ge> 2" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
239 |
assumes p: "prime p" "odd p" and "residue_primroot p g" and "[g ^ (p - 1) \<noteq> 1] (mod p\<^sup>2)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
240 |
shows "[g ^ totient (p ^ (k - 1)) \<noteq> 1] (mod (p ^ k))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
241 |
using assms(1) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
242 |
proof (induction k rule: dec_induct) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
243 |
case base |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
244 |
thus ?case using assms by (simp add: totient_prime) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
245 |
next |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
246 |
case (step k) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
247 |
from p have "p > 2" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
248 |
using prime_gt_1_nat[of p] by (cases "p = 2") auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
249 |
from assms have g: "g > 0" by (auto intro!: Nat.gr0I) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
250 |
have "[g ^ totient (p ^ (k - 1)) = 1] (mod p ^ (k - 1))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
251 |
using assms by (intro euler_theorem) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
252 |
(auto simp: residue_primroot_def totatives_def coprime_commute) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
253 |
from cong_to_1_nat[OF this] |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
254 |
obtain t where *: "g ^ totient (p ^ (k - 1)) - 1 = p ^ (k - 1) * t" by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
255 |
have t: "g ^ totient (p ^ (k - 1)) = p ^ (k - 1) * t + 1" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
256 |
using g by (subst * [symmetric]) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
257 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
258 |
have "\<not>p dvd t" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
259 |
proof |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
260 |
assume "p dvd t" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
261 |
then obtain q where [simp]: "t = p * q" by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
262 |
from t have "g ^ totient (p ^ (k - 1)) = p ^ k * q + 1" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
263 |
using \<open>k \<ge> 2\<close> by (cases k) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
264 |
hence "[g ^ totient (p ^ (k - 1)) = p ^ k * q + 1] (mod p ^ k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
265 |
by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
266 |
also have "[p ^ k * q + 1 = 0 * q + 1] (mod p ^ k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
267 |
by (intro cong_add cong_mult) (auto simp: cong_0_iff) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
268 |
finally have "[g ^ totient (p ^ (k - 1)) = 1] (mod p ^ k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
269 |
by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
270 |
with step.IH show False by contradiction |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
271 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
272 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
273 |
from t have "(g ^ totient (p ^ (k - 1))) ^ p = (p ^ (k - 1) * t + 1) ^ p" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
274 |
by (rule arg_cong) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
275 |
also have "(g ^ totient (p ^ (k - 1))) ^ p = g ^ (p * totient (p ^ (k - 1)))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
276 |
by (simp add: power_mult [symmetric] mult.commute) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
277 |
also have "p * totient (p ^ (k - 1)) = totient (p ^ k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
278 |
using p \<open>k \<ge> 2\<close> by (simp add: totient_prime_power Suc_diff_Suc flip: power_Suc) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
279 |
also have "(p ^ (k - 1) * t + 1) ^ p = (\<Sum>i\<le>p. (p choose i) * t ^ i * p ^ (i * (k - 1)))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
280 |
by (subst binomial) (simp_all add: mult_ac power_mult_distrib power_mult [symmetric]) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
281 |
finally have "[g ^ totient (p ^ k) = (\<Sum>i\<le>p. (p choose i) * t ^ i * p ^ (i * (k - 1)))] |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
282 |
(mod (p ^ Suc k))" (is "[_ = ?rhs] (mod _)") by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
283 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
284 |
also have "[?rhs = (\<Sum>i\<le>p. if i \<le> 1 then (p choose i) * t ^ i * p ^ (i * (k - 1)) else 0)] |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
285 |
(mod (p ^ Suc k))" (is "[sum ?f _ = sum ?g _] (mod _)") |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
286 |
proof (intro cong_sum) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
287 |
fix i assume i: "i \<in> {..p}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
288 |
consider "i \<le> 1" | "i = 2" | "i > 2" by force |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
289 |
thus "[?f i = ?g i] (mod (p ^ Suc k))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
290 |
proof cases |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
291 |
assume i: "i > 2" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
292 |
have "Suc k \<le> 3 * (k - 1)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
293 |
using \<open>k \<ge> 2\<close> by (simp add: algebra_simps) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
294 |
also have "3 * (k - 1) \<le> i * (k - 1)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
295 |
using i by (intro mult_right_mono) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
296 |
finally have "p ^ Suc k dvd ?f i" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
297 |
by (intro dvd_mult le_imp_power_dvd) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
298 |
thus "[?f i = ?g i] (mod (p ^ Suc k))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
299 |
by (simp add: cong_0_iff) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
300 |
next |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
301 |
assume [simp]: "i = 2" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
302 |
have "?f i = p * (p - 1) div 2 * t\<^sup>2 * p ^ (2 * (k - 1))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
303 |
using choose_two[of p] by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
304 |
also have "p * (p - 1) div 2 = (p - 1) div 2 * p" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
305 |
using \<open>odd p\<close> by (auto elim!: oddE) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
306 |
also have "\<dots> * t\<^sup>2 * p ^ (2 * (k - 1)) = (p - 1) div 2 * t\<^sup>2 * (p * p ^ (2 * (k - 1)))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
307 |
by (simp add: algebra_simps) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
308 |
also have "p * p ^ (2 * (k - 1)) = p ^ (2 * k - 1)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
309 |
using \<open>k \<ge> 2\<close> by (cases k) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
310 |
also have "p ^ Suc k dvd (p - 1) div 2 * t\<^sup>2 * p ^ (2 * k - 1)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
311 |
using \<open>k \<ge> 2\<close> by (intro dvd_mult le_imp_power_dvd) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
312 |
finally show "[?f i = ?g i] (mod (p ^ Suc k))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
313 |
by (simp add: cong_0_iff) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
314 |
qed auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
315 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
316 |
also have "(\<Sum>i\<le>p. ?g i) = (\<Sum>i\<le>1. ?f i)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
317 |
using p prime_gt_1_nat[of p] by (intro sum.mono_neutral_cong_right) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
318 |
also have "\<dots> = 1 + t * p ^ k" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
319 |
using choose_two[of p] \<open>k \<ge> 2\<close> by (cases k) simp_all |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
320 |
finally have eq: "[g ^ totient (p ^ k) = 1 + t * p ^ k] (mod p ^ Suc k)" . |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
321 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
322 |
have "[g ^ totient (p ^ k) \<noteq> 1] (mod p ^ Suc k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
323 |
proof |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
324 |
assume "[g ^ totient (p ^ k) = 1] (mod p ^ Suc k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
325 |
hence "[g ^ totient (p ^ k) - g ^ totient (p ^ k) = 1 + t * p ^ k - 1] (mod p ^ Suc k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
326 |
by (intro cong_diff_nat eq) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
327 |
hence "[t * p ^ k = 0] (mod p ^ Suc k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
328 |
by (simp add: cong_sym_eq) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
329 |
hence "p * p ^ k dvd t * p ^ k" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
330 |
by (simp add: cong_0_iff) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
331 |
hence "p dvd t" using \<open>p > 2\<close> by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
332 |
with \<open>\<not>p dvd t\<close> show False by contradiction |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
333 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
334 |
thus ?case by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
335 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
336 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
337 |
text \<open> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
338 |
We can now show that primitive roots modulo \<open>p\<close> with the above condition are |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
339 |
indeed also primitive roots modulo $p^k$. |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
340 |
\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
341 |
proposition residue_primroot_prime_lift_iff: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
342 |
assumes p: "prime p" "odd p" and "residue_primroot p g" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
343 |
shows "(\<forall>k>0. residue_primroot (p ^ k) g) \<longleftrightarrow> [g ^ (p - 1) \<noteq> 1] (mod p\<^sup>2)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
344 |
proof - |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
345 |
from assms have g: "coprime p g" "ord p g = p - 1" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
346 |
by (auto simp: residue_primroot_def totient_prime) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
347 |
show ?thesis |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
348 |
proof |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
349 |
assume "\<forall>k>0. residue_primroot (p ^ k) g" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
350 |
hence "residue_primroot (p\<^sup>2) g" by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
351 |
hence "ord (p\<^sup>2) g = totient (p\<^sup>2)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
352 |
by (simp_all add: residue_primroot_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
353 |
thus "[g ^ (p - 1) \<noteq> 1] (mod p\<^sup>2)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
354 |
using g assms prime_gt_1_nat[of p] |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
355 |
by (auto simp: ord_divides' totient_prime_power) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
356 |
next |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
357 |
assume g': "[g ^ (p - 1) \<noteq> 1] (mod p\<^sup>2)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
358 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
359 |
have "residue_primroot (p ^ k) g" if "k > 0" for k |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
360 |
proof (cases "k = 1") |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
361 |
case False |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
362 |
with that have k: "k > 1" by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
363 |
from g have coprime: "coprime (p ^ k) g" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
364 |
by (auto simp: totatives_def coprime_commute) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
365 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
366 |
define t where "t = ord (p ^ k) g" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
367 |
have "[g ^ t = 1] (mod (p ^ k))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
368 |
by (simp add: t_def ord_divides') |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
369 |
also have "p ^ k = p * p ^ (k - 1)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
370 |
using k by (cases k) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
371 |
finally have "[g ^ t = 1] (mod p)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
372 |
by (rule cong_modulus_mult_nat) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
373 |
hence "totient p dvd t" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
374 |
using g p by (simp add: ord_divides' totient_prime) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
375 |
then obtain q where t: "t = totient p * q" by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
376 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
377 |
have "t dvd totient (p ^ k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
378 |
using coprime by (simp add: t_def order_divides_totient) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
379 |
with t p k have "q dvd p ^ (k - 1)" using prime_gt_1_nat[of p] |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
380 |
by (auto simp: totient_prime totient_prime_power) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
381 |
then obtain b where b: "b \<le> k - 1" "q = p ^ b" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
382 |
using divides_primepow_nat[of p q "k - 1"] p by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
383 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
384 |
have "b = k - 1" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
385 |
proof (rule ccontr) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
386 |
assume "b \<noteq> k - 1" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
387 |
with b have "b < k - 1" by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
388 |
have "t = p ^ b * (p - 1)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
389 |
using b p by (simp add: t totient_prime) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
390 |
also have "\<dots> dvd p ^ (k - 2) * (p - 1)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
391 |
using \<open>b < k - 1\<close> by (intro mult_dvd_mono le_imp_power_dvd) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
392 |
also have "\<dots> = totient (p ^ (k - 1))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
393 |
using k p by (simp add: totient_prime_power numeral_2_eq_2) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
394 |
finally have "[g ^ totient (p ^ (k - 1)) = 1] (mod (p ^ k))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
395 |
by (simp add: ord_divides' t_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
396 |
with residue_primroot_power_prime_power_neq_1[of k p g] p k assms g' show False |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
397 |
by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
398 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
399 |
hence "t = totient (p ^ k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
400 |
using p k by (simp add: t b totient_prime totient_prime_power) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
401 |
thus "residue_primroot (p ^ k) g" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
402 |
using g one_less_power[of p k] prime_gt_1_nat[of p] p k |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
403 |
by (simp add: residue_primroot_def t_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
404 |
qed (use assms in auto) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
405 |
thus "\<forall>k>0. residue_primroot (p ^ k) g" by blast |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
406 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
407 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
408 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
409 |
text \<open> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
410 |
If \<open>p\<close> is an odd prime, there is always a primitive root \<open>g\<close> modulo \<open>p\<close>, and if \<open>g\<close> does not |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
411 |
fulfil the above assumption required for it to be liftable to $p^k$, we can use $g + p$, which |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
412 |
is also a primitive root modulo \<open>p\<close> and \<^emph>\<open>does\<close> fulfil the assumption. |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
413 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
414 |
This shows that any modulus that is a power of an odd prime has a primitive root. |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
415 |
\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
416 |
theorem residue_primroot_odd_prime_power_exists: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
417 |
assumes p: "prime p" "odd p" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
418 |
obtains g where "\<forall>k>0. residue_primroot (p ^ k) g" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
419 |
proof - |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
420 |
obtain g where g: "residue_primroot p g" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
421 |
using prime_primitive_root_exists[of p] assms prime_gt_1_nat[of p] by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
422 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
423 |
have "\<exists>g. residue_primroot p g \<and> [g ^ (p - 1) \<noteq> 1] (mod p\<^sup>2)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
424 |
proof (cases "[g ^ (p - 1) = 1] (mod p\<^sup>2)") |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
425 |
case True |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
426 |
define g' where "g' = p + g" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
427 |
note g |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
428 |
also have "residue_primroot p g \<longleftrightarrow> residue_primroot p g'" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
429 |
unfolding g'_def by (rule residue_primroot_cong) (auto simp: cong_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
430 |
finally have g': "residue_primroot p g'" . |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
431 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
432 |
have "[g' ^ (p - 1) = (\<Sum>k\<le>p-1. ((p-1) choose k) * g ^ (p - Suc k) * p ^ k)] (mod p\<^sup>2)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
433 |
(is "[_ = ?rhs] (mod _)") by (simp add: g'_def binomial mult_ac) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
434 |
also have "[?rhs = (\<Sum>k\<le>p-1. if k \<le> 1 then ((p-1) choose k) * |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
435 |
g ^ (p - Suc k) * p ^ k else 0)] (mod p\<^sup>2)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
436 |
(is "[sum ?f _ = sum ?g _] (mod _)") |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
437 |
proof (intro cong_sum) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
438 |
fix k assume "k \<in> {..p-1}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
439 |
show "[?f k = ?g k] (mod p\<^sup>2)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
440 |
proof (cases "k \<le> 1") |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
441 |
case False |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
442 |
have "p\<^sup>2 dvd ?f k" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
443 |
using False by (intro dvd_mult le_imp_power_dvd) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
444 |
thus ?thesis using False by (simp add: cong_0_iff) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
445 |
qed auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
446 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
447 |
also have "sum ?g {..p-1} = sum ?f {0, 1}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
448 |
using prime_gt_1_nat[of p] p by (intro sum.mono_neutral_cong_right) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
449 |
also have "\<dots> = g ^ (p - 1) + p * (p - 1) * g ^ (p - 2)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
450 |
using p by (simp add: numeral_2_eq_2) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
451 |
also have "[g ^ (p - 1) + p * (p - 1) * g ^ (p - 2) = 1 + p * (p - 1) * g ^ (p - 2)] (mod p\<^sup>2)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
452 |
by (intro cong_add True) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
453 |
finally have "[g' ^ (p - 1) = 1 + p * (p - 1) * g ^ (p - 2)] (mod p\<^sup>2)" . |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
454 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
455 |
moreover have "[1 + p * (p - 1) * g ^ (p - 2) \<noteq> 1] (mod p\<^sup>2)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
456 |
proof |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
457 |
assume "[1 + p * (p - 1) * g ^ (p - 2) = 1] (mod p\<^sup>2)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
458 |
hence "[1 + p * (p - 1) * g ^ (p - 2) - 1 = 1 - 1] (mod p\<^sup>2)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
459 |
by (intro cong_diff_nat) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
460 |
hence "p * p dvd p * ((p - 1) * g ^ (p - 2))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
461 |
by (auto simp: cong_0_iff power2_eq_square) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
462 |
hence "p dvd (p - 1) * g ^ (p - 2)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
463 |
using p by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
464 |
hence "p dvd g ^ (p - 2)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
465 |
using p dvd_imp_le[of p "p - Suc 0"] prime_gt_1_nat[of p] |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
466 |
by (auto simp: prime_dvd_mult_iff) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
467 |
also have "\<dots> dvd g ^ (p - 1)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
468 |
by (intro le_imp_power_dvd) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
469 |
finally have "[g ^ (p - 1) = 0] (mod p)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
470 |
by (simp add: cong_0_iff) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
471 |
hence "[0 = g ^ (p - 1)] (mod p)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
472 |
by (simp add: cong_sym_eq) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
473 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
474 |
also from \<open>residue_primroot p g\<close> have "[g ^ (p - 1) = 1] (mod p)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
475 |
using p by (auto simp: residue_primroot_def ord_divides' totient_prime) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
476 |
finally have "[0 = 1] (mod p)" . |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
477 |
thus False using prime_gt_1_nat[of p] p by (simp add: cong_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
478 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
479 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
480 |
ultimately have "[g' ^ (p - 1) \<noteq> 1] (mod p\<^sup>2)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
481 |
by (simp add: cong_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
482 |
thus "\<exists>g. residue_primroot p g \<and> [g ^ (p - 1) \<noteq> 1] (mod p\<^sup>2)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
483 |
using g' by blast |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
484 |
qed (use g in auto) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
485 |
thus ?thesis |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
486 |
using residue_primroot_prime_lift_iff[OF assms] that by blast |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
487 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
488 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
489 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
490 |
subsection \<open>Carmichael's function\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
491 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
492 |
text \<open> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
493 |
Carmichael's function $\lambda(n)$ gives the LCM of the orders of all elements in the |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
494 |
residue ring modulo $n$ -- or, equivalently, the maximum order, as we will show later. |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
495 |
Algebraically speaking, it is the exponent of the multiplicative group |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
496 |
$(\mathbb{Z}/n\mathbb{Z})^*$. |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
497 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
498 |
It is not to be confused with Liouville's function, which is also denoted by $\lambda(n)$. |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
499 |
\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
500 |
definition Carmichael where |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
501 |
"Carmichael n = (LCM a \<in> totatives n. ord n a)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
502 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
503 |
lemma Carmichael_0 [simp]: "Carmichael 0 = 1" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
504 |
by (simp add: Carmichael_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
505 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
506 |
lemma Carmichael_1 [simp]: "Carmichael 1 = 1" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
507 |
by (simp add: Carmichael_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
508 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
509 |
lemma Carmichael_Suc_0 [simp]: "Carmichael (Suc 0) = 1" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
510 |
by (simp add: Carmichael_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
511 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
512 |
lemma ord_dvd_Carmichael: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
513 |
assumes "n > 1" "coprime n k" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
514 |
shows "ord n k dvd Carmichael n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
515 |
proof - |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
516 |
have "k mod n \<in> totatives n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
517 |
using assms by (auto simp: totatives_def coprime_commute intro!: Nat.gr0I) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
518 |
hence "ord n (k mod n) dvd Carmichael n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
519 |
by (simp add: Carmichael_def del: ord_mod) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
520 |
thus ?thesis by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
521 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
522 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
523 |
lemma Carmichael_divides: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
524 |
assumes "Carmichael n dvd k" "coprime n a" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
525 |
shows "[a ^ k = 1] (mod n)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
526 |
proof (cases "n < 2 \<or> a = 1") |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
527 |
case False |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
528 |
hence "ord n a dvd Carmichael n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
529 |
using False assms by (intro ord_dvd_Carmichael) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
530 |
also have "\<dots> dvd k" by fact |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
531 |
finally have "ord n a dvd k" . |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
532 |
thus ?thesis using ord_divides by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
533 |
next |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
534 |
case True |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
535 |
then consider "a = 1" | "n = 0" | "n = 1" by force |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
536 |
thus ?thesis using assms by cases auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
537 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
538 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
539 |
lemma Carmichael_dvd_totient: "Carmichael n dvd totient n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
540 |
unfolding Carmichael_def |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
541 |
proof (intro Lcm_least, safe) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
542 |
fix a assume "a \<in> totatives n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
543 |
hence "[a ^ totient n = 1] (mod n)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
544 |
by (intro euler_theorem) (auto simp: totatives_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
545 |
thus "ord n a dvd totient n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
546 |
using ord_divides by blast |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
547 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
548 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
549 |
lemma Carmichael_dvd_mono_coprime: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
550 |
assumes "coprime m n" "m > 1" "n > 1" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
551 |
shows "Carmichael m dvd Carmichael (m * n)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
552 |
unfolding Carmichael_def[of m] |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
553 |
proof (intro Lcm_least, safe) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
554 |
fix x assume x: "x \<in> totatives m" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
555 |
from assms have "totatives n \<noteq> {}" by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
556 |
then obtain y where y: "y \<in> totatives n" by blast |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
557 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
558 |
from binary_chinese_remainder_nat[OF assms(1), of x y] |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
559 |
obtain z where z: "[z = x] (mod m)" "[z = y] (mod n)" by blast |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
560 |
have z': "coprime z n" "coprime z m" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
561 |
by (rule cong_imp_coprime; use x y z in \<open>force simp: totatives_def cong_sym_eq\<close>)+ |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
562 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
563 |
from z have "ord m x = ord m z" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
564 |
by (intro ord_cong) (auto simp: cong_sym_eq) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
565 |
also have "ord m z dvd ord (m * n) z" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
566 |
using assms by (auto simp: ord_modulus_mult_coprime) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
567 |
also from z' assms have "\<dots> dvd Carmichael (m * n)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
568 |
by (intro ord_dvd_Carmichael) (auto simp: coprime_commute intro!:one_less_mult) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
569 |
finally show "ord m x dvd Carmichael (m * n)" . |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
570 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
571 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
572 |
text \<open> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
573 |
$\lambda$ distributes over the product of coprime numbers similarly to $\varphi$, but |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
574 |
with LCM instead of multiplication: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
575 |
\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
576 |
lemma Carmichael_mult_coprime: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
577 |
assumes "coprime m n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
578 |
shows "Carmichael (m * n) = lcm (Carmichael m) (Carmichael n)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
579 |
proof (cases "m \<le> 1 \<or> n \<le> 1") |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
580 |
case True |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
581 |
hence "m = 0 \<or> n = 0 \<or> m = 1 \<or> n = 1" by force |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
582 |
thus ?thesis using assms by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
583 |
next |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
584 |
case False |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
585 |
show ?thesis |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
586 |
proof (rule dvd_antisym) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
587 |
show "Carmichael (m * n) dvd lcm (Carmichael m) (Carmichael n)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
588 |
unfolding Carmichael_def[of "m * n"] |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
589 |
proof (intro Lcm_least, safe) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
590 |
fix x assume x: "x \<in> totatives (m * n)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
591 |
have "ord (m * n) x = lcm (ord m x) (ord n x)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
592 |
using assms x by (subst ord_modulus_mult_coprime) (auto simp: coprime_commute totatives_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
593 |
also have "\<dots> dvd lcm (Carmichael m) (Carmichael n)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
594 |
using False x |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
595 |
by (intro lcm_mono ord_dvd_Carmichael) (auto simp: totatives_def coprime_commute) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
596 |
finally show "ord (m * n) x dvd \<dots>" . |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
597 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
598 |
next |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
599 |
show "lcm (Carmichael m) (Carmichael n) dvd Carmichael (m * n)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
600 |
using Carmichael_dvd_mono_coprime[of m n] |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
601 |
Carmichael_dvd_mono_coprime[of n m] assms False |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
602 |
by (auto intro!: lcm_least simp: coprime_commute mult.commute) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
603 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
604 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
605 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
606 |
lemma Carmichael_pos [simp, intro]: "Carmichael n > 0" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
607 |
by (auto simp: Carmichael_def ord_eq_0 totatives_def coprime_commute intro!: Nat.gr0I) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
608 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
609 |
lemma Carmichael_nonzero [simp]: "Carmichael n \<noteq> 0" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
610 |
by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
611 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
612 |
lemma power_Carmichael_eq_1: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
613 |
assumes "n > 1" "coprime n x" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
614 |
shows "[x ^ Carmichael n = 1] (mod n)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
615 |
using ord_dvd_Carmichael[of n x] assms |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
616 |
by (auto simp: ord_divides') |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
617 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
618 |
lemma Carmichael_2 [simp]: "Carmichael 2 = 1" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
619 |
using Carmichael_dvd_totient[of 2] by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
620 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
621 |
lemma Carmichael_4 [simp]: "Carmichael 4 = 2" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
622 |
proof - |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
623 |
have "Carmichael 4 dvd 2" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
624 |
using Carmichael_dvd_totient[of 4] by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
625 |
hence "Carmichael 4 \<le> 2" by (rule dvd_imp_le) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
626 |
moreover have "Carmichael 4 \<noteq> 1" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
627 |
using power_Carmichael_eq_1[of "4::nat" 3] |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
628 |
unfolding coprime_iff_gcd_eq_1 by (auto simp: gcd_non_0_nat cong_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
629 |
ultimately show ?thesis |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
630 |
using Carmichael_pos[of 4] by linarith |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
631 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
632 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
633 |
lemma residue_primroot_Carmichael: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
634 |
assumes "residue_primroot n g" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
635 |
shows "Carmichael n = totient n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
636 |
proof (cases "n = 1") |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
637 |
case False |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
638 |
show ?thesis |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
639 |
proof (intro dvd_antisym Carmichael_dvd_totient) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
640 |
have "ord n g dvd Carmichael n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
641 |
using assms False by (intro ord_dvd_Carmichael) (auto simp: residue_primroot_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
642 |
thus "totient n dvd Carmichael n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
643 |
using assms by (auto simp: residue_primroot_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
644 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
645 |
qed auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
646 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
647 |
lemma Carmichael_odd_prime_power: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
648 |
assumes "prime p" "odd p" "k > 0" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
649 |
shows "Carmichael (p ^ k) = p ^ (k - 1) * (p - 1)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
650 |
proof - |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
651 |
from assms obtain g where "residue_primroot (p ^ k) g" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
652 |
using residue_primroot_odd_prime_power_exists[of p] assms by metis |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
653 |
hence "Carmichael (p ^ k) = totient (p ^ k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
654 |
by (intro residue_primroot_Carmichael[of "p ^ k" g]) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
655 |
with assms show ?thesis by (simp add: totient_prime_power) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
656 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
657 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
658 |
lemma Carmichael_prime: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
659 |
assumes "prime p" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
660 |
shows "Carmichael p = p - 1" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
661 |
proof (cases "even p") |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
662 |
case True |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
663 |
with assms have "p = 2" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
664 |
using primes_dvd_imp_eq two_is_prime_nat by blast |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
665 |
thus ?thesis by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
666 |
next |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
667 |
case False |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
668 |
with Carmichael_odd_prime_power[of p 1] assms show ?thesis by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
669 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
670 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
671 |
lemma Carmichael_twopow_ge_8: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
672 |
assumes "k \<ge> 3" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
673 |
shows "Carmichael (2 ^ k) = 2 ^ (k - 2)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
674 |
proof (intro dvd_antisym) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
675 |
have "2 ^ (k - 2) = ord (2 ^ k) (3 :: nat)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
676 |
using ord_twopow_3_5[of k 3] assms by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
677 |
also have "\<dots> dvd Carmichael (2 ^ k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
678 |
using assms one_less_power[of "2::nat" k] by (intro ord_dvd_Carmichael) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
679 |
finally show "2 ^ (k - 2) dvd \<dots>" . |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
680 |
next |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
681 |
show "Carmichael (2 ^ k) dvd 2 ^ (k - 2)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
682 |
unfolding Carmichael_def |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
683 |
proof (intro Lcm_least, safe) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
684 |
fix x assume "x \<in> totatives (2 ^ k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
685 |
hence "odd x" by (auto simp: totatives_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
686 |
hence "[x ^ 2 ^ (k - 2) = 1] (mod 2 ^ k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
687 |
using assms ord_twopow_aux[of k x] by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
688 |
thus "ord (2 ^ k) x dvd 2 ^ (k - 2)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
689 |
by (simp add: ord_divides') |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
690 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
691 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
692 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
693 |
lemma Carmichael_twopow: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
694 |
"Carmichael (2 ^ k) = (if k \<le> 2 then 2 ^ (k - 1) else 2 ^ (k - 2))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
695 |
proof - |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
696 |
have "k = 0 \<or> k = 1 \<or> k = 2 \<or> k \<ge> 3" by linarith |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
697 |
thus ?thesis by (auto simp: Carmichael_twopow_ge_8) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
698 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
699 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
700 |
lemma Carmichael_prime_power: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
701 |
assumes "prime p" "k > 0" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
702 |
shows "Carmichael (p ^ k) = |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
703 |
(if p = 2 \<and> k > 2 then 2 ^ (k - 2) else p ^ (k - 1) * (p - 1))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
704 |
proof (cases "p = 2") |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
705 |
case True |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
706 |
thus ?thesis by (simp add: Carmichael_twopow) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
707 |
next |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
708 |
case False |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
709 |
with assms have "odd p" "p > 2" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
710 |
using prime_odd_nat[of p] prime_gt_1_nat[of p] by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
711 |
thus ?thesis |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
712 |
using assms Carmichael_odd_prime_power[of p k] by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
713 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
714 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
715 |
lemma Carmichael_prod_coprime: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
716 |
assumes "finite A" "\<And>i j. i \<in> A \<Longrightarrow> j \<in> A \<Longrightarrow> i \<noteq> j \<Longrightarrow> coprime (f i) (f j)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
717 |
shows "Carmichael (\<Prod>i\<in>A. f i) = (LCM i\<in>A. Carmichael (f i))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
718 |
using assms by (induction A rule: finite_induct) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
719 |
(simp, simp, subst Carmichael_mult_coprime[OF prod_coprime_right], auto) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
720 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
721 |
text \<open> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
722 |
Since $\lambda$ distributes over coprime factors and we know the value of $\lambda(p^k)$ |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
723 |
for prime $p$, we can now give a closed formula for $\lambda(n)$ in terms of the prime |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
724 |
factorisation of $n$: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
725 |
\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
726 |
theorem Carmichael_closed_formula: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
727 |
"Carmichael n = |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
728 |
(LCM p\<in>prime_factors n. let k = multiplicity p n |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
729 |
in if p = 2 \<and> k > 2 then 2 ^ (k - 2) else p ^ (k - 1) * (p - 1))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
730 |
(is "_ = Lcm ?A") |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
731 |
proof (cases "n = 0") |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
732 |
case False |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
733 |
hence "n = (\<Prod>p\<in>prime_factors n. p ^ multiplicity p n)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
734 |
using prime_factorization_nat by blast |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
735 |
also have "Carmichael \<dots> = |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
736 |
(LCM p\<in>prime_factors n. Carmichael (p ^ multiplicity p n))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
737 |
by (subst Carmichael_prod_coprime) (auto simp: in_prime_factors_iff primes_coprime) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
738 |
also have "(\<lambda>p. Carmichael (p ^ multiplicity p n)) ` prime_factors n = ?A" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
739 |
by (intro image_cong) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
740 |
(auto simp: Let_def Carmichael_prime_power prime_factors_multiplicity) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
741 |
finally show ?thesis . |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
742 |
qed auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
743 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
744 |
corollary even_Carmichael: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
745 |
assumes "n > 2" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
746 |
shows "even (Carmichael n)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
747 |
proof (cases "\<exists>k. n = 2 ^ k") |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
748 |
case True |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
749 |
then obtain k where [simp]: "n = 2 ^ k" by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
750 |
from assms have "k \<noteq> 0" "k \<noteq> 1" by (auto intro!: Nat.gr0I) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
751 |
hence "k \<ge> 2" by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
752 |
thus ?thesis by (auto simp: Carmichael_twopow) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
753 |
next |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
754 |
case False |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
755 |
from assms have "n \<noteq> 0" by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
756 |
from False have "\<exists>p\<in>prime_factors n. p \<noteq> 2" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
757 |
using assms Ex_other_prime_factor[of n 2] by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
758 |
from divide_out_primepow_ex[OF \<open>n \<noteq> 0\<close> this] guess p k n' . note p = this |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
759 |
from p have coprime: "coprime (p ^ k) n'" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
760 |
using p prime_imp_coprime by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
761 |
have "odd p" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
762 |
using p primes_dvd_imp_eq[of 2 p] by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
763 |
have "even (Carmichael (p ^ k))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
764 |
using p \<open>odd p\<close> by (auto simp: Carmichael_prime_power) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
765 |
with p coprime show ?thesis |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
766 |
by (auto simp: Carmichael_mult_coprime intro!: dvd_lcmI1) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
767 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
768 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
769 |
lemma eval_Carmichael: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
770 |
assumes "prime_factorization n = A" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
771 |
shows "Carmichael n = (LCM p \<in> set_mset A. |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
772 |
let k = count A p in if p = 2 \<and> k > 2 then 2 ^ (k - 2) else p ^ (k - 1) * (p - 1))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
773 |
unfolding assms [symmetric] Carmichael_closed_formula |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
774 |
by (intro arg_cong[where f = Lcm] image_cong) (auto simp: Let_def count_prime_factorization) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
775 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
776 |
text \<open> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
777 |
Any residue ring always contains a $\lambda$-root, i.\,e.\ an element whose |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
778 |
order is $\lambda(n)$. |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
779 |
\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
780 |
theorem Carmichael_root_exists: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
781 |
assumes "n > (0::nat)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
782 |
obtains g where "g \<in> totatives n" and "ord n g = Carmichael n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
783 |
proof (cases "n = 1") |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
784 |
case True |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
785 |
thus ?thesis by (intro that[of 1]) (auto simp: totatives_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
786 |
next |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
787 |
case False |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
788 |
have primepow: "\<exists>g. coprime (p ^ k) g \<and> ord (p ^ k) g = Carmichael (p ^ k)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
789 |
if pk: "prime p" "k > 0" for p k |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
790 |
proof (cases "p = 2") |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
791 |
case [simp]: True |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
792 |
from \<open>k > 0\<close> consider "k = 1" | "k = 2" | "k \<ge> 3" by force |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
793 |
thus ?thesis |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
794 |
proof cases |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
795 |
assume "k = 1" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
796 |
thus ?thesis by (intro exI[of _ 1]) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
797 |
next |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
798 |
assume [simp]: "k = 2" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
799 |
have "coprime 4 (3::nat)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
800 |
by (auto simp: coprime_iff_gcd_eq_1 gcd_non_0_nat) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
801 |
thus ?thesis by (intro exI[of _ 3]) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
802 |
next |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
803 |
assume k: "k \<ge> 3" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
804 |
have "coprime (2 ^ k :: nat) 3" by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
805 |
thus ?thesis using k |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
806 |
by (intro exI[of _ 3]) (auto simp: ord_twopow_3_5 Carmichael_twopow) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
807 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
808 |
next |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
809 |
case False |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
810 |
hence "odd p" using \<open>prime p\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
811 |
using primes_dvd_imp_eq two_is_prime_nat by blast |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
812 |
then obtain g where "residue_primroot (p ^ k) g" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
813 |
using residue_primroot_odd_prime_power_exists[of p] pk by metis |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
814 |
thus ?thesis using False pk |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
815 |
by (intro exI[of _ g]) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
816 |
(auto simp: Carmichael_prime_power residue_primroot_def totient_prime_power) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
817 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
818 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
819 |
define P where "P = prime_factors n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
820 |
define k where "k = (\<lambda>p. multiplicity p n)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
821 |
have "\<forall>p\<in>P. \<exists>g. coprime (p ^ k p) g \<and> ord (p ^ k p) g = Carmichael (p ^ k p)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
822 |
using primepow by (auto simp: P_def k_def prime_factors_multiplicity) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
823 |
hence "\<exists>g. \<forall>p\<in>P. coprime (p ^ k p) (g p) \<and> ord (p ^ k p) (g p) = Carmichael (p ^ k p)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
824 |
by (subst (asm) bchoice_iff) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
825 |
then obtain g where g: "\<And>p. p \<in> P \<Longrightarrow> coprime (p ^ k p) (g p)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
826 |
"\<And>p. p \<in> P \<Longrightarrow> ord (p ^ k p) (g p) = Carmichael (p ^ k p)" by metis |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
827 |
have "\<exists>x. \<forall>i\<in>P. [x = g i] (mod i ^ k i)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
828 |
by (intro chinese_remainder_nat) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
829 |
(auto simp: P_def k_def in_prime_factors_iff primes_coprime) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
830 |
then obtain x where x: "\<And>p. p \<in> P \<Longrightarrow> [x = g p] (mod p ^ k p)" by metis |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
831 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
832 |
have "n = (\<Prod>p\<in>P. p ^ k p)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
833 |
using assms unfolding P_def k_def by (rule prime_factorization_nat) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
834 |
also have "ord \<dots> x = (LCM p\<in>P. ord (p ^ k p) x)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
835 |
by (intro ord_modulus_prod_coprime) (auto simp: P_def in_prime_factors_iff primes_coprime) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
836 |
also have "(\<lambda>p. ord (p ^ k p) x) ` P = (\<lambda>p. ord (p ^ k p) (g p)) ` P" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
837 |
by (intro image_cong ord_cong x) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
838 |
also have "\<dots> = (\<lambda>p. Carmichael (p ^ k p)) ` P" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
839 |
by (intro image_cong g) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
840 |
also have "Lcm \<dots> = Carmichael (\<Prod>p\<in>P. p ^ k p)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
841 |
by (intro Carmichael_prod_coprime [symmetric]) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
842 |
(auto simp: P_def in_prime_factors_iff primes_coprime) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
843 |
also have "(\<Prod>p\<in>P. p ^ k p) = n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
844 |
using assms unfolding P_def k_def by (rule prime_factorization_nat [symmetric]) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
845 |
finally have "ord n x = Carmichael n" . |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
846 |
moreover from this have "coprime n x" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
847 |
by (cases "coprime n x") (auto split: if_splits) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
848 |
ultimately show ?thesis using assms False |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
849 |
by (intro that[of "x mod n"]) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
850 |
(auto simp: totatives_def coprime_commute coprime_absorb_left intro!: Nat.gr0I) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
851 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
852 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
853 |
text \<open> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
854 |
This also means that the Carmichael number is not only the LCM of the orders |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
855 |
of the elements of the residue ring, but indeed the maximum of the orders. |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
856 |
\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
857 |
lemma Carmichael_altdef: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
858 |
"Carmichael n = (if n = 0 then 1 else Max (ord n ` totatives n))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
859 |
proof (cases "n = 0") |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
860 |
case False |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
861 |
have "Carmichael n = Max (ord n ` totatives n)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
862 |
proof (intro antisym Max.boundedI Max.coboundedI) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
863 |
fix k assume k: "k \<in> ord n ` totatives n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
864 |
thus "k \<le> Carmichael n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
865 |
proof (cases "n = 1") |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
866 |
case False |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
867 |
with \<open>n \<noteq> 0\<close> have "n > 1" by linarith |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
868 |
thus ?thesis using k \<open>n \<noteq> 0\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
869 |
by (intro dvd_imp_le) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
870 |
(auto intro!: ord_dvd_Carmichael simp: totatives_def coprime_commute) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
871 |
qed auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
872 |
next |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
873 |
obtain g where "g \<in> totatives n" and "ord n g = Carmichael n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
874 |
using Carmichael_root_exists[of n] \<open>n \<noteq> 0\<close> by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
875 |
thus "Carmichael n \<in> ord n ` totatives n" by force |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
876 |
qed (use \<open>n \<noteq> 0\<close> in auto) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
877 |
thus ?thesis using False by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
878 |
qed auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
879 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
880 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
881 |
subsection \<open>Existence of primitive roots for general moduli\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
882 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
883 |
text \<open> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
884 |
We now related Carmichael's function to the existence of primitive roots and, in the end, |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
885 |
use this to show precisely which moduli have primitive roots and which do not. |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
886 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
887 |
The first criterion for the existence of a primitive root is this: A primitive root modulo $n$ |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
888 |
exists iff $\lambda(n) = \varphi(n)$. |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
889 |
\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
890 |
lemma Carmichael_eq_totient_imp_primroot: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
891 |
assumes "n > 0" and "Carmichael n = totient n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
892 |
shows "\<exists>g. residue_primroot n g" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
893 |
proof - |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
894 |
from \<open>n > 0\<close> obtain g where "g \<in> totatives n" and "ord n g = Carmichael n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
895 |
using Carmichael_root_exists[of n] by metis |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
896 |
with assms show ?thesis by (auto simp: residue_primroot_def totatives_def coprime_commute) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
897 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
898 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
899 |
theorem residue_primroot_iff_Carmichael: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
900 |
"(\<exists>g. residue_primroot n g) \<longleftrightarrow> Carmichael n = totient n \<and> n > 0" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
901 |
proof safe |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
902 |
fix g assume g: "residue_primroot n g" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
903 |
thus "n > 0" by (auto simp: residue_primroot_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
904 |
have "coprime n g" by (rule ccontr) (use g in \<open>auto simp: residue_primroot_def\<close>) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
905 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
906 |
show "Carmichael n = totient n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
907 |
proof (cases "n = 1") |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
908 |
case False |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
909 |
with \<open>n > 0\<close> have "n > 1" by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
910 |
with \<open>coprime n g\<close> have "ord n g dvd Carmichael n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
911 |
by (intro ord_dvd_Carmichael) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
912 |
thus ?thesis using g by (intro dvd_antisym Carmichael_dvd_totient) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
913 |
(auto simp: residue_primroot_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
914 |
qed auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
915 |
qed (use Carmichael_eq_totient_imp_primroot[of n] in auto) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
916 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
917 |
text \<open> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
918 |
Any primitive root modulo $mn$ for coprime $m$, $n$ is also a primitive root modulo $m$ and $n$. |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
919 |
The converse does not hold in general. |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
920 |
\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
921 |
lemma residue_primroot_modulus_mult_coprimeD: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
922 |
assumes "coprime m n" and "residue_primroot (m * n) g" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
923 |
shows "residue_primroot m g" "residue_primroot n g" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
924 |
proof - |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
925 |
have *: "m > 0" "n > 0" "coprime m g" "coprime n g" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
926 |
"lcm (ord m g) (ord n g) = totient m * totient n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
927 |
using assms |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
928 |
by (auto simp: residue_primroot_def ord_modulus_mult_coprime totient_mult_coprime) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
929 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
930 |
define a b where "a = totient m div ord m g" and "b = totient n div ord n g" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
931 |
have ab: "totient m = ord m g * a" "totient n = ord n g * b" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
932 |
using * by (auto simp: a_def b_def order_divides_totient) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
933 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
934 |
have "a = 1" "b = 1" "coprime (ord m g) (ord n g)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
935 |
unfolding coprime_iff_gcd_eq_1 using * by (auto simp: ab lcm_nat_def dvd_div_eq_mult ord_eq_0) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
936 |
with ab and * show "residue_primroot m g" "residue_primroot n g" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
937 |
by (auto simp: residue_primroot_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
938 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
939 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
940 |
text \<open> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
941 |
If a primitive root modulo $mn$ exists for coprime $m, n$, then $\lambda(m)$ and $\lambda(n)$ |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
942 |
must also be coprime. This is helpful in establishing that there are no primitive roots |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
943 |
modulo $mn$ by showing e.\,g.\ that $\lambda(m)$ and $\lambda(n)$ are both even. |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
944 |
\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
945 |
lemma residue_primroot_modulus_mult_coprime_imp_Carmichael_coprime: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
946 |
assumes "coprime m n" and "residue_primroot (m * n) g" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
947 |
shows "coprime (Carmichael m) (Carmichael n)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
948 |
proof - |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
949 |
from residue_primroot_modulus_mult_coprimeD[OF assms] |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
950 |
have *: "residue_primroot m g" "residue_primroot n g" by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
951 |
hence [simp]: "Carmichael m = totient m" "Carmichael n = totient n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
952 |
by (simp_all add: residue_primroot_Carmichael) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
953 |
from * have mn: "m > 0" "n > 0" by (auto simp: residue_primroot_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
954 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
955 |
from assms have "Carmichael (m * n) = totient (m * n) \<and> n > 0" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
956 |
using residue_primroot_iff_Carmichael[of "m * n"] by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
957 |
with assms have "lcm (totient m) (totient n) = totient m * totient n" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
958 |
by (simp add: Carmichael_mult_coprime totient_mult_coprime) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
959 |
thus ?thesis unfolding coprime_iff_gcd_eq_1 using mn |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
960 |
by (simp add: lcm_nat_def dvd_div_eq_mult) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
961 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
962 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
963 |
text \<open> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
964 |
The following moduli are precisely those that have primitive roots. |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
965 |
\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
966 |
definition cyclic_moduli :: "nat set" where |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
967 |
"cyclic_moduli = {1, 2, 4} \<union> {p ^ k |p k. prime p \<and> odd p \<and> k > 0} \<union> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
968 |
{2 * p ^ k |p k. prime p \<and> odd p \<and> k > 0}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
969 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
970 |
theorem residue_primroot_iff_in_cyclic_moduli: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
971 |
"(\<exists>g. residue_primroot m g) \<longleftrightarrow> m \<in> cyclic_moduli" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
972 |
proof - |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
973 |
have "(\<exists>g. residue_primroot m g)" if "m \<in> cyclic_moduli" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
974 |
using that unfolding cyclic_moduli_def |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
975 |
by (intro Carmichael_eq_totient_imp_primroot) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
976 |
(auto dest: prime_gt_0_nat simp: Carmichael_prime_power totient_prime_power |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
977 |
Carmichael_mult_coprime totient_mult_coprime) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
978 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
979 |
moreover have "\<not>(\<exists>g. residue_primroot m g)" if "m \<notin> cyclic_moduli" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
980 |
proof (cases "m = 0") |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
981 |
case False |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
982 |
with that have [simp]: "m > 0" "m \<noteq> 1" by (auto simp: cyclic_moduli_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
983 |
show ?thesis |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
984 |
proof (cases "\<exists>k. m = 2 ^ k") |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
985 |
case True |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
986 |
then obtain k where [simp]: "m = 2 ^ k" by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
987 |
with that have "k \<noteq> 0 \<and> k \<noteq> 1 \<and> k \<noteq> 2" by (auto simp: cyclic_moduli_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
988 |
hence "k \<ge> 3" by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
989 |
thus ?thesis by (subst residue_primroot_iff_Carmichael) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
990 |
(auto simp: Carmichael_twopow totient_prime_power) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
991 |
next |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
992 |
case False |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
993 |
hence "\<exists>p\<in>prime_factors m. p \<noteq> 2" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
994 |
using Ex_other_prime_factor[of m 2] by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
995 |
from divide_out_primepow_ex[OF \<open>m \<noteq> 0\<close> this] |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
996 |
obtain p k m' where p: "p \<noteq> 2" "prime p" "p dvd m" "\<not>p dvd m'" "0 < k" "m = p ^ k * m'" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
997 |
by metis |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
998 |
have "odd p" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
999 |
using p primes_dvd_imp_eq[of 2 p] by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1000 |
from p have coprime: "coprime (p ^ k) m'" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1001 |
using p prime_imp_coprime by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1002 |
have "m \<in> cyclic_moduli" if "m' = 1" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1003 |
using that p \<open>odd p\<close> by (auto simp: cyclic_moduli_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1004 |
moreover have "m \<in> cyclic_moduli" if "m' = 2" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1005 |
proof - |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1006 |
have "m = 2 * p ^ k" using p that by (simp add: mult.commute) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1007 |
thus "m \<in> cyclic_moduli" using p \<open>odd p\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1008 |
unfolding cyclic_moduli_def by blast |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1009 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1010 |
moreover have "m' \<noteq> 0" using p by (intro notI) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1011 |
ultimately have "m' \<noteq> 0 \<and>m' \<noteq> 1 \<and> m' \<noteq> 2" using that by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1012 |
hence "m' > 2" by linarith |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1013 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1014 |
show ?thesis |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1015 |
proof |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1016 |
assume "\<exists>g. residue_primroot m g" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1017 |
with coprime p have coprime': "coprime (p - 1) (Carmichael m')" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1018 |
using residue_primroot_modulus_mult_coprime_imp_Carmichael_coprime[OF coprime] |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1019 |
by (auto simp: Carmichael_prime_power) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1020 |
moreover have "even (p - 1)" "even (Carmichael m')" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1021 |
using \<open>m' > 2\<close> \<open>odd p\<close> by (auto intro!: even_Carmichael) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1022 |
ultimately show False by force |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1023 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1024 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1025 |
qed auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1026 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1027 |
ultimately show ?thesis by metis |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1028 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1029 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1030 |
lemma residue_primroot_is_generator: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1031 |
assumes "m > 1" and "residue_primroot m g" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1032 |
shows "bij_betw (\<lambda>i. g ^ i mod m) {..<totient m} (totatives m)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1033 |
unfolding bij_betw_def |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1034 |
proof |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1035 |
from assms have [simp]: "ord m g = totient m" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1036 |
by (simp add: residue_primroot_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1037 |
from assms have "coprime m g" by (simp add: residue_primroot_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1038 |
hence "inj_on (\<lambda>i. g ^ i mod m) {..<ord m g}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1039 |
by (intro inj_power_mod) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1040 |
thus inj: "inj_on (\<lambda>i. g ^ i mod m) {..<totient m}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1041 |
by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1042 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1043 |
show "(\<lambda>i. g ^ i mod m) ` {..<totient m} = totatives m" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1044 |
proof (rule card_subset_eq) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1045 |
show "card ((\<lambda>i. g ^ i mod m) ` {..<totient m}) = card (totatives m)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1046 |
using inj by (subst card_image) (auto simp: totient_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1047 |
show "(\<lambda>i. g ^ i mod m) ` {..<totient m} \<subseteq> totatives m" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1048 |
using \<open>m > 1\<close> \<open>coprime m g\<close> power_in_totatives[of m g] by blast |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1049 |
qed auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1050 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1051 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1052 |
text \<open> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1053 |
Given one primitive root \<open>g\<close>, all the primitive roots are powers $g^i$ for |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1054 |
$1\leq i \leq \varphi(n)$ with $\text{gcd}(i, \varphi(n)) = 1$. |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1055 |
\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1056 |
theorem residue_primroot_bij_betw_primroots: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1057 |
assumes "m > 1" and "residue_primroot m g" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1058 |
shows "bij_betw (\<lambda>i. g ^ i mod m) (totatives (totient m)) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1059 |
{g\<in>totatives m. residue_primroot m g}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1060 |
proof (cases "m = 2") |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1061 |
case [simp]: True |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1062 |
have [simp]: "totatives 2 = {1}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1063 |
by (auto simp: totatives_def elim!: oddE) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1064 |
from assms have "odd g" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1065 |
by (auto simp: residue_primroot_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1066 |
hence pow_eq: "(\<lambda>i. g ^ i mod m) = (\<lambda>_. 1)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1067 |
by (auto simp: fun_eq_iff mod_2_eq_odd) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1068 |
have "{g \<in> totatives m. residue_primroot m g} = {1}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1069 |
by (auto simp: residue_primroot_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1070 |
thus ?thesis using pow_eq by (auto simp: bij_betw_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1071 |
next |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1072 |
case False |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1073 |
thus ?thesis unfolding bij_betw_def |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1074 |
proof safe |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1075 |
from assms False have "m > 2" by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1076 |
from assms \<open>m > 2\<close> have "totient m > 1" by (intro totient_gt_1) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1077 |
from assms have [simp]: "ord m g = totient m" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1078 |
by (simp add: residue_primroot_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1079 |
from assms have "coprime m g" by (simp add: residue_primroot_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1080 |
hence "inj_on (\<lambda>i. g ^ i mod m) {..<ord m g}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1081 |
by (intro inj_power_mod) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1082 |
thus "inj_on (\<lambda>i. g ^ i mod m) (totatives (totient m))" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1083 |
by (rule inj_on_subset) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1084 |
(use assms \<open>totient m > 1\<close> in \<open>auto simp: totatives_less residue_primroot_def\<close>) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1085 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1086 |
{ |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1087 |
fix i assume i: "i \<in> totatives (totient m)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1088 |
from \<open>coprime m g\<close> and \<open>m > 2\<close> show "g ^ i mod m \<in> totatives m" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1089 |
by (intro power_in_totatives) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1090 |
show "residue_primroot m (g ^ i mod m)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1091 |
using i \<open>m > 2\<close> \<open>coprime m g\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1092 |
by (auto simp: residue_primroot_def coprime_commute ord_power totatives_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1093 |
} |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1094 |
{ |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1095 |
fix x assume x: "x \<in> totatives m" "residue_primroot m x" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1096 |
then obtain i where i: "i < totient m" "x = (g ^ i mod m)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1097 |
using assms residue_primroot_is_generator[of m g] by (auto simp: bij_betw_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1098 |
from i x \<open>m > 2\<close> have "i > 0" by (intro Nat.gr0I) (auto simp: residue_primroot_1_iff) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1099 |
have "totient m div gcd i (totient m) = totient m" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1100 |
using x i \<open>coprime m g\<close> by (auto simp add: residue_primroot_def ord_power) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1101 |
hence "coprime i (totient m)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1102 |
unfolding coprime_iff_gcd_eq_1 using assms by (subst (asm) dvd_div_eq_mult) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1103 |
with i \<open>i > 0\<close> have "i \<in> totatives (totient m)" by (auto simp: totatives_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1104 |
thus "x \<in> (\<lambda>i. g ^ i mod m) ` totatives (totient m)" using i by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1105 |
} |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1106 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1107 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1108 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1109 |
text \<open> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1110 |
It follows from the above statement that any residue ring modulo \<open>n\<close> that \<^emph>\<open>has\<close> primitive roots |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1111 |
has exactly $\varphi(\varphi(n))$ of them. |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1112 |
\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1113 |
corollary card_residue_primroots: |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1114 |
assumes "\<exists>g. residue_primroot m g" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1115 |
shows "card {g\<in>totatives m. residue_primroot m g} = totient (totient m)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1116 |
proof (cases "m = 1") |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1117 |
case [simp]: True |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1118 |
have "{g \<in> totatives m. residue_primroot m g} = {1}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1119 |
by (auto simp: residue_primroot_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1120 |
thus ?thesis by simp |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1121 |
next |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1122 |
case False |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1123 |
from assms obtain g where g: "residue_primroot m g" by auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1124 |
hence "m \<noteq> 0" by (intro notI) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1125 |
with \<open>m \<noteq> 1\<close> have "m > 1" by linarith |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1126 |
from this g have "bij_betw (\<lambda>i. g ^ i mod m) (totatives (totient m)) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1127 |
{g\<in>totatives m. residue_primroot m g}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1128 |
by (rule residue_primroot_bij_betw_primroots) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1129 |
hence "card (totatives (totient m)) = card {g\<in>totatives m. residue_primroot m g}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1130 |
by (rule bij_betw_same_card) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1131 |
thus ?thesis by (simp add: totient_def) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1132 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1133 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1134 |
corollary card_residue_primroots': |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1135 |
"card {g\<in>totatives m. residue_primroot m g} = |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1136 |
(if m \<in> cyclic_moduli then totient (totient m) else 0)" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1137 |
by (simp add: residue_primroot_iff_in_cyclic_moduli [symmetric] card_residue_primroots) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1138 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1139 |
text \<open> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1140 |
As an example, we evaluate $\lambda(122200)$ using the prime factorisation. |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1141 |
\<close> |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1142 |
lemma "Carmichael 122200 = 1380" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1143 |
proof - |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1144 |
have "prime_factorization (2^3 * 5^2 * 13 * 47) = {#2, 2, 2, 5, 5, 13, 47::nat#}" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1145 |
by (intro prime_factorization_eqI) auto |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1146 |
from eval_Carmichael[OF this] show "Carmichael 122200 = 1380" |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1147 |
by (simp add: lcm_nat_def gcd_non_0_nat) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1148 |
qed |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1149 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1150 |
(* TODO: definition of index ("discrete logarithm") *) |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1151 |
|
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
diff
changeset
|
1152 |
end |