77089
|
1 |
theory Left_Coset
|
|
2 |
imports Coset
|
|
3 |
|
|
4 |
(*This instance of Coset.thy but for left cosets is due to Jonas Rädle and has been imported
|
|
5 |
from the AFP entry Orbit_Stabiliser. *)
|
|
6 |
|
|
7 |
begin
|
|
8 |
|
|
9 |
definition
|
|
10 |
LCOSETS :: "[_, 'a set] \<Rightarrow> ('a set)set" ("lcosets\<index> _" [81] 80)
|
|
11 |
where "lcosets\<^bsub>G\<^esub> H = (\<Union>a\<in>carrier G. {a <#\<^bsub>G\<^esub> H})"
|
|
12 |
|
|
13 |
definition
|
|
14 |
LFactGroup :: "[('a,'b) monoid_scheme, 'a set] \<Rightarrow> ('a set) monoid" (infixl "LMod" 65)
|
|
15 |
\<comment> \<open>Actually defined for groups rather than monoids\<close>
|
|
16 |
where "LFactGroup G H = \<lparr>carrier = lcosets\<^bsub>G\<^esub> H, mult = set_mult G, one = H\<rparr>"
|
|
17 |
|
|
18 |
lemma (in group) lcos_self: "[| x \<in> carrier G; subgroup H G |] ==> x \<in> x <# H"
|
|
19 |
by (simp add: group_l_invI subgroup.lcos_module_rev subgroup.one_closed)
|
|
20 |
|
|
21 |
text \<open>Elements of a left coset are in the carrier\<close>
|
|
22 |
lemma (in subgroup) elemlcos_carrier:
|
|
23 |
assumes "group G" "a \<in> carrier G" "a' \<in> a <# H"
|
|
24 |
shows "a' \<in> carrier G"
|
|
25 |
by (meson assms group.l_coset_carrier subgroup_axioms)
|
|
26 |
|
|
27 |
text \<open>Step one for lemma \<open>rcos_module\<close>\<close>
|
|
28 |
lemma (in subgroup) lcos_module_imp:
|
|
29 |
assumes "group G"
|
|
30 |
assumes xcarr: "x \<in> carrier G"
|
|
31 |
and x'cos: "x' \<in> x <# H"
|
|
32 |
shows "(inv x \<otimes> x') \<in> H"
|
|
33 |
proof -
|
|
34 |
interpret group G by fact
|
|
35 |
obtain h
|
|
36 |
where hH: "h \<in> H" and x': "x' = x \<otimes> h" and hcarr: "h \<in> carrier G"
|
|
37 |
using assms by (auto simp: l_coset_def)
|
|
38 |
have "(inv x) \<otimes> x' = (inv x) \<otimes> (x \<otimes> h)"
|
|
39 |
by (simp add: x')
|
|
40 |
have "\<dots> = (x \<otimes> inv x) \<otimes> h"
|
|
41 |
by (metis hcarr inv_closed inv_inv l_inv m_assoc xcarr)
|
|
42 |
also have "\<dots> = h"
|
|
43 |
by (simp add: hcarr xcarr)
|
|
44 |
finally have "(inv x) \<otimes> x' = h"
|
|
45 |
using x' by metis
|
|
46 |
then show "(inv x) \<otimes> x' \<in> H"
|
|
47 |
using hH by force
|
|
48 |
qed
|
|
49 |
|
|
50 |
text \<open>Left cosets are subsets of the carrier.\<close>
|
|
51 |
lemma (in subgroup) lcosets_carrier:
|
|
52 |
assumes "group G"
|
|
53 |
assumes XH: "X \<in> lcosets H"
|
|
54 |
shows "X \<subseteq> carrier G"
|
|
55 |
proof -
|
|
56 |
interpret group G by fact
|
|
57 |
show "X \<subseteq> carrier G"
|
|
58 |
using XH l_coset_subset_G subset by (auto simp: LCOSETS_def)
|
|
59 |
qed
|
|
60 |
|
|
61 |
lemma (in group) lcosets_part_G:
|
|
62 |
assumes "subgroup H G"
|
|
63 |
shows "\<Union>(lcosets H) = carrier G"
|
|
64 |
proof -
|
|
65 |
interpret subgroup H G by fact
|
|
66 |
show ?thesis
|
|
67 |
proof
|
|
68 |
show "\<Union> (lcosets H) \<subseteq> carrier G"
|
|
69 |
by (force simp add: LCOSETS_def l_coset_def)
|
|
70 |
show "carrier G \<subseteq> \<Union> (lcosets H)"
|
|
71 |
by (auto simp add: LCOSETS_def intro: lcos_self assms)
|
|
72 |
qed
|
|
73 |
qed
|
|
74 |
|
|
75 |
lemma (in group) lcosets_subset_PowG:
|
|
76 |
"subgroup H G \<Longrightarrow> lcosets H \<subseteq> Pow(carrier G)"
|
|
77 |
using lcosets_part_G subset_Pow_Union by blast
|
|
78 |
|
|
79 |
lemma (in group) lcos_disjoint:
|
|
80 |
assumes "subgroup H G"
|
|
81 |
assumes p: "a \<in> lcosets H" "b \<in> lcosets H" "a\<noteq>b"
|
|
82 |
shows "a \<inter> b = {}"
|
|
83 |
proof -
|
|
84 |
interpret subgroup H G by fact
|
|
85 |
show ?thesis
|
|
86 |
using p l_repr_independence subgroup_axioms unfolding LCOSETS_def disjoint_iff
|
|
87 |
by blast
|
|
88 |
qed
|
|
89 |
|
|
90 |
text\<open>The next two lemmas support the proof of \<open>card_cosets_equal\<close>.\<close>
|
|
91 |
lemma (in group) inj_on_f':
|
|
92 |
"\<lbrakk>H \<subseteq> carrier G; a \<in> carrier G\<rbrakk> \<Longrightarrow> inj_on (\<lambda>y. y \<otimes> inv a) (a <# H)"
|
|
93 |
by (simp add: inj_on_g l_coset_subset_G)
|
|
94 |
|
|
95 |
lemma (in group) inj_on_f'':
|
|
96 |
"\<lbrakk>H \<subseteq> carrier G; a \<in> carrier G\<rbrakk> \<Longrightarrow> inj_on (\<lambda>y. inv a \<otimes> y) (a <# H)"
|
|
97 |
by (meson inj_on_cmult inv_closed l_coset_subset_G subset_inj_on)
|
|
98 |
|
|
99 |
lemma (in group) inj_on_g':
|
|
100 |
"\<lbrakk>H \<subseteq> carrier G; a \<in> carrier G\<rbrakk> \<Longrightarrow> inj_on (\<lambda>y. a \<otimes> y) H"
|
|
101 |
using inj_on_cmult inj_on_subset by blast
|
|
102 |
|
|
103 |
lemma (in group) l_card_cosets_equal:
|
|
104 |
assumes "c \<in> lcosets H" and H: "H \<subseteq> carrier G" and fin: "finite(carrier G)"
|
|
105 |
shows "card H = card c"
|
|
106 |
proof -
|
|
107 |
obtain x where x: "x \<in> carrier G" and c: "c = x <# H"
|
|
108 |
using assms by (auto simp add: LCOSETS_def)
|
|
109 |
have "inj_on ((\<otimes>) x) H"
|
|
110 |
by (simp add: H group.inj_on_g' x)
|
|
111 |
moreover
|
|
112 |
have "(\<otimes>) x ` H \<subseteq> x <# H"
|
|
113 |
by (force simp add: m_assoc subsetD l_coset_def)
|
|
114 |
moreover
|
|
115 |
have "inj_on ((\<otimes>) (inv x)) (x <# H)"
|
|
116 |
by (simp add: H group.inj_on_f'' x)
|
|
117 |
moreover
|
|
118 |
have "\<And>h. h \<in> H \<Longrightarrow> inv x \<otimes> (x \<otimes> h) \<in> H"
|
|
119 |
by (metis H in_mono inv_solve_left m_closed x)
|
|
120 |
then have "(\<otimes>) (inv x) ` (x <# H) \<subseteq> H"
|
|
121 |
by (auto simp: l_coset_def)
|
|
122 |
ultimately show ?thesis
|
|
123 |
by (metis H fin c card_bij_eq finite_imageD finite_subset)
|
|
124 |
qed
|
|
125 |
|
|
126 |
theorem (in group) l_lagrange:
|
|
127 |
assumes "finite(carrier G)" "subgroup H G"
|
|
128 |
shows "card(lcosets H) * card(H) = order(G)"
|
|
129 |
proof -
|
|
130 |
have "card H * card (lcosets H) = card (\<Union> (lcosets H))"
|
|
131 |
using card_partition
|
|
132 |
by (metis (no_types, lifting) assms finite_UnionD l_card_cosets_equal lcos_disjoint lcosets_part_G subgroup.subset)
|
|
133 |
then show ?thesis
|
|
134 |
by (simp add: assms(2) lcosets_part_G mult.commute order_def)
|
|
135 |
qed
|
|
136 |
|
|
137 |
end
|