author | haftmann |
Thu, 24 Aug 2017 17:41:49 +0200 | |
changeset 66501 | 5a42eddc11c1 |
parent 66453 | cc19f7ca2ed6 |
child 66579 | 2db3fe23fdaf |
permissions | -rw-r--r-- |
35849 | 1 |
(* Title: HOL/Algebra/Group.thy |
2 |
Author: Clemens Ballarin, started 4 February 2003 |
|
13813 | 3 |
|
4 |
Based on work by Florian Kammueller, L C Paulson and Markus Wenzel. |
|
5 |
*) |
|
6 |
||
28823 | 7 |
theory Group |
66501
5a42eddc11c1
swapping of theory dependency yields less pervasive syntax requiring popular symbols \<mu>, \<nu>
haftmann
parents:
66453
diff
changeset
|
8 |
imports Order "HOL-Library.FuncSet" |
28823 | 9 |
begin |
13813 | 10 |
|
61382 | 11 |
section \<open>Monoids and Groups\<close> |
13936 | 12 |
|
61382 | 13 |
subsection \<open>Definitions\<close> |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
19984
diff
changeset
|
14 |
|
61382 | 15 |
text \<open> |
58622 | 16 |
Definitions follow @{cite "Jacobson:1985"}. |
61382 | 17 |
\<close> |
13813 | 18 |
|
14963 | 19 |
record 'a monoid = "'a partial_object" + |
20 |
mult :: "['a, 'a] \<Rightarrow> 'a" (infixl "\<otimes>\<index>" 70) |
|
21 |
one :: 'a ("\<one>\<index>") |
|
13817 | 22 |
|
35847 | 23 |
definition |
14852 | 24 |
m_inv :: "('a, 'b) monoid_scheme => 'a => 'a" ("inv\<index> _" [81] 80) |
35848
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents:
35847
diff
changeset
|
25 |
where "inv\<^bsub>G\<^esub> x = (THE y. y \<in> carrier G & x \<otimes>\<^bsub>G\<^esub> y = \<one>\<^bsub>G\<^esub> & y \<otimes>\<^bsub>G\<^esub> x = \<one>\<^bsub>G\<^esub>)" |
13936 | 26 |
|
35847 | 27 |
definition |
14651 | 28 |
Units :: "_ => 'a set" |
63167 | 29 |
\<comment>\<open>The set of invertible elements\<close> |
35848
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents:
35847
diff
changeset
|
30 |
where "Units G = {y. y \<in> carrier G & (\<exists>x \<in> carrier G. x \<otimes>\<^bsub>G\<^esub> y = \<one>\<^bsub>G\<^esub> & y \<otimes>\<^bsub>G\<^esub> x = \<one>\<^bsub>G\<^esub>)}" |
13936 | 31 |
|
32 |
consts |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46559
diff
changeset
|
33 |
pow :: "[('a, 'm) monoid_scheme, 'a, 'b::semiring_1] => 'a" (infixr "'(^')\<index>" 75) |
35850 | 34 |
|
35 |
overloading nat_pow == "pow :: [_, 'a, nat] => 'a" |
|
36 |
begin |
|
55415 | 37 |
definition "nat_pow G a n = rec_nat \<one>\<^bsub>G\<^esub> (%u b. b \<otimes>\<^bsub>G\<^esub> a) n" |
35850 | 38 |
end |
13936 | 39 |
|
35850 | 40 |
overloading int_pow == "pow :: [_, 'a, int] => 'a" |
41 |
begin |
|
42 |
definition "int_pow G a z = |
|
55415 | 43 |
(let p = rec_nat \<one>\<^bsub>G\<^esub> (%u b. b \<otimes>\<^bsub>G\<^esub> a) |
46559 | 44 |
in if z < 0 then inv\<^bsub>G\<^esub> (p (nat (-z))) else p (nat z))" |
35850 | 45 |
end |
13813 | 46 |
|
61628 | 47 |
lemma int_pow_int: "x (^)\<^bsub>G\<^esub> (int n) = x (^)\<^bsub>G\<^esub> n" |
48 |
by(simp add: int_pow_def nat_pow_def) |
|
49 |
||
19783 | 50 |
locale monoid = |
51 |
fixes G (structure) |
|
13813 | 52 |
assumes m_closed [intro, simp]: |
14963 | 53 |
"\<lbrakk>x \<in> carrier G; y \<in> carrier G\<rbrakk> \<Longrightarrow> x \<otimes> y \<in> carrier G" |
54 |
and m_assoc: |
|
55 |
"\<lbrakk>x \<in> carrier G; y \<in> carrier G; z \<in> carrier G\<rbrakk> |
|
56 |
\<Longrightarrow> (x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)" |
|
57 |
and one_closed [intro, simp]: "\<one> \<in> carrier G" |
|
58 |
and l_one [simp]: "x \<in> carrier G \<Longrightarrow> \<one> \<otimes> x = x" |
|
59 |
and r_one [simp]: "x \<in> carrier G \<Longrightarrow> x \<otimes> \<one> = x" |
|
13817 | 60 |
|
13936 | 61 |
lemma monoidI: |
19783 | 62 |
fixes G (structure) |
13936 | 63 |
assumes m_closed: |
14693 | 64 |
"!!x y. [| x \<in> carrier G; y \<in> carrier G |] ==> x \<otimes> y \<in> carrier G" |
65 |
and one_closed: "\<one> \<in> carrier G" |
|
13936 | 66 |
and m_assoc: |
67 |
"!!x y z. [| x \<in> carrier G; y \<in> carrier G; z \<in> carrier G |] ==> |
|
14693 | 68 |
(x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)" |
69 |
and l_one: "!!x. x \<in> carrier G ==> \<one> \<otimes> x = x" |
|
70 |
and r_one: "!!x. x \<in> carrier G ==> x \<otimes> \<one> = x" |
|
13936 | 71 |
shows "monoid G" |
27714
27b4d7c01f8b
Tuned (for the sake of a meaningless log entry).
ballarin
parents:
27713
diff
changeset
|
72 |
by (fast intro!: monoid.intro intro: assms) |
13936 | 73 |
|
74 |
lemma (in monoid) Units_closed [dest]: |
|
75 |
"x \<in> Units G ==> x \<in> carrier G" |
|
76 |
by (unfold Units_def) fast |
|
77 |
||
78 |
lemma (in monoid) inv_unique: |
|
14693 | 79 |
assumes eq: "y \<otimes> x = \<one>" "x \<otimes> y' = \<one>" |
80 |
and G: "x \<in> carrier G" "y \<in> carrier G" "y' \<in> carrier G" |
|
13936 | 81 |
shows "y = y'" |
82 |
proof - |
|
83 |
from G eq have "y = y \<otimes> (x \<otimes> y')" by simp |
|
84 |
also from G have "... = (y \<otimes> x) \<otimes> y'" by (simp add: m_assoc) |
|
85 |
also from G eq have "... = y'" by simp |
|
86 |
finally show ?thesis . |
|
87 |
qed |
|
88 |
||
27698 | 89 |
lemma (in monoid) Units_m_closed [intro, simp]: |
90 |
assumes x: "x \<in> Units G" and y: "y \<in> Units G" |
|
91 |
shows "x \<otimes> y \<in> Units G" |
|
92 |
proof - |
|
93 |
from x obtain x' where x: "x \<in> carrier G" "x' \<in> carrier G" and xinv: "x \<otimes> x' = \<one>" "x' \<otimes> x = \<one>" |
|
94 |
unfolding Units_def by fast |
|
95 |
from y obtain y' where y: "y \<in> carrier G" "y' \<in> carrier G" and yinv: "y \<otimes> y' = \<one>" "y' \<otimes> y = \<one>" |
|
96 |
unfolding Units_def by fast |
|
97 |
from x y xinv yinv have "y' \<otimes> (x' \<otimes> x) \<otimes> y = \<one>" by simp |
|
98 |
moreover from x y xinv yinv have "x \<otimes> (y \<otimes> y') \<otimes> x' = \<one>" by simp |
|
99 |
moreover note x y |
|
100 |
ultimately show ?thesis unfolding Units_def |
|
63167 | 101 |
\<comment> "Must avoid premature use of \<open>hyp_subst_tac\<close>." |
27698 | 102 |
apply (rule_tac CollectI) |
103 |
apply (rule) |
|
104 |
apply (fast) |
|
105 |
apply (rule bexI [where x = "y' \<otimes> x'"]) |
|
106 |
apply (auto simp: m_assoc) |
|
107 |
done |
|
108 |
qed |
|
109 |
||
13940 | 110 |
lemma (in monoid) Units_one_closed [intro, simp]: |
111 |
"\<one> \<in> Units G" |
|
112 |
by (unfold Units_def) auto |
|
113 |
||
13936 | 114 |
lemma (in monoid) Units_inv_closed [intro, simp]: |
115 |
"x \<in> Units G ==> inv x \<in> carrier G" |
|
13943 | 116 |
apply (unfold Units_def m_inv_def, auto) |
13936 | 117 |
apply (rule theI2, fast) |
13943 | 118 |
apply (fast intro: inv_unique, fast) |
13936 | 119 |
done |
120 |
||
19981 | 121 |
lemma (in monoid) Units_l_inv_ex: |
122 |
"x \<in> Units G ==> \<exists>y \<in> carrier G. y \<otimes> x = \<one>" |
|
123 |
by (unfold Units_def) auto |
|
124 |
||
125 |
lemma (in monoid) Units_r_inv_ex: |
|
126 |
"x \<in> Units G ==> \<exists>y \<in> carrier G. x \<otimes> y = \<one>" |
|
127 |
by (unfold Units_def) auto |
|
128 |
||
27698 | 129 |
lemma (in monoid) Units_l_inv [simp]: |
13936 | 130 |
"x \<in> Units G ==> inv x \<otimes> x = \<one>" |
13943 | 131 |
apply (unfold Units_def m_inv_def, auto) |
13936 | 132 |
apply (rule theI2, fast) |
13943 | 133 |
apply (fast intro: inv_unique, fast) |
13936 | 134 |
done |
135 |
||
27698 | 136 |
lemma (in monoid) Units_r_inv [simp]: |
13936 | 137 |
"x \<in> Units G ==> x \<otimes> inv x = \<one>" |
13943 | 138 |
apply (unfold Units_def m_inv_def, auto) |
13936 | 139 |
apply (rule theI2, fast) |
13943 | 140 |
apply (fast intro: inv_unique, fast) |
13936 | 141 |
done |
142 |
||
143 |
lemma (in monoid) Units_inv_Units [intro, simp]: |
|
144 |
"x \<in> Units G ==> inv x \<in> Units G" |
|
145 |
proof - |
|
146 |
assume x: "x \<in> Units G" |
|
147 |
show "inv x \<in> Units G" |
|
148 |
by (auto simp add: Units_def |
|
149 |
intro: Units_l_inv Units_r_inv x Units_closed [OF x]) |
|
150 |
qed |
|
151 |
||
152 |
lemma (in monoid) Units_l_cancel [simp]: |
|
153 |
"[| x \<in> Units G; y \<in> carrier G; z \<in> carrier G |] ==> |
|
154 |
(x \<otimes> y = x \<otimes> z) = (y = z)" |
|
155 |
proof |
|
156 |
assume eq: "x \<otimes> y = x \<otimes> z" |
|
14693 | 157 |
and G: "x \<in> Units G" "y \<in> carrier G" "z \<in> carrier G" |
13936 | 158 |
then have "(inv x \<otimes> x) \<otimes> y = (inv x \<otimes> x) \<otimes> z" |
27698 | 159 |
by (simp add: m_assoc Units_closed del: Units_l_inv) |
44472 | 160 |
with G show "y = z" by simp |
13936 | 161 |
next |
162 |
assume eq: "y = z" |
|
14693 | 163 |
and G: "x \<in> Units G" "y \<in> carrier G" "z \<in> carrier G" |
13936 | 164 |
then show "x \<otimes> y = x \<otimes> z" by simp |
165 |
qed |
|
166 |
||
167 |
lemma (in monoid) Units_inv_inv [simp]: |
|
168 |
"x \<in> Units G ==> inv (inv x) = x" |
|
169 |
proof - |
|
170 |
assume x: "x \<in> Units G" |
|
27698 | 171 |
then have "inv x \<otimes> inv (inv x) = inv x \<otimes> x" by simp |
172 |
with x show ?thesis by (simp add: Units_closed del: Units_l_inv Units_r_inv) |
|
13936 | 173 |
qed |
174 |
||
175 |
lemma (in monoid) inv_inj_on_Units: |
|
176 |
"inj_on (m_inv G) (Units G)" |
|
177 |
proof (rule inj_onI) |
|
178 |
fix x y |
|
14693 | 179 |
assume G: "x \<in> Units G" "y \<in> Units G" and eq: "inv x = inv y" |
13936 | 180 |
then have "inv (inv x) = inv (inv y)" by simp |
181 |
with G show "x = y" by simp |
|
182 |
qed |
|
183 |
||
13940 | 184 |
lemma (in monoid) Units_inv_comm: |
185 |
assumes inv: "x \<otimes> y = \<one>" |
|
14693 | 186 |
and G: "x \<in> Units G" "y \<in> Units G" |
13940 | 187 |
shows "y \<otimes> x = \<one>" |
188 |
proof - |
|
189 |
from G have "x \<otimes> y \<otimes> x = x \<otimes> \<one>" by (auto simp add: inv Units_closed) |
|
190 |
with G show ?thesis by (simp del: r_one add: m_assoc Units_closed) |
|
191 |
qed |
|
192 |
||
61628 | 193 |
lemma (in monoid) carrier_not_empty: "carrier G \<noteq> {}" |
194 |
by auto |
|
195 |
||
61382 | 196 |
text \<open>Power\<close> |
13936 | 197 |
|
198 |
lemma (in monoid) nat_pow_closed [intro, simp]: |
|
199 |
"x \<in> carrier G ==> x (^) (n::nat) \<in> carrier G" |
|
200 |
by (induct n) (simp_all add: nat_pow_def) |
|
201 |
||
202 |
lemma (in monoid) nat_pow_0 [simp]: |
|
203 |
"x (^) (0::nat) = \<one>" |
|
204 |
by (simp add: nat_pow_def) |
|
205 |
||
206 |
lemma (in monoid) nat_pow_Suc [simp]: |
|
207 |
"x (^) (Suc n) = x (^) n \<otimes> x" |
|
208 |
by (simp add: nat_pow_def) |
|
209 |
||
210 |
lemma (in monoid) nat_pow_one [simp]: |
|
211 |
"\<one> (^) (n::nat) = \<one>" |
|
212 |
by (induct n) simp_all |
|
213 |
||
214 |
lemma (in monoid) nat_pow_mult: |
|
215 |
"x \<in> carrier G ==> x (^) (n::nat) \<otimes> x (^) m = x (^) (n + m)" |
|
216 |
by (induct m) (simp_all add: m_assoc [THEN sym]) |
|
217 |
||
218 |
lemma (in monoid) nat_pow_pow: |
|
219 |
"x \<in> carrier G ==> (x (^) n) (^) m = x (^) (n * m::nat)" |
|
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
57271
diff
changeset
|
220 |
by (induct m) (simp, simp add: nat_pow_mult add.commute) |
13936 | 221 |
|
27698 | 222 |
|
223 |
(* Jacobson defines submonoid here. *) |
|
224 |
(* Jacobson defines the order of a monoid here. *) |
|
225 |
||
226 |
||
61382 | 227 |
subsection \<open>Groups\<close> |
27698 | 228 |
|
61382 | 229 |
text \<open> |
13936 | 230 |
A group is a monoid all of whose elements are invertible. |
61382 | 231 |
\<close> |
13936 | 232 |
|
233 |
locale group = monoid + |
|
234 |
assumes Units: "carrier G <= Units G" |
|
235 |
||
26199 | 236 |
lemma (in group) is_group: "group G" by (rule group_axioms) |
14761 | 237 |
|
13936 | 238 |
theorem groupI: |
19783 | 239 |
fixes G (structure) |
13936 | 240 |
assumes m_closed [simp]: |
14693 | 241 |
"!!x y. [| x \<in> carrier G; y \<in> carrier G |] ==> x \<otimes> y \<in> carrier G" |
242 |
and one_closed [simp]: "\<one> \<in> carrier G" |
|
13936 | 243 |
and m_assoc: |
244 |
"!!x y z. [| x \<in> carrier G; y \<in> carrier G; z \<in> carrier G |] ==> |
|
14693 | 245 |
(x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)" |
246 |
and l_one [simp]: "!!x. x \<in> carrier G ==> \<one> \<otimes> x = x" |
|
14963 | 247 |
and l_inv_ex: "!!x. x \<in> carrier G ==> \<exists>y \<in> carrier G. y \<otimes> x = \<one>" |
13936 | 248 |
shows "group G" |
249 |
proof - |
|
250 |
have l_cancel [simp]: |
|
251 |
"!!x y z. [| x \<in> carrier G; y \<in> carrier G; z \<in> carrier G |] ==> |
|
14693 | 252 |
(x \<otimes> y = x \<otimes> z) = (y = z)" |
13936 | 253 |
proof |
254 |
fix x y z |
|
14693 | 255 |
assume eq: "x \<otimes> y = x \<otimes> z" |
256 |
and G: "x \<in> carrier G" "y \<in> carrier G" "z \<in> carrier G" |
|
13936 | 257 |
with l_inv_ex obtain x_inv where xG: "x_inv \<in> carrier G" |
14693 | 258 |
and l_inv: "x_inv \<otimes> x = \<one>" by fast |
259 |
from G eq xG have "(x_inv \<otimes> x) \<otimes> y = (x_inv \<otimes> x) \<otimes> z" |
|
13936 | 260 |
by (simp add: m_assoc) |
261 |
with G show "y = z" by (simp add: l_inv) |
|
262 |
next |
|
263 |
fix x y z |
|
264 |
assume eq: "y = z" |
|
14693 | 265 |
and G: "x \<in> carrier G" "y \<in> carrier G" "z \<in> carrier G" |
266 |
then show "x \<otimes> y = x \<otimes> z" by simp |
|
13936 | 267 |
qed |
268 |
have r_one: |
|
14693 | 269 |
"!!x. x \<in> carrier G ==> x \<otimes> \<one> = x" |
13936 | 270 |
proof - |
271 |
fix x |
|
272 |
assume x: "x \<in> carrier G" |
|
273 |
with l_inv_ex obtain x_inv where xG: "x_inv \<in> carrier G" |
|
14693 | 274 |
and l_inv: "x_inv \<otimes> x = \<one>" by fast |
275 |
from x xG have "x_inv \<otimes> (x \<otimes> \<one>) = x_inv \<otimes> x" |
|
13936 | 276 |
by (simp add: m_assoc [symmetric] l_inv) |
14693 | 277 |
with x xG show "x \<otimes> \<one> = x" by simp |
13936 | 278 |
qed |
279 |
have inv_ex: |
|
14963 | 280 |
"!!x. x \<in> carrier G ==> \<exists>y \<in> carrier G. y \<otimes> x = \<one> & x \<otimes> y = \<one>" |
13936 | 281 |
proof - |
282 |
fix x |
|
283 |
assume x: "x \<in> carrier G" |
|
284 |
with l_inv_ex obtain y where y: "y \<in> carrier G" |
|
14693 | 285 |
and l_inv: "y \<otimes> x = \<one>" by fast |
286 |
from x y have "y \<otimes> (x \<otimes> y) = y \<otimes> \<one>" |
|
13936 | 287 |
by (simp add: m_assoc [symmetric] l_inv r_one) |
14693 | 288 |
with x y have r_inv: "x \<otimes> y = \<one>" |
13936 | 289 |
by simp |
14963 | 290 |
from x y show "\<exists>y \<in> carrier G. y \<otimes> x = \<one> & x \<otimes> y = \<one>" |
13936 | 291 |
by (fast intro: l_inv r_inv) |
292 |
qed |
|
293 |
then have carrier_subset_Units: "carrier G <= Units G" |
|
294 |
by (unfold Units_def) fast |
|
61169 | 295 |
show ?thesis |
296 |
by standard (auto simp: r_one m_assoc carrier_subset_Units) |
|
13936 | 297 |
qed |
298 |
||
27698 | 299 |
lemma (in monoid) group_l_invI: |
13936 | 300 |
assumes l_inv_ex: |
14963 | 301 |
"!!x. x \<in> carrier G ==> \<exists>y \<in> carrier G. y \<otimes> x = \<one>" |
13936 | 302 |
shows "group G" |
303 |
by (rule groupI) (auto intro: m_assoc l_inv_ex) |
|
304 |
||
305 |
lemma (in group) Units_eq [simp]: |
|
306 |
"Units G = carrier G" |
|
307 |
proof |
|
308 |
show "Units G <= carrier G" by fast |
|
309 |
next |
|
310 |
show "carrier G <= Units G" by (rule Units) |
|
311 |
qed |
|
312 |
||
313 |
lemma (in group) inv_closed [intro, simp]: |
|
314 |
"x \<in> carrier G ==> inv x \<in> carrier G" |
|
315 |
using Units_inv_closed by simp |
|
316 |
||
19981 | 317 |
lemma (in group) l_inv_ex [simp]: |
318 |
"x \<in> carrier G ==> \<exists>y \<in> carrier G. y \<otimes> x = \<one>" |
|
319 |
using Units_l_inv_ex by simp |
|
320 |
||
321 |
lemma (in group) r_inv_ex [simp]: |
|
322 |
"x \<in> carrier G ==> \<exists>y \<in> carrier G. x \<otimes> y = \<one>" |
|
323 |
using Units_r_inv_ex by simp |
|
324 |
||
14963 | 325 |
lemma (in group) l_inv [simp]: |
13936 | 326 |
"x \<in> carrier G ==> inv x \<otimes> x = \<one>" |
327 |
using Units_l_inv by simp |
|
13813 | 328 |
|
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
19984
diff
changeset
|
329 |
|
61382 | 330 |
subsection \<open>Cancellation Laws and Basic Properties\<close> |
13813 | 331 |
|
332 |
lemma (in group) l_cancel [simp]: |
|
333 |
"[| x \<in> carrier G; y \<in> carrier G; z \<in> carrier G |] ==> |
|
334 |
(x \<otimes> y = x \<otimes> z) = (y = z)" |
|
13936 | 335 |
using Units_l_inv by simp |
13940 | 336 |
|
14963 | 337 |
lemma (in group) r_inv [simp]: |
13813 | 338 |
"x \<in> carrier G ==> x \<otimes> inv x = \<one>" |
339 |
proof - |
|
340 |
assume x: "x \<in> carrier G" |
|
341 |
then have "inv x \<otimes> (x \<otimes> inv x) = inv x \<otimes> \<one>" |
|
44472 | 342 |
by (simp add: m_assoc [symmetric]) |
13813 | 343 |
with x show ?thesis by (simp del: r_one) |
344 |
qed |
|
345 |
||
346 |
lemma (in group) r_cancel [simp]: |
|
347 |
"[| x \<in> carrier G; y \<in> carrier G; z \<in> carrier G |] ==> |
|
348 |
(y \<otimes> x = z \<otimes> x) = (y = z)" |
|
349 |
proof |
|
350 |
assume eq: "y \<otimes> x = z \<otimes> x" |
|
14693 | 351 |
and G: "x \<in> carrier G" "y \<in> carrier G" "z \<in> carrier G" |
13813 | 352 |
then have "y \<otimes> (x \<otimes> inv x) = z \<otimes> (x \<otimes> inv x)" |
27698 | 353 |
by (simp add: m_assoc [symmetric] del: r_inv Units_r_inv) |
14963 | 354 |
with G show "y = z" by simp |
13813 | 355 |
next |
356 |
assume eq: "y = z" |
|
14693 | 357 |
and G: "x \<in> carrier G" "y \<in> carrier G" "z \<in> carrier G" |
13813 | 358 |
then show "y \<otimes> x = z \<otimes> x" by simp |
359 |
qed |
|
360 |
||
13854
91c9ab25fece
First distributed version of Group and Ring theory.
ballarin
parents:
13835
diff
changeset
|
361 |
lemma (in group) inv_one [simp]: |
91c9ab25fece
First distributed version of Group and Ring theory.
ballarin
parents:
13835
diff
changeset
|
362 |
"inv \<one> = \<one>" |
91c9ab25fece
First distributed version of Group and Ring theory.
ballarin
parents:
13835
diff
changeset
|
363 |
proof - |
27698 | 364 |
have "inv \<one> = \<one> \<otimes> (inv \<one>)" by (simp del: r_inv Units_r_inv) |
14963 | 365 |
moreover have "... = \<one>" by simp |
13854
91c9ab25fece
First distributed version of Group and Ring theory.
ballarin
parents:
13835
diff
changeset
|
366 |
finally show ?thesis . |
91c9ab25fece
First distributed version of Group and Ring theory.
ballarin
parents:
13835
diff
changeset
|
367 |
qed |
91c9ab25fece
First distributed version of Group and Ring theory.
ballarin
parents:
13835
diff
changeset
|
368 |
|
13813 | 369 |
lemma (in group) inv_inv [simp]: |
370 |
"x \<in> carrier G ==> inv (inv x) = x" |
|
13936 | 371 |
using Units_inv_inv by simp |
372 |
||
373 |
lemma (in group) inv_inj: |
|
374 |
"inj_on (m_inv G) (carrier G)" |
|
375 |
using inv_inj_on_Units by simp |
|
13813 | 376 |
|
13854
91c9ab25fece
First distributed version of Group and Ring theory.
ballarin
parents:
13835
diff
changeset
|
377 |
lemma (in group) inv_mult_group: |
13813 | 378 |
"[| x \<in> carrier G; y \<in> carrier G |] ==> inv (x \<otimes> y) = inv y \<otimes> inv x" |
379 |
proof - |
|
14693 | 380 |
assume G: "x \<in> carrier G" "y \<in> carrier G" |
13813 | 381 |
then have "inv (x \<otimes> y) \<otimes> (x \<otimes> y) = (inv y \<otimes> inv x) \<otimes> (x \<otimes> y)" |
44472 | 382 |
by (simp add: m_assoc) (simp add: m_assoc [symmetric]) |
27698 | 383 |
with G show ?thesis by (simp del: l_inv Units_l_inv) |
13813 | 384 |
qed |
385 |
||
13940 | 386 |
lemma (in group) inv_comm: |
387 |
"[| x \<otimes> y = \<one>; x \<in> carrier G; y \<in> carrier G |] ==> y \<otimes> x = \<one>" |
|
14693 | 388 |
by (rule Units_inv_comm) auto |
13940 | 389 |
|
13944 | 390 |
lemma (in group) inv_equality: |
13943 | 391 |
"[|y \<otimes> x = \<one>; x \<in> carrier G; y \<in> carrier G|] ==> inv x = y" |
392 |
apply (simp add: m_inv_def) |
|
393 |
apply (rule the_equality) |
|
14693 | 394 |
apply (simp add: inv_comm [of y x]) |
395 |
apply (rule r_cancel [THEN iffD1], auto) |
|
13943 | 396 |
done |
397 |
||
57271 | 398 |
(* Contributed by Joachim Breitner *) |
399 |
lemma (in group) inv_solve_left: |
|
400 |
"\<lbrakk> a \<in> carrier G; b \<in> carrier G; c \<in> carrier G \<rbrakk> \<Longrightarrow> a = inv b \<otimes> c \<longleftrightarrow> c = b \<otimes> a" |
|
401 |
by (metis inv_equality l_inv_ex l_one m_assoc r_inv) |
|
402 |
lemma (in group) inv_solve_right: |
|
403 |
"\<lbrakk> a \<in> carrier G; b \<in> carrier G; c \<in> carrier G \<rbrakk> \<Longrightarrow> a = b \<otimes> inv c \<longleftrightarrow> b = a \<otimes> c" |
|
404 |
by (metis inv_equality l_inv_ex l_one m_assoc r_inv) |
|
405 |
||
61382 | 406 |
text \<open>Power\<close> |
13936 | 407 |
|
408 |
lemma (in group) int_pow_def2: |
|
46559 | 409 |
"a (^) (z::int) = (if z < 0 then inv (a (^) (nat (-z))) else a (^) (nat z))" |
13936 | 410 |
by (simp add: int_pow_def nat_pow_def Let_def) |
411 |
||
412 |
lemma (in group) int_pow_0 [simp]: |
|
413 |
"x (^) (0::int) = \<one>" |
|
414 |
by (simp add: int_pow_def2) |
|
415 |
||
416 |
lemma (in group) int_pow_one [simp]: |
|
417 |
"\<one> (^) (z::int) = \<one>" |
|
418 |
by (simp add: int_pow_def2) |
|
419 |
||
57271 | 420 |
(* The following are contributed by Joachim Breitner *) |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
19984
diff
changeset
|
421 |
|
57271 | 422 |
lemma (in group) int_pow_closed [intro, simp]: |
423 |
"x \<in> carrier G ==> x (^) (i::int) \<in> carrier G" |
|
424 |
by (simp add: int_pow_def2) |
|
425 |
||
426 |
lemma (in group) int_pow_1 [simp]: |
|
427 |
"x \<in> carrier G \<Longrightarrow> x (^) (1::int) = x" |
|
428 |
by (simp add: int_pow_def2) |
|
429 |
||
430 |
lemma (in group) int_pow_neg: |
|
431 |
"x \<in> carrier G \<Longrightarrow> x (^) (-i::int) = inv (x (^) i)" |
|
432 |
by (simp add: int_pow_def2) |
|
433 |
||
434 |
lemma (in group) int_pow_mult: |
|
435 |
"x \<in> carrier G \<Longrightarrow> x (^) (i + j::int) = x (^) i \<otimes> x (^) j" |
|
436 |
proof - |
|
437 |
have [simp]: "-i - j = -j - i" by simp |
|
438 |
assume "x : carrier G" then |
|
439 |
show ?thesis |
|
440 |
by (auto simp add: int_pow_def2 inv_solve_left inv_solve_right nat_add_distrib [symmetric] nat_pow_mult ) |
|
441 |
qed |
|
442 |
||
61628 | 443 |
lemma (in group) int_pow_diff: |
444 |
"x \<in> carrier G \<Longrightarrow> x (^) (n - m :: int) = x (^) n \<otimes> inv (x (^) m)" |
|
445 |
by(simp only: diff_conv_add_uminus int_pow_mult int_pow_neg) |
|
446 |
||
447 |
lemma (in group) inj_on_multc: "c \<in> carrier G \<Longrightarrow> inj_on (\<lambda>x. x \<otimes> c) (carrier G)" |
|
448 |
by(simp add: inj_on_def) |
|
449 |
||
450 |
lemma (in group) inj_on_cmult: "c \<in> carrier G \<Longrightarrow> inj_on (\<lambda>x. c \<otimes> x) (carrier G)" |
|
451 |
by(simp add: inj_on_def) |
|
452 |
||
61382 | 453 |
subsection \<open>Subgroups\<close> |
13813 | 454 |
|
19783 | 455 |
locale subgroup = |
456 |
fixes H and G (structure) |
|
14963 | 457 |
assumes subset: "H \<subseteq> carrier G" |
458 |
and m_closed [intro, simp]: "\<lbrakk>x \<in> H; y \<in> H\<rbrakk> \<Longrightarrow> x \<otimes> y \<in> H" |
|
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
19984
diff
changeset
|
459 |
and one_closed [simp]: "\<one> \<in> H" |
14963 | 460 |
and m_inv_closed [intro,simp]: "x \<in> H \<Longrightarrow> inv x \<in> H" |
13813 | 461 |
|
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
19984
diff
changeset
|
462 |
lemma (in subgroup) is_subgroup: |
26199 | 463 |
"subgroup H G" by (rule subgroup_axioms) |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
19984
diff
changeset
|
464 |
|
13813 | 465 |
declare (in subgroup) group.intro [intro] |
13949
0ce528cd6f19
HOL-Algebra complete for release Isabelle2003 (modulo section headers).
ballarin
parents:
13944
diff
changeset
|
466 |
|
14963 | 467 |
lemma (in subgroup) mem_carrier [simp]: |
468 |
"x \<in> H \<Longrightarrow> x \<in> carrier G" |
|
469 |
using subset by blast |
|
13813 | 470 |
|
14963 | 471 |
lemma subgroup_imp_subset: |
472 |
"subgroup H G \<Longrightarrow> H \<subseteq> carrier G" |
|
473 |
by (rule subgroup.subset) |
|
474 |
||
475 |
lemma (in subgroup) subgroup_is_group [intro]: |
|
27611 | 476 |
assumes "group G" |
477 |
shows "group (G\<lparr>carrier := H\<rparr>)" |
|
478 |
proof - |
|
29237 | 479 |
interpret group G by fact |
27611 | 480 |
show ?thesis |
27698 | 481 |
apply (rule monoid.group_l_invI) |
482 |
apply (unfold_locales) [1] |
|
483 |
apply (auto intro: m_assoc l_inv mem_carrier) |
|
484 |
done |
|
27611 | 485 |
qed |
13813 | 486 |
|
61382 | 487 |
text \<open> |
13813 | 488 |
Since @{term H} is nonempty, it contains some element @{term x}. Since |
63167 | 489 |
it is closed under inverse, it contains \<open>inv x\<close>. Since |
490 |
it is closed under product, it contains \<open>x \<otimes> inv x = \<one>\<close>. |
|
61382 | 491 |
\<close> |
13813 | 492 |
|
493 |
lemma (in group) one_in_subset: |
|
494 |
"[| H \<subseteq> carrier G; H \<noteq> {}; \<forall>a \<in> H. inv a \<in> H; \<forall>a\<in>H. \<forall>b\<in>H. a \<otimes> b \<in> H |] |
|
495 |
==> \<one> \<in> H" |
|
44472 | 496 |
by force |
13813 | 497 |
|
61382 | 498 |
text \<open>A characterization of subgroups: closed, non-empty subset.\<close> |
13813 | 499 |
|
500 |
lemma (in group) subgroupI: |
|
501 |
assumes subset: "H \<subseteq> carrier G" and non_empty: "H \<noteq> {}" |
|
14963 | 502 |
and inv: "!!a. a \<in> H \<Longrightarrow> inv a \<in> H" |
503 |
and mult: "!!a b. \<lbrakk>a \<in> H; b \<in> H\<rbrakk> \<Longrightarrow> a \<otimes> b \<in> H" |
|
13813 | 504 |
shows "subgroup H G" |
27714
27b4d7c01f8b
Tuned (for the sake of a meaningless log entry).
ballarin
parents:
27713
diff
changeset
|
505 |
proof (simp add: subgroup_def assms) |
27b4d7c01f8b
Tuned (for the sake of a meaningless log entry).
ballarin
parents:
27713
diff
changeset
|
506 |
show "\<one> \<in> H" by (rule one_in_subset) (auto simp only: assms) |
13813 | 507 |
qed |
508 |
||
13936 | 509 |
declare monoid.one_closed [iff] group.inv_closed [simp] |
510 |
monoid.l_one [simp] monoid.r_one [simp] group.inv_inv [simp] |
|
13813 | 511 |
|
512 |
lemma subgroup_nonempty: |
|
513 |
"~ subgroup {} G" |
|
514 |
by (blast dest: subgroup.one_closed) |
|
515 |
||
516 |
lemma (in subgroup) finite_imp_card_positive: |
|
517 |
"finite (carrier G) ==> 0 < card H" |
|
518 |
proof (rule classical) |
|
41528 | 519 |
assume "finite (carrier G)" and a: "~ 0 < card H" |
14963 | 520 |
then have "finite H" by (blast intro: finite_subset [OF subset]) |
41528 | 521 |
with is_subgroup a have "subgroup {} G" by simp |
13813 | 522 |
with subgroup_nonempty show ?thesis by contradiction |
523 |
qed |
|
524 |
||
13936 | 525 |
(* |
526 |
lemma (in monoid) Units_subgroup: |
|
527 |
"subgroup (Units G) G" |
|
528 |
*) |
|
529 |
||
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
19984
diff
changeset
|
530 |
|
61382 | 531 |
subsection \<open>Direct Products\<close> |
13813 | 532 |
|
35848
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents:
35847
diff
changeset
|
533 |
definition |
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents:
35847
diff
changeset
|
534 |
DirProd :: "_ \<Rightarrow> _ \<Rightarrow> ('a \<times> 'b) monoid" (infixr "\<times>\<times>" 80) where |
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents:
35847
diff
changeset
|
535 |
"G \<times>\<times> H = |
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents:
35847
diff
changeset
|
536 |
\<lparr>carrier = carrier G \<times> carrier H, |
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents:
35847
diff
changeset
|
537 |
mult = (\<lambda>(g, h) (g', h'). (g \<otimes>\<^bsub>G\<^esub> g', h \<otimes>\<^bsub>H\<^esub> h')), |
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents:
35847
diff
changeset
|
538 |
one = (\<one>\<^bsub>G\<^esub>, \<one>\<^bsub>H\<^esub>)\<rparr>" |
13813 | 539 |
|
14963 | 540 |
lemma DirProd_monoid: |
27611 | 541 |
assumes "monoid G" and "monoid H" |
14963 | 542 |
shows "monoid (G \<times>\<times> H)" |
543 |
proof - |
|
30729
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
wenzelm
parents:
29240
diff
changeset
|
544 |
interpret G: monoid G by fact |
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
wenzelm
parents:
29240
diff
changeset
|
545 |
interpret H: monoid H by fact |
27714
27b4d7c01f8b
Tuned (for the sake of a meaningless log entry).
ballarin
parents:
27713
diff
changeset
|
546 |
from assms |
14963 | 547 |
show ?thesis by (unfold monoid_def DirProd_def, auto) |
548 |
qed |
|
13813 | 549 |
|
550 |
||
61382 | 551 |
text\<open>Does not use the previous result because it's easier just to use auto.\<close> |
14963 | 552 |
lemma DirProd_group: |
27611 | 553 |
assumes "group G" and "group H" |
14963 | 554 |
shows "group (G \<times>\<times> H)" |
27611 | 555 |
proof - |
30729
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
wenzelm
parents:
29240
diff
changeset
|
556 |
interpret G: group G by fact |
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
wenzelm
parents:
29240
diff
changeset
|
557 |
interpret H: group H by fact |
27611 | 558 |
show ?thesis by (rule groupI) |
14963 | 559 |
(auto intro: G.m_assoc H.m_assoc G.l_inv H.l_inv |
560 |
simp add: DirProd_def) |
|
27611 | 561 |
qed |
13813 | 562 |
|
14963 | 563 |
lemma carrier_DirProd [simp]: |
564 |
"carrier (G \<times>\<times> H) = carrier G \<times> carrier H" |
|
565 |
by (simp add: DirProd_def) |
|
13944 | 566 |
|
14963 | 567 |
lemma one_DirProd [simp]: |
568 |
"\<one>\<^bsub>G \<times>\<times> H\<^esub> = (\<one>\<^bsub>G\<^esub>, \<one>\<^bsub>H\<^esub>)" |
|
569 |
by (simp add: DirProd_def) |
|
13944 | 570 |
|
14963 | 571 |
lemma mult_DirProd [simp]: |
572 |
"(g, h) \<otimes>\<^bsub>(G \<times>\<times> H)\<^esub> (g', h') = (g \<otimes>\<^bsub>G\<^esub> g', h \<otimes>\<^bsub>H\<^esub> h')" |
|
573 |
by (simp add: DirProd_def) |
|
13944 | 574 |
|
14963 | 575 |
lemma inv_DirProd [simp]: |
27611 | 576 |
assumes "group G" and "group H" |
13944 | 577 |
assumes g: "g \<in> carrier G" |
578 |
and h: "h \<in> carrier H" |
|
14963 | 579 |
shows "m_inv (G \<times>\<times> H) (g, h) = (inv\<^bsub>G\<^esub> g, inv\<^bsub>H\<^esub> h)" |
27611 | 580 |
proof - |
30729
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
wenzelm
parents:
29240
diff
changeset
|
581 |
interpret G: group G by fact |
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
wenzelm
parents:
29240
diff
changeset
|
582 |
interpret H: group H by fact |
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
wenzelm
parents:
29240
diff
changeset
|
583 |
interpret Prod: group "G \<times>\<times> H" |
27714
27b4d7c01f8b
Tuned (for the sake of a meaningless log entry).
ballarin
parents:
27713
diff
changeset
|
584 |
by (auto intro: DirProd_group group.intro group.axioms assms) |
14963 | 585 |
show ?thesis by (simp add: Prod.inv_equality g h) |
586 |
qed |
|
27698 | 587 |
|
14963 | 588 |
|
61382 | 589 |
subsection \<open>Homomorphisms and Isomorphisms\<close> |
13813 | 590 |
|
35847 | 591 |
definition |
592 |
hom :: "_ => _ => ('a => 'b) set" where |
|
35848
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents:
35847
diff
changeset
|
593 |
"hom G H = |
61384 | 594 |
{h. h \<in> carrier G \<rightarrow> carrier H & |
14693 | 595 |
(\<forall>x \<in> carrier G. \<forall>y \<in> carrier G. h (x \<otimes>\<^bsub>G\<^esub> y) = h x \<otimes>\<^bsub>H\<^esub> h y)}" |
13813 | 596 |
|
14761 | 597 |
lemma (in group) hom_compose: |
31754 | 598 |
"[|h \<in> hom G H; i \<in> hom H I|] ==> compose (carrier G) i h \<in> hom G I" |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44655
diff
changeset
|
599 |
by (fastforce simp add: hom_def compose_def) |
13943 | 600 |
|
35848
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents:
35847
diff
changeset
|
601 |
definition |
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents:
35847
diff
changeset
|
602 |
iso :: "_ => _ => ('a => 'b) set" (infixr "\<cong>" 60) |
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents:
35847
diff
changeset
|
603 |
where "G \<cong> H = {h. h \<in> hom G H & bij_betw h (carrier G) (carrier H)}" |
14761 | 604 |
|
14803 | 605 |
lemma iso_refl: "(%x. x) \<in> G \<cong> G" |
31727 | 606 |
by (simp add: iso_def hom_def inj_on_def bij_betw_def Pi_def) |
14761 | 607 |
|
608 |
lemma (in group) iso_sym: |
|
33057 | 609 |
"h \<in> G \<cong> H \<Longrightarrow> inv_into (carrier G) h \<in> H \<cong> G" |
610 |
apply (simp add: iso_def bij_betw_inv_into) |
|
611 |
apply (subgoal_tac "inv_into (carrier G) h \<in> carrier H \<rightarrow> carrier G") |
|
612 |
prefer 2 apply (simp add: bij_betw_imp_funcset [OF bij_betw_inv_into]) |
|
613 |
apply (simp add: hom_def bij_betw_def inv_into_f_eq f_inv_into_f Pi_def) |
|
14761 | 614 |
done |
615 |
||
616 |
lemma (in group) iso_trans: |
|
14803 | 617 |
"[|h \<in> G \<cong> H; i \<in> H \<cong> I|] ==> (compose (carrier G) i h) \<in> G \<cong> I" |
14761 | 618 |
by (auto simp add: iso_def hom_compose bij_betw_compose) |
619 |
||
14963 | 620 |
lemma DirProd_commute_iso: |
621 |
shows "(\<lambda>(x,y). (y,x)) \<in> (G \<times>\<times> H) \<cong> (H \<times>\<times> G)" |
|
31754 | 622 |
by (auto simp add: iso_def hom_def inj_on_def bij_betw_def) |
14761 | 623 |
|
14963 | 624 |
lemma DirProd_assoc_iso: |
625 |
shows "(\<lambda>(x,y,z). (x,(y,z))) \<in> (G \<times>\<times> H \<times>\<times> I) \<cong> (G \<times>\<times> (H \<times>\<times> I))" |
|
31727 | 626 |
by (auto simp add: iso_def hom_def inj_on_def bij_betw_def) |
14761 | 627 |
|
628 |
||
61382 | 629 |
text\<open>Basis for homomorphism proofs: we assume two groups @{term G} and |
630 |
@{term H}, with a homomorphism @{term h} between them\<close> |
|
61565
352c73a689da
Qualifiers in locale expressions default to mandatory regardless of the command.
ballarin
parents:
61384
diff
changeset
|
631 |
locale group_hom = G?: group G + H?: group H for G (structure) and H (structure) + |
29237 | 632 |
fixes h |
13813 | 633 |
assumes homh: "h \<in> hom G H" |
29240 | 634 |
|
635 |
lemma (in group_hom) hom_mult [simp]: |
|
636 |
"[| x \<in> carrier G; y \<in> carrier G |] ==> h (x \<otimes>\<^bsub>G\<^esub> y) = h x \<otimes>\<^bsub>H\<^esub> h y" |
|
637 |
proof - |
|
638 |
assume "x \<in> carrier G" "y \<in> carrier G" |
|
639 |
with homh [unfolded hom_def] show ?thesis by simp |
|
640 |
qed |
|
641 |
||
642 |
lemma (in group_hom) hom_closed [simp]: |
|
643 |
"x \<in> carrier G ==> h x \<in> carrier H" |
|
644 |
proof - |
|
645 |
assume "x \<in> carrier G" |
|
31754 | 646 |
with homh [unfolded hom_def] show ?thesis by auto |
29240 | 647 |
qed |
13813 | 648 |
|
649 |
lemma (in group_hom) one_closed [simp]: |
|
650 |
"h \<one> \<in> carrier H" |
|
651 |
by simp |
|
652 |
||
653 |
lemma (in group_hom) hom_one [simp]: |
|
14693 | 654 |
"h \<one> = \<one>\<^bsub>H\<^esub>" |
13813 | 655 |
proof - |
15076
4b3d280ef06a
New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
14963
diff
changeset
|
656 |
have "h \<one> \<otimes>\<^bsub>H\<^esub> \<one>\<^bsub>H\<^esub> = h \<one> \<otimes>\<^bsub>H\<^esub> h \<one>" |
13813 | 657 |
by (simp add: hom_mult [symmetric] del: hom_mult) |
658 |
then show ?thesis by (simp del: r_one) |
|
659 |
qed |
|
660 |
||
661 |
lemma (in group_hom) inv_closed [simp]: |
|
662 |
"x \<in> carrier G ==> h (inv x) \<in> carrier H" |
|
663 |
by simp |
|
664 |
||
665 |
lemma (in group_hom) hom_inv [simp]: |
|
14693 | 666 |
"x \<in> carrier G ==> h (inv x) = inv\<^bsub>H\<^esub> (h x)" |
13813 | 667 |
proof - |
668 |
assume x: "x \<in> carrier G" |
|
14693 | 669 |
then have "h x \<otimes>\<^bsub>H\<^esub> h (inv x) = \<one>\<^bsub>H\<^esub>" |
14963 | 670 |
by (simp add: hom_mult [symmetric] del: hom_mult) |
14693 | 671 |
also from x have "... = h x \<otimes>\<^bsub>H\<^esub> inv\<^bsub>H\<^esub> (h x)" |
14963 | 672 |
by (simp add: hom_mult [symmetric] del: hom_mult) |
14693 | 673 |
finally have "h x \<otimes>\<^bsub>H\<^esub> h (inv x) = h x \<otimes>\<^bsub>H\<^esub> inv\<^bsub>H\<^esub> (h x)" . |
27698 | 674 |
with x show ?thesis by (simp del: H.r_inv H.Units_r_inv) |
13813 | 675 |
qed |
676 |
||
57271 | 677 |
(* Contributed by Joachim Breitner *) |
678 |
lemma (in group) int_pow_is_hom: |
|
679 |
"x \<in> carrier G \<Longrightarrow> (op(^) x) \<in> hom \<lparr> carrier = UNIV, mult = op +, one = 0::int \<rparr> G " |
|
680 |
unfolding hom_def by (simp add: int_pow_mult) |
|
681 |
||
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
19984
diff
changeset
|
682 |
|
61382 | 683 |
subsection \<open>Commutative Structures\<close> |
13936 | 684 |
|
61382 | 685 |
text \<open> |
13936 | 686 |
Naming convention: multiplicative structures that are commutative |
687 |
are called \emph{commutative}, additive structures are called |
|
688 |
\emph{Abelian}. |
|
61382 | 689 |
\<close> |
13813 | 690 |
|
14963 | 691 |
locale comm_monoid = monoid + |
692 |
assumes m_comm: "\<lbrakk>x \<in> carrier G; y \<in> carrier G\<rbrakk> \<Longrightarrow> x \<otimes> y = y \<otimes> x" |
|
13813 | 693 |
|
14963 | 694 |
lemma (in comm_monoid) m_lcomm: |
695 |
"\<lbrakk>x \<in> carrier G; y \<in> carrier G; z \<in> carrier G\<rbrakk> \<Longrightarrow> |
|
13813 | 696 |
x \<otimes> (y \<otimes> z) = y \<otimes> (x \<otimes> z)" |
697 |
proof - |
|
14693 | 698 |
assume xyz: "x \<in> carrier G" "y \<in> carrier G" "z \<in> carrier G" |
13813 | 699 |
from xyz have "x \<otimes> (y \<otimes> z) = (x \<otimes> y) \<otimes> z" by (simp add: m_assoc) |
700 |
also from xyz have "... = (y \<otimes> x) \<otimes> z" by (simp add: m_comm) |
|
701 |
also from xyz have "... = y \<otimes> (x \<otimes> z)" by (simp add: m_assoc) |
|
702 |
finally show ?thesis . |
|
703 |
qed |
|
704 |
||
14963 | 705 |
lemmas (in comm_monoid) m_ac = m_assoc m_comm m_lcomm |
13813 | 706 |
|
13936 | 707 |
lemma comm_monoidI: |
19783 | 708 |
fixes G (structure) |
13936 | 709 |
assumes m_closed: |
14693 | 710 |
"!!x y. [| x \<in> carrier G; y \<in> carrier G |] ==> x \<otimes> y \<in> carrier G" |
711 |
and one_closed: "\<one> \<in> carrier G" |
|
13936 | 712 |
and m_assoc: |
713 |
"!!x y z. [| x \<in> carrier G; y \<in> carrier G; z \<in> carrier G |] ==> |
|
14693 | 714 |
(x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)" |
715 |
and l_one: "!!x. x \<in> carrier G ==> \<one> \<otimes> x = x" |
|
13936 | 716 |
and m_comm: |
14693 | 717 |
"!!x y. [| x \<in> carrier G; y \<in> carrier G |] ==> x \<otimes> y = y \<otimes> x" |
13936 | 718 |
shows "comm_monoid G" |
719 |
using l_one |
|
14963 | 720 |
by (auto intro!: comm_monoid.intro comm_monoid_axioms.intro monoid.intro |
27714
27b4d7c01f8b
Tuned (for the sake of a meaningless log entry).
ballarin
parents:
27713
diff
changeset
|
721 |
intro: assms simp: m_closed one_closed m_comm) |
13817 | 722 |
|
13936 | 723 |
lemma (in monoid) monoid_comm_monoidI: |
724 |
assumes m_comm: |
|
14693 | 725 |
"!!x y. [| x \<in> carrier G; y \<in> carrier G |] ==> x \<otimes> y = y \<otimes> x" |
13936 | 726 |
shows "comm_monoid G" |
727 |
by (rule comm_monoidI) (auto intro: m_assoc m_comm) |
|
14963 | 728 |
|
14693 | 729 |
(*lemma (in comm_monoid) r_one [simp]: |
13817 | 730 |
"x \<in> carrier G ==> x \<otimes> \<one> = x" |
731 |
proof - |
|
732 |
assume G: "x \<in> carrier G" |
|
733 |
then have "x \<otimes> \<one> = \<one> \<otimes> x" by (simp add: m_comm) |
|
734 |
also from G have "... = x" by simp |
|
735 |
finally show ?thesis . |
|
14693 | 736 |
qed*) |
14963 | 737 |
|
13936 | 738 |
lemma (in comm_monoid) nat_pow_distr: |
739 |
"[| x \<in> carrier G; y \<in> carrier G |] ==> |
|
740 |
(x \<otimes> y) (^) (n::nat) = x (^) n \<otimes> y (^) n" |
|
741 |
by (induct n) (simp, simp add: m_ac) |
|
742 |
||
743 |
locale comm_group = comm_monoid + group |
|
744 |
||
745 |
lemma (in group) group_comm_groupI: |
|
746 |
assumes m_comm: "!!x y. [| x \<in> carrier G; y \<in> carrier G |] ==> |
|
14693 | 747 |
x \<otimes> y = y \<otimes> x" |
13936 | 748 |
shows "comm_group G" |
61169 | 749 |
by standard (simp_all add: m_comm) |
13817 | 750 |
|
13936 | 751 |
lemma comm_groupI: |
19783 | 752 |
fixes G (structure) |
13936 | 753 |
assumes m_closed: |
14693 | 754 |
"!!x y. [| x \<in> carrier G; y \<in> carrier G |] ==> x \<otimes> y \<in> carrier G" |
755 |
and one_closed: "\<one> \<in> carrier G" |
|
13936 | 756 |
and m_assoc: |
757 |
"!!x y z. [| x \<in> carrier G; y \<in> carrier G; z \<in> carrier G |] ==> |
|
14693 | 758 |
(x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)" |
13936 | 759 |
and m_comm: |
14693 | 760 |
"!!x y. [| x \<in> carrier G; y \<in> carrier G |] ==> x \<otimes> y = y \<otimes> x" |
761 |
and l_one: "!!x. x \<in> carrier G ==> \<one> \<otimes> x = x" |
|
14963 | 762 |
and l_inv_ex: "!!x. x \<in> carrier G ==> \<exists>y \<in> carrier G. y \<otimes> x = \<one>" |
13936 | 763 |
shows "comm_group G" |
27714
27b4d7c01f8b
Tuned (for the sake of a meaningless log entry).
ballarin
parents:
27713
diff
changeset
|
764 |
by (fast intro: group.group_comm_groupI groupI assms) |
13936 | 765 |
|
766 |
lemma (in comm_group) inv_mult: |
|
13854
91c9ab25fece
First distributed version of Group and Ring theory.
ballarin
parents:
13835
diff
changeset
|
767 |
"[| x \<in> carrier G; y \<in> carrier G |] ==> inv (x \<otimes> y) = inv x \<otimes> inv y" |
13936 | 768 |
by (simp add: m_ac inv_mult_group) |
13854
91c9ab25fece
First distributed version of Group and Ring theory.
ballarin
parents:
13835
diff
changeset
|
769 |
|
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
19984
diff
changeset
|
770 |
|
61382 | 771 |
subsection \<open>The Lattice of Subgroups of a Group\<close> |
14751
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
772 |
|
61382 | 773 |
text_raw \<open>\label{sec:subgroup-lattice}\<close> |
14751
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
774 |
|
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
775 |
theorem (in group) subgroups_partial_order: |
55926 | 776 |
"partial_order \<lparr>carrier = {H. subgroup H G}, eq = op =, le = op \<subseteq>\<rparr>" |
61169 | 777 |
by standard simp_all |
14751
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
778 |
|
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
779 |
lemma (in group) subgroup_self: |
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
780 |
"subgroup (carrier G) G" |
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
781 |
by (rule subgroupI) auto |
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
782 |
|
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
783 |
lemma (in group) subgroup_imp_group: |
55926 | 784 |
"subgroup H G ==> group (G\<lparr>carrier := H\<rparr>)" |
26199 | 785 |
by (erule subgroup.subgroup_is_group) (rule group_axioms) |
14751
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
786 |
|
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
787 |
lemma (in group) is_monoid [intro, simp]: |
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
788 |
"monoid G" |
14963 | 789 |
by (auto intro: monoid.intro m_assoc) |
14751
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
790 |
|
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
791 |
lemma (in group) subgroup_inv_equality: |
55926 | 792 |
"[| subgroup H G; x \<in> H |] ==> m_inv (G \<lparr>carrier := H\<rparr>) x = inv x" |
14751
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
793 |
apply (rule_tac inv_equality [THEN sym]) |
14761 | 794 |
apply (rule group.l_inv [OF subgroup_imp_group, simplified], assumption+) |
795 |
apply (rule subsetD [OF subgroup.subset], assumption+) |
|
796 |
apply (rule subsetD [OF subgroup.subset], assumption) |
|
797 |
apply (rule_tac group.inv_closed [OF subgroup_imp_group, simplified], assumption+) |
|
14751
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
798 |
done |
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
799 |
|
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
800 |
theorem (in group) subgroups_Inter: |
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
801 |
assumes subgr: "(!!H. H \<in> A ==> subgroup H G)" |
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
802 |
and not_empty: "A ~= {}" |
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
803 |
shows "subgroup (\<Inter>A) G" |
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
804 |
proof (rule subgroupI) |
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
805 |
from subgr [THEN subgroup.subset] and not_empty |
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
806 |
show "\<Inter>A \<subseteq> carrier G" by blast |
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
807 |
next |
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
808 |
from subgr [THEN subgroup.one_closed] |
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
809 |
show "\<Inter>A ~= {}" by blast |
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
810 |
next |
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
811 |
fix x assume "x \<in> \<Inter>A" |
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
812 |
with subgr [THEN subgroup.m_inv_closed] |
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
813 |
show "inv x \<in> \<Inter>A" by blast |
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
814 |
next |
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
815 |
fix x y assume "x \<in> \<Inter>A" "y \<in> \<Inter>A" |
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
816 |
with subgr [THEN subgroup.m_closed] |
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
817 |
show "x \<otimes> y \<in> \<Inter>A" by blast |
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
818 |
qed |
0d7850e27fed
Change of theory hierarchy: Group is now based in Lattice.
ballarin
parents:
14706
diff
changeset
|
819 |
|
13813 | 820 |
end |