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