author | wenzelm |
Wed, 13 Nov 2024 20:10:34 +0100 | |
changeset 81438 | 95c9af7483b1 |
parent 80914 | d97fdabd9e2b |
permissions | -rw-r--r-- |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1 |
(* Title: HOL/Algebra/Multiplicative_Group.thy |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
2 |
Author: Simon Wimmer |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
3 |
Author: Lars Noschinski |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
4 |
*) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
5 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
6 |
theory Multiplicative_Group |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
7 |
imports |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
8 |
Complex_Main |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
9 |
Group |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
10 |
Coset |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
11 |
UnivPoly |
68575 | 12 |
Generated_Groups |
72630
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
13 |
Elementary_Groups |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
14 |
begin |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
15 |
|
67226 | 16 |
section \<open>Simplification Rules for Polynomials\<close> |
17 |
text_raw \<open>\label{sec:simp-rules}\<close> |
|
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
18 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
19 |
lemma (in ring_hom_cring) hom_sub[simp]: |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
20 |
assumes "x \<in> carrier R" "y \<in> carrier R" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
21 |
shows "h (x \<ominus> y) = h x \<ominus>\<^bsub>S\<^esub> h y" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
22 |
using assms by (simp add: R.minus_eq S.minus_eq) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
23 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
24 |
context UP_ring begin |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
25 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
26 |
lemma deg_nzero_nzero: |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
27 |
assumes deg_p_nzero: "deg R p \<noteq> 0" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
28 |
shows "p \<noteq> \<zero>\<^bsub>P\<^esub>" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
29 |
using deg_zero deg_p_nzero by auto |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
30 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
31 |
lemma deg_add_eq: |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
32 |
assumes c: "p \<in> carrier P" "q \<in> carrier P" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
33 |
assumes "deg R q \<noteq> deg R p" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
34 |
shows "deg R (p \<oplus>\<^bsub>P\<^esub> q) = max (deg R p) (deg R q)" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
35 |
proof - |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
36 |
let ?m = "max (deg R p) (deg R q)" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
37 |
from assms have "coeff P p ?m = \<zero> \<longleftrightarrow> coeff P q ?m \<noteq> \<zero>" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
38 |
by (metis deg_belowI lcoeff_nonzero[OF deg_nzero_nzero] linear max.absorb_iff2 max.absorb1) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
39 |
then have "coeff P (p \<oplus>\<^bsub>P\<^esub> q) ?m \<noteq> \<zero>" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
40 |
using assms by auto |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
41 |
then have "deg R (p \<oplus>\<^bsub>P\<^esub> q) \<ge> ?m" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
42 |
using assms by (blast intro: deg_belowI) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
43 |
with deg_add[OF c] show ?thesis by arith |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
44 |
qed |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
45 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
46 |
lemma deg_minus_eq: |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
47 |
assumes "p \<in> carrier P" "q \<in> carrier P" "deg R q \<noteq> deg R p" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
48 |
shows "deg R (p \<ominus>\<^bsub>P\<^esub> q) = max (deg R p) (deg R q)" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
49 |
using assms by (simp add: deg_add_eq a_minus_def) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
50 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
51 |
end |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
52 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
53 |
context UP_cring begin |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
54 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
55 |
lemma evalRR_add: |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
56 |
assumes "p \<in> carrier P" "q \<in> carrier P" |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
57 |
assumes x: "x \<in> carrier R" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
58 |
shows "eval R R id x (p \<oplus>\<^bsub>P\<^esub> q) = eval R R id x p \<oplus> eval R R id x q" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
59 |
proof - |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
60 |
interpret UP_pre_univ_prop R R id by unfold_locales simp |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
61 |
interpret ring_hom_cring P R "eval R R id x" by unfold_locales (rule eval_ring_hom[OF x]) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
62 |
show ?thesis using assms by simp |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
63 |
qed |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
64 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
65 |
lemma evalRR_sub: |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
66 |
assumes "p \<in> carrier P" "q \<in> carrier P" |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
67 |
assumes x: "x \<in> carrier R" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
68 |
shows "eval R R id x (p \<ominus>\<^bsub>P\<^esub> q) = eval R R id x p \<ominus> eval R R id x q" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
69 |
proof - |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
70 |
interpret UP_pre_univ_prop R R id by unfold_locales simp |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
71 |
interpret ring_hom_cring P R "eval R R id x" by unfold_locales (rule eval_ring_hom[OF x]) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
72 |
show ?thesis using assms by simp |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
73 |
qed |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
74 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
75 |
lemma evalRR_mult: |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
76 |
assumes "p \<in> carrier P" "q \<in> carrier P" |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
77 |
assumes x: "x \<in> carrier R" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
78 |
shows "eval R R id x (p \<otimes>\<^bsub>P\<^esub> q) = eval R R id x p \<otimes> eval R R id x q" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
79 |
proof - |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
80 |
interpret UP_pre_univ_prop R R id by unfold_locales simp |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
81 |
interpret ring_hom_cring P R "eval R R id x" by unfold_locales (rule eval_ring_hom[OF x]) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
82 |
show ?thesis using assms by simp |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
83 |
qed |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
84 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
85 |
lemma evalRR_monom: |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
86 |
assumes a: "a \<in> carrier R" and x: "x \<in> carrier R" |
67341
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents:
67299
diff
changeset
|
87 |
shows "eval R R id x (monom P a d) = a \<otimes> x [^] d" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
88 |
proof - |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
89 |
interpret UP_pre_univ_prop R R id by unfold_locales simp |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
90 |
show ?thesis using assms by (simp add: eval_monom) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
91 |
qed |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
92 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
93 |
lemma evalRR_one: |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
94 |
assumes x: "x \<in> carrier R" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
95 |
shows "eval R R id x \<one>\<^bsub>P\<^esub> = \<one>" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
96 |
proof - |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
97 |
interpret UP_pre_univ_prop R R id by unfold_locales simp |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
98 |
interpret ring_hom_cring P R "eval R R id x" by unfold_locales (rule eval_ring_hom[OF x]) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
99 |
show ?thesis using assms by simp |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
100 |
qed |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
101 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
102 |
lemma carrier_evalRR: |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
103 |
assumes x: "x \<in> carrier R" and "p \<in> carrier P" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
104 |
shows "eval R R id x p \<in> carrier R" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
105 |
proof - |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
106 |
interpret UP_pre_univ_prop R R id by unfold_locales simp |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
107 |
interpret ring_hom_cring P R "eval R R id x" by unfold_locales (rule eval_ring_hom[OF x]) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
108 |
show ?thesis using assms by simp |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
109 |
qed |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
110 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
111 |
lemmas evalRR_simps = evalRR_add evalRR_sub evalRR_mult evalRR_monom evalRR_one carrier_evalRR |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
112 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
113 |
end |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
114 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
115 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
116 |
|
67226 | 117 |
section \<open>Properties of the Euler \<open>\<phi>\<close>-function\<close> |
118 |
text_raw \<open>\label{sec:euler-phi}\<close> |
|
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
119 |
|
67226 | 120 |
text\<open> |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
121 |
In this section we prove that for every positive natural number the equation |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
122 |
$\sum_{d | n}^n \varphi(d) = n$ holds. |
67226 | 123 |
\<close> |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
124 |
|
68575 | 125 |
lemma dvd_div_ge_1: |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
126 |
fixes a b :: nat |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
127 |
assumes "a \<ge> 1" "b dvd a" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
128 |
shows "a div b \<ge> 1" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
129 |
proof - |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
130 |
from \<open>b dvd a\<close> obtain c where "a = b * c" .. |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
131 |
with \<open>a \<ge> 1\<close> show ?thesis by simp |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
132 |
qed |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
133 |
|
68575 | 134 |
lemma dvd_nat_bounds: |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
135 |
fixes n p :: nat |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
136 |
assumes "p > 0" "n dvd p" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
137 |
shows "n > 0 \<and> n \<le> p" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
138 |
using assms by (simp add: dvd_pos_nat dvd_imp_le) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
139 |
|
69785
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69749
diff
changeset
|
140 |
(* TODO FIXME: This is the "totient" function from HOL-Number_Theory, but since part of |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69749
diff
changeset
|
141 |
HOL-Number_Theory depends on HOL-Algebra.Multiplicative_Group, there would be a cyclic |
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents:
69749
diff
changeset
|
142 |
dependency. *) |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
143 |
definition phi' :: "nat => nat" |
67051 | 144 |
where "phi' m = card {x. 1 \<le> x \<and> x \<le> m \<and> coprime x m}" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
145 |
|
66500 | 146 |
notation (latex output) |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
77061
diff
changeset
|
147 |
phi' (\<open>\<phi> _\<close>) |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
148 |
|
68575 | 149 |
lemma phi'_nonzero: |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
150 |
assumes "m > 0" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
151 |
shows "phi' m > 0" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
152 |
proof - |
67051 | 153 |
have "1 \<in> {x. 1 \<le> x \<and> x \<le> m \<and> coprime x m}" using assms by simp |
154 |
hence "card {x. 1 \<le> x \<and> x \<le> m \<and> coprime x m} > 0" by (auto simp: card_gt_0_iff) |
|
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
155 |
thus ?thesis unfolding phi'_def by simp |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
156 |
qed |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
157 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
158 |
lemma dvd_div_eq_1: |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
159 |
fixes a b c :: nat |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
160 |
assumes "c dvd a" "c dvd b" "a div c = b div c" |
67226 | 161 |
shows "a = b" using assms dvd_mult_div_cancel[OF \<open>c dvd a\<close>] dvd_mult_div_cancel[OF \<open>c dvd b\<close>] |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
162 |
by presburger |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
163 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
164 |
lemma dvd_div_eq_2: |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
165 |
fixes a b c :: nat |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
166 |
assumes "c>0" "a dvd c" "b dvd c" "c div a = c div b" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
167 |
shows "a = b" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
168 |
proof - |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
169 |
have "a > 0" "a \<le> c" using dvd_nat_bounds[OF assms(1-2)] by auto |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
170 |
have "a*(c div a) = c" using assms dvd_mult_div_cancel by fastforce |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
171 |
also have "\<dots> = b*(c div a)" using assms dvd_mult_div_cancel by fastforce |
67226 | 172 |
finally show "a = b" using \<open>c>0\<close> dvd_div_ge_1[OF _ \<open>a dvd c\<close>] by fastforce |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
173 |
qed |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
174 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
175 |
lemma div_mult_mono: |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
176 |
fixes a b c :: nat |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
177 |
assumes "a > 0" "a\<le>d" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
178 |
shows "a * b div d \<le> b" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
179 |
proof - |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
180 |
have "a*b div d \<le> b*a div a" using assms div_le_mono2 mult.commute[of a b] by presburger |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
181 |
thus ?thesis using assms by force |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
182 |
qed |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
183 |
|
67226 | 184 |
text\<open> |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
185 |
We arrive at the main result of this section: |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
186 |
For every positive natural number the equation $\sum_{d | n}^n \varphi(d) = n$ holds. |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
187 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
188 |
The outline of the proof for this lemma is as follows: |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
189 |
We count the $n$ fractions $1/n$, $\ldots$, $(n-1)/n$, $n/n$. |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
190 |
We analyze the reduced form $a/d = m/n$ for any of those fractions. |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
191 |
We want to know how many fractions $m/n$ have the reduced form denominator $d$. |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
192 |
The condition $1 \leq m \leq n$ is equivalent to the condition $1 \leq a \leq d$. |
69597 | 193 |
Therefore we want to know how many $a$ with $1 \leq a \leq d$ exist, s.t. \<^term>\<open>gcd a d = 1\<close>. |
194 |
This number is exactly \<^term>\<open>phi' d\<close>. |
|
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
195 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
196 |
Finally, by counting the fractions $m/n$ according to their reduced form denominator, |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
197 |
we get: @{term [display] "(\<Sum>d | d dvd n . phi' d) = n"}. |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
198 |
To formalize this proof in Isabelle, we analyze for an arbitrary divisor $d$ of $n$ |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
199 |
\begin{itemize} |
69597 | 200 |
\item the set of reduced form numerators \<^term>\<open>{a. (1::nat) \<le> a \<and> a \<le> d \<and> coprime a d}\<close> |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
201 |
\item the set of numerators $m$, for which $m/n$ has the reduced form denominator $d$, |
69597 | 202 |
i.e. the set \<^term>\<open>{m \<in> {1::nat .. n}. n div gcd m n = d}\<close> |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
203 |
\end{itemize} |
69597 | 204 |
We show that \<^term>\<open>\<lambda>a. a*n div d\<close> with the inverse \<^term>\<open>\<lambda>a. a div gcd a n\<close> is |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
205 |
a bijection between theses sets, thus yielding the equality |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
206 |
@{term [display] "phi' d = card {m \<in> {1 .. n}. n div gcd m n = d}"} |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
207 |
This gives us |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
208 |
@{term [display] "(\<Sum>d | d dvd n . phi' d) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
209 |
= card (\<Union>d \<in> {d. d dvd n}. {m \<in> {1 .. n}. n div gcd m n = d})"} |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
210 |
and by showing |
69597 | 211 |
\<^term>\<open>(\<Union>d \<in> {d. d dvd n}. {m \<in> {1::nat .. n}. n div gcd m n = d}) \<supseteq> {1 .. n}\<close> |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
212 |
(this is our counting argument) the thesis follows. |
67226 | 213 |
\<close> |
68575 | 214 |
lemma sum_phi'_factors: |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
215 |
fixes n :: nat |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
216 |
assumes "n > 0" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
217 |
shows "(\<Sum>d | d dvd n. phi' d) = n" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
218 |
proof - |
81438 | 219 |
have "card {a. 1 \<le> a \<and> a \<le> d \<and> coprime a d} = card {m \<in> {1 .. n}. n div gcd m n = d}" |
220 |
(is "card ?RF = card ?F") |
|
221 |
if "d dvd n" for d |
|
222 |
proof (rule card_bij_eq) |
|
223 |
from that obtain q where q: "n = d * q" .. |
|
224 |
have "a = b" if "a * n div d = b * n div d" for a b |
|
225 |
proof - |
|
226 |
from that have "a * (n div d) = b * (n div d)" |
|
227 |
using dvd_div_mult[OF \<open>d dvd n\<close>] by (fastforce simp add: mult.commute) |
|
228 |
then show ?thesis using dvd_div_ge_1[OF _ \<open>d dvd n\<close>] \<open>n>0\<close> |
|
229 |
by (simp add: mult.commute nat_mult_eq_cancel1) |
|
230 |
qed |
|
231 |
thus "inj_on (\<lambda>a. a*n div d) ?RF" unfolding inj_on_def by blast |
|
232 |
have "a * n div d \<in> ?F" if a: "a\<in>?RF" for a |
|
233 |
proof - |
|
234 |
from that have "a * (n div d) \<ge> 1" using \<open>n>0\<close> dvd_div_ge_1[OF _ \<open>d dvd n\<close>] by simp |
|
235 |
hence ge_1: "a * n div d \<ge> 1" by (simp add: \<open>d dvd n\<close> div_mult_swap) |
|
236 |
have le_n: "a * n div d \<le> n" using div_mult_mono a by simp |
|
237 |
have "gcd (a * n div d) n = n div d * gcd a d" |
|
238 |
by (simp add: gcd_mult_distrib_nat q ac_simps) |
|
239 |
hence "n div gcd (a * n div d) n = d*n div (d*(n div d))" using a by simp |
|
240 |
then show ?thesis |
|
241 |
using ge_1 le_n by (fastforce simp add: \<open>d dvd n\<close>) |
|
242 |
qed |
|
243 |
thus "(\<lambda>a. a*n div d) ` ?RF \<subseteq> ?F" by blast |
|
244 |
have "m = l" if A: "m \<in> ?F" "l \<in> ?F" "m div gcd m n = l div gcd l n" for m l |
|
245 |
proof - |
|
246 |
from that have "gcd m n = gcd l n" using dvd_div_eq_2[OF assms] by fastforce |
|
247 |
then show ?thesis using dvd_div_eq_1[of "gcd m n" m l] A(3) by fastforce |
|
248 |
qed |
|
249 |
thus "inj_on (\<lambda>a. a div gcd a n) ?F" unfolding inj_on_def by blast |
|
250 |
have "m div gcd m n \<in> ?RF" if "m \<in> ?F" for m |
|
251 |
using that dvd_div_ge_1 by (fastforce simp add: div_le_mono div_gcd_coprime) |
|
252 |
thus "(\<lambda>a. a div gcd a n) ` ?F \<subseteq> ?RF" by blast |
|
253 |
qed force+ |
|
254 |
hence phi'_eq: "\<And>d. d dvd n \<Longrightarrow> phi' d = card {m \<in> {1 .. n}. n div gcd m n = d}" |
|
255 |
unfolding phi'_def by presburger |
|
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
256 |
have fin: "finite {d. d dvd n}" using dvd_nat_bounds[OF \<open>n>0\<close>] by force |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
257 |
have "(\<Sum>d | d dvd n. phi' d) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
258 |
= card (\<Union>d \<in> {d. d dvd n}. {m \<in> {1 .. n}. n div gcd m n = d})" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
259 |
using card_UN_disjoint[OF fin, of "(\<lambda>d. {m \<in> {1 .. n}. n div gcd m n = d})"] phi'_eq |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
260 |
by fastforce |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
261 |
also have "(\<Union>d \<in> {d. d dvd n}. {m \<in> {1 .. n}. n div gcd m n = d}) = {1 .. n}" (is "?L = ?R") |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
262 |
proof |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
263 |
show "?L \<supseteq> ?R" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
264 |
proof |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
265 |
fix m assume m: "m \<in> ?R" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
266 |
thus "m \<in> ?L" using dvd_triv_right[of "n div gcd m n" "gcd m n"] |
67051 | 267 |
by simp |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
268 |
qed |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
269 |
qed fastforce |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
270 |
finally show ?thesis by force |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
271 |
qed |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
272 |
|
72630
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
273 |
|
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
274 |
|
67226 | 275 |
section \<open>Order of an Element of a Group\<close> |
276 |
text_raw \<open>\label{sec:order-elem}\<close> |
|
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
277 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
278 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
279 |
context group begin |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
280 |
|
70030
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
281 |
definition (in group) ord :: "'a \<Rightarrow> nat" where |
70131 | 282 |
"ord x \<equiv> (@d. \<forall>n::nat. x [^] n = \<one> \<longleftrightarrow> d dvd n)" |
70030
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
283 |
|
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
284 |
lemma (in group) pow_eq_id: |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
285 |
assumes "x \<in> carrier G" |
70131 | 286 |
shows "x [^] n = \<one> \<longleftrightarrow> (ord x) dvd n" |
70030
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
287 |
proof (cases "\<forall>n::nat. pow G x n = one G \<longrightarrow> n = 0") |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
288 |
case True |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
289 |
show ?thesis |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
290 |
unfolding ord_def |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
291 |
by (rule someI2 [where a=0]) (auto simp: True) |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
292 |
next |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
293 |
case False |
70030
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
294 |
define N where "N \<equiv> LEAST n::nat. x [^] n = \<one> \<and> n > 0" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
295 |
have N: "x [^] N = \<one> \<and> N > 0" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
296 |
using False |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
297 |
apply (simp add: N_def) |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
298 |
by (metis (mono_tags, lifting) LeastI) |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
299 |
have eq0: "n = 0" if "x [^] n = \<one>" "n < N" for n |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
300 |
using N_def not_less_Least that by fastforce |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
301 |
show ?thesis |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
302 |
unfolding ord_def |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
303 |
proof (rule someI2 [where a = N], rule allI) |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
304 |
fix n :: "nat" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
305 |
show "(x [^] n = \<one>) \<longleftrightarrow> (N dvd n)" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
306 |
proof (cases "n = 0") |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
307 |
case False |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
308 |
show ?thesis |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
309 |
unfolding dvd_def |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
310 |
proof safe |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
311 |
assume 1: "x [^] n = \<one>" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
312 |
have "x [^] n = x [^] (n mod N + N * (n div N))" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
313 |
by simp |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
314 |
also have "\<dots> = x [^] (n mod N) \<otimes> x [^] (N * (n div N))" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
315 |
by (simp add: assms nat_pow_mult) |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
316 |
also have "\<dots> = x [^] (n mod N)" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
317 |
by (metis N assms l_cancel_one nat_pow_closed nat_pow_one nat_pow_pow) |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
318 |
finally have "x [^] (n mod N) = \<one>" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
319 |
by (simp add: "1") |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
320 |
then have "n mod N = 0" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
321 |
using N eq0 mod_less_divisor by blast |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
322 |
then show "\<exists>k. n = N * k" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
323 |
by blast |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
324 |
next |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
325 |
fix k :: "nat" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
326 |
assume "n = N * k" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
327 |
with N show "x [^] (N * k) = \<one>" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
328 |
by (metis assms nat_pow_one nat_pow_pow) |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
329 |
qed |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
330 |
qed simp |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
331 |
qed blast |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
332 |
qed |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
333 |
|
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
334 |
lemma (in group) pow_ord_eq_1 [simp]: |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
335 |
"x \<in> carrier G \<Longrightarrow> x [^] ord x = \<one>" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
336 |
by (simp add: pow_eq_id) |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
337 |
|
70030
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
338 |
lemma (in group) int_pow_eq_id: |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
339 |
assumes "x \<in> carrier G" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
340 |
shows "(pow G x i = one G \<longleftrightarrow> int (ord x) dvd i)" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
341 |
proof (cases i rule: int_cases2) |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
342 |
case (nonneg n) |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
343 |
then show ?thesis |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
344 |
by (simp add: int_pow_int pow_eq_id assms) |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
345 |
next |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
346 |
case (nonpos n) |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
347 |
then have "x [^] i = inv (x [^] n)" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
348 |
by (simp add: assms int_pow_int int_pow_neg) |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
349 |
then show ?thesis |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
350 |
by (simp add: assms pow_eq_id nonpos) |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
351 |
qed |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
352 |
|
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
353 |
lemma (in group) int_pow_eq: |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
354 |
"x \<in> carrier G \<Longrightarrow> (x [^] m = x [^] n) \<longleftrightarrow> int (ord x) dvd (n - m)" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
355 |
apply (simp flip: int_pow_eq_id) |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
356 |
by (metis int_pow_closed int_pow_diff inv_closed r_inv right_cancel) |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
357 |
|
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
358 |
lemma (in group) ord_eq_0: |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
359 |
"x \<in> carrier G \<Longrightarrow> (ord x = 0 \<longleftrightarrow> (\<forall>n::nat. n \<noteq> 0 \<longrightarrow> x [^] n \<noteq> \<one>))" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
360 |
by (auto simp: pow_eq_id) |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
361 |
|
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
362 |
lemma (in group) ord_unique: |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
363 |
"x \<in> carrier G \<Longrightarrow> ord x = d \<longleftrightarrow> (\<forall>n. pow G x n = one G \<longleftrightarrow> d dvd n)" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
364 |
by (meson dvd_antisym dvd_refl pow_eq_id) |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
365 |
|
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
366 |
lemma (in group) ord_eq_1: |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
367 |
"x \<in> carrier G \<Longrightarrow> (ord x = 1 \<longleftrightarrow> x = \<one>)" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
368 |
by (metis pow_eq_id nat_dvd_1_iff_1 nat_pow_eone) |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
369 |
|
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
370 |
lemma (in group) ord_id [simp]: |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
371 |
"ord (one G) = 1" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
372 |
using ord_eq_1 by blast |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
373 |
|
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
374 |
lemma (in group) ord_inv [simp]: |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
375 |
"x \<in> carrier G |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
376 |
\<Longrightarrow> ord (m_inv G x) = ord x" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
377 |
by (simp add: ord_unique pow_eq_id nat_pow_inv) |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
378 |
|
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
379 |
lemma (in group) ord_pow: |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
380 |
assumes "x \<in> carrier G" "k dvd ord x" "k \<noteq> 0" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
381 |
shows "ord (pow G x k) = ord x div k" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
382 |
proof - |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
383 |
have "(x [^] k) [^] (ord x div k) = \<one>" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
384 |
using assms by (simp add: nat_pow_pow) |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
385 |
moreover have "ord x dvd k * ord (x [^] k)" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
386 |
by (metis assms(1) pow_ord_eq_1 pow_eq_id nat_pow_closed nat_pow_pow) |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
387 |
ultimately show ?thesis |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
388 |
by (metis assms div_dvd_div dvd_antisym dvd_triv_left pow_eq_id nat_pow_closed nonzero_mult_div_cancel_left) |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
389 |
qed |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
390 |
|
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
391 |
lemma (in group) ord_mul_divides: |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
392 |
assumes eq: "x \<otimes> y = y \<otimes> x" and xy: "x \<in> carrier G" "y \<in> carrier G" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
393 |
shows "ord (x \<otimes> y) dvd (ord x * ord y)" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
394 |
apply (simp add: xy flip: pow_eq_id eq) |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
395 |
by (metis dvd_triv_left dvd_triv_right eq pow_eq_id one_closed pow_mult_distrib r_one xy) |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
396 |
|
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
397 |
lemma (in comm_group) abelian_ord_mul_divides: |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
398 |
"\<lbrakk>x \<in> carrier G; y \<in> carrier G\<rbrakk> |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
399 |
\<Longrightarrow> ord (x \<otimes> y) dvd (ord x * ord y)" |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
400 |
by (simp add: ord_mul_divides m_comm) |
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
401 |
|
68575 | 402 |
lemma ord_inj: |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
403 |
assumes a: "a \<in> carrier G" |
67341
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents:
67299
diff
changeset
|
404 |
shows "inj_on (\<lambda> x . a [^] x) {0 .. ord a - 1}" |
70131 | 405 |
proof - |
406 |
let ?M = "Max (ord ` carrier G)" |
|
407 |
have "finite {d \<in> {..?M}. a [^] d = \<one>}" by auto |
|
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
408 |
|
70131 | 409 |
have *: False if A: "x < y" "x \<in> {0 .. ord a - 1}" "y \<in> {0 .. ord a - 1}" |
410 |
"a [^] x = a [^] y" for x y |
|
411 |
proof - |
|
412 |
have "y - x < ord a" using that by auto |
|
413 |
moreover have "a [^] (y-x) = \<one>" using a A by (simp add: pow_eq_div2) |
|
414 |
ultimately have "min (y - x) (ord a) = ord a" |
|
415 |
using A(1) a pow_eq_id by auto |
|
416 |
with \<open>y - x < ord a\<close> show False by linarith |
|
417 |
qed |
|
418 |
show ?thesis |
|
419 |
unfolding inj_on_def by (metis nat_neq_iff *) |
|
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
420 |
qed |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
421 |
|
68575 | 422 |
lemma ord_inj': |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
423 |
assumes a: "a \<in> carrier G" |
67341
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents:
67299
diff
changeset
|
424 |
shows "inj_on (\<lambda> x . a [^] x) {1 .. ord a}" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
425 |
proof (rule inj_onI, rule ccontr) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
426 |
fix x y :: nat |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
427 |
assume A: "x \<in> {1 .. ord a}" "y \<in> {1 .. ord a}" "a [^] x = a [^] y" "x\<noteq>y" |
81438 | 428 |
then consider "x < ord a" "y < ord a" | "x = ord a" "y < ord a" | "y = ord a" "x < ord a" |
429 |
by force |
|
430 |
then show False |
|
431 |
proof cases |
|
432 |
case 1 |
|
433 |
then show ?thesis using ord_inj[OF assms] A unfolding inj_on_def by fastforce |
|
434 |
next |
|
435 |
case 2 |
|
70030
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
436 |
hence "a [^] y = a [^] (0::nat)" using pow_ord_eq_1 A by (simp add: a) |
67226 | 437 |
hence "y=0" using ord_inj[OF assms] \<open>y < ord a\<close> unfolding inj_on_def by force |
81438 | 438 |
with A show ?thesis by fastforce |
439 |
next |
|
440 |
case 3 |
|
70030
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
441 |
hence "a [^] x = a [^] (0::nat)" using pow_ord_eq_1 A by (simp add: a) |
67226 | 442 |
hence "x=0" using ord_inj[OF assms] \<open>x < ord a\<close> unfolding inj_on_def by force |
81438 | 443 |
with A show ?thesis by fastforce |
444 |
qed |
|
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
445 |
qed |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
446 |
|
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
447 |
lemma (in group) ord_ge_1: |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
448 |
assumes finite: "finite (carrier G)" and a: "a \<in> carrier G" |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
449 |
shows "ord a \<ge> 1" |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
450 |
proof - |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
451 |
have "((\<lambda>n::nat. a [^] n) ` {0<..}) \<subseteq> carrier G" |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
452 |
using a by blast |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
453 |
then have "finite ((\<lambda>n::nat. a [^] n) ` {0<..})" |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
454 |
using finite_subset finite by auto |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
455 |
then have "\<not> inj_on (\<lambda>n::nat. a [^] n) {0<..}" |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
456 |
using finite_imageD infinite_Ioi by blast |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
457 |
then obtain i j::nat where "i \<noteq> j" "a [^] i = a [^] j" |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
458 |
by (auto simp: inj_on_def) |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
459 |
then have "\<exists>n::nat. n>0 \<and> a [^] n = \<one>" |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
460 |
by (metis a diffs0_imp_equal pow_eq_div2 neq0_conv) |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
461 |
then have "ord a \<noteq> 0" |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
462 |
by (simp add: ord_eq_0 [OF a]) |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
463 |
then show ?thesis |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
464 |
by simp |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
465 |
qed |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
466 |
|
68575 | 467 |
lemma ord_elems: |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
468 |
assumes "finite (carrier G)" "a \<in> carrier G" |
67341
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents:
67299
diff
changeset
|
469 |
shows "{a[^]x | x. x \<in> (UNIV :: nat set)} = {a[^]x | x. x \<in> {0 .. ord a - 1}}" (is "?L = ?R") |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
470 |
proof |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
471 |
show "?R \<subseteq> ?L" by blast |
81438 | 472 |
have "y \<in> {a[^]x | x. x \<in> {0 .. ord a - 1}}" if "y \<in> ?L" for y |
473 |
proof - |
|
474 |
from that obtain x::nat where x: "y = a[^]x" by auto |
|
68157 | 475 |
define r q where "r = x mod ord a" and "q = x div ord a" |
476 |
then have "x = q * ord a + r" |
|
477 |
by (simp add: div_mult_mod_eq) |
|
67341
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents:
67299
diff
changeset
|
478 |
hence "y = (a[^]ord a)[^]q \<otimes> a[^]r" |
70030
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
479 |
using x assms by (metis mult.commute nat_pow_mult nat_pow_pow) |
67341
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents:
67299
diff
changeset
|
480 |
hence "y = a[^]r" using assms by (simp add: pow_ord_eq_1) |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
481 |
have "r < ord a" using ord_ge_1[OF assms] by (simp add: r_def) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
482 |
hence "r \<in> {0 .. ord a - 1}" by (force simp: r_def) |
81438 | 483 |
thus ?thesis using \<open>y=a[^]r\<close> by blast |
484 |
qed |
|
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
485 |
thus "?L \<subseteq> ?R" by auto |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
486 |
qed |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
487 |
|
72630
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
488 |
lemma (in group) |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
489 |
assumes "x \<in> carrier G" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
490 |
shows finite_cyclic_subgroup: |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
491 |
"finite(carrier(subgroup_generated G {x})) \<longleftrightarrow> (\<exists>n::nat. n \<noteq> 0 \<and> x [^] n = \<one>)" (is "?fin \<longleftrightarrow> ?nat1") |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
492 |
and infinite_cyclic_subgroup: |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
493 |
"infinite(carrier(subgroup_generated G {x})) \<longleftrightarrow> (\<forall>m n::nat. x [^] m = x [^] n \<longrightarrow> m = n)" (is "\<not> ?fin \<longleftrightarrow> ?nateq") |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
494 |
and finite_cyclic_subgroup_int: |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
495 |
"finite(carrier(subgroup_generated G {x})) \<longleftrightarrow> (\<exists>i::int. i \<noteq> 0 \<and> x [^] i = \<one>)" (is "?fin \<longleftrightarrow> ?int1") |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
496 |
and infinite_cyclic_subgroup_int: |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
497 |
"infinite(carrier(subgroup_generated G {x})) \<longleftrightarrow> (\<forall>i j::int. x [^] i = x [^] j \<longrightarrow> i = j)" (is "\<not> ?fin \<longleftrightarrow> ?inteq") |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
498 |
proof - |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
499 |
have 1: "\<not> ?fin" if ?nateq |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
500 |
proof - |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
501 |
have "infinite (range (\<lambda>n::nat. x [^] n))" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
502 |
using that range_inj_infinite [of "(\<lambda>n::nat. x [^] n)"] by (auto simp: inj_on_def) |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
503 |
moreover have "range (\<lambda>n::nat. x [^] n) \<subseteq> range (\<lambda>i::int. x [^] i)" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
504 |
apply clarify |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
505 |
by (metis assms group.int_pow_neg int_pow_closed int_pow_neg_int is_group local.inv_equality nat_pow_closed r_inv rangeI) |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
506 |
ultimately show ?thesis |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
507 |
using carrier_subgroup_generated_by_singleton [OF assms] finite_subset by auto |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
508 |
qed |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
509 |
have 2: "m = n" if mn: "x [^] m = x [^] n" and eq [rule_format]: "?inteq" for m n::nat |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
510 |
using eq [of "int m" "int n"] |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
511 |
by (simp add: int_pow_int mn) |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
512 |
have 3: ?nat1 if non: "\<not> ?inteq" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
513 |
proof - |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
514 |
obtain i j::int where eq: "x [^] i = x [^] j" and "i \<noteq> j" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
515 |
using non by auto |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
516 |
show ?thesis |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
517 |
proof (cases i j rule: linorder_cases) |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
518 |
case less |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
519 |
then have [simp]: "x [^] (j - i) = \<one>" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
520 |
by (simp add: eq assms int_pow_diff) |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
521 |
show ?thesis |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
522 |
using less by (rule_tac x="nat (j-i)" in exI) auto |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
523 |
next |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
524 |
case greater |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
525 |
then have [simp]: "x [^] (i - j) = \<one>" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
526 |
by (simp add: eq assms int_pow_diff) |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
527 |
then show ?thesis |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
528 |
using greater by (rule_tac x="nat (i-j)" in exI) auto |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
529 |
qed (use \<open>i \<noteq> j\<close> in auto) |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
530 |
qed |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
531 |
have 4: "\<exists>i::int. (i \<noteq> 0) \<and> x [^] i = \<one>" if "n \<noteq> 0" "x [^] n = \<one>" for n::nat |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
532 |
apply (rule_tac x="int n" in exI) |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
533 |
by (simp add: int_pow_int that) |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
534 |
have 5: "finite (carrier (subgroup_generated G {x}))" if "i \<noteq> 0" and 1: "x [^] i = \<one>" for i::int |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
535 |
proof - |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
536 |
obtain n::nat where n: "n > 0" "x [^] n = \<one>" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
537 |
using "1" "3" \<open>i \<noteq> 0\<close> by fastforce |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
538 |
have "x [^] a \<in> ([^]) x ` {0..<n}" for a::int |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
539 |
proof |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
540 |
show "x [^] a = x [^] nat (a mod int n)" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
541 |
using n |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
542 |
by simp (metis (no_types, lifting) assms dvd_minus_mod dvd_trans int_pow_eq int_pow_eq_id int_pow_int) |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
543 |
show "nat (a mod int n) \<in> {0..<n}" |
77061
5de3772609ea
generalized theory name: euclidean division denotes one particular division definition on integers
haftmann
parents:
76987
diff
changeset
|
544 |
using n by (simp add: nat_less_iff) |
72630
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
545 |
qed |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
546 |
then have "carrier (subgroup_generated G {x}) \<subseteq> ([^]) x ` {0..<n}" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
547 |
using carrier_subgroup_generated_by_singleton [OF assms] by auto |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
548 |
then show ?thesis |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
549 |
using finite_surj by blast |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
550 |
qed |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
551 |
show "?fin \<longleftrightarrow> ?nat1" "\<not> ?fin \<longleftrightarrow> ?nateq" "?fin \<longleftrightarrow> ?int1" "\<not> ?fin \<longleftrightarrow> ?inteq" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
552 |
using 1 2 3 4 5 by meson+ |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
553 |
qed |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
554 |
|
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
555 |
lemma (in group) finite_cyclic_subgroup_order: |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
556 |
"x \<in> carrier G \<Longrightarrow> finite(carrier(subgroup_generated G {x})) \<longleftrightarrow> ord x \<noteq> 0" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
557 |
by (simp add: finite_cyclic_subgroup ord_eq_0) |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
558 |
|
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
559 |
lemma (in group) infinite_cyclic_subgroup_order: |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
560 |
"x \<in> carrier G \<Longrightarrow> infinite (carrier(subgroup_generated G {x})) \<longleftrightarrow> ord x = 0" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
561 |
by (simp add: finite_cyclic_subgroup_order) |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
562 |
|
69895
6b03a8cf092d
more formal contributors (with the help of the history);
wenzelm
parents:
69785
diff
changeset
|
563 |
lemma generate_pow_on_finite_carrier: \<^marker>\<open>contributor \<open>Paulo Emílio de Vilhena\<close>\<close> |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
564 |
assumes "finite (carrier G)" and a: "a \<in> carrier G" |
68575 | 565 |
shows "generate G { a } = { a [^] k | k. k \<in> (UNIV :: nat set) }" |
566 |
proof |
|
567 |
show "{ a [^] k | k. k \<in> (UNIV :: nat set) } \<subseteq> generate G { a }" |
|
568 |
proof |
|
569 |
fix b assume "b \<in> { a [^] k | k. k \<in> (UNIV :: nat set) }" |
|
570 |
then obtain k :: nat where "b = a [^] k" by blast |
|
571 |
hence "b = a [^] (int k)" |
|
69749
10e48c47a549
some new results in group theory
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
572 |
by (simp add: int_pow_int) |
68575 | 573 |
thus "b \<in> generate G { a }" |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
574 |
unfolding generate_pow[OF a] by blast |
68575 | 575 |
qed |
576 |
next |
|
577 |
show "generate G { a } \<subseteq> { a [^] k | k. k \<in> (UNIV :: nat set) }" |
|
578 |
proof |
|
579 |
fix b assume "b \<in> generate G { a }" |
|
580 |
then obtain k :: int where k: "b = a [^] k" |
|
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
581 |
unfolding generate_pow[OF a] by blast |
68575 | 582 |
show "b \<in> { a [^] k | k. k \<in> (UNIV :: nat set) }" |
583 |
proof (cases "k < 0") |
|
584 |
assume "\<not> k < 0" |
|
585 |
hence "b = a [^] (nat k)" |
|
70027
94494b92d8d0
some new group theory results: integer group, trivial group, etc.
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
586 |
by (simp add: k) |
68575 | 587 |
thus ?thesis by blast |
588 |
next |
|
589 |
assume "k < 0" |
|
590 |
hence b: "b = inv (a [^] (nat (- k)))" |
|
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
591 |
using k a by (auto simp: int_pow_neg) |
68575 | 592 |
obtain m where m: "ord a * m \<ge> nat (- k)" |
593 |
by (metis assms mult.left_neutral mult_le_mono1 ord_ge_1) |
|
594 |
hence "a [^] (ord a * m) = \<one>" |
|
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
595 |
by (metis a nat_pow_one nat_pow_pow pow_ord_eq_1) |
68575 | 596 |
then obtain k' :: nat where "(a [^] (nat (- k))) \<otimes> (a [^] k') = \<one>" |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
597 |
using m a nat_le_iff_add nat_pow_mult by auto |
68575 | 598 |
hence "b = a [^] k'" |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
599 |
using b a by (metis inv_unique' nat_pow_closed nat_pow_comm) |
68575 | 600 |
thus "b \<in> { a [^] k | k. k \<in> (UNIV :: nat set) }" by blast |
601 |
qed |
|
602 |
qed |
|
603 |
qed |
|
604 |
||
72630
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
605 |
lemma ord_elems_inf_carrier: |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
606 |
assumes "a \<in> carrier G" "ord a \<noteq> 0" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
607 |
shows "{a[^]x | x. x \<in> (UNIV :: nat set)} = {a[^]x | x. x \<in> {0 .. ord a - 1}}" (is "?L = ?R") |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
608 |
proof |
81438 | 609 |
show "?R \<subseteq> ?L" by blast |
610 |
have "y \<in> {a[^]x | x. x \<in> {0 .. ord a - 1}}" if "y \<in> ?L" for y |
|
611 |
proof - |
|
612 |
from that obtain x::nat where x: "y = a[^]x" by auto |
|
613 |
define r q where "r = x mod ord a" and "q = x div ord a" |
|
614 |
then have "x = q * ord a + r" |
|
615 |
by (simp add: div_mult_mod_eq) |
|
616 |
hence "y = (a[^]ord a)[^]q \<otimes> a[^]r" |
|
617 |
using x assms by (metis mult.commute nat_pow_mult nat_pow_pow) |
|
618 |
hence "y = a[^]r" using assms by simp |
|
619 |
have "r < ord a" using assms by (simp add: r_def) |
|
620 |
hence "r \<in> {0 .. ord a - 1}" by (force simp: r_def) |
|
621 |
then show ?thesis using \<open>y=a[^]r\<close> by blast |
|
622 |
qed |
|
623 |
thus "?L \<subseteq> ?R" by auto |
|
68575 | 624 |
qed |
625 |
||
72630
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
626 |
lemma generate_pow_nat: |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
627 |
assumes a: "a \<in> carrier G" and "ord a \<noteq> 0" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
628 |
shows "generate G { a } = { a [^] k | k. k \<in> (UNIV :: nat set) }" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
629 |
proof |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
630 |
show "{ a [^] k | k. k \<in> (UNIV :: nat set) } \<subseteq> generate G { a }" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
631 |
proof |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
632 |
fix b assume "b \<in> { a [^] k | k. k \<in> (UNIV :: nat set) }" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
633 |
then obtain k :: nat where "b = a [^] k" by blast |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
634 |
hence "b = a [^] (int k)" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
635 |
by (simp add: int_pow_int) |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
636 |
thus "b \<in> generate G { a }" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
637 |
unfolding generate_pow[OF a] by blast |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
638 |
qed |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
639 |
next |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
640 |
show "generate G { a } \<subseteq> { a [^] k | k. k \<in> (UNIV :: nat set) }" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
641 |
proof |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
642 |
fix b assume "b \<in> generate G { a }" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
643 |
then obtain k :: int where k: "b = a [^] k" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
644 |
unfolding generate_pow[OF a] by blast |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
645 |
show "b \<in> { a [^] k | k. k \<in> (UNIV :: nat set) }" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
646 |
proof (cases "k < 0") |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
647 |
assume "\<not> k < 0" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
648 |
hence "b = a [^] (nat k)" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
649 |
by (simp add: k) |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
650 |
thus ?thesis by blast |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
651 |
next |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
652 |
assume "k < 0" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
653 |
hence b: "b = inv (a [^] (nat (- k)))" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
654 |
using k a by (auto simp: int_pow_neg) |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
655 |
obtain m where m: "ord a * m \<ge> nat (- k)" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
656 |
by (metis assms(2) dvd_imp_le dvd_triv_right le_zero_eq mult_eq_0_iff not_gr_zero) |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
657 |
hence "a [^] (ord a * m) = \<one>" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
658 |
by (metis a nat_pow_one nat_pow_pow pow_ord_eq_1) |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
659 |
then obtain k' :: nat where "(a [^] (nat (- k))) \<otimes> (a [^] k') = \<one>" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
660 |
using m a nat_le_iff_add nat_pow_mult by auto |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
661 |
hence "b = a [^] k'" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
662 |
using b a by (metis inv_unique' nat_pow_closed nat_pow_comm) |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
663 |
thus "b \<in> { a [^] k | k. k \<in> (UNIV :: nat set) }" by blast |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
664 |
qed |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
665 |
qed |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
666 |
qed |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
667 |
|
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
668 |
lemma generate_pow_card: |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
669 |
assumes a: "a \<in> carrier G" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
670 |
shows "ord a = card (generate G { a })" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
671 |
proof (cases "ord a = 0") |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
672 |
case True |
72630
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
673 |
then have "infinite (carrier (subgroup_generated G {a}))" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
674 |
using infinite_cyclic_subgroup_order[OF a] by auto |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
675 |
then have "infinite (generate G {a})" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
676 |
unfolding subgroup_generated_def |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
677 |
using a by simp |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
678 |
then show ?thesis |
73102 | 679 |
using \<open>ord a = 0\<close> by auto |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
680 |
next |
81438 | 681 |
case finite_subgroup: False |
72630
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
682 |
then have "generate G { a } = (([^]) a) ` {0..ord a - 1}" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
683 |
using generate_pow_nat ord_elems_inf_carrier a by auto |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
684 |
hence "card (generate G {a}) = card {0..ord a - 1}" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
685 |
using ord_inj[OF a] card_image by metis |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
686 |
also have "... = ord a" using finite_subgroup by auto |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
687 |
finally show ?thesis.. |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
688 |
qed |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
689 |
|
72630
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
690 |
lemma (in group) cyclic_order_is_ord: |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
691 |
assumes "g \<in> carrier G" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
692 |
shows "ord g = order (subgroup_generated G {g})" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
693 |
unfolding order_def subgroup_generated_def |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
694 |
using assms generate_pow_card by simp |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
695 |
|
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
696 |
lemma ord_dvd_group_order: |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
697 |
assumes "a \<in> carrier G" shows "(ord a) dvd (order G)" |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
698 |
using lagrange[OF generate_is_subgroup[of "{a}"]] assms |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
699 |
unfolding generate_pow_card[OF assms] |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
700 |
by (metis dvd_triv_right empty_subsetI insert_subset) |
4167d3d3d478
Jakub Kądziołka's stronger version of generate_pow_card (required some restructuring)
paulson <lp15@cam.ac.uk>
parents:
71392
diff
changeset
|
701 |
|
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
702 |
lemma (in group) pow_order_eq_1: |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
703 |
assumes "a \<in> carrier G" shows "a [^] order G = \<one>" |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
704 |
using assms by (metis nat_pow_pow ord_dvd_group_order pow_ord_eq_1 dvdE nat_pow_one) |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
705 |
|
70131 | 706 |
lemma dvd_gcd: |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
707 |
fixes a b :: nat |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
708 |
obtains q where "a * (b div gcd a b) = b*q" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
709 |
proof |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
710 |
have "a * (b div gcd a b) = (a div gcd a b) * b" by (simp add: div_mult_swap dvd_div_mult) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
711 |
also have "\<dots> = b * (a div gcd a b)" by simp |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
712 |
finally show "a * (b div gcd a b) = b * (a div gcd a b) " . |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
713 |
qed |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
714 |
|
70131 | 715 |
lemma (in group) ord_le_group_order: |
716 |
assumes finite: "finite (carrier G)" and a: "a \<in> carrier G" |
|
717 |
shows "ord a \<le> order G" |
|
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
718 |
by (simp add: a dvd_imp_le local.finite ord_dvd_group_order order_gt_0_iff_finite) |
70131 | 719 |
|
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
720 |
lemma (in group) ord_pow_gen: |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
721 |
assumes "x \<in> carrier G" |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
722 |
shows "ord (pow G x k) = (if k = 0 then 1 else ord x div gcd (ord x) k)" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
723 |
proof - |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
724 |
have "ord (x [^] k) = ord x div gcd (ord x) k" |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
725 |
if "0 < k" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
726 |
proof - |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
727 |
have "(d dvd k * n) = (d div gcd (d) k dvd n)" for d n |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
728 |
using that by (simp add: div_dvd_iff_mult gcd_mult_distrib_nat mult.commute) |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
729 |
then show ?thesis |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
730 |
using that by (auto simp add: assms ord_unique nat_pow_pow pow_eq_id) |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
731 |
qed |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
732 |
then show ?thesis by auto |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
733 |
qed |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
734 |
|
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
735 |
lemma (in group) |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
736 |
assumes finite': "finite (carrier G)" "a \<in> carrier G" |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
737 |
shows pow_ord_eq_ord_iff: "group.ord G (a [^] k) = ord a \<longleftrightarrow> coprime k (ord a)" (is "?L \<longleftrightarrow> ?R") |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
738 |
using assms ord_ge_1 [OF assms] |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
739 |
by (auto simp: div_eq_dividend_iff ord_pow_gen coprime_iff_gcd_eq_1 gcd.commute split: if_split_asm) |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
740 |
|
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
741 |
lemma element_generates_subgroup: |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
742 |
assumes finite[simp]: "finite (carrier G)" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
743 |
assumes a[simp]: "a \<in> carrier G" |
67341
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents:
67299
diff
changeset
|
744 |
shows "subgroup {a [^] i | i. i \<in> {0 .. ord a - 1}} G" |
68575 | 745 |
using generate_is_subgroup[of "{ a }"] assms(2) |
746 |
generate_pow_on_finite_carrier[OF assms] |
|
747 |
unfolding ord_elems[OF assms] by auto |
|
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
748 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
749 |
end |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
750 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
751 |
|
67226 | 752 |
section \<open>Number of Roots of a Polynomial\<close> |
753 |
text_raw \<open>\label{sec:number-roots}\<close> |
|
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
754 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
755 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
756 |
definition mult_of :: "('a, 'b) ring_scheme \<Rightarrow> 'a monoid" where |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
757 |
"mult_of R \<equiv> \<lparr> carrier = carrier R - {\<zero>\<^bsub>R\<^esub>}, mult = mult R, one = \<one>\<^bsub>R\<^esub>\<rparr>" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
758 |
|
68583 | 759 |
lemma carrier_mult_of [simp]: "carrier (mult_of R) = carrier R - {\<zero>\<^bsub>R\<^esub>}" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
760 |
by (simp add: mult_of_def) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
761 |
|
68583 | 762 |
lemma mult_mult_of [simp]: "mult (mult_of R) = mult R" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
763 |
by (simp add: mult_of_def) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
764 |
|
67399 | 765 |
lemma nat_pow_mult_of: "([^]\<^bsub>mult_of R\<^esub>) = (([^]\<^bsub>R\<^esub>) :: _ \<Rightarrow> nat \<Rightarrow> _)" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
766 |
by (simp add: mult_of_def fun_eq_iff nat_pow_def) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
767 |
|
68583 | 768 |
lemma one_mult_of [simp]: "\<one>\<^bsub>mult_of R\<^esub> = \<one>\<^bsub>R\<^esub>" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
769 |
by (simp add: mult_of_def) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
770 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
771 |
lemmas mult_of_simps = carrier_mult_of mult_mult_of nat_pow_mult_of one_mult_of |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
772 |
|
68575 | 773 |
context field |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
68445
diff
changeset
|
774 |
begin |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
68445
diff
changeset
|
775 |
|
68575 | 776 |
lemma mult_of_is_Units: "mult_of R = units_of R" |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
68445
diff
changeset
|
777 |
unfolding mult_of_def units_of_def using field_Units by auto |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
778 |
|
70131 | 779 |
lemma m_inv_mult_of: |
68561 | 780 |
"\<And>x. x \<in> carrier (mult_of R) \<Longrightarrow> m_inv (mult_of R) x = m_inv R x" |
781 |
using mult_of_is_Units units_of_inv unfolding units_of_def |
|
68575 | 782 |
by simp |
68561 | 783 |
|
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
784 |
lemma (in field) field_mult_group: "group (mult_of R)" |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
785 |
proof (rule groupI) |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
786 |
show "\<exists>y\<in>carrier (mult_of R). y \<otimes>\<^bsub>mult_of R\<^esub> x = \<one>\<^bsub>mult_of R\<^esub>" |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
787 |
if "x \<in> carrier (mult_of R)" for x |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
788 |
using group.l_inv_ex mult_of_is_Units that units_group by fastforce |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
789 |
qed (auto simp: m_assoc dest: integral) |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
790 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
791 |
lemma finite_mult_of: "finite (carrier R) \<Longrightarrow> finite (carrier (mult_of R))" |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
792 |
by simp |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
793 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
794 |
lemma order_mult_of: "finite (carrier R) \<Longrightarrow> order (mult_of R) = order R - 1" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
795 |
unfolding order_def carrier_mult_of by (simp add: card.remove) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
796 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
797 |
end |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
798 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
799 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
800 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
801 |
lemma (in monoid) Units_pow_closed : |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
802 |
fixes d :: nat |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
803 |
assumes "x \<in> Units G" |
67341
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents:
67299
diff
changeset
|
804 |
shows "x [^] d \<in> Units G" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
805 |
by (metis assms group.is_monoid monoid.nat_pow_closed units_group units_of_carrier units_of_pow) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
806 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
807 |
lemma (in ring) r_right_minus_eq[simp]: |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
808 |
assumes "a \<in> carrier R" "b \<in> carrier R" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
809 |
shows "a \<ominus> b = \<zero> \<longleftrightarrow> a = b" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
810 |
using assms by (metis a_minus_def add.inv_closed minus_equality r_neg) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
811 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
812 |
context UP_cring begin |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
813 |
|
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
814 |
lemma is_UP_cring: "UP_cring R" by (unfold_locales) |
70131 | 815 |
lemma is_UP_ring: |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
816 |
shows "UP_ring R" by (unfold_locales) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
817 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
818 |
end |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
819 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
820 |
context UP_domain begin |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
821 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
822 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
823 |
lemma roots_bound: |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
824 |
assumes f [simp]: "f \<in> carrier P" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
825 |
assumes f_not_zero: "f \<noteq> \<zero>\<^bsub>P\<^esub>" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
826 |
assumes finite: "finite (carrier R)" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
827 |
shows "finite {a \<in> carrier R . eval R R id a f = \<zero>} \<and> |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
828 |
card {a \<in> carrier R . eval R R id a f = \<zero>} \<le> deg R f" using f f_not_zero |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
829 |
proof (induction "deg R f" arbitrary: f) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
830 |
case 0 |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
831 |
have "\<And>x. eval R R id x f \<noteq> \<zero>" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
832 |
proof - |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
833 |
fix x |
67341
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents:
67299
diff
changeset
|
834 |
have "(\<Oplus>i\<in>{..deg R f}. id (coeff P f i) \<otimes> x [^] i) \<noteq> \<zero>" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
835 |
using 0 lcoeff_nonzero_nonzero[where p = f] by simp |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
836 |
thus "eval R R id x f \<noteq> \<zero>" using 0 unfolding eval_def P_def by simp |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
837 |
qed |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
838 |
then have *: "{a \<in> carrier R. eval R R (\<lambda>a. a) a f = \<zero>} = {}" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
839 |
by (auto simp: id_def) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
840 |
show ?case by (simp add: *) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
841 |
next |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
842 |
case (Suc x) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
843 |
show ?case |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
844 |
proof (cases "\<exists> a \<in> carrier R . eval R R id a f = \<zero>") |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
845 |
case True |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
846 |
then obtain a where a_carrier[simp]: "a \<in> carrier R" and a_root: "eval R R id a f = \<zero>" by blast |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
847 |
have R_not_triv: "carrier R \<noteq> {\<zero>}" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
848 |
by (metis R.one_zeroI R.zero_not_one) |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
849 |
obtain q where q: "(q \<in> carrier P)" and |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
850 |
f: "f = (monom P \<one>\<^bsub>R\<^esub> 1 \<ominus>\<^bsub> P\<^esub> monom P a 0) \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> monom P (eval R R id a f) 0" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
851 |
using remainder_theorem[OF Suc.prems(1) a_carrier R_not_triv] by auto |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
852 |
hence lin_fac: "f = (monom P \<one>\<^bsub>R\<^esub> 1 \<ominus>\<^bsub> P\<^esub> monom P a 0) \<otimes>\<^bsub>P\<^esub> q" using q by (simp add: a_root) |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
853 |
have deg: "deg R (monom P \<one>\<^bsub>R\<^esub> 1 \<ominus>\<^bsub> P\<^esub> monom P a 0) = 1" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
854 |
using a_carrier by (simp add: deg_minus_eq) |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
855 |
hence mon_not_zero: "(monom P \<one>\<^bsub>R\<^esub> 1 \<ominus>\<^bsub> P\<^esub> monom P a 0) \<noteq> \<zero>\<^bsub>P\<^esub>" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
856 |
by (fastforce simp del: r_right_minus_eq) |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
857 |
have q_not_zero: "q \<noteq> \<zero>\<^bsub>P\<^esub>" using Suc by (auto simp add : lin_fac) |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
858 |
hence "deg R q = x" using Suc deg deg_mult[OF mon_not_zero q_not_zero _ q] |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
859 |
by (simp add : lin_fac) |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
860 |
hence q_IH: "finite {a \<in> carrier R . eval R R id a q = \<zero>} |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
861 |
\<and> card {a \<in> carrier R . eval R R id a q = \<zero>} \<le> x" using Suc q q_not_zero by blast |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
862 |
have subs: "{a \<in> carrier R . eval R R id a f = \<zero>} |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
863 |
\<subseteq> {a \<in> carrier R . eval R R id a q = \<zero>} \<union> {a}" (is "?L \<subseteq> ?R \<union> {a}") |
67226 | 864 |
using a_carrier \<open>q \<in> _\<close> |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
865 |
by (auto simp: evalRR_simps lin_fac R.integral_iff) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
866 |
have "{a \<in> carrier R . eval R R id a f = \<zero>} \<subseteq> insert a {a \<in> carrier R . eval R R id a q = \<zero>}" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
867 |
using subs by auto |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
868 |
hence "card {a \<in> carrier R . eval R R id a f = \<zero>} \<le> |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
869 |
card (insert a {a \<in> carrier R . eval R R id a q = \<zero>})" using q_IH by (blast intro: card_mono) |
67226 | 870 |
also have "\<dots> \<le> deg R f" using q_IH \<open>Suc x = _\<close> |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
871 |
by (simp add: card_insert_if) |
67226 | 872 |
finally show ?thesis using q_IH \<open>Suc x = _\<close> using finite by force |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
873 |
next |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
874 |
case False |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
875 |
hence "card {a \<in> carrier R. eval R R id a f = \<zero>} = 0" using finite by auto |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
876 |
also have "\<dots> \<le> deg R f" by simp |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
877 |
finally show ?thesis using finite by auto |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
878 |
qed |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
879 |
qed |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
880 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
881 |
end |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
882 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
883 |
lemma (in domain) num_roots_le_deg : |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
884 |
fixes p d :: nat |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
885 |
assumes finite: "finite (carrier R)" |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
886 |
assumes d_neq_zero: "d \<noteq> 0" |
67341
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents:
67299
diff
changeset
|
887 |
shows "card {x \<in> carrier R. x [^] d = \<one>} \<le> d" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
888 |
proof - |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
889 |
let ?f = "monom (UP R) \<one>\<^bsub>R\<^esub> d \<ominus>\<^bsub> (UP R)\<^esub> monom (UP R) \<one>\<^bsub>R\<^esub> 0" |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
890 |
have one_in_carrier: "\<one> \<in> carrier R" by simp |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
891 |
interpret R: UP_domain R "UP R" by (unfold_locales) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
892 |
have "deg R ?f = d" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
893 |
using d_neq_zero by (simp add: R.deg_minus_eq) |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
894 |
hence f_not_zero: "?f \<noteq> \<zero>\<^bsub>UP R\<^esub>" using d_neq_zero by (auto simp add : R.deg_nzero_nzero) |
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
895 |
have roots_bound: "finite {a \<in> carrier R . eval R R id a ?f = \<zero>} \<and> |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
896 |
card {a \<in> carrier R . eval R R id a ?f = \<zero>} \<le> deg R ?f" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
897 |
using finite by (intro R.roots_bound[OF _ f_not_zero]) simp |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
898 |
have subs: "{x \<in> carrier R. x [^] d = \<one>} \<subseteq> {a \<in> carrier R . eval R R id a ?f = \<zero>}" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
899 |
by (auto simp: R.evalRR_simps) |
67341
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents:
67299
diff
changeset
|
900 |
then have "card {x \<in> carrier R. x [^] d = \<one>} \<le> |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
901 |
card {a \<in> carrier R. eval R R id a ?f = \<zero>}" using finite by (simp add : card_mono) |
67226 | 902 |
thus ?thesis using \<open>deg R ?f = d\<close> roots_bound by linarith |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
903 |
qed |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
904 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
905 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
906 |
|
67226 | 907 |
section \<open>The Multiplicative Group of a Field\<close> |
908 |
text_raw \<open>\label{sec:mult-group}\<close> |
|
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
909 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
910 |
|
67226 | 911 |
text \<open> |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
912 |
In this section we show that the multiplicative group of a finite field |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
913 |
is generated by a single element, i.e. it is cyclic. The proof is inspired |
76987 | 914 |
by the first proof given in the survey~\<^cite>\<open>"conrad-cyclicity"\<close>. |
67226 | 915 |
\<close> |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
916 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
917 |
context field begin |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
918 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
919 |
lemma num_elems_of_ord_eq_phi': |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
920 |
assumes finite: "finite (carrier R)" and dvd: "d dvd order (mult_of R)" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
921 |
and exists: "\<exists>a\<in>carrier (mult_of R). group.ord (mult_of R) a = d" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
922 |
shows "card {a \<in> carrier (mult_of R). group.ord (mult_of R) a = d} = phi' d" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
923 |
proof - |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
924 |
note mult_of_simps[simp] |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
925 |
have finite': "finite (carrier (mult_of R))" using finite by (rule finite_mult_of) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
926 |
|
67399 | 927 |
interpret G:group "mult_of R" rewrites "([^]\<^bsub>mult_of R\<^esub>) = (([^]) :: _ \<Rightarrow> nat \<Rightarrow> _)" and "\<one>\<^bsub>mult_of R\<^esub> = \<one>" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
928 |
by (rule field_mult_group) simp_all |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
929 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
930 |
from exists |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
931 |
obtain a where a: "a \<in> carrier (mult_of R)" and ord_a: "group.ord (mult_of R) a = d" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
932 |
by (auto simp add: card_gt_0_iff) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
933 |
|
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
934 |
have set_eq1: "{a[^]n| n. n \<in> {1 .. d}} = {x \<in> carrier (mult_of R). x [^] d = \<one>}" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
935 |
proof (rule card_seteq) |
67341
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents:
67299
diff
changeset
|
936 |
show "finite {x \<in> carrier (mult_of R). x [^] d = \<one>}" using finite by auto |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
937 |
|
67341
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents:
67299
diff
changeset
|
938 |
show "{a[^]n| n. n \<in> {1 ..d}} \<subseteq> {x \<in> carrier (mult_of R). x[^]d = \<one>}" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
939 |
proof |
67341
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents:
67299
diff
changeset
|
940 |
fix x assume "x \<in> {a[^]n | n. n \<in> {1 .. d}}" |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
941 |
then obtain n where n: "x = a[^]n \<and> n \<in> {1 .. d}" by auto |
67341
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents:
67299
diff
changeset
|
942 |
have "x[^]d =(a[^]d)[^]n" using n a ord_a by (simp add:nat_pow_pow mult.commute) |
70030
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
943 |
hence "x[^]d = \<one>" using ord_a G.pow_ord_eq_1[OF a] by fastforce |
67341
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents:
67299
diff
changeset
|
944 |
thus "x \<in> {x \<in> carrier (mult_of R). x[^]d = \<one>}" using G.nat_pow_closed[OF a] n by blast |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
945 |
qed |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
946 |
|
67341
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents:
67299
diff
changeset
|
947 |
show "card {x \<in> carrier (mult_of R). x [^] d = \<one>} \<le> card {a[^]n | n. n \<in> {1 .. d}}" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
948 |
proof - |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
949 |
have *: "{a[^]n | n. n \<in> {1 .. d }} = ((\<lambda> n. a[^]n) ` {1 .. d})" by auto |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
950 |
have "0 < order (mult_of R)" unfolding order_mult_of[OF finite] |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
951 |
using card_mono[OF finite, of "{\<zero>, \<one>}"] by (simp add: order_def) |
67341
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents:
67299
diff
changeset
|
952 |
have "card {x \<in> carrier (mult_of R). x [^] d = \<one>} \<le> card {x \<in> carrier R. x [^] d = \<one>}" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
953 |
using finite by (auto intro: card_mono) |
67226 | 954 |
also have "\<dots> \<le> d" using \<open>0 < order (mult_of R)\<close> num_roots_le_deg[OF finite, of d] |
955 |
by (simp add : dvd_pos_nat[OF _ \<open>d dvd order (mult_of R)\<close>]) |
|
70131 | 956 |
finally show ?thesis using G.ord_inj'[OF a] ord_a * by (simp add: card_image) |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
957 |
qed |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
958 |
qed |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
959 |
|
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
960 |
have set_eq2: "{x \<in> carrier (mult_of R) . group.ord (mult_of R) x = d} |
67341
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents:
67299
diff
changeset
|
961 |
= (\<lambda> n . a[^]n) ` {n \<in> {1 .. d}. group.ord (mult_of R) (a[^]n) = d}" (is "?L = ?R") |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
962 |
proof |
81438 | 963 |
have "x \<in> ?R" if x: "x \<in> (carrier (mult_of R)) \<and> group.ord (mult_of R) x = d" for x |
964 |
proof - |
|
965 |
from that have "x \<in> {x \<in> carrier (mult_of R). x [^] d = \<one>}" |
|
70030
042ae6ca2c40
The order of a group now follows the HOL Light definition, which is more general
paulson <lp15@cam.ac.uk>
parents:
70027
diff
changeset
|
966 |
by (simp add: G.pow_ord_eq_1[of x, symmetric]) |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
967 |
then obtain n where n: "x = a[^]n \<and> n \<in> {1 .. d}" using set_eq1 by blast |
81438 | 968 |
then show ?thesis using x by fast |
969 |
qed |
|
970 |
thus "?L \<subseteq> ?R" by blast |
|
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
971 |
show "?R \<subseteq> ?L" using a by (auto simp add: carrier_mult_of[symmetric] simp del: carrier_mult_of) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
972 |
qed |
67341
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents:
67299
diff
changeset
|
973 |
have "inj_on (\<lambda> n . a[^]n) {n \<in> {1 .. d}. group.ord (mult_of R) (a[^]n) = d}" |
70131 | 974 |
using G.ord_inj'[OF a, unfolded ord_a] unfolding inj_on_def by fast |
67341
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents:
67299
diff
changeset
|
975 |
hence "card ((\<lambda>n. a[^]n) ` {n \<in> {1 .. d}. group.ord (mult_of R) (a[^]n) = d}) |
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents:
67299
diff
changeset
|
976 |
= card {k \<in> {1 .. d}. group.ord (mult_of R) (a[^]k) = d}" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
977 |
using card_image by blast |
67226 | 978 |
thus ?thesis using set_eq2 G.pow_ord_eq_ord_iff[OF finite' \<open>a \<in> _\<close>, unfolded ord_a] |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
979 |
by (simp add: phi'_def) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
980 |
qed |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
981 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
982 |
end |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
983 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
984 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
985 |
theorem (in field) finite_field_mult_group_has_gen : |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
986 |
assumes finite: "finite (carrier R)" |
67341
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents:
67299
diff
changeset
|
987 |
shows "\<exists> a \<in> carrier (mult_of R) . carrier (mult_of R) = {a[^]i | i::nat . i \<in> UNIV}" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
988 |
proof - |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
989 |
note mult_of_simps[simp] |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
990 |
have finite': "finite (carrier (mult_of R))" using finite by (rule finite_mult_of) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
991 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
992 |
interpret G: group "mult_of R" rewrites |
67399 | 993 |
"([^]\<^bsub>mult_of R\<^esub>) = (([^]) :: _ \<Rightarrow> nat \<Rightarrow> _)" and "\<one>\<^bsub>mult_of R\<^esub> = \<one>" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
994 |
by (rule field_mult_group) (simp_all add: fun_eq_iff nat_pow_def) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
995 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
996 |
let ?N = "\<lambda> x . card {a \<in> carrier (mult_of R). group.ord (mult_of R) a = x}" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
997 |
have "0 < order R - 1" unfolding order_def using card_mono[OF finite, of "{\<zero>, \<one>}"] by simp |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
998 |
then have *: "0 < order (mult_of R)" using assms by (simp add: order_mult_of) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
999 |
have fin: "finite {d. d dvd order (mult_of R) }" using dvd_nat_bounds[OF *] by force |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1000 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1001 |
have "(\<Sum>d | d dvd order (mult_of R). ?N d) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1002 |
= card (UN d:{d . d dvd order (mult_of R) }. {a \<in> carrier (mult_of R). group.ord (mult_of R) a = d})" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1003 |
(is "_ = card ?U") |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1004 |
using fin finite by (subst card_UN_disjoint) auto |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1005 |
also have "?U = carrier (mult_of R)" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1006 |
proof |
81438 | 1007 |
have "x \<in> ?U" if x: "x \<in> carrier (mult_of R)" for x |
1008 |
proof - |
|
1009 |
from that have x': "x\<in>carrier (mult_of R)" by simp |
|
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1010 |
then have "group.ord (mult_of R) x dvd order (mult_of R)" |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
1011 |
using G.ord_dvd_group_order by blast |
81438 | 1012 |
then show ?thesis |
1013 |
using dvd_nat_bounds[of "order (mult_of R)" "group.ord (mult_of R) x"] x by blast |
|
1014 |
qed |
|
1015 |
thus "carrier (mult_of R) \<subseteq> ?U" by blast |
|
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1016 |
qed auto |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1017 |
also have "card ... = order (mult_of R)" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1018 |
using order_mult_of finite' by (simp add: order_def) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1019 |
finally have sum_Ns_eq: "(\<Sum>d | d dvd order (mult_of R). ?N d) = order (mult_of R)" . |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1020 |
|
81438 | 1021 |
have "card {a \<in> carrier (mult_of R). group.ord (mult_of R) a = d} \<le> phi' d" |
1022 |
if d: "d dvd order (mult_of R)" for d |
|
1023 |
proof (cases "card {a \<in> carrier (mult_of R). group.ord (mult_of R) a = d} = 0") |
|
1024 |
case True |
|
1025 |
thus ?thesis by presburger |
|
1026 |
next |
|
1027 |
case False |
|
1028 |
hence "\<exists>a \<in> carrier (mult_of R). group.ord (mult_of R) a = d" by (auto simp: card_eq_0_iff) |
|
1029 |
thus ?thesis using num_elems_of_ord_eq_phi'[OF finite d] by auto |
|
1030 |
qed |
|
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
1031 |
hence all_le: "\<And>i. i \<in> {d. d dvd order (mult_of R) } |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1032 |
\<Longrightarrow> (\<lambda>i. card {a \<in> carrier (mult_of R). group.ord (mult_of R) a = i}) i \<le> (\<lambda>i. phi' i) i" by fast |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
1033 |
hence le: "(\<Sum>i | i dvd order (mult_of R). ?N i) |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1034 |
\<le> (\<Sum>i | i dvd order (mult_of R). phi' i)" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1035 |
using sum_mono[of "{d . d dvd order (mult_of R)}" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1036 |
"\<lambda>i. card {a \<in> carrier (mult_of R). group.ord (mult_of R) a = i}"] by presburger |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1037 |
have "order (mult_of R) = (\<Sum>d | d dvd order (mult_of R). phi' d)" using * |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1038 |
by (simp add: sum_phi'_factors) |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
1039 |
hence eq: "(\<Sum>i | i dvd order (mult_of R). ?N i) |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1040 |
= (\<Sum>i | i dvd order (mult_of R). phi' i)" using le sum_Ns_eq by presburger |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1041 |
have "\<And>i. i \<in> {d. d dvd order (mult_of R) } \<Longrightarrow> ?N i = (\<lambda>i. phi' i) i" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1042 |
proof (rule ccontr) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1043 |
fix i |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
1044 |
assume i1: "i \<in> {d. d dvd order (mult_of R)}" and "?N i \<noteq> phi' i" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1045 |
hence "?N i = 0" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1046 |
using num_elems_of_ord_eq_phi'[OF finite, of i] by (auto simp: card_eq_0_iff) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1047 |
moreover have "0 < i" using * i1 by (simp add: dvd_nat_bounds[of "order (mult_of R)" i]) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1048 |
ultimately have "?N i < phi' i" using phi'_nonzero by presburger |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1049 |
hence "(\<Sum>i | i dvd order (mult_of R). ?N i) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1050 |
< (\<Sum>i | i dvd order (mult_of R). phi' i)" |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1051 |
using sum_strict_mono_ex1[OF fin, of "?N" "\<lambda> i . phi' i"] |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1052 |
i1 all_le by auto |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1053 |
thus False using eq by force |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1054 |
qed |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1055 |
hence "?N (order (mult_of R)) > 0" using * by (simp add: phi'_nonzero) |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
1056 |
then obtain a where a: "a \<in> carrier (mult_of R)" and a_ord: "group.ord (mult_of R) a = order (mult_of R)" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1057 |
by (auto simp add: card_gt_0_iff) |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
1058 |
hence set_eq: "{a[^]i | i::nat. i \<in> UNIV} = (\<lambda>x. a[^]x) ` {0 .. group.ord (mult_of R) a - 1}" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1059 |
using G.ord_elems[OF finite'] by auto |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
1060 |
have card_eq: "card ((\<lambda>x. a[^]x) ` {0 .. group.ord (mult_of R) a - 1}) = card {0 .. group.ord (mult_of R) a - 1}" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1061 |
by (intro card_image G.ord_inj finite' a) |
67341
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents:
67299
diff
changeset
|
1062 |
hence "card ((\<lambda> x . a[^]x) ` {0 .. group.ord (mult_of R) a - 1}) = card {0 ..order (mult_of R) - 1}" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1063 |
using assms by (simp add: card_eq a_ord) |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
1064 |
hence card_R_minus_1: "card {a[^]i | i::nat. i \<in> UNIV} = order (mult_of R)" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1065 |
using * by (subst set_eq) auto |
70133
4f19b92ab6d7
tidying up messy proofs about group element order
paulson <lp15@cam.ac.uk>
parents:
70131
diff
changeset
|
1066 |
have **: "{a[^]i | i::nat. i \<in> UNIV} \<subseteq> carrier (mult_of R)" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1067 |
using G.nat_pow_closed[OF a] by auto |
67341
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents:
67299
diff
changeset
|
1068 |
with _ have "carrier (mult_of R) = {a[^]i|i::nat. i \<in> UNIV}" |
65416
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1069 |
by (rule card_seteq[symmetric]) (simp_all add: card_R_minus_1 finite order_def del: UNIV_I) |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1070 |
thus ?thesis using a by blast |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1071 |
qed |
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1072 |
|
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
haftmann
parents:
diff
changeset
|
1073 |
end |