--- a/src/HOL/Algebra/AbelCoset.thy Thu Jun 14 10:51:12 2018 +0200
+++ b/src/HOL/Algebra/AbelCoset.thy Thu Jun 14 14:23:48 2018 +0100
@@ -17,45 +17,42 @@
definition
a_r_coset :: "[_, 'a set, 'a] \<Rightarrow> 'a set" (infixl "+>\<index>" 60)
- where "a_r_coset G = r_coset \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
+ where "a_r_coset G = r_coset (add_monoid G)"
definition
a_l_coset :: "[_, 'a, 'a set] \<Rightarrow> 'a set" (infixl "<+\<index>" 60)
- where "a_l_coset G = l_coset \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
+ where "a_l_coset G = l_coset (add_monoid G)"
definition
A_RCOSETS :: "[_, 'a set] \<Rightarrow> ('a set)set" ("a'_rcosets\<index> _" [81] 80)
- where "A_RCOSETS G H = RCOSETS \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> H"
+ where "A_RCOSETS G H = RCOSETS (add_monoid G) H"
definition
set_add :: "[_, 'a set ,'a set] \<Rightarrow> 'a set" (infixl "<+>\<index>" 60)
- where "set_add G = set_mult \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
+ where "set_add G = set_mult (add_monoid G)"
definition
A_SET_INV :: "[_,'a set] \<Rightarrow> 'a set" ("a'_set'_inv\<index> _" [81] 80)
- where "A_SET_INV G H = SET_INV \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> H"
+ where "A_SET_INV G H = SET_INV (add_monoid G) H"
definition
a_r_congruent :: "[('a,'b)ring_scheme, 'a set] \<Rightarrow> ('a*'a)set" ("racong\<index>")
- where "a_r_congruent G = r_congruent \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
+ where "a_r_congruent G = r_congruent (add_monoid G)"
definition
A_FactGroup :: "[('a,'b) ring_scheme, 'a set] \<Rightarrow> ('a set) monoid" (infixl "A'_Mod" 65)
\<comment> \<open>Actually defined for groups rather than monoids\<close>
- where "A_FactGroup G H = FactGroup \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> H"
+ where "A_FactGroup G H = FactGroup (add_monoid G) H"
definition
a_kernel :: "('a, 'm) ring_scheme \<Rightarrow> ('b, 'n) ring_scheme \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a set"
\<comment> \<open>the kernel of a homomorphism (additive)\<close>
- where "a_kernel G H h =
- kernel \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>
- \<lparr>carrier = carrier H, mult = add H, one = zero H\<rparr> h"
+ where "a_kernel G H h = kernel (add_monoid G) (add_monoid H) h"
locale abelian_group_hom = G?: abelian_group G + H?: abelian_group H
for G (structure) and H (structure) +
fixes h
- assumes a_group_hom: "group_hom \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>
- \<lparr>carrier = carrier H, mult = add H, one = zero H\<rparr> h"
+ assumes a_group_hom: "group_hom (add_monoid G) (add_monoid H) h"
lemmas a_r_coset_defs =
a_r_coset_def r_coset_def
@@ -63,8 +60,7 @@
lemma a_r_coset_def':
fixes G (structure)
shows "H +> a \<equiv> \<Union>h\<in>H. {h \<oplus> a}"
-unfolding a_r_coset_defs
-by simp
+ unfolding a_r_coset_defs by simp
lemmas a_l_coset_defs =
a_l_coset_def l_coset_def
@@ -72,8 +68,7 @@
lemma a_l_coset_def':
fixes G (structure)
shows "a <+ H \<equiv> \<Union>h\<in>H. {a \<oplus> h}"
-unfolding a_l_coset_defs
-by simp
+ unfolding a_l_coset_defs by simp
lemmas A_RCOSETS_defs =
A_RCOSETS_def RCOSETS_def
@@ -81,8 +76,7 @@
lemma A_RCOSETS_def':
fixes G (structure)
shows "a_rcosets H \<equiv> \<Union>a\<in>carrier G. {H +> a}"
-unfolding A_RCOSETS_defs
-by (fold a_r_coset_def, simp)
+ unfolding A_RCOSETS_defs by (fold a_r_coset_def, simp)
lemmas set_add_defs =
set_add_def set_mult_def
@@ -90,8 +84,7 @@
lemma set_add_def':
fixes G (structure)
shows "H <+> K \<equiv> \<Union>h\<in>H. \<Union>k\<in>K. {h \<oplus> k}"
-unfolding set_add_defs
-by simp
+ unfolding set_add_defs by simp
lemmas A_SET_INV_defs =
A_SET_INV_def SET_INV_def
@@ -99,18 +92,53 @@
lemma A_SET_INV_def':
fixes G (structure)
shows "a_set_inv H \<equiv> \<Union>h\<in>H. {\<ominus> h}"
-unfolding A_SET_INV_defs
-by (fold a_inv_def)
+ unfolding A_SET_INV_defs by (fold a_inv_def)
subsubsection \<open>Cosets\<close>
+sublocale abelian_group <
+ add: group "(add_monoid G)"
+ rewrites "carrier (add_monoid G) = carrier G"
+ and " mult (add_monoid G) = add G"
+ and " one (add_monoid G) = zero G"
+ and " m_inv (add_monoid G) = a_inv G"
+ and "finprod (add_monoid G) = finsum G"
+ and "r_coset (add_monoid G) = a_r_coset G"
+ and "l_coset (add_monoid G) = a_l_coset G"
+ and "(\<lambda>a k. pow (add_monoid G) a k) = (\<lambda>a k. add_pow G k a)"
+ by (rule a_group)
+ (auto simp: m_inv_def a_inv_def finsum_def a_r_coset_def a_l_coset_def add_pow_def)
+
+context abelian_group
+begin
+
+thm add.coset_mult_assoc
+lemmas a_repr_independence' = add.repr_independence
+
+(*
+lemmas a_coset_add_assoc = add.coset_mult_assoc
+lemmas a_coset_add_zero [simp] = add.coset_mult_one
+lemmas a_coset_add_inv1 = add.coset_mult_inv1
+lemmas a_coset_add_inv2 = add.coset_mult_inv2
+lemmas a_coset_join1 = add.coset_join1
+lemmas a_coset_join2 = add.coset_join2
+lemmas a_solve_equation = add.solve_equation
+lemmas a_repr_independence = add.repr_independence
+lemmas a_rcosI = add.rcosI
+lemmas a_rcosetsI = add.rcosetsI
+*)
+
+end
+
lemma (in abelian_group) a_coset_add_assoc:
"[| M \<subseteq> carrier G; g \<in> carrier G; h \<in> carrier G |]
==> (M +> g) +> h = M +> (g \<oplus> h)"
by (rule group.coset_mult_assoc [OF a_group,
folded a_r_coset_def, simplified monoid_record_simps])
+thm abelian_group.a_coset_add_assoc
+
lemma (in abelian_group) a_coset_add_zero [simp]:
"M \<subseteq> carrier G ==> M +> \<zero> = M"
by (rule group.coset_mult_one [OF a_group,
@@ -129,22 +157,22 @@
folded a_r_coset_def a_inv_def, simplified monoid_record_simps])
lemma (in abelian_group) a_coset_join1:
- "[| H +> x = H; x \<in> carrier G; subgroup H \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> |] ==> x \<in> H"
+ "[| H +> x = H; x \<in> carrier G; subgroup H (add_monoid G) |] ==> x \<in> H"
by (rule group.coset_join1 [OF a_group,
folded a_r_coset_def, simplified monoid_record_simps])
lemma (in abelian_group) a_solve_equation:
- "\<lbrakk>subgroup H \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>; x \<in> H; y \<in> H\<rbrakk> \<Longrightarrow> \<exists>h\<in>H. y = h \<oplus> x"
+ "\<lbrakk>subgroup H (add_monoid G); x \<in> H; y \<in> H\<rbrakk> \<Longrightarrow> \<exists>h\<in>H. y = h \<oplus> x"
by (rule group.solve_equation [OF a_group,
folded a_r_coset_def, simplified monoid_record_simps])
lemma (in abelian_group) a_repr_independence:
- "\<lbrakk>y \<in> H +> x; x \<in> carrier G; subgroup H \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> \<rbrakk> \<Longrightarrow> H +> x = H +> y"
-by (rule group.repr_independence [OF a_group,
- folded a_r_coset_def, simplified monoid_record_simps])
+ "\<lbrakk> y \<in> H +> x; x \<in> carrier G; subgroup H (add_monoid G) \<rbrakk> \<Longrightarrow>
+ H +> x = H +> y"
+ using a_repr_independence' by (simp add: a_r_coset_def)
lemma (in abelian_group) a_coset_join2:
- "\<lbrakk>x \<in> carrier G; subgroup H \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>; x\<in>H\<rbrakk> \<Longrightarrow> H +> x = H"
+ "\<lbrakk>x \<in> carrier G; subgroup H (add_monoid G); x\<in>H\<rbrakk> \<Longrightarrow> H +> x = H"
by (rule group.coset_join2 [OF a_group,
folded a_r_coset_def, simplified monoid_record_simps])
@@ -167,23 +195,15 @@
lemma (in abelian_group) a_transpose_inv:
"[| x \<oplus> y = z; x \<in> carrier G; y \<in> carrier G; z \<in> carrier G |]
==> (\<ominus> x) \<oplus> z = y"
-by (rule group.transpose_inv [OF a_group,
- folded a_r_coset_def a_inv_def, simplified monoid_record_simps])
+ using r_neg1 by blast
-(*
---"duplicate"
-lemma (in abelian_group) a_rcos_self:
- "[| x \<in> carrier G; subgroup H \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> |] ==> x \<in> H +> x"
-by (rule group.rcos_self [OF a_group,
- folded a_r_coset_def, simplified monoid_record_simps])
-*)
subsubsection \<open>Subgroups\<close>
locale additive_subgroup =
fixes H and G (structure)
- assumes a_subgroup: "subgroup H \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
+ assumes a_subgroup: "subgroup H (add_monoid G)"
lemma (in additive_subgroup) is_additive_subgroup:
shows "additive_subgroup H G"
@@ -191,7 +211,7 @@
lemma additive_subgroupI:
fixes G (structure)
- assumes a_subgroup: "subgroup H \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
+ assumes a_subgroup: "subgroup H (add_monoid G)"
shows "additive_subgroup H G"
by (rule additive_subgroup.intro) (rule a_subgroup)
@@ -221,18 +241,18 @@
text \<open>Every subgroup of an \<open>abelian_group\<close> is normal\<close>
locale abelian_subgroup = additive_subgroup + abelian_group G +
- assumes a_normal: "normal H \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
+ assumes a_normal: "normal H (add_monoid G)"
lemma (in abelian_subgroup) is_abelian_subgroup:
shows "abelian_subgroup H G"
by (rule abelian_subgroup_axioms)
lemma abelian_subgroupI:
- assumes a_normal: "normal H \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
+ assumes a_normal: "normal H (add_monoid G)"
and a_comm: "!!x y. [| x \<in> carrier G; y \<in> carrier G |] ==> x \<oplus>\<^bsub>G\<^esub> y = y \<oplus>\<^bsub>G\<^esub> x"
shows "abelian_subgroup H G"
proof -
- interpret normal "H" "\<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
+ interpret normal "H" "(add_monoid G)"
by (rule a_normal)
show "abelian_subgroup H G"
@@ -241,13 +261,13 @@
lemma abelian_subgroupI2:
fixes G (structure)
- assumes a_comm_group: "comm_group \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
- and a_subgroup: "subgroup H \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
+ assumes a_comm_group: "comm_group (add_monoid G)"
+ and a_subgroup: "subgroup H (add_monoid G)"
shows "abelian_subgroup H G"
proof -
- interpret comm_group "\<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
+ interpret comm_group "(add_monoid G)"
by (rule a_comm_group)
- interpret subgroup "H" "\<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
+ interpret subgroup "H" "(add_monoid G)"
by (rule a_subgroup)
show "abelian_subgroup H G"
@@ -264,13 +284,10 @@
lemma abelian_subgroupI3:
fixes G (structure)
- assumes asg: "additive_subgroup H G"
- and ag: "abelian_group G"
+ assumes "additive_subgroup H G"
+ and "abelian_group G"
shows "abelian_subgroup H G"
-apply (rule abelian_subgroupI2)
- apply (rule abelian_group.a_comm_group[OF ag])
-apply (rule additive_subgroup.a_subgroup[OF asg])
-done
+ using assms abelian_subgroupI2 abelian_group.a_comm_group additive_subgroup_def by blast
lemma (in abelian_subgroup) a_coset_eq:
"(\<forall>x \<in> carrier G. H +> x = x <+ H)"
@@ -289,15 +306,14 @@
text\<open>Alternative characterization of normal subgroups\<close>
lemma (in abelian_group) a_normal_inv_iff:
- "(N \<lhd> \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>) =
- (subgroup N \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> \<and> (\<forall>x \<in> carrier G. \<forall>h \<in> N. x \<oplus> h \<oplus> (\<ominus> x) \<in> N))"
+ "(N \<lhd> (add_monoid G)) =
+ (subgroup N (add_monoid G) & (\<forall>x \<in> carrier G. \<forall>h \<in> N. x \<oplus> h \<oplus> (\<ominus> x) \<in> N))"
(is "_ = ?rhs")
by (rule group.normal_inv_iff [OF a_group,
folded a_inv_def, simplified monoid_record_simps])
lemma (in abelian_group) a_lcos_m_assoc:
- "[| M \<subseteq> carrier G; g \<in> carrier G; h \<in> carrier G |]
- ==> g <+ (h <+ M) = (g \<oplus> h) <+ M"
+ "\<lbrakk> M \<subseteq> carrier G; g \<in> carrier G; h \<in> carrier G \<rbrakk> \<Longrightarrow> g <+ (h <+ M) = (g \<oplus> h) <+ M"
by (rule group.lcos_m_assoc [OF a_group,
folded a_l_coset_def, simplified monoid_record_simps])
@@ -308,33 +324,28 @@
lemma (in abelian_group) a_l_coset_subset_G:
- "[| H \<subseteq> carrier G; x \<in> carrier G |] ==> x <+ H \<subseteq> carrier G"
+ "\<lbrakk> H \<subseteq> carrier G; x \<in> carrier G \<rbrakk> \<Longrightarrow> x <+ H \<subseteq> carrier G"
by (rule group.l_coset_subset_G [OF a_group,
folded a_l_coset_def, simplified monoid_record_simps])
lemma (in abelian_group) a_l_coset_swap:
- "\<lbrakk>y \<in> x <+ H; x \<in> carrier G; subgroup H \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>\<rbrakk> \<Longrightarrow> x \<in> y <+ H"
+ "\<lbrakk>y \<in> x <+ H; x \<in> carrier G; subgroup H (add_monoid G)\<rbrakk> \<Longrightarrow> x \<in> y <+ H"
by (rule group.l_coset_swap [OF a_group,
folded a_l_coset_def, simplified monoid_record_simps])
lemma (in abelian_group) a_l_coset_carrier:
- "[| y \<in> x <+ H; x \<in> carrier G; subgroup H \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> |] ==> y \<in> carrier G"
+ "[| y \<in> x <+ H; x \<in> carrier G; subgroup H (add_monoid G) |] ==> y \<in> carrier G"
by (rule group.l_coset_carrier [OF a_group,
folded a_l_coset_def, simplified monoid_record_simps])
lemma (in abelian_group) a_l_repr_imp_subset:
- assumes y: "y \<in> x <+ H" and x: "x \<in> carrier G" and sb: "subgroup H \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
+ assumes "y \<in> x <+ H" "x \<in> carrier G" "subgroup H (add_monoid G)"
shows "y <+ H \<subseteq> x <+ H"
-apply (rule group.l_repr_imp_subset [OF a_group,
- folded a_l_coset_def, simplified monoid_record_simps])
-apply (rule y)
-apply (rule x)
-apply (rule sb)
-done
+ by (metis (full_types) a_l_coset_defs(1) add.l_repr_independence assms set_eq_subset)
lemma (in abelian_group) a_l_repr_independence:
- assumes y: "y \<in> x <+ H" and x: "x \<in> carrier G" and sb: "subgroup H \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
+ assumes y: "y \<in> x <+ H" and x: "x \<in> carrier G" and sb: "subgroup H (add_monoid G)"
shows "x <+ H = y <+ H"
apply (rule group.l_repr_independence [OF a_group,
folded a_l_coset_def, simplified monoid_record_simps])
@@ -348,7 +359,7 @@
by (rule group.setmult_subset_G [OF a_group,
folded set_add_def, simplified monoid_record_simps])
-lemma (in abelian_group) subgroup_add_id: "subgroup H \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> \<Longrightarrow> H <+> H = H"
+lemma (in abelian_group) subgroup_add_id: "subgroup H (add_monoid G) \<Longrightarrow> H <+> H = H"
by (rule group.subgroup_mult_id [OF a_group,
folded set_add_def, simplified monoid_record_simps])
@@ -427,8 +438,7 @@
lemma (in abelian_group) a_card_cosets_equal:
"\<lbrakk>c \<in> a_rcosets H; H \<subseteq> carrier G; finite(carrier G)\<rbrakk>
\<Longrightarrow> card c = card H"
-by (rule group.card_cosets_equal [OF a_group,
- folded A_RCOSETS_def, simplified monoid_record_simps])
+ by (simp add: A_RCOSETS_defs(1) add.card_rcosets_equal)
lemma (in abelian_group) rcosets_subset_PowG:
"additive_subgroup H G \<Longrightarrow> a_rcosets H \<subseteq> Pow(carrier G)"
@@ -509,7 +519,7 @@
text\<open>The coset map is a homomorphism from @{term G} to the quotient group
@{term "G Mod H"}\<close>
lemma (in abelian_subgroup) a_r_coset_hom_A_Mod:
- "(\<lambda>a. H +> a) \<in> hom \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> (G A_Mod H)"
+ "(\<lambda>a. H +> a) \<in> hom (add_monoid G) (G A_Mod H)"
by (rule normal.r_coset_hom_Mod [OF a_normal,
folded A_FactGroup_def a_r_coset_def, simplified monoid_record_simps])
@@ -535,8 +545,8 @@
lemma abelian_group_homI:
assumes "abelian_group G"
assumes "abelian_group H"
- assumes a_group_hom: "group_hom \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>
- \<lparr>carrier = carrier H, mult = add H, one = zero H\<rparr> h"
+ assumes a_group_hom: "group_hom (add_monoid G)
+ (add_monoid H) h"
shows "abelian_group_hom G H h"
proof -
interpret G: abelian_group G by fact
@@ -614,7 +624,7 @@
lemma (in abelian_group_hom) A_FactGroup_hom:
"(\<lambda>X. the_elem (h`X)) \<in> hom (G A_Mod (a_kernel G H h))
- \<lparr>carrier = carrier H, mult = add H, one = zero H\<rparr>"
+ (add_monoid H)"
by (rule group_hom.FactGroup_hom[OF a_group_hom,
folded a_kernel_def A_FactGroup_def, simplified ring_record_simps])
@@ -633,13 +643,16 @@
text\<open>If @{term h} is a homomorphism from @{term G} onto @{term H}, then the
quotient group @{term "G Mod (kernel G H h)"} is isomorphic to @{term H}.\<close>
-theorem (in abelian_group_hom) A_FactGroup_iso:
+theorem (in abelian_group_hom) A_FactGroup_iso_set:
"h ` carrier G = carrier H
- \<Longrightarrow> (\<lambda>X. the_elem (h`X)) \<in> (G A_Mod (a_kernel G H h)) \<cong>
- \<lparr>carrier = carrier H, mult = add H, one = zero H\<rparr>"
-by (rule group_hom.FactGroup_iso[OF a_group_hom,
+ \<Longrightarrow> (\<lambda>X. the_elem (h`X)) \<in> iso (G A_Mod (a_kernel G H h)) (add_monoid H)"
+by (rule group_hom.FactGroup_iso_set[OF a_group_hom,
folded a_kernel_def A_FactGroup_def, simplified ring_record_simps])
+corollary (in abelian_group_hom) A_FactGroup_iso :
+ "h ` carrier G = carrier H
+ \<Longrightarrow> (G A_Mod (a_kernel G H h)) \<cong> (add_monoid H)"
+ using A_FactGroup_iso_set unfolding is_iso_def by auto
subsubsection \<open>Cosets\<close>
--- a/src/HOL/Algebra/Congruence.thy Thu Jun 14 10:51:12 2018 +0200
+++ b/src/HOL/Algebra/Congruence.thy Thu Jun 14 14:23:48 2018 +0100
@@ -43,6 +43,10 @@
where "class_of\<^bsub>S\<^esub> x = {y \<in> carrier S. x .=\<^bsub>S\<^esub> y}"
definition
+ eq_classes :: "_ \<Rightarrow> ('a set) set" ("classes\<index>")
+ where "classes\<^bsub>S\<^esub> = {class_of\<^bsub>S\<^esub> x | x. x \<in> carrier S}"
+
+definition
eq_closure_of :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a set" ("closure'_of\<index>")
where "closure_of\<^bsub>S\<^esub> A = {y \<in> carrier S. y .\<in>\<^bsub>S\<^esub> A}"
@@ -69,235 +73,148 @@
and trans [trans]:
"\<lbrakk> x .= y; y .= z; x \<in> carrier S; y \<in> carrier S; z \<in> carrier S \<rbrakk> \<Longrightarrow> x .= z"
+lemma equivalenceI:
+ fixes P :: "'a \<Rightarrow> 'a \<Rightarrow> bool" and E :: "'a set"
+ assumes refl: "\<And>x. \<lbrakk> x \<in> E \<rbrakk> \<Longrightarrow> P x x"
+ and sym: "\<And>x y. \<lbrakk> x \<in> E; y \<in> E \<rbrakk> \<Longrightarrow> P x y \<Longrightarrow> P y x"
+ and trans: "\<And>x y z. \<lbrakk> x \<in> E; y \<in> E; z \<in> E \<rbrakk> \<Longrightarrow> P x y \<Longrightarrow> P y z \<Longrightarrow> P x z"
+ shows "equivalence \<lparr> carrier = E, eq = P \<rparr>"
+ unfolding equivalence_def using assms
+ by (metis eq_object.select_convs(1) partial_object.select_convs(1))
+
+locale partition =
+ fixes A :: "'a set" and B :: "('a set) set"
+ assumes unique_class: "\<And>a. a \<in> A \<Longrightarrow> \<exists>!b \<in> B. a \<in> b"
+ and incl: "\<And>b. b \<in> B \<Longrightarrow> b \<subseteq> A"
+
+lemma equivalence_subset:
+ assumes "equivalence L" "A \<subseteq> carrier L"
+ shows "equivalence (L\<lparr> carrier := A \<rparr>)"
+proof -
+ interpret L: equivalence L
+ by (simp add: assms)
+ show ?thesis
+ by (unfold_locales, simp_all add: L.sym assms rev_subsetD, meson L.trans assms(2) contra_subsetD)
+qed
+
+
(* Lemmas by Stephan Hohe *)
lemma elemI:
fixes R (structure)
- assumes "a' \<in> A" and "a .= a'"
+ assumes "a' \<in> A" "a .= a'"
shows "a .\<in> A"
-unfolding elem_def
-using assms
-by fast
+ unfolding elem_def using assms by auto
lemma (in equivalence) elem_exact:
- assumes "a \<in> carrier S" and "a \<in> A"
+ assumes "a \<in> carrier S" "a \<in> A"
shows "a .\<in> A"
-using assms
-by (fast intro: elemI)
+ unfolding elem_def using assms by auto
lemma elemE:
fixes S (structure)
assumes "a .\<in> A"
and "\<And>a'. \<lbrakk>a' \<in> A; a .= a'\<rbrakk> \<Longrightarrow> P"
shows "P"
-using assms
-unfolding elem_def
-by fast
+ using assms unfolding elem_def by auto
lemma (in equivalence) elem_cong_l [trans]:
- assumes cong: "a' .= a"
- and a: "a .\<in> A"
- and carr: "a \<in> carrier S" "a' \<in> carrier S"
- and Acarr: "A \<subseteq> carrier S"
+ assumes "a \<in> carrier S" "a' \<in> carrier S" "A \<subseteq> carrier S"
+ and "a' .= a" "a .\<in> A"
shows "a' .\<in> A"
-using a
-apply (elim elemE, intro elemI)
-proof assumption
- fix b
- assume bA: "b \<in> A"
- note [simp] = carr bA[THEN subsetD[OF Acarr]]
- note cong
- also assume "a .= b"
- finally show "a' .= b" by simp
-qed
+ using assms by (meson elem_def trans subsetCE)
lemma (in equivalence) elem_subsetD:
- assumes "A \<subseteq> B"
- and aA: "a .\<in> A"
+ assumes "A \<subseteq> B" "a .\<in> A"
shows "a .\<in> B"
-using assms
-by (fast intro: elemI elim: elemE dest: subsetD)
+ using assms by (fast intro: elemI elim: elemE dest: subsetD)
lemma (in equivalence) mem_imp_elem [simp, intro]:
- "\<lbrakk> x \<in> A; x \<in> carrier S \<rbrakk> \<Longrightarrow> x .\<in> A"
- unfolding elem_def by blast
+ assumes "x \<in> carrier S"
+ shows "x \<in> A \<Longrightarrow> x .\<in> A"
+ using assms unfolding elem_def by blast
lemma set_eqI:
fixes R (structure)
- assumes ltr: "\<And>a. a \<in> A \<Longrightarrow> a .\<in> B"
- and rtl: "\<And>b. b \<in> B \<Longrightarrow> b .\<in> A"
+ assumes "\<And>a. a \<in> A \<Longrightarrow> a .\<in> B"
+ and "\<And>b. b \<in> B \<Longrightarrow> b .\<in> A"
shows "A {.=} B"
-unfolding set_eq_def
-by (fast intro: ltr rtl)
+ using assms unfolding set_eq_def by auto
lemma set_eqI2:
fixes R (structure)
- assumes ltr: "\<And>a b. a \<in> A \<Longrightarrow> \<exists>b\<in>B. a .= b"
- and rtl: "\<And>b. b \<in> B \<Longrightarrow> \<exists>a\<in>A. b .= a"
+ assumes "\<And>a. a \<in> A \<Longrightarrow> \<exists>b \<in> B. a .= b"
+ and "\<And>b. b \<in> B \<Longrightarrow> \<exists>a \<in> A. b .= a"
shows "A {.=} B"
- by (intro set_eqI, unfold elem_def) (fast intro: ltr rtl)+
+ using assms by (simp add: set_eqI elem_def)
lemma set_eqD1:
fixes R (structure)
- assumes AA': "A {.=} A'"
- and "a \<in> A"
+ assumes "A {.=} A'" and "a \<in> A"
shows "\<exists>a'\<in>A'. a .= a'"
-using assms
-unfolding set_eq_def elem_def
-by fast
+ using assms by (simp add: set_eq_def elem_def)
lemma set_eqD2:
fixes R (structure)
- assumes AA': "A {.=} A'"
- and "a' \<in> A'"
+ assumes "A {.=} A'" and "a' \<in> A'"
shows "\<exists>a\<in>A. a' .= a"
-using assms
-unfolding set_eq_def elem_def
-by fast
+ using assms by (simp add: set_eq_def elem_def)
lemma set_eqE:
fixes R (structure)
- assumes AB: "A {.=} B"
- and r: "\<lbrakk>\<forall>a\<in>A. a .\<in> B; \<forall>b\<in>B. b .\<in> A\<rbrakk> \<Longrightarrow> P"
+ assumes "A {.=} B"
+ and "\<lbrakk> \<forall>a \<in> A. a .\<in> B; \<forall>b \<in> B. b .\<in> A \<rbrakk> \<Longrightarrow> P"
shows "P"
-using AB
-unfolding set_eq_def
-by (blast dest: r)
+ using assms unfolding set_eq_def by blast
lemma set_eqE2:
fixes R (structure)
- assumes AB: "A {.=} B"
- and r: "\<lbrakk>\<forall>a\<in>A. (\<exists>b\<in>B. a .= b); \<forall>b\<in>B. (\<exists>a\<in>A. b .= a)\<rbrakk> \<Longrightarrow> P"
+ assumes "A {.=} B"
+ and "\<lbrakk> \<forall>a \<in> A. \<exists>b \<in> B. a .= b; \<forall>b \<in> B. \<exists>a \<in> A. b .= a \<rbrakk> \<Longrightarrow> P"
shows "P"
-using AB
-unfolding set_eq_def elem_def
-by (blast dest: r)
+ using assms unfolding set_eq_def by (simp add: elem_def)
lemma set_eqE':
fixes R (structure)
- assumes AB: "A {.=} B"
- and aA: "a \<in> A" and bB: "b \<in> B"
- and r: "\<And>a' b'. \<lbrakk>a' \<in> A; b .= a'; b' \<in> B; a .= b'\<rbrakk> \<Longrightarrow> P"
+ assumes "A {.=} B" "a \<in> A" "b \<in> B"
+ and "\<And>a' b'. \<lbrakk> a' \<in> A; b' \<in> B \<rbrakk> \<Longrightarrow> b .= a' \<Longrightarrow> a .= b' \<Longrightarrow> P"
shows "P"
-proof -
- from AB aA
- have "\<exists>b'\<in>B. a .= b'" by (rule set_eqD1)
- from this obtain b'
- where b': "b' \<in> B" "a .= b'" by auto
-
- from AB bB
- have "\<exists>a'\<in>A. b .= a'" by (rule set_eqD2)
- from this obtain a'
- where a': "a' \<in> A" "b .= a'" by auto
-
- from a' b'
- show "P" by (rule r)
-qed
+ using assms by (meson set_eqE2)
lemma (in equivalence) eq_elem_cong_r [trans]:
- assumes a: "a .\<in> A"
- and cong: "A {.=} A'"
- and carr: "a \<in> carrier S"
- and Carr: "A \<subseteq> carrier S" "A' \<subseteq> carrier S"
- shows "a .\<in> A'"
-using a cong
-proof (elim elemE set_eqE)
- fix b
- assume bA: "b \<in> A"
- and inA': "\<forall>b\<in>A. b .\<in> A'"
- note [simp] = carr Carr Carr[THEN subsetD] bA
- assume "a .= b"
- also from bA inA'
- have "b .\<in> A'" by fast
- finally
- show "a .\<in> A'" by simp
-qed
+ assumes "A \<subseteq> carrier S" "A' \<subseteq> carrier S" "A {.=} A'"
+ shows "\<lbrakk> a \<in> carrier S \<rbrakk> \<Longrightarrow> a .\<in> A \<Longrightarrow> a .\<in> A'"
+ using assms by (meson elemE elem_cong_l set_eqE subset_eq)
lemma (in equivalence) set_eq_sym [sym]:
- assumes "A {.=} B"
- shows "B {.=} A"
-using assms
-unfolding set_eq_def elem_def
-by fast
+ assumes "A \<subseteq> carrier S" "B \<subseteq> carrier S"
+ shows "A {.=} B \<Longrightarrow> B {.=} A"
+ using assms unfolding set_eq_def elem_def by auto
lemma (in equivalence) equal_set_eq_trans [trans]:
- assumes AB: "A = B" and BC: "B {.=} C"
- shows "A {.=} C"
- using AB BC by simp
+ "\<lbrakk> A = B; B {.=} C \<rbrakk> \<Longrightarrow> A {.=} C"
+ by simp
lemma (in equivalence) set_eq_equal_trans [trans]:
- assumes AB: "A {.=} B" and BC: "B = C"
- shows "A {.=} C"
- using AB BC by simp
+ "\<lbrakk> A {.=} B; B = C \<rbrakk> \<Longrightarrow> A {.=} C"
+ by simp
-lemma (in equivalence) set_eq_trans [trans]:
- assumes AB: "A {.=} B" and BC: "B {.=} C"
- and carr: "A \<subseteq> carrier S" "B \<subseteq> carrier S" "C \<subseteq> carrier S"
+lemma (in equivalence) set_eq_trans_aux:
+ assumes "A \<subseteq> carrier S" "B \<subseteq> carrier S" "C \<subseteq> carrier S"
+ and "A {.=} B" "B {.=} C"
+ shows "\<And>a. a \<in> A \<Longrightarrow> a .\<in> C"
+ using assms by (simp add: eq_elem_cong_r subset_iff)
+
+corollary (in equivalence) set_eq_trans [trans]:
+ assumes "A \<subseteq> carrier S" "B \<subseteq> carrier S" "C \<subseteq> carrier S"
+ and "A {.=} B" " B {.=} C"
shows "A {.=} C"
proof (intro set_eqI)
- fix a
- assume aA: "a \<in> A"
- with carr have "a \<in> carrier S" by fast
- note [simp] = carr this
-
- from aA
- have "a .\<in> A" by (simp add: elem_exact)
- also note AB
- also note BC
- finally
- show "a .\<in> C" by simp
+ show "\<And>a. a \<in> A \<Longrightarrow> a .\<in> C" using set_eq_trans_aux assms by blast
next
- fix c
- assume cC: "c \<in> C"
- with carr have "c \<in> carrier S" by fast
- note [simp] = carr this
-
- from cC
- have "c .\<in> C" by (simp add: elem_exact)
- also note BC[symmetric]
- also note AB[symmetric]
- finally
- show "c .\<in> A" by simp
+ show "\<And>b. b \<in> C \<Longrightarrow> b .\<in> A" using set_eq_trans_aux set_eq_sym assms by blast
qed
-lemma (in equivalence) set_eq_insert_aux:
- assumes x: "x .= x'"
- and carr: "x \<in> carrier S" "x' \<in> carrier S" "A \<subseteq> carrier S"
- shows "\<forall>a \<in> (insert x A). a .\<in> (insert x' A)"
-proof
- fix a
- show "a \<in> insert x A \<Longrightarrow> a .\<in> insert x' A"
- proof cases
- assume "a \<in> A"
- thus "a .\<in> insert x' A"
- using carr(3) by blast
- next
- assume "a \<in> insert x A" "a \<notin> A"
- hence "a = x"
- by blast
- thus "a .\<in> insert x' A"
- by (meson x elemI insertI1)
- qed
-qed
-
-lemma (in equivalence) set_eq_insert:
- assumes x: "x .= x'"
- and carr: "x \<in> carrier S" "x' \<in> carrier S" "A \<subseteq> carrier S"
- shows "insert x A {.=} insert x' A"
-proof-
- have "(\<forall>a \<in> (insert x A). a .\<in> (insert x' A)) \<and>
- (\<forall>a \<in> (insert x' A). a .\<in> (insert x A))"
- using set_eq_insert_aux carr x sym by blast
- thus "insert x A {.=} insert x' A"
- using set_eq_def by blast
-qed
-
-lemma (in equivalence) set_eq_pairI:
- assumes xx': "x .= x'"
- and carr: "x \<in> carrier S" "x' \<in> carrier S" "y \<in> carrier S"
- shows "{x, y} {.=} {x', y}"
- using assms set_eq_insert by simp
-
lemma (in equivalence) is_closedI:
assumes closed: "\<And>x y. \<lbrakk>x .= y; x \<in> A; y \<in> carrier S\<rbrakk> \<Longrightarrow> y \<in> A"
and S: "A \<subseteq> carrier S"
@@ -307,29 +224,28 @@
by (blast dest: closed sym)
lemma (in equivalence) closure_of_eq:
- "\<lbrakk>x .= x'; A \<subseteq> carrier S; x \<in> closure_of A; x' \<in> carrier S\<rbrakk> \<Longrightarrow> x' \<in> closure_of A"
- unfolding eq_closure_of_def elem_def
- by (blast intro: trans sym)
+ assumes "A \<subseteq> carrier S" "x \<in> closure_of A"
+ shows "\<lbrakk> x' \<in> carrier S; x .= x' \<rbrakk> \<Longrightarrow> x' \<in> closure_of A"
+ using assms elem_cong_l sym unfolding eq_closure_of_def by blast
lemma (in equivalence) is_closed_eq [dest]:
- "\<lbrakk>x .= x'; x \<in> A; is_closed A; x \<in> carrier S; x' \<in> carrier S\<rbrakk> \<Longrightarrow> x' \<in> A"
- unfolding eq_is_closed_def
- using closure_of_eq [where A = A]
- by simp
+ assumes "is_closed A" "x \<in> A"
+ shows "\<lbrakk> x .= x'; x' \<in> carrier S \<rbrakk> \<Longrightarrow> x' \<in> A"
+ using assms closure_of_eq [where A = A] unfolding eq_is_closed_def by simp
-lemma (in equivalence) is_closed_eq_rev [dest]:
- "\<lbrakk>x .= x'; x' \<in> A; is_closed A; x \<in> carrier S; x' \<in> carrier S\<rbrakk> \<Longrightarrow> x \<in> A"
- by (meson subsetD eq_is_closed_def is_closed_eq sym)
+corollary (in equivalence) is_closed_eq_rev [dest]:
+ assumes "is_closed A" "x' \<in> A"
+ shows "\<lbrakk> x .= x'; x \<in> carrier S \<rbrakk> \<Longrightarrow> x \<in> A"
+ using sym is_closed_eq assms by (meson contra_subsetD eq_is_closed_def)
lemma closure_of_closed [simp, intro]:
fixes S (structure)
shows "closure_of A \<subseteq> carrier S"
-unfolding eq_closure_of_def
-by fast
+ unfolding eq_closure_of_def by auto
lemma closure_of_memI:
fixes S (structure)
- assumes "a .\<in> A" and "a \<in> carrier S"
+ assumes "a .\<in> A" "a \<in> carrier S"
shows "a \<in> closure_of A"
by (simp add: assms eq_closure_of_def)
@@ -351,67 +267,193 @@
assumes "a \<in> closure_of A"
and "\<And>a'. \<lbrakk>a \<in> carrier S; a' \<in> A; a .= a'\<rbrakk> \<Longrightarrow> P"
shows "P"
- by (meson closure_of_memE elemE assms)
+ using assms by (meson closure_of_memE elemE)
+
+
+(* Lemmas by Paulo EmÃlio de Vilhena *)
+
+lemma (in partition) equivalence_from_partition:
+ "equivalence \<lparr> carrier = A, eq = (\<lambda>x y. y \<in> (THE b. b \<in> B \<and> x \<in> b))\<rparr>"
+ unfolding partition_def equivalence_def
+proof (auto)
+ let ?f = "\<lambda>x. THE b. b \<in> B \<and> x \<in> b"
+ show "\<And>x. x \<in> A \<Longrightarrow> x \<in> ?f x"
+ using unique_class by (metis (mono_tags, lifting) theI')
+ show "\<And>x y. \<lbrakk> x \<in> A; y \<in> A \<rbrakk> \<Longrightarrow> y \<in> ?f x \<Longrightarrow> x \<in> ?f y"
+ using unique_class by (metis (mono_tags, lifting) the_equality)
+ show "\<And>x y z. \<lbrakk> x \<in> A; y \<in> A; z \<in> A \<rbrakk> \<Longrightarrow> y \<in> ?f x \<Longrightarrow> z \<in> ?f y \<Longrightarrow> z \<in> ?f x"
+ using unique_class by (metis (mono_tags, lifting) the_equality)
+qed
+
+lemma (in partition) partition_coverture: "\<Union>B = A"
+ by (meson Sup_le_iff UnionI unique_class incl subsetI subset_antisym)
+
+lemma (in partition) disjoint_union:
+ assumes "b1 \<in> B" "b2 \<in> B"
+ and "b1 \<inter> b2 \<noteq> {}"
+ shows "b1 = b2"
+proof (rule ccontr)
+ assume "b1 \<noteq> b2"
+ obtain a where "a \<in> A" "a \<in> b1" "a \<in> b2"
+ using assms(2-3) incl by blast
+ thus False using unique_class \<open>b1 \<noteq> b2\<close> assms(1) assms(2) by blast
+qed
+
+lemma partitionI:
+ fixes A :: "'a set" and B :: "('a set) set"
+ assumes "\<Union>B = A"
+ and "\<And>b1 b2. \<lbrakk> b1 \<in> B; b2 \<in> B \<rbrakk> \<Longrightarrow> b1 \<inter> b2 \<noteq> {} \<Longrightarrow> b1 = b2"
+ shows "partition A B"
+proof
+ show "\<And>a. a \<in> A \<Longrightarrow> \<exists>!b. b \<in> B \<and> a \<in> b"
+ proof (rule ccontr)
+ fix a assume "a \<in> A" "\<nexists>!b. b \<in> B \<and> a \<in> b"
+ then obtain b1 b2 where "b1 \<in> B" "a \<in> b1"
+ and "b2 \<in> B" "a \<in> b2" "b1 \<noteq> b2" using assms(1) by blast
+ thus False using assms(2) by blast
+ qed
+next
+ show "\<And>b. b \<in> B \<Longrightarrow> b \<subseteq> A" using assms(1) by blast
+qed
+
+lemma (in partition) remove_elem:
+ assumes "b \<in> B"
+ shows "partition (A - b) (B - {b})"
+proof
+ show "\<And>a. a \<in> A - b \<Longrightarrow> \<exists>!b'. b' \<in> B - {b} \<and> a \<in> b'"
+ using unique_class by fastforce
+next
+ show "\<And>b'. b' \<in> B - {b} \<Longrightarrow> b' \<subseteq> A - b"
+ using assms unique_class incl partition_axioms partition_coverture by fastforce
+qed
+
+lemma disjoint_sum:
+ "\<lbrakk> finite B; finite A; partition A B\<rbrakk> \<Longrightarrow> (\<Sum>b\<in>B. \<Sum>a\<in>b. f a) = (\<Sum>a\<in>A. f a)"
+proof (induct arbitrary: A set: finite)
+ case empty thus ?case using partition.unique_class by fastforce
+next
+ case (insert b B')
+ have "(\<Sum>b\<in>(insert b B'). \<Sum>a\<in>b. f a) = (\<Sum>a\<in>b. f a) + (\<Sum>b\<in>B'. \<Sum>a\<in>b. f a)"
+ by (simp add: insert.hyps(1) insert.hyps(2))
+ also have "... = (\<Sum>a\<in>b. f a) + (\<Sum>a\<in>(A - b). f a)"
+ using partition.remove_elem[of A "insert b B'" b] insert.hyps insert.prems
+ by (metis Diff_insert_absorb finite_Diff insert_iff)
+ finally show "(\<Sum>b\<in>(insert b B'). \<Sum>a\<in>b. f a) = (\<Sum>a\<in>A. f a)"
+ using partition.remove_elem[of A "insert b B'" b] insert.prems
+ by (metis add.commute insert_iff partition.incl sum.subset_diff)
+qed
+
+lemma (in partition) disjoint_sum:
+ assumes "finite A"
+ shows "(\<Sum>b\<in>B. \<Sum>a\<in>b. f a) = (\<Sum>a\<in>A. f a)"
+proof -
+ have "finite B"
+ by (simp add: assms finite_UnionD partition_coverture)
+ thus ?thesis using disjoint_sum assms partition_axioms by blast
+qed
+
+lemma (in equivalence) set_eq_insert_aux:
+ assumes "A \<subseteq> carrier S"
+ and "x \<in> carrier S" "x' \<in> carrier S" "x .= x'"
+ and "y \<in> insert x A"
+ shows "y .\<in> insert x' A"
+ by (metis assms(1) assms(4) assms(5) contra_subsetD elemI elem_exact insert_iff)
+
+corollary (in equivalence) set_eq_insert:
+ assumes "A \<subseteq> carrier S"
+ and "x \<in> carrier S" "x' \<in> carrier S" "x .= x'"
+ shows "insert x A {.=} insert x' A"
+ by (meson set_eqI assms set_eq_insert_aux sym equivalence_axioms)
+
+lemma (in equivalence) set_eq_pairI:
+ assumes xx': "x .= x'"
+ and carr: "x \<in> carrier S" "x' \<in> carrier S" "y \<in> carrier S"
+ shows "{x, y} {.=} {x', y}"
+ using assms set_eq_insert by simp
lemma (in equivalence) closure_inclusion:
assumes "A \<subseteq> B"
shows "closure_of A \<subseteq> closure_of B"
- unfolding eq_closure_of_def
-proof
- fix x
- assume "x \<in> {y \<in> carrier S. y .\<in> A}"
- hence "x \<in> carrier S \<and> x .\<in> A"
- by blast
- hence "x \<in> carrier S \<and> x .\<in> B"
- using assms elem_subsetD by blast
- thus "x \<in> {y \<in> carrier S. y .\<in> B}"
- by simp
-qed
+ unfolding eq_closure_of_def using assms elem_subsetD by auto
lemma (in equivalence) classes_small:
assumes "is_closed B"
and "A \<subseteq> B"
shows "closure_of A \<subseteq> B"
-proof-
- have "closure_of A \<subseteq> closure_of B"
- using closure_inclusion assms by simp
- thus "closure_of A \<subseteq> B"
- using assms(1) eq_is_closed_def by fastforce
-qed
+ by (metis assms closure_inclusion eq_is_closed_def)
lemma (in equivalence) classes_eq:
assumes "A \<subseteq> carrier S"
shows "A {.=} closure_of A"
-using assms
-by (blast intro: set_eqI elem_exact closure_of_memI elim: closure_of_memE)
+ using assms by (blast intro: set_eqI elem_exact closure_of_memI elim: closure_of_memE)
lemma (in equivalence) complete_classes:
- assumes c: "is_closed A"
+ assumes "is_closed A"
shows "A = closure_of A"
using assms by (simp add: eq_is_closed_def)
-lemma (in equivalence) closure_idemp_weak:
+lemma (in equivalence) closure_idem_weak:
"closure_of (closure_of A) {.=} closure_of A"
by (simp add: classes_eq set_eq_sym)
-lemma (in equivalence) closure_idemp_strong:
+lemma (in equivalence) closure_idem_strong:
assumes "A \<subseteq> carrier S"
shows "closure_of (closure_of A) = closure_of A"
using assms closure_of_eq complete_classes is_closedI by auto
-lemma (in equivalence) complete_classes2:
+lemma (in equivalence) classes_consistent:
assumes "A \<subseteq> carrier S"
shows "is_closed (closure_of A)"
- using closure_idemp_strong by (simp add: assms eq_is_closed_def)
+ using closure_idem_strong by (simp add: assms eq_is_closed_def)
+
+lemma (in equivalence) classes_coverture:
+ "\<Union>classes = carrier S"
+proof
+ show "\<Union>classes \<subseteq> carrier S"
+ unfolding eq_classes_def eq_class_of_def by blast
+next
+ show "carrier S \<subseteq> \<Union>classes" unfolding eq_classes_def eq_class_of_def
+ proof
+ fix x assume "x \<in> carrier S"
+ hence "x \<in> {y \<in> carrier S. x .= y}" using refl by simp
+ thus "x \<in> \<Union>{{y \<in> carrier S. x .= y} |x. x \<in> carrier S}" by blast
+ qed
+qed
-lemma equivalence_subset:
- assumes "equivalence L" "A \<subseteq> carrier L"
- shows "equivalence (L\<lparr> carrier := A \<rparr>)"
+lemma (in equivalence) disjoint_union:
+ assumes "class1 \<in> classes" "class2 \<in> classes"
+ and "class1 \<inter> class2 \<noteq> {}"
+ shows "class1 = class2"
proof -
- interpret L: equivalence L
- by (simp add: assms)
- show ?thesis
- by (unfold_locales, simp_all add: L.sym assms rev_subsetD, meson L.trans assms(2) contra_subsetD)
+ obtain x y where x: "x \<in> carrier S" "class1 = class_of x"
+ and y: "y \<in> carrier S" "class2 = class_of y"
+ using assms(1-2) unfolding eq_classes_def by blast
+ obtain z where z: "z \<in> carrier S" "z \<in> class1 \<inter> class2"
+ using assms classes_coverture by fastforce
+ hence "x .= z \<and> y .= z" using x y unfolding eq_class_of_def by blast
+ hence "x .= y" using x y z trans sym by meson
+ hence "class_of x = class_of y"
+ unfolding eq_class_of_def using local.sym trans x y by blast
+ thus ?thesis using x y by simp
+qed
+
+lemma (in equivalence) partition_from_equivalence:
+ "partition (carrier S) classes"
+proof (intro partitionI)
+ show "\<Union>classes = carrier S" using classes_coverture by simp
+next
+ show "\<And>class1 class2. \<lbrakk> class1 \<in> classes; class2 \<in> classes \<rbrakk> \<Longrightarrow>
+ class1 \<inter> class2 \<noteq> {} \<Longrightarrow> class1 = class2"
+ using disjoint_union by simp
+qed
+
+lemma (in equivalence) disjoint_sum:
+ assumes "finite (carrier S)"
+ shows "(\<Sum>c\<in>classes. \<Sum>x\<in>c. f x) = (\<Sum>x\<in>(carrier S). f x)"
+proof -
+ have "finite classes"
+ unfolding eq_classes_def using assms by auto
+ thus ?thesis using disjoint_sum assms partition_from_equivalence by blast
qed
end
--- a/src/HOL/Algebra/Coset.thy Thu Jun 14 10:51:12 2018 +0200
+++ b/src/HOL/Algebra/Coset.thy Thu Jun 14 14:23:48 2018 +0100
@@ -1,7 +1,6 @@
(* Title: HOL/Algebra/Coset.thy
- Author: Florian Kammueller
- Author: L C Paulson
- Author: Stephan Hohe
+ Authors: Florian Kammueller, L C Paulson, Stephan Hohe
+With additional contributions from Martin Baillon and Paulo EmÃlio de Vilhena.
*)
theory Coset
@@ -38,346 +37,321 @@
normal_rel :: "['a set, ('a, 'b) monoid_scheme] \<Rightarrow> bool" (infixl "\<lhd>" 60) where
"H \<lhd> G \<equiv> normal H G"
+(* ************************************************************************** *)
+(* Next two lemmas contributed by Martin Baillon. *)
+
+lemma l_coset_eq_set_mult:
+ fixes G (structure)
+ shows "x <# H = {x} <#> H"
+ unfolding l_coset_def set_mult_def by simp
+
+lemma r_coset_eq_set_mult:
+ fixes G (structure)
+ shows "H #> x = H <#> {x}"
+ unfolding r_coset_def set_mult_def by simp
+
+(* ************************************************************************** *)
+
+
+(* ************************************************************************** *)
+(* Next five lemmas contributed by Paulo EmÃlio de Vilhena. *)
+
+lemma (in subgroup) rcosets_not_empty:
+ assumes "R \<in> rcosets H"
+ shows "R \<noteq> {}"
+proof -
+ obtain g where "g \<in> carrier G" "R = H #> g"
+ using assms unfolding RCOSETS_def by blast
+ hence "\<one> \<otimes> g \<in> R"
+ using one_closed unfolding r_coset_def by blast
+ thus ?thesis by blast
+qed
+
+lemma (in group) diff_neutralizes:
+ assumes "subgroup H G" "R \<in> rcosets H"
+ shows "\<And>r1 r2. \<lbrakk> r1 \<in> R; r2 \<in> R \<rbrakk> \<Longrightarrow> r1 \<otimes> (inv r2) \<in> H"
+proof -
+ fix r1 r2 assume r1: "r1 \<in> R" and r2: "r2 \<in> R"
+ obtain g where g: "g \<in> carrier G" "R = H #> g"
+ using assms unfolding RCOSETS_def by blast
+ then obtain h1 h2 where h1: "h1 \<in> H" "r1 = h1 \<otimes> g"
+ and h2: "h2 \<in> H" "r2 = h2 \<otimes> g"
+ using r1 r2 unfolding r_coset_def by blast
+ hence "r1 \<otimes> (inv r2) = (h1 \<otimes> g) \<otimes> ((inv g) \<otimes> (inv h2))"
+ using inv_mult_group is_group assms(1) g(1) subgroup.mem_carrier by fastforce
+ also have " ... = (h1 \<otimes> (g \<otimes> inv g) \<otimes> inv h2)"
+ using h1 h2 assms(1) g(1) inv_closed m_closed monoid.m_assoc
+ monoid_axioms subgroup.mem_carrier by smt
+ finally have "r1 \<otimes> inv r2 = h1 \<otimes> inv h2"
+ using assms(1) g(1) h1(1) subgroup.mem_carrier by fastforce
+ thus "r1 \<otimes> inv r2 \<in> H" by (metis assms(1) h1(1) h2(1) subgroup_def)
+qed
+
+
+subsection \<open>Stable Operations for Subgroups\<close>
+
+lemma (in group) subgroup_set_mult_equality[simp]:
+ assumes "subgroup H G" "I \<subseteq> H" "J \<subseteq> H"
+ shows "I <#>\<^bsub>G \<lparr> carrier := H \<rparr>\<^esub> J = I <#> J"
+ unfolding set_mult_def subgroup_mult_equality[OF assms(1)] by auto
+
+lemma (in group) subgroup_rcos_equality[simp]:
+ assumes "subgroup H G" "I \<subseteq> H" "h \<in> H"
+ shows "I #>\<^bsub>G \<lparr> carrier := H \<rparr>\<^esub> h = I #> h"
+ using subgroup_set_mult_equality by (simp add: r_coset_eq_set_mult assms)
+
+lemma (in group) subgroup_lcos_equality[simp]:
+ assumes "subgroup H G" "I \<subseteq> H" "h \<in> H"
+ shows "h <#\<^bsub>G \<lparr> carrier := H \<rparr>\<^esub> I = h <# I"
+ using subgroup_set_mult_equality by (simp add: l_coset_eq_set_mult assms)
+
+
+
+subsection \<open>Basic Properties of set multiplication\<close>
+
+lemma (in group) setmult_subset_G:
+ assumes "H \<subseteq> carrier G" "K \<subseteq> carrier G"
+ shows "H <#> K \<subseteq> carrier G" using assms
+ by (auto simp add: set_mult_def subsetD)
+
+lemma (in monoid) set_mult_closed:
+ assumes "H \<subseteq> carrier G" "K \<subseteq> carrier G"
+ shows "H <#> K \<subseteq> carrier G"
+ using assms by (auto simp add: set_mult_def subsetD)
+
+(* Next lemma contributed by Martin Baillon.*)
+lemma (in group) set_mult_assoc:
+ assumes "M \<subseteq> carrier G" "H \<subseteq> carrier G" "K \<subseteq> carrier G"
+ shows "(M <#> H) <#> K = M <#> (H <#> K)"
+proof
+ show "(M <#> H) <#> K \<subseteq> M <#> (H <#> K)"
+ proof
+ fix x assume "x \<in> (M <#> H) <#> K"
+ then obtain m h k where x: "m \<in> M" "h \<in> H" "k \<in> K" "x = (m \<otimes> h) \<otimes> k"
+ unfolding set_mult_def by blast
+ hence "x = m \<otimes> (h \<otimes> k)"
+ using assms m_assoc by blast
+ thus "x \<in> M <#> (H <#> K)"
+ unfolding set_mult_def using x by blast
+ qed
+next
+ show "M <#> (H <#> K) \<subseteq> (M <#> H) <#> K"
+ proof
+ fix x assume "x \<in> M <#> (H <#> K)"
+ then obtain m h k where x: "m \<in> M" "h \<in> H" "k \<in> K" "x = m \<otimes> (h \<otimes> k)"
+ unfolding set_mult_def by blast
+ hence "x = (m \<otimes> h) \<otimes> k"
+ using assms m_assoc rev_subsetD by metis
+ thus "x \<in> (M <#> H) <#> K"
+ unfolding set_mult_def using x by blast
+ qed
+qed
+
+
subsection \<open>Basic Properties of Cosets\<close>
lemma (in group) coset_mult_assoc:
- "[| M \<subseteq> carrier G; g \<in> carrier G; h \<in> carrier G |]
- ==> (M #> g) #> h = M #> (g \<otimes> h)"
-by (force simp add: r_coset_def m_assoc)
+ assumes "M \<subseteq> carrier G" "g \<in> carrier G" "h \<in> carrier G"
+ shows "(M #> g) #> h = M #> (g \<otimes> h)"
+ using assms by (force simp add: r_coset_def m_assoc)
+
+lemma (in group) coset_assoc:
+ assumes "x \<in> carrier G" "y \<in> carrier G" "H \<subseteq> carrier G"
+ shows "x <# (H #> y) = (x <# H) #> y"
+ using set_mult_assoc[of "{x}" H "{y}"]
+ by (simp add: l_coset_eq_set_mult r_coset_eq_set_mult assms)
lemma (in group) coset_mult_one [simp]: "M \<subseteq> carrier G ==> M #> \<one> = M"
by (force simp add: r_coset_def)
lemma (in group) coset_mult_inv1:
- "[| M #> (x \<otimes> (inv y)) = M; x \<in> carrier G ; y \<in> carrier G;
- M \<subseteq> carrier G |] ==> M #> x = M #> y"
-apply (erule subst [of concl: "%z. M #> x = z #> y"])
-apply (simp add: coset_mult_assoc m_assoc)
-done
+ assumes "M #> (x \<otimes> (inv y)) = M"
+ and "x \<in> carrier G" "y \<in> carrier G" "M \<subseteq> carrier G"
+ shows "M #> x = M #> y" using assms
+ by (metis coset_mult_assoc group.inv_solve_right is_group subgroup_def subgroup_self)
lemma (in group) coset_mult_inv2:
- "[| M #> x = M #> y; x \<in> carrier G; y \<in> carrier G; M \<subseteq> carrier G |]
- ==> M #> (x \<otimes> (inv y)) = M "
-apply (simp add: coset_mult_assoc [symmetric])
-apply (simp add: coset_mult_assoc)
-done
+ assumes "M #> x = M #> y"
+ and "x \<in> carrier G" "y \<in> carrier G" "M \<subseteq> carrier G"
+ shows "M #> (x \<otimes> (inv y)) = M " using assms
+ by (metis group.coset_mult_assoc group.coset_mult_one inv_closed is_group r_inv)
lemma (in group) coset_join1:
- "[| H #> x = H; x \<in> carrier G; subgroup H G |] ==> x \<in> H"
-apply (erule subst)
-apply (simp add: r_coset_def)
-apply (blast intro: l_one subgroup.one_closed sym)
-done
+ assumes "H #> x = H"
+ and "x \<in> carrier G" "subgroup H G"
+ shows "x \<in> H"
+ using assms r_coset_def l_one subgroup.one_closed sym by fastforce
lemma (in group) solve_equation:
- "\<lbrakk>subgroup H G; x \<in> H; y \<in> H\<rbrakk> \<Longrightarrow> \<exists>h\<in>H. y = h \<otimes> x"
-apply (rule bexI [of _ "y \<otimes> (inv x)"])
-apply (auto simp add: subgroup.m_closed subgroup.m_inv_closed m_assoc
- subgroup.subset [THEN subsetD])
-done
+ assumes "subgroup H G" "x \<in> H" "y \<in> H"
+ shows "\<exists>h \<in> H. y = h \<otimes> x"
+proof -
+ have "y = (y \<otimes> (inv x)) \<otimes> x" using assms
+ by (simp add: m_assoc subgroup.mem_carrier)
+ moreover have "y \<otimes> (inv x) \<in> H" using assms
+ by (simp add: subgroup_def)
+ ultimately show ?thesis by blast
+qed
lemma (in group) repr_independence:
- "\<lbrakk>y \<in> H #> x; x \<in> carrier G; subgroup H G\<rbrakk> \<Longrightarrow> H #> x = H #> y"
+ assumes "y \<in> H #> x" "x \<in> carrier G" "subgroup H G"
+ shows "H #> x = H #> y" using assms
by (auto simp add: r_coset_def m_assoc [symmetric]
subgroup.subset [THEN subsetD]
subgroup.m_closed solve_equation)
lemma (in group) coset_join2:
- "\<lbrakk>x \<in> carrier G; subgroup H G; x\<in>H\<rbrakk> \<Longrightarrow> H #> x = H"
+ assumes "x \<in> carrier G" "subgroup H G" "x \<in> H"
+ shows "H #> x = H" using assms
\<comment> \<open>Alternative proof is to put @{term "x=\<one>"} in \<open>repr_independence\<close>.\<close>
by (force simp add: subgroup.m_closed r_coset_def solve_equation)
+lemma (in group) coset_join3:
+ assumes "x \<in> carrier G" "subgroup H G" "x \<in> H"
+ shows "x <# H = H"
+proof
+ have "\<And>h. h \<in> H \<Longrightarrow> x \<otimes> h \<in> H" using assms
+ by (simp add: subgroup.m_closed)
+ thus "x <# H \<subseteq> H" unfolding l_coset_def by blast
+next
+ have "\<And>h. h \<in> H \<Longrightarrow> x \<otimes> ((inv x) \<otimes> h) = h"
+ by (smt assms inv_closed l_one m_assoc r_inv subgroup.mem_carrier)
+ moreover have "\<And>h. h \<in> H \<Longrightarrow> (inv x) \<otimes> h \<in> H"
+ by (simp add: assms subgroup.m_closed subgroup.m_inv_closed)
+ ultimately show "H \<subseteq> x <# H" unfolding l_coset_def by blast
+qed
+
lemma (in monoid) r_coset_subset_G:
- "[| H \<subseteq> carrier G; x \<in> carrier G |] ==> H #> x \<subseteq> carrier G"
+ "\<lbrakk> H \<subseteq> carrier G; x \<in> carrier G \<rbrakk> \<Longrightarrow> H #> x \<subseteq> carrier G"
by (auto simp add: r_coset_def)
lemma (in group) rcosI:
- "[| h \<in> H; H \<subseteq> carrier G; x \<in> carrier G|] ==> h \<otimes> x \<in> H #> x"
+ "\<lbrakk> h \<in> H; H \<subseteq> carrier G; x \<in> carrier G \<rbrakk> \<Longrightarrow> h \<otimes> x \<in> H #> x"
by (auto simp add: r_coset_def)
lemma (in group) rcosetsI:
"\<lbrakk>H \<subseteq> carrier G; x \<in> carrier G\<rbrakk> \<Longrightarrow> H #> x \<in> rcosets H"
by (auto simp add: RCOSETS_def)
-text\<open>Really needed?\<close>
-lemma (in group) transpose_inv:
- "[| x \<otimes> y = z; x \<in> carrier G; y \<in> carrier G; z \<in> carrier G |]
- ==> (inv x) \<otimes> z = y"
-by (force simp add: m_assoc [symmetric])
-
-lemma (in group) rcos_self: "[| x \<in> carrier G; subgroup H G |] ==> x \<in> H #> x"
-apply (simp add: r_coset_def)
-apply (blast intro: sym l_one subgroup.subset [THEN subsetD]
- subgroup.one_closed)
-done
+lemma (in group) rcos_self:
+ "\<lbrakk> x \<in> carrier G; subgroup H G \<rbrakk> \<Longrightarrow> x \<in> H #> x"
+ by (metis l_one rcosI subgroup_def)
text (in group) \<open>Opposite of @{thm [source] "repr_independence"}\<close>
lemma (in group) repr_independenceD:
- assumes "subgroup H G"
- assumes ycarr: "y \<in> carrier G"
- and repr: "H #> x = H #> y"
+ assumes "subgroup H G" "y \<in> carrier G"
+ and "H #> x = H #> y"
shows "y \<in> H #> x"
-proof -
- interpret subgroup H G by fact
- show ?thesis apply (subst repr)
- apply (intro rcos_self)
- apply (rule ycarr)
- apply (rule is_subgroup)
- done
-qed
+ using assms by (simp add: rcos_self)
text \<open>Elements of a right coset are in the carrier\<close>
lemma (in subgroup) elemrcos_carrier:
- assumes "group G"
- assumes acarr: "a \<in> carrier G"
- and a': "a' \<in> H #> a"
+ assumes "group G" "a \<in> carrier G"
+ and "a' \<in> H #> a"
shows "a' \<in> carrier G"
-proof -
- interpret group G by fact
- from subset and acarr
- have "H #> a \<subseteq> carrier G" by (rule r_coset_subset_G)
- from this and a'
- show "a' \<in> carrier G"
- by fast
-qed
+ by (meson assms group.is_monoid monoid.r_coset_subset_G subset subsetCE)
lemma (in subgroup) rcos_const:
- assumes "group G"
- assumes hH: "h \<in> H"
+ assumes "group G" "h \<in> H"
shows "H #> h = H"
-proof -
- interpret group G by fact
- show ?thesis apply (unfold r_coset_def)
- apply rule
- apply rule
- apply clarsimp
- apply (intro subgroup.m_closed)
- apply (rule is_subgroup)
- apply assumption
- apply (rule hH)
- apply rule
- apply simp
- proof -
- fix h'
- assume h'H: "h' \<in> H"
- note carr = hH[THEN mem_carrier] h'H[THEN mem_carrier]
- from carr
- have a: "h' = (h' \<otimes> inv h) \<otimes> h" by (simp add: m_assoc)
- from h'H hH
- have "h' \<otimes> inv h \<in> H" by simp
- from this and a
- show "\<exists>x\<in>H. h' = x \<otimes> h" by fast
- qed
-qed
+ using group.coset_join2[OF assms(1), of h H]
+ by (simp add: assms(2) subgroup_axioms)
-text \<open>Step one for lemma \<open>rcos_module\<close>\<close>
lemma (in subgroup) rcos_module_imp:
- assumes "group G"
- assumes xcarr: "x \<in> carrier G"
- and x'cos: "x' \<in> H #> x"
+ assumes "group G" "x \<in> carrier G"
+ and "x' \<in> H #> x"
shows "(x' \<otimes> inv x) \<in> H"
proof -
- interpret group G by fact
- from xcarr x'cos
- have x'carr: "x' \<in> carrier G"
- by (rule elemrcos_carrier[OF is_group])
- from xcarr
- have ixcarr: "inv x \<in> carrier G"
- by simp
- from x'cos
- have "\<exists>h\<in>H. x' = h \<otimes> x"
- unfolding r_coset_def
- by fast
- from this
- obtain h
- where hH: "h \<in> H"
- and x': "x' = h \<otimes> x"
- by auto
- from hH and subset
- have hcarr: "h \<in> carrier G" by fast
- note carr = xcarr x'carr hcarr
- from x' and carr
- have "x' \<otimes> (inv x) = (h \<otimes> x) \<otimes> (inv x)" by fast
- also from carr
- have "\<dots> = h \<otimes> (x \<otimes> inv x)" by (simp add: m_assoc)
- also from carr
- have "\<dots> = h \<otimes> \<one>" by simp
- also from carr
- have "\<dots> = h" by simp
- finally
- have "x' \<otimes> (inv x) = h" by simp
- from hH this
- show "x' \<otimes> (inv x) \<in> H" by simp
+ obtain h where h: "h \<in> H" "x' = h \<otimes> x"
+ using assms(3) unfolding r_coset_def by blast
+ hence "x' \<otimes> inv x = h"
+ by (metis assms elemrcos_carrier group.inv_solve_right mem_carrier)
+ thus ?thesis using h by blast
qed
-text \<open>Step two for lemma \<open>rcos_module\<close>\<close>
lemma (in subgroup) rcos_module_rev:
- assumes "group G"
- assumes carr: "x \<in> carrier G" "x' \<in> carrier G"
- and xixH: "(x' \<otimes> inv x) \<in> H"
+ assumes "group G" "x \<in> carrier G" "x' \<in> carrier G"
+ and "(x' \<otimes> inv x) \<in> H"
shows "x' \<in> H #> x"
proof -
- interpret group G by fact
- from xixH
- have "\<exists>h\<in>H. x' \<otimes> (inv x) = h" by fast
- from this
- obtain h
- where hH: "h \<in> H"
- and hsym: "x' \<otimes> (inv x) = h"
- by fast
- from hH subset have hcarr: "h \<in> carrier G" by simp
- note carr = carr hcarr
- from hsym[symmetric] have "h \<otimes> x = x' \<otimes> (inv x) \<otimes> x" by fast
- also from carr
- have "\<dots> = x' \<otimes> ((inv x) \<otimes> x)" by (simp add: m_assoc)
- also from carr
- have "\<dots> = x' \<otimes> \<one>" by simp
- also from carr
- have "\<dots> = x'" by simp
- finally
- have "h \<otimes> x = x'" by simp
- from this[symmetric] and hH
- show "x' \<in> H #> x"
- unfolding r_coset_def
- by fast
+ obtain h where h: "h \<in> H" "x' \<otimes> inv x = h"
+ using assms(4) unfolding r_coset_def by blast
+ hence "x' = h \<otimes> x"
+ by (metis assms group.inv_solve_right mem_carrier)
+ thus ?thesis using h unfolding r_coset_def by blast
qed
text \<open>Module property of right cosets\<close>
lemma (in subgroup) rcos_module:
- assumes "group G"
- assumes carr: "x \<in> carrier G" "x' \<in> carrier G"
+ assumes "group G" "x \<in> carrier G" "x' \<in> carrier G"
shows "(x' \<in> H #> x) = (x' \<otimes> inv x \<in> H)"
-proof -
- interpret group G by fact
- show ?thesis proof assume "x' \<in> H #> x"
- from this and carr
- show "x' \<otimes> inv x \<in> H"
- by (intro rcos_module_imp[OF is_group])
- next
- assume "x' \<otimes> inv x \<in> H"
- from this and carr
- show "x' \<in> H #> x"
- by (intro rcos_module_rev[OF is_group])
- qed
-qed
+ using rcos_module_rev rcos_module_imp assms by blast
text \<open>Right cosets are subsets of the carrier.\<close>
lemma (in subgroup) rcosets_carrier:
- assumes "group G"
- assumes XH: "X \<in> rcosets H"
+ assumes "group G" "X \<in> rcosets H"
shows "X \<subseteq> carrier G"
-proof -
- interpret group G by fact
- from XH have "\<exists>x\<in> carrier G. X = H #> x"
- unfolding RCOSETS_def
- by fast
- from this
- obtain x
- where xcarr: "x\<in> carrier G"
- and X: "X = H #> x"
- by fast
- from subset and xcarr
- show "X \<subseteq> carrier G"
- unfolding X
- by (rule r_coset_subset_G)
-qed
+ using assms elemrcos_carrier singletonD
+ subset_eq unfolding RCOSETS_def by force
+
text \<open>Multiplication of general subsets\<close>
-lemma (in monoid) set_mult_closed:
- assumes Acarr: "A \<subseteq> carrier G"
- and Bcarr: "B \<subseteq> carrier G"
- shows "A <#> B \<subseteq> carrier G"
-apply rule apply (simp add: set_mult_def, clarsimp)
-proof -
- fix a b
- assume "a \<in> A"
- from this and Acarr
- have acarr: "a \<in> carrier G" by fast
-
- assume "b \<in> B"
- from this and Bcarr
- have bcarr: "b \<in> carrier G" by fast
-
- from acarr bcarr
- show "a \<otimes> b \<in> carrier G" by (rule m_closed)
-qed
lemma (in comm_group) mult_subgroups:
- assumes subH: "subgroup H G"
- and subK: "subgroup K G"
+ assumes "subgroup H G" and "subgroup K G"
shows "subgroup (H <#> K) G"
-apply (rule subgroup.intro)
- apply (intro set_mult_closed subgroup.subset[OF subH] subgroup.subset[OF subK])
- apply (simp add: set_mult_def) apply clarsimp defer 1
- apply (simp add: set_mult_def) defer 1
- apply (simp add: set_mult_def, clarsimp) defer 1
-proof -
- fix ha hb ka kb
- assume haH: "ha \<in> H" and hbH: "hb \<in> H" and kaK: "ka \<in> K" and kbK: "kb \<in> K"
- note carr = haH[THEN subgroup.mem_carrier[OF subH]] hbH[THEN subgroup.mem_carrier[OF subH]]
- kaK[THEN subgroup.mem_carrier[OF subK]] kbK[THEN subgroup.mem_carrier[OF subK]]
- from carr
- have "(ha \<otimes> ka) \<otimes> (hb \<otimes> kb) = ha \<otimes> (ka \<otimes> hb) \<otimes> kb" by (simp add: m_assoc)
- also from carr
- have "\<dots> = ha \<otimes> (hb \<otimes> ka) \<otimes> kb" by (simp add: m_comm)
- also from carr
- have "\<dots> = (ha \<otimes> hb) \<otimes> (ka \<otimes> kb)" by (simp add: m_assoc)
- finally
- have eq: "(ha \<otimes> ka) \<otimes> (hb \<otimes> kb) = (ha \<otimes> hb) \<otimes> (ka \<otimes> kb)" .
-
- from haH hbH have hH: "ha \<otimes> hb \<in> H" by (simp add: subgroup.m_closed[OF subH])
- from kaK kbK have kK: "ka \<otimes> kb \<in> K" by (simp add: subgroup.m_closed[OF subK])
-
- from hH and kK and eq
- show "\<exists>h'\<in>H. \<exists>k'\<in>K. (ha \<otimes> ka) \<otimes> (hb \<otimes> kb) = h' \<otimes> k'" by fast
+proof (rule subgroup.intro)
+ show "H <#> K \<subseteq> carrier G"
+ by (simp add: setmult_subset_G assms subgroup_imp_subset)
+next
+ have "\<one> \<otimes> \<one> \<in> H <#> K"
+ unfolding set_mult_def using assms subgroup.one_closed by blast
+ thus "\<one> \<in> H <#> K" by simp
next
- have "\<one> = \<one> \<otimes> \<one>" by simp
- from subgroup.one_closed[OF subH] subgroup.one_closed[OF subK] this
- show "\<exists>h\<in>H. \<exists>k\<in>K. \<one> = h \<otimes> k" by fast
+ show "\<And>x. x \<in> H <#> K \<Longrightarrow> inv x \<in> H <#> K"
+ proof -
+ fix x assume "x \<in> H <#> K"
+ then obtain h k where hk: "h \<in> H" "k \<in> K" "x = h \<otimes> k"
+ unfolding set_mult_def by blast
+ hence "inv x = (inv k) \<otimes> (inv h)"
+ by (meson inv_mult_group assms subgroup.mem_carrier)
+ hence "inv x = (inv h) \<otimes> (inv k)"
+ by (metis hk inv_mult assms subgroup.mem_carrier)
+ thus "inv x \<in> H <#> K"
+ unfolding set_mult_def using hk assms
+ by (metis (no_types, lifting) UN_iff singletonI subgroup_def)
+ qed
next
- fix h k
- assume hH: "h \<in> H"
- and kK: "k \<in> K"
-
- from hH[THEN subgroup.mem_carrier[OF subH]] kK[THEN subgroup.mem_carrier[OF subK]]
- have "inv (h \<otimes> k) = inv h \<otimes> inv k" by (simp add: inv_mult_group m_comm)
-
- from subgroup.m_inv_closed[OF subH hH] and subgroup.m_inv_closed[OF subK kK] and this
- show "\<exists>ha\<in>H. \<exists>ka\<in>K. inv (h \<otimes> k) = ha \<otimes> ka" by fast
+ show "\<And>x y. x \<in> H <#> K \<Longrightarrow> y \<in> H <#> K \<Longrightarrow> x \<otimes> y \<in> H <#> K"
+ proof -
+ fix x y assume "x \<in> H <#> K" "y \<in> H <#> K"
+ then obtain h1 k1 h2 k2 where h1k1: "h1 \<in> H" "k1 \<in> K" "x = h1 \<otimes> k1"
+ and h2k2: "h2 \<in> H" "k2 \<in> K" "y = h2 \<otimes> k2"
+ unfolding set_mult_def by blast
+ hence "x \<otimes> y = (h1 \<otimes> k1) \<otimes> (h2 \<otimes> k2)" by simp
+ also have " ... = h1 \<otimes> (k1 \<otimes> h2) \<otimes> k2"
+ by (smt h1k1 h2k2 m_assoc m_closed assms subgroup.mem_carrier)
+ also have " ... = h1 \<otimes> (h2 \<otimes> k1) \<otimes> k2"
+ by (metis (no_types, hide_lams) assms m_comm h1k1(2) h2k2(1) subgroup.mem_carrier)
+ finally have "x \<otimes> y = (h1 \<otimes> h2) \<otimes> (k1 \<otimes> k2)"
+ by (smt assms h1k1 h2k2 m_assoc monoid.m_closed monoid_axioms subgroup.mem_carrier)
+ thus "x \<otimes> y \<in> H <#> K" unfolding set_mult_def
+ using subgroup.m_closed[OF assms(1) h1k1(1) h2k2(1)]
+ subgroup.m_closed[OF assms(2) h1k1(2) h2k2(2)] by blast
+ qed
qed
lemma (in subgroup) lcos_module_rev:
- assumes "group G"
- assumes carr: "x \<in> carrier G" "x' \<in> carrier G"
- and xixH: "(inv x \<otimes> x') \<in> H"
+ assumes "group G" "x \<in> carrier G" "x' \<in> carrier G"
+ and "(inv x \<otimes> x') \<in> H"
shows "x' \<in> x <# H"
proof -
- interpret group G by fact
- from xixH
- have "\<exists>h\<in>H. (inv x) \<otimes> x' = h" by fast
- from this
- obtain h
- where hH: "h \<in> H"
- and hsym: "(inv x) \<otimes> x' = h"
- by fast
-
- from hH subset have hcarr: "h \<in> carrier G" by simp
- note carr = carr hcarr
- from hsym[symmetric] have "x \<otimes> h = x \<otimes> ((inv x) \<otimes> x')" by fast
- also from carr
- have "\<dots> = (x \<otimes> (inv x)) \<otimes> x'" by (simp add: m_assoc[symmetric])
- also from carr
- have "\<dots> = \<one> \<otimes> x'" by simp
- also from carr
- have "\<dots> = x'" by simp
- finally
- have "x \<otimes> h = x'" by simp
-
- from this[symmetric] and hH
- show "x' \<in> x <# H"
- unfolding l_coset_def
- by fast
+ obtain h where h: "h \<in> H" "inv x \<otimes> x' = h"
+ using assms(4) unfolding l_coset_def by blast
+ hence "x' = x \<otimes> h"
+ by (metis assms group.inv_solve_left mem_carrier)
+ thus ?thesis using h unfolding l_coset_def by blast
qed
@@ -391,21 +365,21 @@
by (simp add: normal_def normal_axioms_def is_group)
lemma (in normal) inv_op_closed1:
- "\<lbrakk>x \<in> carrier G; h \<in> H\<rbrakk> \<Longrightarrow> (inv x) \<otimes> h \<otimes> x \<in> H"
-apply (insert coset_eq)
-apply (auto simp add: l_coset_def r_coset_def)
-apply (drule bspec, assumption)
-apply (drule equalityD1 [THEN subsetD], blast, clarify)
-apply (simp add: m_assoc)
-apply (simp add: m_assoc [symmetric])
-done
+ assumes "x \<in> carrier G" and "h \<in> H"
+ shows "(inv x) \<otimes> h \<otimes> x \<in> H"
+proof -
+ have "h \<otimes> x \<in> x <# H"
+ using assms coset_eq assms(1) unfolding r_coset_def by blast
+ then obtain h' where "h' \<in> H" "h \<otimes> x = x \<otimes> h'"
+ unfolding l_coset_def by blast
+ thus ?thesis by (metis assms inv_closed l_inv l_one m_assoc mem_carrier)
+qed
lemma (in normal) inv_op_closed2:
- "\<lbrakk>x \<in> carrier G; h \<in> H\<rbrakk> \<Longrightarrow> x \<otimes> h \<otimes> (inv x) \<in> H"
-apply (subgoal_tac "inv (inv x) \<otimes> h \<otimes> (inv x) \<in> H")
-apply (simp add: )
-apply (blast intro: inv_op_closed1)
-done
+ assumes "x \<in> carrier G" and "h \<in> H"
+ shows "x \<otimes> h \<otimes> (inv x) \<in> H"
+ using assms inv_op_closed1 by (metis inv_closed inv_inv)
+
text\<open>Alternative characterization of normal subgroups\<close>
lemma (in group) normal_inv_iff:
@@ -455,74 +429,81 @@
qed
qed
+corollary (in group) normal_invI:
+ assumes "subgroup N G" and "\<And>x h. \<lbrakk> x \<in> carrier G; h \<in> N \<rbrakk> \<Longrightarrow> x \<otimes> h \<otimes> inv x \<in> N"
+ shows "N \<lhd> G"
+ using assms normal_inv_iff by blast
-subsection\<open>More Properties of Cosets\<close>
+corollary (in group) normal_invE:
+ assumes "N \<lhd> G"
+ shows "subgroup N G" and "\<And>x h. \<lbrakk> x \<in> carrier G; h \<in> N \<rbrakk> \<Longrightarrow> x \<otimes> h \<otimes> inv x \<in> N"
+ using assms normal_inv_iff apply blast
+ by (simp add: assms normal.inv_op_closed2)
+
+
+lemma (in group) one_is_normal :
+ "{\<one>} \<lhd> G"
+proof(intro normal_invI )
+ show "subgroup {\<one>} G"
+ by (simp add: subgroup_def)
+ show "\<And>x h. x \<in> carrier G \<Longrightarrow> h \<in> {\<one>} \<Longrightarrow> x \<otimes> h \<otimes> inv x \<in> {\<one>}" by simp
+qed
+
+
+subsection\<open>More Properties of Left Cosets\<close>
+
+lemma (in group) l_repr_independence:
+ assumes "y \<in> x <# H" "x \<in> carrier G" "subgroup H G"
+ shows "x <# H = y <# H"
+proof -
+ obtain h' where h': "h' \<in> H" "y = x \<otimes> h'"
+ using assms(1) unfolding l_coset_def by blast
+ hence "\<And> h. h \<in> H \<Longrightarrow> x \<otimes> h = y \<otimes> ((inv h') \<otimes> h)"
+ by (smt assms(2-3) inv_closed inv_solve_right m_assoc m_closed subgroup.mem_carrier)
+ hence "\<And> xh. xh \<in> x <# H \<Longrightarrow> xh \<in> y <# H"
+ unfolding l_coset_def by (metis (no_types, lifting) UN_iff assms(3) h'(1) subgroup_def)
+ moreover have "\<And> h. h \<in> H \<Longrightarrow> y \<otimes> h = x \<otimes> (h' \<otimes> h)"
+ using h' by (meson assms(2) assms(3) m_assoc subgroup.mem_carrier)
+ hence "\<And> yh. yh \<in> y <# H \<Longrightarrow> yh \<in> x <# H"
+ unfolding l_coset_def using subgroup.m_closed[OF assms(3) h'(1)] by blast
+ ultimately show ?thesis by blast
+qed
lemma (in group) lcos_m_assoc:
- "[| M \<subseteq> carrier G; g \<in> carrier G; h \<in> carrier G |]
- ==> g <# (h <# M) = (g \<otimes> h) <# M"
+ "\<lbrakk> M \<subseteq> carrier G; g \<in> carrier G; h \<in> carrier G \<rbrakk> \<Longrightarrow> g <# (h <# M) = (g \<otimes> h) <# M"
by (force simp add: l_coset_def m_assoc)
-lemma (in group) lcos_mult_one: "M \<subseteq> carrier G ==> \<one> <# M = M"
+lemma (in group) lcos_mult_one: "M \<subseteq> carrier G \<Longrightarrow> \<one> <# M = M"
by (force simp add: l_coset_def)
lemma (in group) l_coset_subset_G:
- "[| H \<subseteq> carrier G; x \<in> carrier G |] ==> x <# H \<subseteq> carrier G"
+ "\<lbrakk> H \<subseteq> carrier G; x \<in> carrier G \<rbrakk> \<Longrightarrow> x <# H \<subseteq> carrier G"
by (auto simp add: l_coset_def subsetD)
-lemma (in group) l_coset_swap:
- "\<lbrakk>y \<in> x <# H; x \<in> carrier G; subgroup H G\<rbrakk> \<Longrightarrow> x \<in> y <# H"
-proof (simp add: l_coset_def)
- assume "\<exists>h\<in>H. y = x \<otimes> h"
- and x: "x \<in> carrier G"
- and sb: "subgroup H G"
- then obtain h' where h': "h' \<in> H \<and> x \<otimes> h' = y" by blast
- show "\<exists>h\<in>H. x = y \<otimes> h"
- proof
- show "x = y \<otimes> inv h'" using h' x sb
- by (auto simp add: m_assoc subgroup.subset [THEN subsetD])
- show "inv h' \<in> H" using h' sb
- by (auto simp add: subgroup.subset [THEN subsetD] subgroup.m_inv_closed)
- qed
-qed
-
lemma (in group) l_coset_carrier:
- "[| y \<in> x <# H; x \<in> carrier G; subgroup H G |] ==> y \<in> carrier G"
-by (auto simp add: l_coset_def m_assoc
- subgroup.subset [THEN subsetD] subgroup.m_closed)
+ "\<lbrakk> y \<in> x <# H; x \<in> carrier G; subgroup H G \<rbrakk> \<Longrightarrow> y \<in> carrier G"
+ by (auto simp add: l_coset_def m_assoc subgroup.subset [THEN subsetD] subgroup.m_closed)
-lemma (in group) l_repr_imp_subset:
- assumes y: "y \<in> x <# H" and x: "x \<in> carrier G" and sb: "subgroup H G"
- shows "y <# H \<subseteq> x <# H"
-proof -
- from y
- obtain h' where "h' \<in> H" "x \<otimes> h' = y" by (auto simp add: l_coset_def)
- thus ?thesis using x sb
- by (auto simp add: l_coset_def m_assoc
- subgroup.subset [THEN subsetD] subgroup.m_closed)
-qed
+lemma (in group) l_coset_swap:
+ assumes "y \<in> x <# H" "x \<in> carrier G" "subgroup H G"
+ shows "x \<in> y <# H"
+ using assms(2) l_repr_independence[OF assms] subgroup.one_closed[OF assms(3)]
+ unfolding l_coset_def by fastforce
-lemma (in group) l_repr_independence:
- assumes y: "y \<in> x <# H" and x: "x \<in> carrier G" and sb: "subgroup H G"
- shows "x <# H = y <# H"
+lemma (in group) subgroup_mult_id:
+ assumes "subgroup H G"
+ shows "H <#> H = H"
proof
- show "x <# H \<subseteq> y <# H"
- by (rule l_repr_imp_subset,
- (blast intro: l_coset_swap l_coset_carrier y x sb)+)
- show "y <# H \<subseteq> x <# H" by (rule l_repr_imp_subset [OF y x sb])
+ show "H <#> H \<subseteq> H"
+ unfolding set_mult_def using subgroup.m_closed[OF assms] by (simp add: UN_subset_iff)
+ show "H \<subseteq> H <#> H"
+ proof
+ fix x assume x: "x \<in> H" thus "x \<in> H <#> H" unfolding set_mult_def
+ using subgroup.m_closed[OF assms subgroup.one_closed[OF assms] x] subgroup.one_closed[OF assms]
+ by (smt UN_iff assms coset_join3 l_coset_def subgroup.mem_carrier)
+ qed
qed
-lemma (in group) setmult_subset_G:
- "\<lbrakk>H \<subseteq> carrier G; K \<subseteq> carrier G\<rbrakk> \<Longrightarrow> H <#> K \<subseteq> carrier G"
-by (auto simp add: set_mult_def subsetD)
-
-lemma (in group) subgroup_mult_id: "subgroup H G \<Longrightarrow> H <#> H = H"
-apply (auto simp add: subgroup.m_closed set_mult_def Sigma_def)
-apply (rule_tac x = x in bexI)
-apply (rule bexI [of _ "\<one>"])
-apply (auto simp add: subgroup.one_closed subgroup.subset [THEN subsetD])
-done
-
subsubsection \<open>Set of Inverses of an \<open>r_coset\<close>.\<close>
@@ -552,20 +533,21 @@
subsubsection \<open>Theorems for \<open><#>\<close> with \<open>#>\<close> or \<open><#\<close>.\<close>
lemma (in group) setmult_rcos_assoc:
- "\<lbrakk>H \<subseteq> carrier G; K \<subseteq> carrier G; x \<in> carrier G\<rbrakk>
- \<Longrightarrow> H <#> (K #> x) = (H <#> K) #> x"
-by (force simp add: r_coset_def set_mult_def m_assoc)
+ "\<lbrakk>H \<subseteq> carrier G; K \<subseteq> carrier G; x \<in> carrier G\<rbrakk> \<Longrightarrow>
+ H <#> (K #> x) = (H <#> K) #> x"
+ using set_mult_assoc[of H K "{x}"] by (simp add: r_coset_eq_set_mult)
lemma (in group) rcos_assoc_lcos:
- "\<lbrakk>H \<subseteq> carrier G; K \<subseteq> carrier G; x \<in> carrier G\<rbrakk>
- \<Longrightarrow> (H #> x) <#> K = H <#> (x <# K)"
-by (force simp add: r_coset_def l_coset_def set_mult_def m_assoc)
+ "\<lbrakk>H \<subseteq> carrier G; K \<subseteq> carrier G; x \<in> carrier G\<rbrakk> \<Longrightarrow>
+ (H #> x) <#> K = H <#> (x <# K)"
+ using set_mult_assoc[of H "{x}" K]
+ by (simp add: l_coset_eq_set_mult r_coset_eq_set_mult)
lemma (in normal) rcos_mult_step1:
- "\<lbrakk>x \<in> carrier G; y \<in> carrier G\<rbrakk>
- \<Longrightarrow> (H #> x) <#> (H #> y) = (H <#> (x <# H)) #> y"
-by (simp add: setmult_rcos_assoc subset
- r_coset_subset_G l_coset_subset_G rcos_assoc_lcos)
+ "\<lbrakk>x \<in> carrier G; y \<in> carrier G\<rbrakk> \<Longrightarrow>
+ (H #> x) <#> (H #> y) = (H <#> (x <# H)) #> y"
+ by (simp add: setmult_rcos_assoc r_coset_subset_G
+ subset l_coset_subset_G rcos_assoc_lcos)
lemma (in normal) rcos_mult_step2:
"\<lbrakk>x \<in> carrier G; y \<in> carrier G\<rbrakk>
@@ -645,7 +627,7 @@
lemma (in subgroup) l_coset_eq_rcong:
assumes "group G"
assumes a: "a \<in> carrier G"
- shows "a <# H = rcong H `` {a}"
+ shows "a <# H = (rcong H) `` {a}"
proof -
interpret group G by fact
show ?thesis by (force simp add: r_congruent_def l_coset_def m_assoc [symmetric] a )
@@ -661,9 +643,7 @@
proof -
interpret subgroup H G by fact
from p show ?thesis apply (rule_tac UN_I [of "hb \<otimes> ((inv ha) \<otimes> h)"])
- apply (simp add: )
- apply (simp add: m_assoc transpose_inv)
- done
+ apply blast by (simp add: inv_solve_left m_assoc)
qed
lemma (in group) rcos_disjoint:
@@ -793,28 +773,47 @@
"\<lbrakk>H \<subseteq> carrier G; a \<in> carrier G\<rbrakk> \<Longrightarrow> inj_on (\<lambda>y. y \<otimes> a) H"
by (force simp add: inj_on_def subsetD)
+(* ************************************************************************** *)
+
lemma (in group) card_cosets_equal:
- "\<lbrakk>c \<in> rcosets H; H \<subseteq> carrier G; finite(carrier G)\<rbrakk>
- \<Longrightarrow> card c = card H"
-apply (auto simp add: RCOSETS_def)
-apply (rule card_bij_eq)
- apply (rule inj_on_f, assumption+)
- apply (force simp add: m_assoc subsetD r_coset_def)
- apply (rule inj_on_g, assumption+)
- apply (force simp add: m_assoc subsetD r_coset_def)
- txt\<open>The sets @{term "H #> a"} and @{term "H"} are finite.\<close>
- apply (simp add: r_coset_subset_G [THEN finite_subset])
-apply (blast intro: finite_subset)
-done
+ assumes "R \<in> rcosets H" "H \<subseteq> carrier G"
+ shows "\<exists>f. bij_betw f H R"
+proof -
+ obtain g where g: "g \<in> carrier G" "R = H #> g"
+ using assms(1) unfolding RCOSETS_def by blast
+
+ let ?f = "\<lambda>h. h \<otimes> g"
+ have "\<And>r. r \<in> R \<Longrightarrow> \<exists>h \<in> H. ?f h = r"
+ proof -
+ fix r assume "r \<in> R"
+ then obtain h where "h \<in> H" "r = h \<otimes> g"
+ using g unfolding r_coset_def by blast
+ thus "\<exists>h \<in> H. ?f h = r" by blast
+ qed
+ hence "R \<subseteq> ?f ` H" by blast
+ moreover have "?f ` H \<subseteq> R"
+ using g unfolding r_coset_def by blast
+ ultimately show ?thesis using inj_on_g unfolding bij_betw_def
+ using assms(2) g(1) by auto
+qed
+
+corollary (in group) card_rcosets_equal:
+ assumes "R \<in> rcosets H" "H \<subseteq> carrier G"
+ shows "card H = card R"
+ using card_cosets_equal assms bij_betw_same_card by blast
+
+corollary (in group) rcosets_finite:
+ assumes "R \<in> rcosets H" "H \<subseteq> carrier G" "finite H"
+ shows "finite R"
+ using card_cosets_equal assms bij_betw_finite is_group by blast
+
+(* ************************************************************************** *)
lemma (in group) rcosets_subset_PowG:
"subgroup H G \<Longrightarrow> rcosets H \<subseteq> Pow(carrier G)"
-apply (simp add: RCOSETS_def)
-apply (blast dest: r_coset_subset_G subgroup.subset)
-done
+ using rcosets_part_G by auto
-
-theorem (in group) lagrange:
+proposition (in group) lagrange_finite:
"\<lbrakk>finite(carrier G); subgroup H G\<rbrakk>
\<Longrightarrow> card(rcosets H) * card(H) = order(G)"
apply (simp (no_asm_simp) add: order_def rcosets_part_G [symmetric])
@@ -822,10 +821,42 @@
apply (rule card_partition)
apply (simp add: rcosets_subset_PowG [THEN finite_subset])
apply (simp add: rcosets_part_G)
- apply (simp add: card_cosets_equal subgroup.subset)
+ apply (simp add: card_rcosets_equal subgroup_imp_subset)
apply (simp add: rcos_disjoint)
done
+theorem (in group) lagrange:
+ assumes "subgroup H G"
+ shows "card (rcosets H) * card H = order G"
+proof (cases "finite (carrier G)")
+ case True thus ?thesis using lagrange_finite assms by simp
+next
+ case False note inf_G = this
+ thus ?thesis
+ proof (cases "finite H")
+ case False thus ?thesis using inf_G by (simp add: order_def)
+ next
+ case True note finite_H = this
+ have "infinite (rcosets H)"
+ proof (rule ccontr)
+ assume "\<not> infinite (rcosets H)"
+ hence finite_rcos: "finite (rcosets H)" by simp
+ hence "card (\<Union>(rcosets H)) = (\<Sum>R\<in>(rcosets H). card R)"
+ using card_Union_disjoint[of "rcosets H"] finite_H rcos_disjoint[OF assms(1)]
+ rcosets_finite[where ?H = H] by (simp add: assms subgroup_imp_subset)
+ hence "order G = (\<Sum>R\<in>(rcosets H). card R)"
+ by (simp add: assms order_def rcosets_part_G)
+ hence "order G = (\<Sum>R\<in>(rcosets H). card H)"
+ using card_rcosets_equal by (simp add: assms subgroup_imp_subset)
+ hence "order G = (card H) * (card (rcosets H))" by simp
+ hence "order G \<noteq> 0" using finite_rcos finite_H assms ex_in_conv
+ rcosets_part_G subgroup.one_closed by fastforce
+ thus False using inf_G order_gt_0_iff_finite by blast
+ qed
+ thus ?thesis using inf_G by (simp add: order_def)
+ qed
+qed
+
subsection \<open>Quotient Groups: Factorization of a Group\<close>
@@ -845,7 +876,7 @@
lemma (in normal) rcosets_assoc:
"\<lbrakk>M1 \<in> rcosets H; M2 \<in> rcosets H; M3 \<in> rcosets H\<rbrakk>
\<Longrightarrow> M1 <#> M2 <#> M3 = M1 <#> (M2 <#> M3)"
-by (auto simp add: RCOSETS_def rcos_sum m_assoc)
+ by (simp add: group.set_mult_assoc is_group rcosets_carrier)
lemma (in subgroup) subgroup_in_rcosets:
assumes "group G"
@@ -1016,10 +1047,253 @@
text\<open>If @{term h} is a homomorphism from @{term G} onto @{term H}, then the
quotient group @{term "G Mod (kernel G H h)"} is isomorphic to @{term H}.\<close>
-theorem (in group_hom) FactGroup_iso:
+theorem (in group_hom) FactGroup_iso_set:
"h ` carrier G = carrier H
- \<Longrightarrow> (\<lambda>X. the_elem (h`X)) \<in> (G Mod (kernel G H h)) \<cong> H"
+ \<Longrightarrow> (\<lambda>X. the_elem (h`X)) \<in> iso (G Mod (kernel G H h)) H"
by (simp add: iso_def FactGroup_hom FactGroup_inj_on bij_betw_def
FactGroup_onto)
+corollary (in group_hom) FactGroup_iso :
+ "h ` carrier G = carrier H
+ \<Longrightarrow> (G Mod (kernel G H h))\<cong> H"
+ using FactGroup_iso_set unfolding is_iso_def by auto
+
+
+(* Next two lemmas contributed by Paulo EmÃlio de Vilhena. *)
+
+lemma (in group_hom) trivial_hom_iff:
+ "(h ` (carrier G) = { \<one>\<^bsub>H\<^esub> }) = (kernel G H h = carrier G)"
+ unfolding kernel_def using one_closed by force
+
+lemma (in group_hom) trivial_ker_imp_inj:
+ assumes "kernel G H h = { \<one> }"
+ shows "inj_on h (carrier G)"
+proof (rule inj_onI)
+ fix g1 g2 assume A: "g1 \<in> carrier G" "g2 \<in> carrier G" "h g1 = h g2"
+ hence "h (g1 \<otimes> (inv g2)) = \<one>\<^bsub>H\<^esub>" by simp
+ hence "g1 \<otimes> (inv g2) = \<one>"
+ using A assms unfolding kernel_def by blast
+ thus "g1 = g2"
+ using A G.inv_equality G.inv_inv by blast
+qed
+
+
+(* Next subsection contributed by Martin Baillon. *)
+
+subsection \<open>Theorems about Factor Groups and Direct product\<close>
+
+lemma (in group) DirProd_normal :
+ assumes "group K"
+ and "H \<lhd> G"
+ and "N \<lhd> K"
+ shows "H \<times> N \<lhd> G \<times>\<times> K"
+proof (intro group.normal_invI[OF DirProd_group[OF group_axioms assms(1)]])
+ show sub : "subgroup (H \<times> N) (G \<times>\<times> K)"
+ using DirProd_subgroups[OF group_axioms normal_imp_subgroup[OF assms(2)]assms(1)
+ normal_imp_subgroup[OF assms(3)]].
+ show "\<And>x h. x \<in> carrier (G\<times>\<times>K) \<Longrightarrow> h \<in> H\<times>N \<Longrightarrow> x \<otimes>\<^bsub>G\<times>\<times>K\<^esub> h \<otimes>\<^bsub>G\<times>\<times>K\<^esub> inv\<^bsub>G\<times>\<times>K\<^esub> x \<in> H\<times>N"
+ proof-
+ fix x h assume xGK : "x \<in> carrier (G \<times>\<times> K)" and hHN : " h \<in> H \<times> N"
+ hence hGK : "h \<in> carrier (G \<times>\<times> K)" using subgroup_imp_subset[OF sub] by auto
+ from xGK obtain x1 x2 where x1x2 :"x1 \<in> carrier G" "x2 \<in> carrier K" "x = (x1,x2)"
+ unfolding DirProd_def by fastforce
+ from hHN obtain h1 h2 where h1h2 : "h1 \<in> H" "h2 \<in> N" "h = (h1,h2)"
+ unfolding DirProd_def by fastforce
+ hence h1h2GK : "h1 \<in> carrier G" "h2 \<in> carrier K"
+ using normal_imp_subgroup subgroup_imp_subset assms apply blast+.
+ have "inv\<^bsub>G \<times>\<times> K\<^esub> x = (inv\<^bsub>G\<^esub> x1,inv\<^bsub>K\<^esub> x2)"
+ using inv_DirProd[OF group_axioms assms(1) x1x2(1)x1x2(2)] x1x2 by auto
+ hence "x \<otimes>\<^bsub>G \<times>\<times> K\<^esub> h \<otimes>\<^bsub>G \<times>\<times> K\<^esub> inv\<^bsub>G \<times>\<times> K\<^esub> x = (x1 \<otimes> h1 \<otimes> inv x1,x2 \<otimes>\<^bsub>K\<^esub> h2 \<otimes>\<^bsub>K\<^esub> inv\<^bsub>K\<^esub> x2)"
+ using h1h2 x1x2 h1h2GK by auto
+ moreover have "x1 \<otimes> h1 \<otimes> inv x1 \<in> H" "x2 \<otimes>\<^bsub>K\<^esub> h2 \<otimes>\<^bsub>K\<^esub> inv\<^bsub>K\<^esub> x2 \<in> N"
+ using normal_invE group.normal_invE[OF assms(1)] assms x1x2 h1h2 apply auto.
+ hence "(x1 \<otimes> h1 \<otimes> inv x1, x2 \<otimes>\<^bsub>K\<^esub> h2 \<otimes>\<^bsub>K\<^esub> inv\<^bsub>K\<^esub> x2)\<in> H \<times> N" by auto
+ ultimately show " x \<otimes>\<^bsub>G \<times>\<times> K\<^esub> h \<otimes>\<^bsub>G \<times>\<times> K\<^esub> inv\<^bsub>G \<times>\<times> K\<^esub> x \<in> H \<times> N" by auto
+ qed
+qed
+
+lemma (in group) FactGroup_DirProd_multiplication_iso_set :
+ assumes "group K"
+ and "H \<lhd> G"
+ and "N \<lhd> K"
+ shows "(\<lambda> (X, Y). X \<times> Y) \<in> iso ((G Mod H) \<times>\<times> (K Mod N)) (G \<times>\<times> K Mod H \<times> N)"
+
+proof-
+ have R :"(\<lambda>(X, Y). X \<times> Y) \<in> carrier (G Mod H) \<times> carrier (K Mod N) \<rightarrow> carrier (G \<times>\<times> K Mod H \<times> N)"
+ unfolding r_coset_def Sigma_def DirProd_def FactGroup_def RCOSETS_def apply simp by blast
+ moreover have "(\<forall>x\<in>carrier (G Mod H). \<forall>y\<in>carrier (K Mod N). \<forall>xa\<in>carrier (G Mod H).
+ \<forall>ya\<in>carrier (K Mod N). (x <#> xa) \<times> (y <#>\<^bsub>K\<^esub> ya) = x \<times> y <#>\<^bsub>G \<times>\<times> K\<^esub> xa \<times> ya)"
+ unfolding set_mult_def apply auto apply blast+.
+ moreover have "(\<forall>x\<in>carrier (G Mod H). \<forall>y\<in>carrier (K Mod N). \<forall>xa\<in>carrier (G Mod H).
+ \<forall>ya\<in>carrier (K Mod N). x \<times> y = xa \<times> ya \<longrightarrow> x = xa \<and> y = ya)"
+ unfolding FactGroup_def using times_eq_iff subgroup.rcosets_not_empty
+ by (metis assms(2) assms(3) normal_def partial_object.select_convs(1))
+ moreover have "(\<lambda>(X, Y). X \<times> Y) ` (carrier (G Mod H) \<times> carrier (K Mod N)) =
+ carrier (G \<times>\<times> K Mod H \<times> N)"
+ unfolding image_def apply auto using R apply force
+ unfolding DirProd_def FactGroup_def RCOSETS_def r_coset_def apply auto apply force.
+ ultimately show ?thesis
+ unfolding iso_def hom_def bij_betw_def inj_on_def by simp
+qed
+
+corollary (in group) FactGroup_DirProd_multiplication_iso_1 :
+ assumes "group K"
+ and "H \<lhd> G"
+ and "N \<lhd> K"
+ shows " ((G Mod H) \<times>\<times> (K Mod N)) \<cong> (G \<times>\<times> K Mod H \<times> N)"
+ unfolding is_iso_def using FactGroup_DirProd_multiplication_iso_set assms by auto
+
+corollary (in group) FactGroup_DirProd_multiplication_iso_2 :
+ assumes "group K"
+ and "H \<lhd> G"
+ and "N \<lhd> K"
+ shows "(G \<times>\<times> K Mod H \<times> N) \<cong> ((G Mod H) \<times>\<times> (K Mod N))"
+ using FactGroup_DirProd_multiplication_iso_1 group.iso_sym assms
+ DirProd_group[OF normal.factorgroup_is_group normal.factorgroup_is_group]
+ by blast
+
+subsubsection "More Lemmas about set multiplication"
+
+(*A group multiplied by a subgroup stays the same*)
+lemma (in group) set_mult_carrier_idem:
+ assumes "subgroup H G"
+ shows "(carrier G) <#> H = carrier G"
+proof
+ show "(carrier G)<#>H \<subseteq> carrier G"
+ unfolding set_mult_def using subgroup_imp_subset assms by blast
+next
+ have " (carrier G) #> \<one> = carrier G" unfolding set_mult_def r_coset_def group_axioms by simp
+ moreover have "(carrier G) #> \<one> \<subseteq> (carrier G) <#> H" unfolding set_mult_def r_coset_def
+ using assms subgroup.one_closed[OF assms] by blast
+ ultimately show "carrier G \<subseteq> (carrier G) <#> H" by simp
+qed
+
+(*Same lemma as above, but everything is included in a subgroup*)
+lemma (in group) set_mult_subgroup_idem:
+ assumes "subgroup H G"
+ and "subgroup N (G\<lparr>carrier:=H\<rparr>)"
+ shows "H<#>N = H"
+ using group.set_mult_carrier_idem[OF subgroup_imp_group] subgroup_imp_subset assms
+ by (metis monoid.cases_scheme order_refl partial_object.simps(1)
+ partial_object.update_convs(1) subgroup_set_mult_equality)
+
+(*A normal subgroup is commutative with set_mult*)
+lemma (in group) commut_normal:
+ assumes "subgroup H G"
+ and "N\<lhd>G"
+ shows "H<#>N = N<#>H"
+proof-
+ have aux1: "{H <#> N} = {\<Union>h\<in>H. h <# N }" unfolding set_mult_def l_coset_def by auto
+ also have "... = {\<Union>h\<in>H. N #> h }" using assms normal.coset_eq subgroup.mem_carrier by fastforce
+ moreover have aux2: "{N <#> H} = {\<Union>h\<in>H. N #> h }"unfolding set_mult_def r_coset_def by auto
+ ultimately show "H<#>N = N<#>H" by simp
+qed
+
+(*Same lemma as above, but everything is included in a subgroup*)
+lemma (in group) commut_normal_subgroup:
+ assumes "subgroup H G"
+ and "N\<lhd>(G\<lparr>carrier:=H\<rparr>)"
+ and "subgroup K (G\<lparr>carrier:=H\<rparr>)"
+ shows "K<#>N = N<#>K"
+proof-
+ have "N \<subseteq> carrier (G\<lparr>carrier := H\<rparr>)" using assms normal_imp_subgroup subgroup_imp_subset by blast
+ hence NH : "N \<subseteq> H" by simp
+ have "K \<subseteq> carrier(G\<lparr>carrier := H\<rparr>)" using subgroup_imp_subset assms by blast
+ hence KH : "K \<subseteq> H" by simp
+ have Egal : "K <#>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> N = N <#>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> K"
+ using group.commut_normal[where ?G = "G\<lparr>carrier :=H\<rparr>", of K N,OF subgroup_imp_group[OF assms(1)]
+ assms(3) assms(2)] by auto
+ also have "... = N <#> K" using subgroup_set_mult_equality[of H N K, OF assms(1) NH KH] by auto
+ moreover have "K <#>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> N = K <#> N"
+ using subgroup_set_mult_equality[of H K N, OF assms(1) KH NH] by auto
+ ultimately show "K<#>N = N<#>K" by auto
+qed
+
+
+
+subsubsection "Lemmas about intersection and normal subgroups"
+
+lemma (in group) normal_inter:
+ assumes "subgroup H G"
+ and "subgroup K G"
+ and "H1\<lhd>G\<lparr>carrier := H\<rparr>"
+ shows " (H1\<inter>K)\<lhd>(G\<lparr>carrier:= (H\<inter>K)\<rparr>)"
+proof-
+ define HK and H1K and GH and GHK
+ where "HK = H\<inter>K" and "H1K=H1\<inter>K" and "GH =G\<lparr>carrier := H\<rparr>" and "GHK = (G\<lparr>carrier:= (H\<inter>K)\<rparr>)"
+ show "H1K\<lhd>GHK"
+ proof (intro group.normal_invI[of GHK H1K])
+ show "Group.group GHK"
+ using GHK_def subgroups_Inter_pair subgroup_imp_group assms by blast
+
+ next
+ have H1K_incl:"subgroup H1K (G\<lparr>carrier:= (H\<inter>K)\<rparr>)"
+ proof(intro subgroup_incl)
+ show "subgroup H1K G"
+ using assms normal_imp_subgroup subgroups_Inter_pair incl_subgroup H1K_def by blast
+ next
+ show "subgroup (H\<inter>K) G" using HK_def subgroups_Inter_pair assms by auto
+ next
+ have "H1 \<subseteq> (carrier (G\<lparr>carrier:=H\<rparr>))"
+ using assms(3) normal_imp_subgroup subgroup_imp_subset by blast
+ also have "... \<subseteq> H" by simp
+ thus "H1K \<subseteq>H\<inter>K"
+ using H1K_def calculation by auto
+ qed
+ thus "subgroup H1K GHK" using GHK_def by simp
+ next
+ show "\<And> x h. x\<in>carrier GHK \<Longrightarrow> h\<in>H1K \<Longrightarrow> x \<otimes>\<^bsub>GHK\<^esub> h \<otimes>\<^bsub>GHK\<^esub> inv\<^bsub>GHK\<^esub> x\<in> H1K"
+ proof-
+ have invHK: "\<lbrakk>y\<in>HK\<rbrakk> \<Longrightarrow> inv\<^bsub>GHK\<^esub> y = inv\<^bsub>GH\<^esub> y"
+ using subgroup_inv_equality assms HK_def GH_def GHK_def subgroups_Inter_pair by simp
+ have multHK : "\<lbrakk>x\<in>HK;y\<in>HK\<rbrakk> \<Longrightarrow> x \<otimes>\<^bsub>(G\<lparr>carrier:=HK\<rparr>)\<^esub> y = x \<otimes> y"
+ using HK_def by simp
+ fix x assume p: "x\<in>carrier GHK"
+ fix h assume p2 : "h:H1K"
+ have "carrier(GHK)\<subseteq>HK"
+ using GHK_def HK_def by simp
+ hence xHK:"x\<in>HK" using p by auto
+ hence invx:"inv\<^bsub>GHK\<^esub> x = inv\<^bsub>GH\<^esub> x"
+ using invHK assms GHK_def HK_def GH_def subgroup_inv_equality subgroups_Inter_pair by simp
+ have "H1\<subseteq>carrier(GH)"
+ using assms GH_def normal_imp_subgroup subgroup_imp_subset by blast
+ hence hHK:"h\<in>HK"
+ using p2 H1K_def HK_def GH_def by auto
+ hence xhx_egal : "x \<otimes>\<^bsub>GHK\<^esub> h \<otimes>\<^bsub>GHK\<^esub> inv\<^bsub>GHK\<^esub>x = x \<otimes>\<^bsub>GH\<^esub> h \<otimes>\<^bsub>GH\<^esub> inv\<^bsub>GH\<^esub> x"
+ using invx invHK multHK GHK_def GH_def by auto
+ have xH:"x\<in>carrier(GH)"
+ using xHK HK_def GH_def by auto
+ have hH:"h\<in>carrier(GH)"
+ using hHK HK_def GH_def by auto
+ have "(\<forall>x\<in>carrier (GH). \<forall>h\<in>H1. x \<otimes>\<^bsub>GH\<^esub> h \<otimes>\<^bsub>GH\<^esub> inv\<^bsub>GH\<^esub> x \<in> H1)"
+ using assms normal_invE GH_def normal.inv_op_closed2 by fastforce
+ hence INCL_1 : "x \<otimes>\<^bsub>GH\<^esub> h \<otimes>\<^bsub>GH\<^esub> inv\<^bsub>GH\<^esub> x \<in> H1"
+ using xH H1K_def p2 by blast
+ have " x \<otimes>\<^bsub>GH\<^esub> h \<otimes>\<^bsub>GH\<^esub> inv\<^bsub>GH\<^esub> x \<in> HK"
+ using assms HK_def subgroups_Inter_pair hHK xHK
+ by (metis GH_def inf.cobounded1 subgroup_def subgroup_incl)
+ hence " x \<otimes>\<^bsub>GH\<^esub> h \<otimes>\<^bsub>GH\<^esub> inv\<^bsub>GH\<^esub> x \<in> K" using HK_def by simp
+ hence " x \<otimes>\<^bsub>GH\<^esub> h \<otimes>\<^bsub>GH\<^esub> inv\<^bsub>GH\<^esub> x \<in> H1K" using INCL_1 H1K_def by auto
+ thus "x \<otimes>\<^bsub>GHK\<^esub> h \<otimes>\<^bsub>GHK\<^esub> inv\<^bsub>GHK\<^esub> x \<in> H1K" using xhx_egal by simp
+ qed
+ qed
+qed
+
+
+lemma (in group) normal_inter_subgroup:
+ assumes "subgroup H G"
+ and "N \<lhd> G"
+ shows "(N\<inter>H) \<lhd> (G\<lparr>carrier := H\<rparr>)"
+proof -
+ define K where "K = carrier G"
+ have "G\<lparr>carrier := K\<rparr> = G" using K_def by auto
+ moreover have "subgroup K G" using K_def subgroup_self by blast
+ moreover have "normal N (G \<lparr>carrier :=K\<rparr>)" using assms K_def by simp
+ ultimately have "N \<inter> H \<lhd> G\<lparr>carrier := K \<inter> H\<rparr>"
+ using normal_inter[of K H N] assms(1) by blast
+ moreover have "K \<inter> H = H" using K_def assms subgroup_imp_subset by blast
+ ultimately show "normal (N\<inter>H) (G\<lparr>carrier := H\<rparr>)" by auto
+qed
+
+
end
--- a/src/HOL/Algebra/FiniteProduct.thy Thu Jun 14 10:51:12 2018 +0200
+++ b/src/HOL/Algebra/FiniteProduct.thy Thu Jun 14 14:23:48 2018 +0100
@@ -524,4 +524,49 @@
end
+(* Jeremy Avigad. This should be generalized to arbitrary groups, not just commutative
+ ones, using Lagrange's theorem. *)
+
+lemma (in comm_group) power_order_eq_one:
+ assumes fin [simp]: "finite (carrier G)"
+ and a [simp]: "a \<in> carrier G"
+ shows "a [^] card(carrier G) = one G"
+proof -
+ have "(\<Otimes>x\<in>carrier G. x) = (\<Otimes>x\<in>carrier G. a \<otimes> x)"
+ by (subst (2) finprod_reindex [symmetric],
+ auto simp add: Pi_def inj_on_const_mult surj_const_mult)
+ also have "\<dots> = (\<Otimes>x\<in>carrier G. a) \<otimes> (\<Otimes>x\<in>carrier G. x)"
+ by (auto simp add: finprod_multf Pi_def)
+ also have "(\<Otimes>x\<in>carrier G. a) = a [^] card(carrier G)"
+ by (auto simp add: finprod_const)
+ finally show ?thesis
+(* uses the preceeding lemma *)
+ by auto
+qed
+
+
+lemma (in comm_monoid) finprod_UN_disjoint:
+ "finite I \<Longrightarrow> (\<forall>i\<in>I. finite (A i)) \<longrightarrow> (\<forall>i\<in>I. \<forall>j\<in>I. i \<noteq> j \<longrightarrow> A i \<inter> A j = {}) \<longrightarrow>
+ (\<forall>i\<in>I. \<forall>x \<in> A i. g x \<in> carrier G) \<longrightarrow>
+ finprod G g (UNION I A) = finprod G (\<lambda>i. finprod G g (A i)) I"
+ apply (induct set: finite)
+ apply force
+ apply clarsimp
+ apply (subst finprod_Un_disjoint)
+ apply blast
+ apply (erule finite_UN_I)
+ apply blast
+ apply (fastforce)
+ apply (auto intro!: funcsetI finprod_closed)
+ done
+
+lemma (in comm_monoid) finprod_Union_disjoint:
+ "finite C \<Longrightarrow>
+ \<forall>A\<in>C. finite A \<and> (\<forall>x\<in>A. f x \<in> carrier G) \<Longrightarrow>
+ \<forall>A\<in>C. \<forall>B\<in>C. A \<noteq> B \<longrightarrow> A \<inter> B = {} \<Longrightarrow>
+ finprod G f (\<Union>C) = finprod G (finprod G f) C"
+ apply (frule finprod_UN_disjoint [of C id f])
+ apply auto
+ done
+
end
--- a/src/HOL/Algebra/Group.thy Thu Jun 14 10:51:12 2018 +0200
+++ b/src/HOL/Algebra/Group.thy Thu Jun 14 14:23:48 2018 +0100
@@ -2,6 +2,7 @@
Author: Clemens Ballarin, started 4 February 2003
Based on work by Florian Kammueller, L C Paulson and Markus Wenzel.
+With additional contributions from Martin Baillon and Paulo EmÃlio de Vilhena.
*)
theory Group
@@ -52,7 +53,7 @@
assumes m_closed [intro, simp]:
"\<lbrakk>x \<in> carrier G; y \<in> carrier G\<rbrakk> \<Longrightarrow> x \<otimes> y \<in> carrier G"
and m_assoc:
- "\<lbrakk>x \<in> carrier G; y \<in> carrier G; z \<in> carrier G\<rbrakk>
+ "\<lbrakk>x \<in> carrier G; y \<in> carrier G; z \<in> carrier G\<rbrakk>
\<Longrightarrow> (x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)"
and one_closed [intro, simp]: "\<one> \<in> carrier G"
and l_one [simp]: "x \<in> carrier G \<Longrightarrow> \<one> \<otimes> x = x"
@@ -75,6 +76,12 @@
"x \<in> Units G ==> x \<in> carrier G"
by (unfold Units_def) fast
+lemma (in monoid) one_unique:
+ assumes "u \<in> carrier G"
+ and "\<And>x. x \<in> carrier G \<Longrightarrow> u \<otimes> x = x"
+ shows "u = \<one>"
+ using assms(2)[OF one_closed] r_one[OF assms(1)] by simp
+
lemma (in monoid) inv_unique:
assumes eq: "y \<otimes> x = \<one>" "x \<otimes> y' = \<one>"
and G: "x \<in> carrier G" "y \<in> carrier G" "y' \<in> carrier G"
@@ -86,7 +93,7 @@
finally show ?thesis .
qed
-lemma (in monoid) Units_m_closed [intro, simp]:
+lemma (in monoid) Units_m_closed [simp, intro]:
assumes x: "x \<in> Units G" and y: "y \<in> Units G"
shows "x \<otimes> y \<in> Units G"
proof -
@@ -98,13 +105,7 @@
moreover from x y xinv yinv have "x \<otimes> (y \<otimes> y') \<otimes> x' = \<one>" by simp
moreover note x y
ultimately show ?thesis unfolding Units_def
- \<comment> \<open>Must avoid premature use of \<open>hyp_subst_tac\<close>.\<close>
- apply (rule_tac CollectI)
- apply (rule)
- apply (fast)
- apply (rule bexI [where x = "y' \<otimes> x'"])
- apply (auto simp: m_assoc)
- done
+ by simp (metis m_assoc m_closed)
qed
lemma (in monoid) Units_one_closed [intro, simp]:
@@ -140,6 +141,10 @@
apply (fast intro: inv_unique, fast)
done
+lemma (in monoid) inv_one [simp]:
+ "inv \<one> = \<one>"
+ by (metis Units_one_closed Units_r_inv l_one monoid.Units_inv_closed monoid_axioms)
+
lemma (in monoid) Units_inv_Units [intro, simp]:
"x \<in> Units G ==> inv x \<in> Units G"
proof -
@@ -215,10 +220,23 @@
"x \<in> carrier G ==> x [^] (n::nat) \<otimes> x [^] m = x [^] (n + m)"
by (induct m) (simp_all add: m_assoc [THEN sym])
+lemma (in monoid) nat_pow_comm:
+ "x \<in> carrier G \<Longrightarrow> (x [^] (n::nat)) \<otimes> (x [^] (m :: nat)) = (x [^] m) \<otimes> (x [^] n)"
+ using nat_pow_mult[of x n m] nat_pow_mult[of x m n] by (simp add: add.commute)
+
+lemma (in monoid) nat_pow_Suc2:
+ "x \<in> carrier G \<Longrightarrow> x [^] (Suc n) = x \<otimes> (x [^] n)"
+ using nat_pow_mult[of x 1 n] Suc_eq_plus1[of n]
+ by (metis One_nat_def Suc_eq_plus1_left l_one nat.rec(1) nat_pow_Suc nat_pow_def)
+
lemma (in monoid) nat_pow_pow:
"x \<in> carrier G ==> (x [^] n) [^] m = x [^] (n * m::nat)"
by (induct m) (simp, simp add: nat_pow_mult add.commute)
+lemma (in monoid) nat_pow_consistent:
+ "x [^] (n :: nat) = x [^]\<^bsub>(G \<lparr> carrier := H \<rparr>)\<^esub> n"
+ unfolding nat_pow_def by simp
+
(* Jacobson defines submonoid here. *)
(* Jacobson defines the order of a monoid here. *)
@@ -338,14 +356,6 @@
(y \<otimes> x = z \<otimes> x) = (y = z)"
by (metis inv_closed m_assoc r_inv r_one)
-lemma (in group) inv_one [simp]:
- "inv \<one> = \<one>"
-proof -
- have "inv \<one> = \<one> \<otimes> (inv \<one>)" by (simp del: r_inv Units_r_inv)
- moreover have "... = \<one>" by simp
- finally show ?thesis .
-qed
-
lemma (in group) inv_inv [simp]:
"x \<in> carrier G ==> inv (inv x) = x"
using Units_inv_inv by simp
@@ -416,6 +426,57 @@
by (auto simp add: int_pow_def2 inv_solve_left inv_solve_right nat_add_distrib [symmetric] nat_pow_mult )
qed
+lemma (in group) nat_pow_inv:
+ "x \<in> carrier G \<Longrightarrow> (inv x) [^] (i :: nat) = inv (x [^] i)"
+proof (induction i)
+ case 0 thus ?case by simp
+next
+ case (Suc i)
+ have "(inv x) [^] Suc i = ((inv x) [^] i) \<otimes> inv x"
+ by simp
+ also have " ... = (inv (x [^] i)) \<otimes> inv x"
+ by (simp add: Suc.IH Suc.prems)
+ also have " ... = inv (x \<otimes> (x [^] i))"
+ using inv_mult_group[OF Suc.prems nat_pow_closed[OF Suc.prems, of i]] by simp
+ also have " ... = inv (x [^] (Suc i))"
+ using Suc.prems nat_pow_Suc2 by auto
+ finally show ?case .
+qed
+
+lemma (in group) int_pow_inv:
+ "x \<in> carrier G \<Longrightarrow> (inv x) [^] (i :: int) = inv (x [^] i)"
+ by (simp add: nat_pow_inv int_pow_def2)
+
+lemma (in group) int_pow_pow:
+ assumes "x \<in> carrier G"
+ shows "(x [^] (n :: int)) [^] (m :: int) = x [^] (n * m :: int)"
+proof (cases)
+ assume n_ge: "n \<ge> 0" thus ?thesis
+ proof (cases)
+ assume m_ge: "m \<ge> 0" thus ?thesis
+ using n_ge nat_pow_pow[OF assms, of "nat n" "nat m"] int_pow_def2
+ by (simp add: mult_less_0_iff nat_mult_distrib)
+ next
+ assume m_lt: "\<not> m \<ge> 0" thus ?thesis
+ using n_ge int_pow_def2 nat_pow_pow[OF assms, of "nat n" "nat (- m)"]
+ by (smt assms group.int_pow_neg is_group mult_minus_right nat_mult_distrib split_mult_neg_le)
+ qed
+next
+ assume n_lt: "\<not> n \<ge> 0" thus ?thesis
+ proof (cases)
+ assume m_ge: "m \<ge> 0" thus ?thesis
+ using n_lt nat_pow_pow[OF assms, of "nat (- n)" "nat m"]
+ nat_pow_inv[of "x [^] nat (- n)" "nat m"] int_pow_def2
+ by (smt assms group.int_pow_closed group.int_pow_neg is_group mult_minus_right
+ mult_nonpos_nonpos nat_mult_distrib_neg)
+ next
+ assume m_lt: "\<not> m \<ge> 0" thus ?thesis
+ using n_lt nat_pow_pow[OF assms, of "nat (- n)" "nat (- m)"]
+ nat_pow_inv[of "x [^] nat (- n)" "nat (- m)"] int_pow_def2
+ by (smt assms inv_inv mult_nonpos_nonpos nat_mult_distrib_neg nat_pow_closed)
+ qed
+qed
+
lemma (in group) int_pow_diff:
"x \<in> carrier G \<Longrightarrow> x [^] (n - m :: int) = x [^] n \<otimes> inv (x [^] m)"
by(simp only: diff_conv_add_uminus int_pow_mult int_pow_neg)
@@ -426,6 +487,70 @@
lemma (in group) inj_on_cmult: "c \<in> carrier G \<Longrightarrow> inj_on (\<lambda>x. c \<otimes> x) (carrier G)"
by(simp add: inj_on_def)
+(*Following subsection contributed by Martin Baillon*)
+subsection \<open>Submonoids\<close>
+
+locale submonoid =
+ fixes H and G (structure)
+ assumes subset: "H \<subseteq> carrier G"
+ and m_closed [intro, simp]: "\<lbrakk>x \<in> H; y \<in> H\<rbrakk> \<Longrightarrow> x \<otimes> y \<in> H"
+ and one_closed [simp]: "\<one> \<in> H"
+
+lemma (in submonoid) is_submonoid:
+ "submonoid H G" by (rule submonoid_axioms)
+
+lemma (in submonoid) mem_carrier [simp]:
+ "x \<in> H \<Longrightarrow> x \<in> carrier G"
+ using subset by blast
+
+lemma submonoid_imp_subset:
+ "submonoid H G \<Longrightarrow> H \<subseteq> carrier G"
+ by (rule submonoid.subset)
+
+lemma (in submonoid) submonoid_is_monoid [intro]:
+ assumes "monoid G"
+ shows "monoid (G\<lparr>carrier := H\<rparr>)"
+proof -
+ interpret monoid G by fact
+ show ?thesis
+ by (simp add: monoid_def m_assoc)
+qed
+
+lemma (in monoid) submonoidE:
+ assumes "submonoid H G"
+ shows "H \<subseteq> carrier G"
+ and "H \<noteq> {}"
+ and "\<And>a b. \<lbrakk>a \<in> H; b \<in> H\<rbrakk> \<Longrightarrow> a \<otimes> b \<in> H"
+ using assms submonoid_imp_subset apply blast
+ using assms submonoid_def apply auto[1]
+ by (simp add: assms submonoid.m_closed)+
+
+lemma submonoid_nonempty:
+ "~ submonoid {} G"
+ by (blast dest: submonoid.one_closed)
+
+lemma (in submonoid) finite_monoid_imp_card_positive:
+ "finite (carrier G) ==> 0 < card H"
+proof (rule classical)
+ assume "finite (carrier G)" and a: "~ 0 < card H"
+ then have "finite H" by (blast intro: finite_subset [OF subset])
+ with is_submonoid a have "submonoid {} G" by simp
+ with submonoid_nonempty show ?thesis by contradiction
+qed
+
+
+lemma (in monoid) monoid_incl_imp_submonoid :
+ assumes "H \<subseteq> carrier G"
+and "monoid (G\<lparr>carrier := H\<rparr>)"
+shows "submonoid H G"
+proof (intro submonoid.intro[OF assms(1)])
+ have ab_eq : "\<And> a b. a \<in> H \<Longrightarrow> b \<in> H \<Longrightarrow> a \<otimes>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> b = a \<otimes> b" using assms by simp
+ have "\<And>a b. a \<in> H \<Longrightarrow> b \<in> H \<Longrightarrow> a \<otimes> b \<in> carrier (G\<lparr>carrier := H\<rparr>) "
+ using assms ab_eq unfolding group_def using monoid.m_closed by fastforce
+ thus "\<And>a b. a \<in> H \<Longrightarrow> b \<in> H \<Longrightarrow> a \<otimes> b \<in> H" by simp
+ show "\<one> \<in> H " using monoid.one_closed[OF assms(2)] assms by simp
+qed
+
subsection \<open>Subgroups\<close>
locale subgroup =
@@ -444,6 +569,7 @@
"x \<in> H \<Longrightarrow> x \<in> carrier G"
using subset by blast
+(*DELETE*)
lemma subgroup_imp_subset:
"subgroup H G \<Longrightarrow> H \<subseteq> carrier G"
by (rule subgroup.subset)
@@ -460,6 +586,40 @@
done
qed
+lemma (in group) subgroup_inv_equality:
+ assumes "subgroup H G" "x \<in> H"
+ shows "m_inv (G \<lparr>carrier := H\<rparr>) x = inv x"
+ unfolding m_inv_def apply auto
+ using subgroup.m_inv_closed[OF assms] inv_equality
+ by (metis (no_types, hide_lams) assms subgroup.mem_carrier)
+
+lemma (in group) int_pow_consistent: (* by Paulo *)
+ assumes "subgroup H G" "x \<in> H"
+ shows "x [^] (n :: int) = x [^]\<^bsub>(G \<lparr> carrier := H \<rparr>)\<^esub> n"
+proof (cases)
+ assume ge: "n \<ge> 0"
+ hence "x [^] n = x [^] (nat n)"
+ using int_pow_def2 by auto
+ also have " ... = x [^]\<^bsub>(G \<lparr> carrier := H \<rparr>)\<^esub> (nat n)"
+ using nat_pow_consistent by simp
+ also have " ... = x [^]\<^bsub>(G \<lparr> carrier := H \<rparr>)\<^esub> n"
+ using group.int_pow_def2[OF subgroup.subgroup_is_group[OF assms(1) is_group]] ge by auto
+ finally show ?thesis .
+next
+ assume "\<not> n \<ge> 0" hence lt: "n < 0" by simp
+ hence "x [^] n = inv (x [^] (nat (- n)))"
+ using int_pow_def2 by auto
+ also have " ... = (inv x) [^] (nat (- n))"
+ by (metis assms nat_pow_inv subgroup.mem_carrier)
+ also have " ... = (inv\<^bsub>(G \<lparr> carrier := H \<rparr>)\<^esub> x) [^]\<^bsub>(G \<lparr> carrier := H \<rparr>)\<^esub> (nat (- n))"
+ using subgroup_inv_equality[OF assms] nat_pow_consistent by auto
+ also have " ... = inv\<^bsub>(G \<lparr> carrier := H \<rparr>)\<^esub> (x [^]\<^bsub>(G \<lparr> carrier := H \<rparr>)\<^esub> (nat (- n)))"
+ using group.nat_pow_inv[OF subgroup.subgroup_is_group[OF assms(1) is_group]] assms(2) by auto
+ also have " ... = x [^]\<^bsub>(G \<lparr> carrier := H \<rparr>)\<^esub> n"
+ using group.int_pow_def2[OF subgroup.subgroup_is_group[OF assms(1) is_group]] lt by auto
+ finally show ?thesis .
+qed
+
text \<open>
Since @{term H} is nonempty, it contains some element @{term x}. Since
it is closed under inverse, it contains \<open>inv x\<close>. Since
@@ -482,6 +642,17 @@
show "\<one> \<in> H" by (rule one_in_subset) (auto simp only: assms)
qed
+
+lemma (in group) subgroupE:
+ assumes "subgroup H G"
+ shows "H \<subseteq> carrier G"
+ and "H \<noteq> {}"
+ and "\<And>a. a \<in> H \<Longrightarrow> inv a \<in> H"
+ and "\<And>a b. \<lbrakk>a \<in> H; b \<in> H\<rbrakk> \<Longrightarrow> a \<otimes> b \<in> H"
+ using assms subgroup.subset apply blast
+ using assms subgroup_def apply auto[1]
+ by (simp add: assms subgroup.m_closed subgroup.m_inv_closed)+
+
declare monoid.one_closed [iff] group.inv_closed [simp]
monoid.l_one [simp] monoid.r_one [simp] group.inv_inv [simp]
@@ -498,10 +669,38 @@
with subgroup_nonempty show ?thesis by contradiction
qed
-(*
-lemma (in monoid) Units_subgroup:
- "subgroup (Units G) G"
-*)
+(*Following 3 lemmas contributed by Martin Baillon*)
+
+lemma (in subgroup) subgroup_is_submonoid :
+ "submonoid H G"
+ by (simp add: submonoid.intro subset)
+
+lemma (in group) submonoid_subgroupI :
+ assumes "submonoid H G"
+ and "\<And>a. a \<in> H \<Longrightarrow> inv a \<in> H"
+ shows "subgroup H G"
+ by (metis assms subgroup_def submonoid_def)
+
+lemma (in group) group_incl_imp_subgroup:
+ assumes "H \<subseteq> carrier G"
+ and "group (G\<lparr>carrier := H\<rparr>)"
+ shows "subgroup H G"
+proof (intro submonoid_subgroupI[OF monoid_incl_imp_submonoid[OF assms(1)]])
+ show "monoid (G\<lparr>carrier := H\<rparr>)" using group_def assms by blast
+ have ab_eq : "\<And> a b. a \<in> H \<Longrightarrow> b \<in> H \<Longrightarrow> a \<otimes>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> b = a \<otimes> b" using assms by simp
+ fix a assume aH : "a \<in> H"
+ have " inv\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> a \<in> carrier G"
+ using assms aH group.inv_closed[OF assms(2)] by auto
+ moreover have "\<one>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> = \<one>" using assms monoid.one_closed ab_eq one_def by simp
+ hence "a \<otimes>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> inv\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> a= \<one>"
+ using assms ab_eq aH group.r_inv[OF assms(2)] by simp
+ hence "a \<otimes> inv\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> a= \<one>"
+ using aH assms group.inv_closed[OF assms(2)] ab_eq by simp
+ ultimately have "inv\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> a = inv a"
+ by (smt aH assms(1) contra_subsetD group.inv_inv is_group local.inv_equality)
+ moreover have "inv\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> a \<in> H" using aH group.inv_closed[OF assms(2)] by auto
+ ultimately show "inv a \<in> H" by auto
+qed
subsection \<open>Direct Products\<close>
@@ -520,7 +719,7 @@
interpret G: monoid G by fact
interpret H: monoid H by fact
from assms
- show ?thesis by (unfold monoid_def DirProd_def, auto)
+ show ?thesis by (unfold monoid_def DirProd_def, auto)
qed
@@ -548,6 +747,10 @@
"(g, h) \<otimes>\<^bsub>(G \<times>\<times> H)\<^esub> (g', h') = (g \<otimes>\<^bsub>G\<^esub> g', h \<otimes>\<^bsub>H\<^esub> h')"
by (simp add: DirProd_def)
+lemma DirProd_assoc :
+"(G \<times>\<times> H \<times>\<times> I) = (G \<times>\<times> (H \<times>\<times> I))"
+ by auto
+
lemma inv_DirProd [simp]:
assumes "group G" and "group H"
assumes g: "g \<in> carrier G"
@@ -561,6 +764,22 @@
show ?thesis by (simp add: Prod.inv_equality g h)
qed
+lemma DirProd_subgroups :
+ assumes "group G"
+ and "subgroup H G"
+ and "group K"
+ and "subgroup I K"
+ shows "subgroup (H \<times> I) (G \<times>\<times> K)"
+proof (intro group.group_incl_imp_subgroup[OF DirProd_group[OF assms(1)assms(3)]])
+ have "H \<subseteq> carrier G" "I \<subseteq> carrier K" using subgroup.subset assms apply blast+.
+ thus "(H \<times> I) \<subseteq> carrier (G \<times>\<times> K)" unfolding DirProd_def by auto
+ have "Group.group ((G\<lparr>carrier := H\<rparr>) \<times>\<times> (K\<lparr>carrier := I\<rparr>))"
+ using DirProd_group[OF subgroup.subgroup_is_group[OF assms(2)assms(1)]
+ subgroup.subgroup_is_group[OF assms(4)assms(3)]].
+ moreover have "((G\<lparr>carrier := H\<rparr>) \<times>\<times> (K\<lparr>carrier := I\<rparr>)) = ((G \<times>\<times> K)\<lparr>carrier := H \<times> I\<rparr>)"
+ unfolding DirProd_def using assms apply simp.
+ ultimately show "Group.group ((G \<times>\<times> K)\<lparr>carrier := H \<times> I\<rparr>)" by simp
+qed
subsection \<open>Homomorphisms and Isomorphisms\<close>
@@ -575,31 +794,203 @@
by (fastforce simp add: hom_def compose_def)
definition
- iso :: "_ => _ => ('a => 'b) set" (infixr "\<cong>" 60)
- where "G \<cong> H = {h. h \<in> hom G H \<and> bij_betw h (carrier G) (carrier H)}"
+ iso :: "_ => _ => ('a => 'b) set"
+ where "iso G H = {h. h \<in> hom G H \<and> bij_betw h (carrier G) (carrier H)}"
-lemma iso_refl: "(\<lambda>x. x) \<in> G \<cong> G"
-by (simp add: iso_def hom_def inj_on_def bij_betw_def Pi_def)
+definition
+ is_iso :: "_ \<Rightarrow> _ \<Rightarrow> bool" (infixr "\<cong>" 60)
+ where "G \<cong> H = (iso G H \<noteq> {})"
-lemma (in group) iso_sym:
- "h \<in> G \<cong> H \<Longrightarrow> inv_into (carrier G) h \<in> H \<cong> G"
-apply (simp add: iso_def bij_betw_inv_into)
-apply (subgoal_tac "inv_into (carrier G) h \<in> carrier H \<rightarrow> carrier G")
- prefer 2 apply (simp add: bij_betw_imp_funcset [OF bij_betw_inv_into])
+lemma iso_set_refl: "(\<lambda>x. x) \<in> iso G G"
+ by (simp add: iso_def hom_def inj_on_def bij_betw_def Pi_def)
+
+corollary iso_refl : "G \<cong> G"
+ using iso_set_refl unfolding is_iso_def by auto
+
+lemma (in group) iso_set_sym:
+ "h \<in> iso G H \<Longrightarrow> inv_into (carrier G) h \<in> (iso H G)"
+apply (simp add: iso_def bij_betw_inv_into)
+apply (subgoal_tac "inv_into (carrier G) h \<in> carrier H \<rightarrow> carrier G")
+ prefer 2 apply (simp add: bij_betw_imp_funcset [OF bij_betw_inv_into])
apply (simp add: hom_def bij_betw_def inv_into_f_eq f_inv_into_f Pi_def)
done
-lemma (in group) iso_trans:
- "[|h \<in> G \<cong> H; i \<in> H \<cong> I|] ==> (compose (carrier G) i h) \<in> G \<cong> I"
+corollary (in group) iso_sym :
+"G \<cong> H \<Longrightarrow> H \<cong> G"
+ using iso_set_sym unfolding is_iso_def by auto
+
+lemma (in group) iso_set_trans:
+ "[|h \<in> iso G H; i \<in> iso H I|] ==> (compose (carrier G) i h) \<in> iso G I"
by (auto simp add: iso_def hom_compose bij_betw_compose)
-lemma DirProd_commute_iso:
- shows "(\<lambda>(x,y). (y,x)) \<in> (G \<times>\<times> H) \<cong> (H \<times>\<times> G)"
+corollary (in group) iso_trans :
+"\<lbrakk>G \<cong> H ; H \<cong> I\<rbrakk> \<Longrightarrow> G \<cong> I"
+ using iso_set_trans unfolding is_iso_def by blast
+
+(* Next four lemmas contributed by Paulo. *)
+
+lemma (in monoid) hom_imp_img_monoid:
+ assumes "h \<in> hom G H"
+ shows "monoid (H \<lparr> carrier := h ` (carrier G), one := h \<one>\<^bsub>G\<^esub> \<rparr>)" (is "monoid ?h_img")
+proof (rule monoidI)
+ show "\<one>\<^bsub>?h_img\<^esub> \<in> carrier ?h_img"
+ by auto
+next
+ fix x y z assume "x \<in> carrier ?h_img" "y \<in> carrier ?h_img" "z \<in> carrier ?h_img"
+ then obtain g1 g2 g3
+ where g1: "g1 \<in> carrier G" "x = h g1"
+ and g2: "g2 \<in> carrier G" "y = h g2"
+ and g3: "g3 \<in> carrier G" "z = h g3"
+ using image_iff[where ?f = h and ?A = "carrier G"] by auto
+ have aux_lemma:
+ "\<And>a b. \<lbrakk> a \<in> carrier G; b \<in> carrier G \<rbrakk> \<Longrightarrow> h a \<otimes>\<^bsub>(?h_img)\<^esub> h b = h (a \<otimes> b)"
+ using assms unfolding hom_def by auto
+
+ show "x \<otimes>\<^bsub>(?h_img)\<^esub> \<one>\<^bsub>(?h_img)\<^esub> = x"
+ using aux_lemma[OF g1(1) one_closed] g1(2) r_one[OF g1(1)] by simp
+
+ show "\<one>\<^bsub>(?h_img)\<^esub> \<otimes>\<^bsub>(?h_img)\<^esub> x = x"
+ using aux_lemma[OF one_closed g1(1)] g1(2) l_one[OF g1(1)] by simp
+
+ have "x \<otimes>\<^bsub>(?h_img)\<^esub> y = h (g1 \<otimes> g2)"
+ using aux_lemma g1 g2 by auto
+ thus "x \<otimes>\<^bsub>(?h_img)\<^esub> y \<in> carrier ?h_img"
+ using g1(1) g2(1) by simp
+
+ have "(x \<otimes>\<^bsub>(?h_img)\<^esub> y) \<otimes>\<^bsub>(?h_img)\<^esub> z = h ((g1 \<otimes> g2) \<otimes> g3)"
+ using aux_lemma g1 g2 g3 by auto
+ also have " ... = h (g1 \<otimes> (g2 \<otimes> g3))"
+ using m_assoc[OF g1(1) g2(1) g3(1)] by simp
+ also have " ... = x \<otimes>\<^bsub>(?h_img)\<^esub> (y \<otimes>\<^bsub>(?h_img)\<^esub> z)"
+ using aux_lemma g1 g2 g3 by auto
+ finally show "(x \<otimes>\<^bsub>(?h_img)\<^esub> y) \<otimes>\<^bsub>(?h_img)\<^esub> z = x \<otimes>\<^bsub>(?h_img)\<^esub> (y \<otimes>\<^bsub>(?h_img)\<^esub> z)" .
+qed
+
+lemma (in group) hom_imp_img_group:
+ assumes "h \<in> hom G H"
+ shows "group (H \<lparr> carrier := h ` (carrier G), one := h \<one>\<^bsub>G\<^esub> \<rparr>)" (is "group ?h_img")
+proof -
+ interpret monoid ?h_img
+ using hom_imp_img_monoid[OF assms] .
+
+ show ?thesis
+ proof (unfold_locales)
+ show "carrier ?h_img \<subseteq> Units ?h_img"
+ proof (auto simp add: Units_def)
+ have aux_lemma:
+ "\<And>g1 g2. \<lbrakk> g1 \<in> carrier G; g2 \<in> carrier G \<rbrakk> \<Longrightarrow> h g1 \<otimes>\<^bsub>H\<^esub> h g2 = h (g1 \<otimes> g2)"
+ using assms unfolding hom_def by auto
+
+ fix g1 assume g1: "g1 \<in> carrier G"
+ thus "\<exists>g2 \<in> carrier G. (h g2) \<otimes>\<^bsub>H\<^esub> (h g1) = h \<one> \<and> (h g1) \<otimes>\<^bsub>H\<^esub> (h g2) = h \<one>"
+ using aux_lemma[OF g1 inv_closed[OF g1]]
+ aux_lemma[OF inv_closed[OF g1] g1]
+ inv_closed by auto
+ qed
+ qed
+qed
+
+lemma (in group) iso_imp_group:
+ assumes "G \<cong> H" and "monoid H"
+ shows "group H"
+proof -
+ obtain \<phi> where phi: "\<phi> \<in> iso G H" "inv_into (carrier G) \<phi> \<in> iso H G"
+ using iso_set_sym assms unfolding is_iso_def by blast
+ define \<psi> where psi_def: "\<psi> = inv_into (carrier G) \<phi>"
+
+ from phi
+ have surj: "\<phi> ` (carrier G) = (carrier H)" "\<psi> ` (carrier H) = (carrier G)"
+ and inj: "inj_on \<phi> (carrier G)" "inj_on \<psi> (carrier H)"
+ and phi_hom: "\<And>g1 g2. \<lbrakk> g1 \<in> carrier G; g2 \<in> carrier G \<rbrakk> \<Longrightarrow> \<phi> (g1 \<otimes> g2) = (\<phi> g1) \<otimes>\<^bsub>H\<^esub> (\<phi> g2)"
+ and psi_hom: "\<And>h1 h2. \<lbrakk> h1 \<in> carrier H; h2 \<in> carrier H \<rbrakk> \<Longrightarrow> \<psi> (h1 \<otimes>\<^bsub>H\<^esub> h2) = (\<psi> h1) \<otimes> (\<psi> h2)"
+ using psi_def unfolding iso_def bij_betw_def hom_def by auto
+
+ have phi_one: "\<phi> \<one> = \<one>\<^bsub>H\<^esub>"
+ proof -
+ have "(\<phi> \<one>) \<otimes>\<^bsub>H\<^esub> \<one>\<^bsub>H\<^esub> = (\<phi> \<one>) \<otimes>\<^bsub>H\<^esub> (\<phi> \<one>)"
+ by (metis assms(2) image_eqI monoid.r_one one_closed phi_hom r_one surj(1))
+ thus ?thesis
+ by (metis (no_types, hide_lams) Units_eq Units_one_closed assms(2) f_inv_into_f imageI
+ monoid.l_one monoid.one_closed phi_hom psi_def r_one surj)
+ qed
+
+ have "carrier H \<subseteq> Units H"
+ proof
+ fix h assume h: "h \<in> carrier H"
+ let ?inv_h = "\<phi> (inv (\<psi> h))"
+ have "h \<otimes>\<^bsub>H\<^esub> ?inv_h = \<phi> (\<psi> h) \<otimes>\<^bsub>H\<^esub> ?inv_h"
+ by (simp add: f_inv_into_f h psi_def surj(1))
+ also have " ... = \<phi> ((\<psi> h) \<otimes> inv (\<psi> h))"
+ by (metis h imageI inv_closed phi_hom surj(2))
+ also have " ... = \<phi> \<one>"
+ by (simp add: h inv_into_into psi_def surj(1))
+ finally have 1: "h \<otimes>\<^bsub>H\<^esub> ?inv_h = \<one>\<^bsub>H\<^esub>"
+ using phi_one by simp
+
+ have "?inv_h \<otimes>\<^bsub>H\<^esub> h = ?inv_h \<otimes>\<^bsub>H\<^esub> \<phi> (\<psi> h)"
+ by (simp add: f_inv_into_f h psi_def surj(1))
+ also have " ... = \<phi> (inv (\<psi> h) \<otimes> (\<psi> h))"
+ by (metis h imageI inv_closed phi_hom surj(2))
+ also have " ... = \<phi> \<one>"
+ by (simp add: h inv_into_into psi_def surj(1))
+ finally have 2: "?inv_h \<otimes>\<^bsub>H\<^esub> h = \<one>\<^bsub>H\<^esub>"
+ using phi_one by simp
+
+ thus "h \<in> Units H" unfolding Units_def using 1 2 h surj by fastforce
+ qed
+ thus ?thesis unfolding group_def group_axioms_def using assms(2) by simp
+qed
+
+corollary (in group) iso_imp_img_group:
+ assumes "h \<in> iso G H"
+ shows "group (H \<lparr> one := h \<one> \<rparr>)"
+proof -
+ let ?h_img = "H \<lparr> carrier := h ` (carrier G), one := h \<one> \<rparr>"
+ have "h \<in> iso G ?h_img"
+ using assms unfolding iso_def hom_def bij_betw_def by auto
+ hence "G \<cong> ?h_img"
+ unfolding is_iso_def by auto
+ hence "group ?h_img"
+ using iso_imp_group[of ?h_img] hom_imp_img_monoid[of h H] assms unfolding iso_def by simp
+ moreover have "carrier H = carrier ?h_img"
+ using assms unfolding iso_def bij_betw_def by simp
+ hence "H \<lparr> one := h \<one> \<rparr> = ?h_img"
+ by simp
+ ultimately show ?thesis by simp
+qed
+
+lemma DirProd_commute_iso_set:
+ shows "(\<lambda>(x,y). (y,x)) \<in> iso (G \<times>\<times> H) (H \<times>\<times> G)"
+ by (auto simp add: iso_def hom_def inj_on_def bij_betw_def)
+
+corollary DirProd_commute_iso :
+"(G \<times>\<times> H) \<cong> (H \<times>\<times> G)"
+ using DirProd_commute_iso_set unfolding is_iso_def by blast
+
+lemma DirProd_assoc_iso_set:
+ shows "(\<lambda>(x,y,z). (x,(y,z))) \<in> iso (G \<times>\<times> H \<times>\<times> I) (G \<times>\<times> (H \<times>\<times> I))"
by (auto simp add: iso_def hom_def inj_on_def bij_betw_def)
-lemma DirProd_assoc_iso:
- shows "(\<lambda>(x,y,z). (x,(y,z))) \<in> (G \<times>\<times> H \<times>\<times> I) \<cong> (G \<times>\<times> (H \<times>\<times> I))"
-by (auto simp add: iso_def hom_def inj_on_def bij_betw_def)
+lemma (in group) DirProd_iso_set_trans:
+ assumes "g \<in> iso G G2"
+ and "h \<in> iso H I"
+ shows "(\<lambda>(x,y). (g x, h y)) \<in> iso (G \<times>\<times> H) (G2 \<times>\<times> I)"
+proof-
+ have "(\<lambda>(x,y). (g x, h y)) \<in> hom (G \<times>\<times> H) (G2 \<times>\<times> I)"
+ using assms unfolding iso_def hom_def by auto
+ moreover have " inj_on (\<lambda>(x,y). (g x, h y)) (carrier (G \<times>\<times> H))"
+ using assms unfolding iso_def DirProd_def bij_betw_def inj_on_def by auto
+ moreover have "(\<lambda>(x, y). (g x, h y)) ` carrier (G \<times>\<times> H) = carrier (G2 \<times>\<times> I)"
+ using assms unfolding iso_def bij_betw_def image_def DirProd_def by fastforce
+ ultimately show "(\<lambda>(x,y). (g x, h y)) \<in> iso (G \<times>\<times> H) (G2 \<times>\<times> I)"
+ unfolding iso_def bij_betw_def by auto
+qed
+
+corollary (in group) DirProd_iso_trans :
+ assumes "G \<cong> G2"
+ and "H \<cong> I"
+ shows "G \<times>\<times> H \<cong> G2 \<times>\<times> I"
+ using DirProd_iso_set_trans assms unfolding is_iso_def by blast
text\<open>Basis for homomorphism proofs: we assume two groups @{term G} and
@@ -655,6 +1046,56 @@
"x \<in> carrier G \<Longrightarrow> (([^]) x) \<in> hom \<lparr> carrier = UNIV, mult = (+), one = 0::int \<rparr> G "
unfolding hom_def by (simp add: int_pow_mult)
+(* Next six lemmas contributed by Paulo. *)
+
+lemma (in group_hom) img_is_subgroup: "subgroup (h ` (carrier G)) H"
+ apply (rule subgroupI)
+ apply (auto simp add: image_subsetI)
+ apply (metis (no_types, hide_lams) G.inv_closed hom_inv image_iff)
+ apply (smt G.monoid_axioms hom_mult image_iff monoid.m_closed)
+ done
+
+lemma (in group_hom) subgroup_img_is_subgroup:
+ assumes "subgroup I G"
+ shows "subgroup (h ` I) H"
+proof -
+ have "h \<in> hom (G \<lparr> carrier := I \<rparr>) H"
+ using G.subgroupE[OF assms] subgroup.mem_carrier[OF assms] homh
+ unfolding hom_def by auto
+ hence "group_hom (G \<lparr> carrier := I \<rparr>) H h"
+ using subgroup.subgroup_is_group[OF assms G.is_group] is_group
+ unfolding group_hom_def group_hom_axioms_def by simp
+ thus ?thesis
+ using group_hom.img_is_subgroup[of "G \<lparr> carrier := I \<rparr>" H h] by simp
+qed
+
+lemma (in group_hom) induced_group_hom:
+ assumes "subgroup I G"
+ shows "group_hom (G \<lparr> carrier := I \<rparr>) (H \<lparr> carrier := h ` I \<rparr>) h"
+proof -
+ have "h \<in> hom (G \<lparr> carrier := I \<rparr>) (H \<lparr> carrier := h ` I \<rparr>)"
+ using homh subgroup.mem_carrier[OF assms] unfolding hom_def by auto
+ thus ?thesis
+ unfolding group_hom_def group_hom_axioms_def
+ using subgroup.subgroup_is_group[OF assms G.is_group]
+ subgroup.subgroup_is_group[OF subgroup_img_is_subgroup[OF assms] is_group] by simp
+qed
+
+lemma (in group) canonical_inj_is_hom:
+ assumes "subgroup H G"
+ shows "group_hom (G \<lparr> carrier := H \<rparr>) G id"
+ unfolding group_hom_def group_hom_axioms_def hom_def
+ using subgroup.subgroup_is_group[OF assms is_group]
+ is_group subgroup.subset[OF assms] by auto
+
+lemma (in group_hom) nat_pow_hom:
+ "x \<in> carrier G \<Longrightarrow> h (x [^] (n :: nat)) = (h x) [^]\<^bsub>H\<^esub> n"
+ by (induction n) auto
+
+lemma (in group_hom) int_pow_hom:
+ "x \<in> carrier G \<Longrightarrow> h (x [^] (n :: int)) = (h x) [^]\<^bsub>H\<^esub> n"
+ using int_pow_def2 nat_pow_hom by (simp add: G.int_pow_def2)
+
subsection \<open>Commutative Structures\<close>
@@ -693,7 +1134,7 @@
"!!x y. [| x \<in> carrier G; y \<in> carrier G |] ==> x \<otimes> y = y \<otimes> x"
shows "comm_monoid G"
using l_one
- by (auto intro!: comm_monoid.intro comm_monoid_axioms.intro monoid.intro
+ by (auto intro!: comm_monoid.intro comm_monoid_axioms.intro monoid.intro
intro: assms simp: m_closed one_closed m_comm)
lemma (in monoid) monoid_comm_monoidI:
@@ -716,6 +1157,18 @@
(x \<otimes> y) [^] (n::nat) = x [^] n \<otimes> y [^] n"
by (induct n) (simp, simp add: m_ac)
+lemma (in comm_monoid) submonoid_is_comm_monoid :
+ assumes "submonoid H G"
+ shows "comm_monoid (G\<lparr>carrier := H\<rparr>)"
+proof (intro monoid.monoid_comm_monoidI)
+ show "monoid (G\<lparr>carrier := H\<rparr>)"
+ using submonoid.submonoid_is_monoid assms comm_monoid_axioms comm_monoid_def by blast
+ show "\<And>x y. x \<in> carrier (G\<lparr>carrier := H\<rparr>) \<Longrightarrow> y \<in> carrier (G\<lparr>carrier := H\<rparr>)
+ \<Longrightarrow> x \<otimes>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> y = y \<otimes>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> x" apply simp
+ using assms comm_monoid_axioms_def submonoid.mem_carrier
+ by (metis m_comm)
+qed
+
locale comm_group = comm_monoid + group
lemma (in group) group_comm_groupI:
@@ -739,10 +1192,109 @@
shows "comm_group G"
by (fast intro: group.group_comm_groupI groupI assms)
+lemma comm_groupE:
+ fixes G (structure)
+ assumes "comm_group G"
+ shows "\<And>x y. \<lbrakk> x \<in> carrier G; y \<in> carrier G \<rbrakk> \<Longrightarrow> x \<otimes> y \<in> carrier G"
+ and "\<one> \<in> carrier G"
+ and "\<And>x y z. \<lbrakk> x \<in> carrier G; y \<in> carrier G; z \<in> carrier G \<rbrakk> \<Longrightarrow> (x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)"
+ and "\<And>x y. \<lbrakk> x \<in> carrier G; y \<in> carrier G \<rbrakk> \<Longrightarrow> x \<otimes> y = y \<otimes> x"
+ and "\<And>x. x \<in> carrier G \<Longrightarrow> \<one> \<otimes> x = x"
+ and "\<And>x. x \<in> carrier G \<Longrightarrow> \<exists>y \<in> carrier G. y \<otimes> x = \<one>"
+ apply (simp_all add: group.axioms assms comm_group.axioms comm_monoid.m_comm comm_monoid.m_ac(1))
+ by (simp_all add: Group.group.axioms(1) assms comm_group.axioms(2) monoid.m_closed group.r_inv_ex)
+
lemma (in comm_group) inv_mult:
"[| x \<in> carrier G; y \<in> carrier G |] ==> inv (x \<otimes> y) = inv x \<otimes> inv y"
by (simp add: m_ac inv_mult_group)
+(* Next three lemmas contributed by Paulo. *)
+
+lemma (in comm_monoid) hom_imp_img_comm_monoid:
+ assumes "h \<in> hom G H"
+ shows "comm_monoid (H \<lparr> carrier := h ` (carrier G), one := h \<one>\<^bsub>G\<^esub> \<rparr>)" (is "comm_monoid ?h_img")
+proof (rule monoid.monoid_comm_monoidI)
+ show "monoid ?h_img"
+ using hom_imp_img_monoid[OF assms] .
+next
+ fix x y assume "x \<in> carrier ?h_img" "y \<in> carrier ?h_img"
+ then obtain g1 g2
+ where g1: "g1 \<in> carrier G" "x = h g1"
+ and g2: "g2 \<in> carrier G" "y = h g2"
+ by auto
+ have "x \<otimes>\<^bsub>(?h_img)\<^esub> y = h (g1 \<otimes> g2)"
+ using g1 g2 assms unfolding hom_def by auto
+ also have " ... = h (g2 \<otimes> g1)"
+ using m_comm[OF g1(1) g2(1)] by simp
+ also have " ... = y \<otimes>\<^bsub>(?h_img)\<^esub> x"
+ using g1 g2 assms unfolding hom_def by auto
+ finally show "x \<otimes>\<^bsub>(?h_img)\<^esub> y = y \<otimes>\<^bsub>(?h_img)\<^esub> x" .
+qed
+
+lemma (in comm_group) iso_imp_img_comm_group:
+ assumes "h \<in> iso G H"
+ shows "comm_group (H \<lparr> one := h \<one>\<^bsub>G\<^esub> \<rparr>)"
+proof -
+ let ?h_img = "H \<lparr> carrier := h ` (carrier G), one := h \<one> \<rparr>"
+ have "comm_monoid ?h_img"
+ using hom_imp_img_comm_monoid[of h H] assms unfolding iso_def by simp
+ moreover have "carrier H = carrier ?h_img"
+ using assms unfolding iso_def bij_betw_def by simp
+ hence "H \<lparr> one := h \<one> \<rparr> = ?h_img"
+ by simp
+ ultimately have "comm_monoid (H \<lparr> one := h \<one>\<^bsub>G\<^esub> \<rparr>)"
+ by simp
+ thus ?thesis
+ unfolding comm_group_def using iso_imp_img_group[OF assms] by simp
+qed
+
+lemma (in comm_group) iso_imp_comm_group:
+ assumes "G \<cong> H" "monoid H"
+ shows "comm_group H"
+proof -
+ obtain h where h: "h \<in> iso G H"
+ using assms(1) unfolding is_iso_def by auto
+ hence comm_gr: "comm_group (H \<lparr> one := h \<one> \<rparr>)"
+ using iso_imp_img_comm_group[of h H] by simp
+ hence "\<And>x. x \<in> carrier H \<Longrightarrow> h \<one> \<otimes>\<^bsub>H\<^esub> x = x"
+ using monoid.l_one[of "H \<lparr> one := h \<one> \<rparr>"] unfolding comm_group_def comm_monoid_def by simp
+ moreover have "h \<one> \<in> carrier H"
+ using h one_closed unfolding iso_def hom_def by auto
+ ultimately have "h \<one> = \<one>\<^bsub>H\<^esub>"
+ using monoid.one_unique[OF assms(2), of "h \<one>"] by simp
+ hence "H = H \<lparr> one := h \<one> \<rparr>"
+ by simp
+ thus ?thesis
+ using comm_gr by simp
+qed
+
+(*A subgroup of a subgroup is a subgroup of the group*)
+lemma (in group) incl_subgroup:
+ assumes "subgroup J G"
+ and "subgroup I (G\<lparr>carrier:=J\<rparr>)"
+ shows "subgroup I G" unfolding subgroup_def
+proof
+ have H1: "I \<subseteq> carrier (G\<lparr>carrier:=J\<rparr>)" using assms(2) subgroup_imp_subset by blast
+ also have H2: "...\<subseteq>J" by simp
+ also have "...\<subseteq>(carrier G)" by (simp add: assms(1) subgroup_imp_subset)
+ finally have H: "I \<subseteq> carrier G" by simp
+ have "(\<And>x y. \<lbrakk>x \<in> I ; y \<in> I\<rbrakk> \<Longrightarrow> x \<otimes> y \<in> I)" using assms(2) by (auto simp add: subgroup_def)
+ thus "I \<subseteq> carrier G \<and> (\<forall>x y. x \<in> I \<longrightarrow> y \<in> I \<longrightarrow> x \<otimes> y \<in> I)" using H by blast
+ have K: "\<one> \<in> I" using assms(2) by (auto simp add: subgroup_def)
+ have "(\<And>x. x \<in> I \<Longrightarrow> inv x \<in> I)" using assms subgroup.m_inv_closed H
+ by (metis H1 H2 subgroup_inv_equality subsetCE)
+ thus "\<one> \<in> I \<and> (\<forall>x. x \<in> I \<longrightarrow> inv x \<in> I)" using K by blast
+qed
+
+(*A subgroup included in another subgroup is a subgroup of the subgroup*)
+lemma (in group) subgroup_incl:
+ assumes "subgroup I G"
+ and "subgroup J G"
+ and "I\<subseteq>J"
+ shows "subgroup I (G\<lparr>carrier:=J\<rparr>)"using assms subgroup_inv_equality
+ by (auto simp add: subgroup_def)
+
+
subsection \<open>The Lattice of Subgroups of a Group\<close>
@@ -762,16 +1314,11 @@
lemma (in group) is_monoid [intro, simp]:
"monoid G"
- by (auto intro: monoid.intro m_assoc)
+ by (auto intro: monoid.intro m_assoc)
-lemma (in group) subgroup_inv_equality:
- "[| subgroup H G; x \<in> H |] ==> m_inv (G \<lparr>carrier := H\<rparr>) x = inv x"
-apply (rule_tac inv_equality [THEN sym])
- apply (rule group.l_inv [OF subgroup_imp_group, simplified], assumption+)
- apply (rule subsetD [OF subgroup.subset], assumption+)
-apply (rule subsetD [OF subgroup.subset], assumption)
-apply (rule_tac group.inv_closed [OF subgroup_imp_group, simplified], assumption+)
-done
+lemma (in group) subgroup_mult_equality:
+ "\<lbrakk> subgroup H G; h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<otimes>\<^bsub>G \<lparr> carrier := H \<rparr>\<^esub> h2 = h1 \<otimes> h2"
+ unfolding subgroup_def by simp
theorem (in group) subgroups_Inter:
assumes subgr: "(\<And>H. H \<in> A \<Longrightarrow> subgroup H G)"
@@ -793,6 +1340,11 @@
show "x \<otimes> y \<in> \<Inter>A" by blast
qed
+lemma (in group) subgroups_Inter_pair :
+ assumes "subgroup I G"
+ and "subgroup J G"
+ shows "subgroup (I\<inter>J) G" using subgroups_Inter[ where ?A = "{I,J}"] assms by auto
+
theorem (in group) subgroups_complete_lattice:
"complete_lattice \<lparr>carrier = {H. subgroup H G}, eq = (=), le = (\<subseteq>)\<rparr>"
(is "complete_lattice ?L")
@@ -831,4 +1383,90 @@
then show "\<exists>I. greatest ?L I (Lower ?L A)" ..
qed
+subsection\<open>Jeremy Avigad's @{text"More_Group"} material\<close>
+
+text \<open>
+ Show that the units in any monoid give rise to a group.
+
+ The file Residues.thy provides some infrastructure to use
+ facts about the unit group within the ring locale.
+\<close>
+
+definition units_of :: "('a, 'b) monoid_scheme \<Rightarrow> 'a monoid"
+ where "units_of G =
+ \<lparr>carrier = Units G, Group.monoid.mult = Group.monoid.mult G, one = one G\<rparr>"
+
+lemma (in monoid) units_group: "group (units_of G)"
+ apply (unfold units_of_def)
+ apply (rule groupI)
+ apply auto
+ apply (subst m_assoc)
+ apply auto
+ apply (rule_tac x = "inv x" in bexI)
+ apply auto
+ done
+
+lemma (in comm_monoid) units_comm_group: "comm_group (units_of G)"
+ apply (rule group.group_comm_groupI)
+ apply (rule units_group)
+ apply (insert comm_monoid_axioms)
+ apply (unfold units_of_def Units_def comm_monoid_def comm_monoid_axioms_def)
+ apply auto
+ done
+
+lemma units_of_carrier: "carrier (units_of G) = Units G"
+ by (auto simp: units_of_def)
+
+lemma units_of_mult: "mult (units_of G) = mult G"
+ by (auto simp: units_of_def)
+
+lemma units_of_one: "one (units_of G) = one G"
+ by (auto simp: units_of_def)
+
+lemma (in monoid) units_of_inv: "x \<in> Units G \<Longrightarrow> m_inv (units_of G) x = m_inv G x"
+ apply (rule sym)
+ apply (subst m_inv_def)
+ apply (rule the1_equality)
+ apply (rule ex_ex1I)
+ apply (subst (asm) Units_def)
+ apply auto
+ apply (erule inv_unique)
+ apply auto
+ apply (rule Units_closed)
+ apply (simp_all only: units_of_carrier [symmetric])
+ apply (insert units_group)
+ apply auto
+ apply (subst units_of_mult [symmetric])
+ apply (subst units_of_one [symmetric])
+ apply (erule group.r_inv, assumption)
+ apply (subst units_of_mult [symmetric])
+ apply (subst units_of_one [symmetric])
+ apply (erule group.l_inv, assumption)
+ done
+
+lemma (in group) inj_on_const_mult: "a \<in> carrier G \<Longrightarrow> inj_on (\<lambda>x. a \<otimes> x) (carrier G)"
+ unfolding inj_on_def by auto
+
+lemma (in group) surj_const_mult: "a \<in> carrier G \<Longrightarrow> (\<lambda>x. a \<otimes> x) ` carrier G = carrier G"
+ apply (auto simp add: image_def)
+ apply (rule_tac x = "(m_inv G a) \<otimes> x" in bexI)
+ apply auto
+(* auto should get this. I suppose we need "comm_monoid_simprules"
+ for ac_simps rewriting. *)
+ apply (subst m_assoc [symmetric])
+ apply auto
+ done
+
+lemma (in group) l_cancel_one [simp]: "x \<in> carrier G \<Longrightarrow> a \<in> carrier G \<Longrightarrow> x \<otimes> a = x \<longleftrightarrow> a = one G"
+ by (metis Units_eq Units_l_cancel monoid.r_one monoid_axioms one_closed)
+
+lemma (in group) r_cancel_one [simp]: "x \<in> carrier G \<Longrightarrow> a \<in> carrier G \<Longrightarrow> a \<otimes> x = x \<longleftrightarrow> a = one G"
+ by (metis monoid.l_one monoid_axioms one_closed right_cancel)
+
+lemma (in group) l_cancel_one' [simp]: "x \<in> carrier G \<Longrightarrow> a \<in> carrier G \<Longrightarrow> x = x \<otimes> a \<longleftrightarrow> a = one G"
+ using l_cancel_one by fastforce
+
+lemma (in group) r_cancel_one' [simp]: "x \<in> carrier G \<Longrightarrow> a \<in> carrier G \<Longrightarrow> x = a \<otimes> x \<longleftrightarrow> a = one G"
+ using r_cancel_one by fastforce
+
end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Algebra/Group_Action.thy Thu Jun 14 14:23:48 2018 +0100
@@ -0,0 +1,925 @@
+(* Title: Group_Action.thy *)
+(* Author: Paulo EmÃlio de Vilhena *)
+
+theory Group_Action
+imports Bij Coset Congruence
+
+begin
+
+section \<open>Group Actions\<close>
+
+locale group_action =
+ fixes G (structure) and E and \<phi>
+ assumes group_hom: "group_hom G (BijGroup E) \<phi>"
+
+definition
+ orbit :: "[_, 'a \<Rightarrow> 'b \<Rightarrow> 'b, 'b] \<Rightarrow> 'b set"
+ where "orbit G \<phi> x = {(\<phi> g) x | g. g \<in> carrier G}"
+
+definition
+ orbits :: "[_, 'b set, 'a \<Rightarrow> 'b \<Rightarrow> 'b] \<Rightarrow> ('b set) set"
+ where "orbits G E \<phi> = {orbit G \<phi> x | x. x \<in> E}"
+
+definition
+ stabilizer :: "[_, 'a \<Rightarrow> 'b \<Rightarrow> 'b, 'b] \<Rightarrow> 'a set"
+ where "stabilizer G \<phi> x = {g \<in> carrier G. (\<phi> g) x = x}"
+
+definition
+ invariants :: "['b set, 'a \<Rightarrow> 'b \<Rightarrow> 'b, 'a] \<Rightarrow> 'b set"
+ where "invariants E \<phi> g = {x \<in> E. (\<phi> g) x = x}"
+
+definition
+ normalizer :: "[_, 'a set] \<Rightarrow> 'a set"
+ where "normalizer G H =
+ stabilizer G (\<lambda>g. \<lambda>H \<in> {H. H \<subseteq> carrier G}. g <#\<^bsub>G\<^esub> H #>\<^bsub>G\<^esub> (inv\<^bsub>G\<^esub> g)) H"
+
+locale faithful_action = group_action +
+ assumes faithful: "inj_on \<phi> (carrier G)"
+
+locale transitive_action = group_action +
+ assumes unique_orbit: "\<lbrakk> x \<in> E; y \<in> E \<rbrakk> \<Longrightarrow> \<exists>g \<in> carrier G. (\<phi> g) x = y"
+
+
+
+subsection \<open>Prelimineries\<close>
+
+text \<open>Some simple lemmas to make group action's properties more explicit\<close>
+
+lemma (in group_action) id_eq_one: "(\<lambda>x \<in> E. x) = \<phi> \<one>"
+ by (metis BijGroup_def group_hom group_hom.hom_one select_convs(2))
+
+lemma (in group_action) bij_prop0:
+ "\<And> g. g \<in> carrier G \<Longrightarrow> (\<phi> g) \<in> Bij E"
+ by (metis BijGroup_def group_hom group_hom.hom_closed partial_object.select_convs(1))
+
+lemma (in group_action) surj_prop:
+ "\<And> g. g \<in> carrier G \<Longrightarrow> (\<phi> g) ` E = E"
+ using bij_prop0 by (simp add: Bij_def bij_betw_def)
+
+lemma (in group_action) inj_prop:
+ "\<And> g. g \<in> carrier G \<Longrightarrow> inj_on (\<phi> g) E"
+ using bij_prop0 by (simp add: Bij_def bij_betw_def)
+
+lemma (in group_action) bij_prop1:
+ "\<And> g y. \<lbrakk> g \<in> carrier G; y \<in> E \<rbrakk> \<Longrightarrow> \<exists>!x \<in> E. (\<phi> g) x = y"
+proof -
+ fix g y assume "g \<in> carrier G" "y \<in> E"
+ hence "\<exists>x \<in> E. (\<phi> g) x = y"
+ using surj_prop by force
+ moreover have "\<And> x1 x2. \<lbrakk> x1 \<in> E; x2 \<in> E \<rbrakk> \<Longrightarrow> (\<phi> g) x1 = (\<phi> g) x2 \<Longrightarrow> x1 = x2"
+ using inj_prop by (meson \<open>g \<in> carrier G\<close> inj_on_eq_iff)
+ ultimately show "\<exists>!x \<in> E. (\<phi> g) x = y"
+ by blast
+qed
+
+lemma (in group_action) composition_rule:
+ assumes "x \<in> E" "g1 \<in> carrier G" "g2 \<in> carrier G"
+ shows "\<phi> (g1 \<otimes> g2) x = (\<phi> g1) (\<phi> g2 x)"
+proof -
+ have "\<phi> (g1 \<otimes> g2) x = ((\<phi> g1) \<otimes>\<^bsub>BijGroup E\<^esub> (\<phi> g2)) x"
+ using assms(2) assms(3) group_hom group_hom.hom_mult by fastforce
+ also have " ... = (compose E (\<phi> g1) (\<phi> g2)) x"
+ unfolding BijGroup_def by (simp add: assms bij_prop0)
+ finally show "\<phi> (g1 \<otimes> g2) x = (\<phi> g1) (\<phi> g2 x)"
+ by (simp add: assms(1) compose_eq)
+qed
+
+lemma (in group_action) element_image:
+ assumes "g \<in> carrier G" and "x \<in> E" and "(\<phi> g) x = y"
+ shows "y \<in> E"
+ using surj_prop assms by blast
+
+
+
+subsection \<open>Orbits\<close>
+
+text\<open>We prove here that orbits form an equivalence relation\<close>
+
+lemma (in group_action) orbit_sym_aux:
+ assumes "g \<in> carrier G"
+ and "x \<in> E"
+ and "(\<phi> g) x = y"
+ shows "(\<phi> (inv g)) y = x"
+proof -
+ interpret group G
+ using group_hom group_hom.axioms(1) by auto
+ have "y \<in> E"
+ using element_image assms by simp
+ have "inv g \<in> carrier G"
+ by (simp add: assms(1))
+
+ have "(\<phi> (inv g)) y = (\<phi> (inv g)) ((\<phi> g) x)"
+ using assms(3) by simp
+ also have " ... = compose E (\<phi> (inv g)) (\<phi> g) x"
+ by (simp add: assms(2) compose_eq)
+ also have " ... = ((\<phi> (inv g)) \<otimes>\<^bsub>BijGroup E\<^esub> (\<phi> g)) x"
+ by (simp add: BijGroup_def assms(1) bij_prop0)
+ also have " ... = (\<phi> ((inv g) \<otimes> g)) x"
+ by (metis \<open>inv g \<in> carrier G\<close> assms(1) group_hom group_hom.hom_mult)
+ finally show "(\<phi> (inv g)) y = x"
+ by (metis assms(1) assms(2) id_eq_one l_inv restrict_apply)
+qed
+
+lemma (in group_action) orbit_refl:
+ "x \<in> E \<Longrightarrow> x \<in> orbit G \<phi> x"
+proof -
+ assume "x \<in> E" hence "(\<phi> \<one>) x = x"
+ using id_eq_one by (metis restrict_apply')
+ thus "x \<in> orbit G \<phi> x" unfolding orbit_def
+ using group.is_monoid group_hom group_hom.axioms(1) by force
+qed
+
+lemma (in group_action) orbit_sym:
+ assumes "x \<in> E" and "y \<in> E" and "y \<in> orbit G \<phi> x"
+ shows "x \<in> orbit G \<phi> y"
+proof -
+ have "\<exists> g \<in> carrier G. (\<phi> g) x = y"
+ by (smt assms(3) mem_Collect_eq orbit_def)
+ then obtain g where g: "g \<in> carrier G \<and> (\<phi> g) x = y" by blast
+ hence "(\<phi> (inv g)) y = x"
+ using orbit_sym_aux by (simp add: assms(1))
+ thus ?thesis
+ using g group_hom group_hom.axioms(1) orbit_def by fastforce
+qed
+
+lemma (in group_action) orbit_trans:
+ assumes "x \<in> E" "y \<in> E" "z \<in> E"
+ and "y \<in> orbit G \<phi> x" "z \<in> orbit G \<phi> y"
+ shows "z \<in> orbit G \<phi> x"
+proof -
+ interpret group G
+ using group_hom group_hom.axioms(1) by auto
+
+ have "\<exists> g1 \<in> carrier G. (\<phi> g1) x = y"
+ by (smt assms mem_Collect_eq orbit_def)
+ then obtain g1 where g1: "g1 \<in> carrier G \<and> (\<phi> g1) x = y" by blast
+
+ have "\<exists> g2 \<in> carrier G. (\<phi> g2) y = z"
+ by (smt assms mem_Collect_eq orbit_def)
+ then obtain g2 where g2: "g2 \<in> carrier G \<and> (\<phi> g2) y = z" by blast
+
+ have "(\<phi> (g2 \<otimes> g1)) x = ((\<phi> g2) \<otimes>\<^bsub>BijGroup E\<^esub> (\<phi> g1)) x"
+ using g1 g2 group_hom group_hom.hom_mult by fastforce
+ also have " ... = (\<phi> g2) ((\<phi> g1) x)"
+ using composition_rule assms(1) calculation g1 g2 by auto
+ finally have "(\<phi> (g2 \<otimes> g1)) x = z"
+ by (simp add: g1 g2)
+ thus ?thesis
+ using g1 g2 orbit_def by force
+qed
+
+lemma (in group_action) orbits_as_classes:
+ "classes\<^bsub>\<lparr> carrier = E, eq = \<lambda>x. \<lambda>y. y \<in> orbit G \<phi> x \<rparr>\<^esub> = orbits G E \<phi>"
+ unfolding eq_classes_def eq_class_of_def orbits_def apply simp
+proof -
+ have "\<And>x. x \<in> E \<Longrightarrow> {y \<in> E. y \<in> orbit G \<phi> x} = orbit G \<phi> x"
+ by (smt Collect_cong element_image mem_Collect_eq orbit_def)
+ thus "{{y \<in> E. y \<in> orbit G \<phi> x} |x. x \<in> E} = {orbit G \<phi> x |x. x \<in> E}" by blast
+qed
+
+theorem (in group_action) orbit_partition:
+ "partition E (orbits G E \<phi>)"
+proof -
+ have "equivalence \<lparr> carrier = E, eq = \<lambda>x. \<lambda>y. y \<in> orbit G \<phi> x \<rparr>"
+ unfolding equivalence_def apply simp
+ using orbit_refl orbit_sym orbit_trans by blast
+ thus ?thesis using equivalence.partition_from_equivalence orbits_as_classes by fastforce
+qed
+
+corollary (in group_action) orbits_coverture:
+ "\<Union> (orbits G E \<phi>) = E"
+ using partition.partition_coverture[OF orbit_partition] by simp
+
+corollary (in group_action) disjoint_union:
+ assumes "orb1 \<in> (orbits G E \<phi>)" "orb2 \<in> (orbits G E \<phi>)"
+ shows "(orb1 = orb2) \<or> (orb1 \<inter> orb2) = {}"
+ using partition.disjoint_union[OF orbit_partition] assms by auto
+
+corollary (in group_action) disjoint_sum:
+ assumes "finite E"
+ shows "(\<Sum>orb\<in>(orbits G E \<phi>). \<Sum>x\<in>orb. f x) = (\<Sum>x\<in>E. f x)"
+ using partition.disjoint_sum[OF orbit_partition] assms by auto
+
+
+subsubsection \<open>Transitive Actions\<close>
+
+text \<open>Transitive actions have only one orbit\<close>
+
+lemma (in transitive_action) all_equivalent:
+ "\<lbrakk> x \<in> E; y \<in> E \<rbrakk> \<Longrightarrow> x .=\<^bsub>\<lparr>carrier = E, eq = \<lambda>x y. y \<in> orbit G \<phi> x\<rparr>\<^esub> y"
+proof -
+ assume "x \<in> E" "y \<in> E"
+ hence "\<exists> g \<in> carrier G. (\<phi> g) x = y"
+ using unique_orbit by blast
+ hence "y \<in> orbit G \<phi> x"
+ using orbit_def by fastforce
+ thus "x .=\<^bsub>\<lparr>carrier = E, eq = \<lambda>x y. y \<in> orbit G \<phi> x\<rparr>\<^esub> y" by simp
+qed
+
+proposition (in transitive_action) one_orbit:
+ assumes "E \<noteq> {}"
+ shows "card (orbits G E \<phi>) = 1"
+proof -
+ have "orbits G E \<phi> \<noteq> {}"
+ using assms orbits_coverture by auto
+ moreover have "\<And> orb1 orb2. \<lbrakk> orb1 \<in> (orbits G E \<phi>); orb2 \<in> (orbits G E \<phi>) \<rbrakk> \<Longrightarrow> orb1 = orb2"
+ proof -
+ fix orb1 orb2 assume orb1: "orb1 \<in> (orbits G E \<phi>)"
+ and orb2: "orb2 \<in> (orbits G E \<phi>)"
+ then obtain x y where x: "orb1 = orbit G \<phi> x" and x_E: "x \<in> E"
+ and y: "orb2 = orbit G \<phi> y" and y_E: "y \<in> E"
+ unfolding orbits_def by blast
+ hence "x \<in> orbit G \<phi> y" using all_equivalent by auto
+ hence "orb1 \<inter> orb2 \<noteq> {}" using x y x_E orbit_refl by auto
+ thus "orb1 = orb2" using disjoint_union[of orb1 orb2] orb1 orb2 by auto
+ qed
+ ultimately show "card (orbits G E \<phi>) = 1"
+ by (meson is_singletonI' is_singleton_altdef)
+qed
+
+
+
+subsection \<open>Stabilizers\<close>
+
+text \<open>We show that stabilizers are subgroups from the acting group\<close>
+
+lemma (in group_action) stabilizer_subset:
+ "stabilizer G \<phi> x \<subseteq> carrier G"
+ by (metis (no_types, lifting) mem_Collect_eq stabilizer_def subsetI)
+
+lemma (in group_action) stabilizer_m_closed:
+ assumes "x \<in> E" "g1 \<in> (stabilizer G \<phi> x)" "g2 \<in> (stabilizer G \<phi> x)"
+ shows "(g1 \<otimes> g2) \<in> (stabilizer G \<phi> x)"
+proof -
+ interpret group G
+ using group_hom group_hom.axioms(1) by auto
+
+ have "\<phi> g1 x = x"
+ using assms stabilizer_def by fastforce
+ moreover have "\<phi> g2 x = x"
+ using assms stabilizer_def by fastforce
+ moreover have g1: "g1 \<in> carrier G"
+ by (meson assms contra_subsetD stabilizer_subset)
+ moreover have g2: "g2 \<in> carrier G"
+ by (meson assms contra_subsetD stabilizer_subset)
+ ultimately have "\<phi> (g1 \<otimes> g2) x = x"
+ using composition_rule assms by simp
+
+ thus ?thesis
+ by (simp add: g1 g2 stabilizer_def)
+qed
+
+lemma (in group_action) stabilizer_one_closed:
+ assumes "x \<in> E"
+ shows "\<one> \<in> (stabilizer G \<phi> x)"
+proof -
+ have "\<phi> \<one> x = x"
+ by (metis assms id_eq_one restrict_apply')
+ thus ?thesis
+ using group_def group_hom group_hom.axioms(1) stabilizer_def by fastforce
+qed
+
+lemma (in group_action) stabilizer_m_inv_closed:
+ assumes "x \<in> E" "g \<in> (stabilizer G \<phi> x)"
+ shows "(inv g) \<in> (stabilizer G \<phi> x)"
+proof -
+ interpret group G
+ using group_hom group_hom.axioms(1) by auto
+
+ have "\<phi> g x = x"
+ using assms(2) stabilizer_def by fastforce
+ moreover have g: "g \<in> carrier G"
+ using assms(2) stabilizer_subset by blast
+ moreover have inv_g: "inv g \<in> carrier G"
+ by (simp add: g)
+ ultimately have "\<phi> (inv g) x = x"
+ using assms(1) orbit_sym_aux by blast
+
+ thus ?thesis by (simp add: inv_g stabilizer_def)
+qed
+
+theorem (in group_action) stabilizer_subgroup:
+ assumes "x \<in> E"
+ shows "subgroup (stabilizer G \<phi> x) G"
+ unfolding subgroup_def
+ using stabilizer_subset stabilizer_m_closed stabilizer_one_closed
+ stabilizer_m_inv_closed assms by simp
+
+
+
+subsection \<open>The Orbit-Stabilizer Theorem\<close>
+
+text \<open>In this subsection, we prove the Orbit-Stabilizer theorem.
+ Our approach is to show the existence of a bijection between
+ "rcosets (stabilizer G phi x)" and "orbit G phi x". Then we use
+ Lagrange's theorem to find the cardinal of the first set.\<close>
+
+subsubsection \<open>Rcosets - Supporting Lemmas\<close>
+
+corollary (in group_action) stab_rcosets_not_empty:
+ assumes "x \<in> E" "R \<in> rcosets (stabilizer G \<phi> x)"
+ shows "R \<noteq> {}"
+ using subgroup.rcosets_not_empty[OF stabilizer_subgroup[OF assms(1)] assms(2)] by simp
+
+corollary (in group_action) diff_stabilizes:
+ assumes "x \<in> E" "R \<in> rcosets (stabilizer G \<phi> x)"
+ shows "\<And>g1 g2. \<lbrakk> g1 \<in> R; g2 \<in> R \<rbrakk> \<Longrightarrow> g1 \<otimes> (inv g2) \<in> stabilizer G \<phi> x"
+ using group.diff_neutralizes[of G "stabilizer G \<phi> x" R] stabilizer_subgroup[OF assms(1)]
+ assms(2) group_hom group_hom.axioms(1) by blast
+
+
+subsubsection \<open>Bijection Between Rcosets and an Orbit - Definition and Supporting Lemmas\<close>
+
+(* This definition could be easier if lcosets were available, and it's indeed a considerable
+ modification at Coset theory, since we could derive it easily from the definition of rcosets
+ following the same idea we use here: f: rcosets \<rightarrow> lcosets, s.t. f R = (\<lambda>g. inv g) ` R
+ is a bijection. *)
+
+definition
+ orb_stab_fun :: "[_, ('a \<Rightarrow> 'b \<Rightarrow> 'b), 'a set, 'b] \<Rightarrow> 'b"
+ where "orb_stab_fun G \<phi> R x = (\<phi> (inv\<^bsub>G\<^esub> (SOME h. h \<in> R))) x"
+
+lemma (in group_action) orbit_stab_fun_is_well_defined0:
+ assumes "x \<in> E" "R \<in> rcosets (stabilizer G \<phi> x)"
+ shows "\<And>g1 g2. \<lbrakk> g1 \<in> R; g2 \<in> R \<rbrakk> \<Longrightarrow> (\<phi> (inv g1)) x = (\<phi> (inv g2)) x"
+proof -
+ fix g1 g2 assume g1: "g1 \<in> R" and g2: "g2 \<in> R"
+ have R_carr: "R \<subseteq> carrier G"
+ using subgroup.rcosets_carrier[OF stabilizer_subgroup[OF assms(1)]]
+ assms(2) group_hom group_hom.axioms(1) by auto
+ from R_carr have g1_carr: "g1 \<in> carrier G" using g1 by blast
+ from R_carr have g2_carr: "g2 \<in> carrier G" using g2 by blast
+
+ have "g1 \<otimes> (inv g2) \<in> stabilizer G \<phi> x"
+ using diff_stabilizes[of x R g1 g2] assms g1 g2 by blast
+ hence "\<phi> (g1 \<otimes> (inv g2)) x = x"
+ by (simp add: stabilizer_def)
+ hence "(\<phi> (inv g1)) x = (\<phi> (inv g1)) (\<phi> (g1 \<otimes> (inv g2)) x)" by simp
+ also have " ... = \<phi> ((inv g1) \<otimes> (g1 \<otimes> (inv g2))) x"
+ using group_def assms(1) composition_rule g1_carr g2_carr
+ group_hom group_hom.axioms(1) monoid.m_closed by fastforce
+ also have " ... = \<phi> (((inv g1) \<otimes> g1) \<otimes> (inv g2)) x"
+ using group_def g1_carr g2_carr group_hom group_hom.axioms(1) monoid.m_assoc by fastforce
+ finally show "(\<phi> (inv g1)) x = (\<phi> (inv g2)) x"
+ using group_def g1_carr g2_carr group.l_inv group_hom group_hom.axioms(1) by fastforce
+qed
+
+lemma (in group_action) orbit_stab_fun_is_well_defined1:
+ assumes "x \<in> E" "R \<in> rcosets (stabilizer G \<phi> x)"
+ shows "\<And>g. g \<in> R \<Longrightarrow> (\<phi> (inv (SOME h. h \<in> R))) x = (\<phi> (inv g)) x"
+ by (meson assms orbit_stab_fun_is_well_defined0 someI_ex)
+
+lemma (in group_action) orbit_stab_fun_is_inj:
+ assumes "x \<in> E"
+ and "R1 \<in> rcosets (stabilizer G \<phi> x)"
+ and "R2 \<in> rcosets (stabilizer G \<phi> x)"
+ and "\<phi> (inv (SOME h. h \<in> R1)) x = \<phi> (inv (SOME h. h \<in> R2)) x"
+ shows "R1 = R2"
+proof -
+ have "(\<exists>g1. g1 \<in> R1) \<and> (\<exists>g2. g2 \<in> R2)"
+ using assms(1-3) stab_rcosets_not_empty by auto
+ then obtain g1 g2 where g1: "g1 \<in> R1" and g2: "g2 \<in> R2" by blast
+ hence g12_carr: "g1 \<in> carrier G \<and> g2 \<in> carrier G"
+ using subgroup.rcosets_carrier assms(1-3) group_hom
+ group_hom.axioms(1) stabilizer_subgroup by blast
+
+ then obtain r1 r2 where r1: "r1 \<in> carrier G" "R1 = (stabilizer G \<phi> x) #> r1"
+ and r2: "r2 \<in> carrier G" "R2 = (stabilizer G \<phi> x) #> r2"
+ using assms(1-3) unfolding RCOSETS_def by blast
+ then obtain s1 s2 where s1: "s1 \<in> (stabilizer G \<phi> x)" "g1 = s1 \<otimes> r1"
+ and s2: "s2 \<in> (stabilizer G \<phi> x)" "g2 = s2 \<otimes> r2"
+ using g1 g2 unfolding r_coset_def by blast
+
+ have "\<phi> (inv g1) x = \<phi> (inv (SOME h. h \<in> R1)) x"
+ using orbit_stab_fun_is_well_defined1[OF assms(1) assms(2) g1] by simp
+ also have " ... = \<phi> (inv (SOME h. h \<in> R2)) x"
+ using assms(4) by simp
+ finally have "\<phi> (inv g1) x = \<phi> (inv g2) x"
+ using orbit_stab_fun_is_well_defined1[OF assms(1) assms(3) g2] by simp
+
+ hence "\<phi> g2 (\<phi> (inv g1) x) = \<phi> g2 (\<phi> (inv g2) x)" by simp
+ also have " ... = \<phi> (g2 \<otimes> (inv g2)) x"
+ using assms(1) composition_rule g12_carr group_hom group_hom.axioms(1) by fastforce
+ finally have "\<phi> g2 (\<phi> (inv g1) x) = x"
+ using g12_carr assms(1) group.r_inv group_hom group_hom.axioms(1)
+ id_eq_one restrict_apply by metis
+ hence "\<phi> (g2 \<otimes> (inv g1)) x = x"
+ using assms(1) composition_rule g12_carr group_hom group_hom.axioms(1) by fastforce
+ hence "g2 \<otimes> (inv g1) \<in> (stabilizer G \<phi> x)"
+ using g12_carr group.subgroup_self group_hom group_hom.axioms(1)
+ mem_Collect_eq stabilizer_def subgroup_def by (metis (mono_tags, lifting))
+ then obtain s where s: "s \<in> (stabilizer G \<phi> x)" "s = g2 \<otimes> (inv g1)" by blast
+
+ let ?h = "s \<otimes> g1"
+ have "?h = s \<otimes> (s1 \<otimes> r1)" by (simp add: s1)
+ hence "?h = (s \<otimes> s1) \<otimes> r1"
+ using stabilizer_subgroup[OF assms(1)] group_def group_hom
+ group_hom.axioms(1) monoid.m_assoc r1 s s1 subgroup.mem_carrier by fastforce
+ hence inR1: "?h \<in> (stabilizer G \<phi> x) #> r1" unfolding r_coset_def
+ using stabilizer_subgroup[OF assms(1)] assms(1) s s1 stabilizer_m_closed by auto
+
+ have "?h = g2" using s stabilizer_subgroup[OF assms(1)] g12_carr group.inv_solve_right
+ group_hom group_hom.axioms(1) subgroup.mem_carrier by metis
+ hence inR2: "?h \<in> (stabilizer G \<phi> x) #> r2"
+ using g2 r2 by blast
+
+ have "R1 \<inter> R2 \<noteq> {}" using inR1 inR2 r1 r2 by blast
+ thus ?thesis using stabilizer_subgroup group.rcos_disjoint[of G "stabilizer G \<phi> x" R1 R2]
+ assms group_hom group_hom.axioms(1) by blast
+qed
+
+lemma (in group_action) orbit_stab_fun_is_surj:
+ assumes "x \<in> E" "y \<in> orbit G \<phi> x"
+ shows "\<exists>R \<in> rcosets (stabilizer G \<phi> x). \<phi> (inv (SOME h. h \<in> R)) x = y"
+proof -
+ have "\<exists>g \<in> carrier G. (\<phi> g) x = y"
+ using assms(2) unfolding orbit_def by blast
+ then obtain g where g: "g \<in> carrier G \<and> (\<phi> g) x = y" by blast
+
+ let ?R = "(stabilizer G \<phi> x) #> (inv g)"
+ have R: "?R \<in> rcosets (stabilizer G \<phi> x)"
+ unfolding RCOSETS_def using g group_hom group_hom.axioms(1) by fastforce
+ moreover have "\<one> \<otimes> (inv g) \<in> ?R"
+ unfolding r_coset_def using assms(1) stabilizer_one_closed by auto
+ ultimately have "\<phi> (inv (SOME h. h \<in> ?R)) x = \<phi> (inv (\<one> \<otimes> (inv g))) x"
+ using orbit_stab_fun_is_well_defined1[OF assms(1)] by simp
+ also have " ... = (\<phi> g) x"
+ using group_def g group_hom group_hom.axioms(1) monoid.l_one by fastforce
+ finally have "\<phi> (inv (SOME h. h \<in> ?R)) x = y"
+ using g by simp
+ thus ?thesis using R by blast
+qed
+
+proposition (in group_action) orbit_stab_fun_is_bij:
+ assumes "x \<in> E"
+ shows "bij_betw (\<lambda>R. (\<phi> (inv (SOME h. h \<in> R))) x) (rcosets (stabilizer G \<phi> x)) (orbit G \<phi> x)"
+ unfolding bij_betw_def
+proof
+ show "inj_on (\<lambda>R. \<phi> (inv (SOME h. h \<in> R)) x) (rcosets stabilizer G \<phi> x)"
+ using orbit_stab_fun_is_inj[OF assms(1)] by (simp add: inj_on_def)
+next
+ have "\<And>R. R \<in> (rcosets stabilizer G \<phi> x) \<Longrightarrow> \<phi> (inv (SOME h. h \<in> R)) x \<in> orbit G \<phi> x "
+ proof -
+ fix R assume R: "R \<in> (rcosets stabilizer G \<phi> x)"
+ then obtain g where g: "g \<in> R"
+ using assms stab_rcosets_not_empty by auto
+ hence "\<phi> (inv (SOME h. h \<in> R)) x = \<phi> (inv g) x"
+ using R assms orbit_stab_fun_is_well_defined1 by blast
+ thus "\<phi> (inv (SOME h. h \<in> R)) x \<in> orbit G \<phi> x" unfolding orbit_def
+ using subgroup.rcosets_carrier group_hom group_hom.axioms(1)
+ g R assms stabilizer_subgroup by fastforce
+ qed
+ moreover have "orbit G \<phi> x \<subseteq> (\<lambda>R. \<phi> (inv (SOME h. h \<in> R)) x) ` (rcosets stabilizer G \<phi> x)"
+ using assms orbit_stab_fun_is_surj by fastforce
+ ultimately show "(\<lambda>R. \<phi> (inv (SOME h. h \<in> R)) x) ` (rcosets stabilizer G \<phi> x) = orbit G \<phi> x "
+ using assms set_eq_subset by blast
+qed
+
+
+subsubsection \<open>The Theorem\<close>
+
+theorem (in group_action) orbit_stabilizer_theorem:
+ assumes "x \<in> E"
+ shows "card (orbit G \<phi> x) * card (stabilizer G \<phi> x) = order G"
+proof -
+ have "card (rcosets stabilizer G \<phi> x) = card (orbit G \<phi> x)"
+ using orbit_stab_fun_is_bij[OF assms(1)] bij_betw_same_card by blast
+ moreover have "card (rcosets stabilizer G \<phi> x) * card (stabilizer G \<phi> x) = order G"
+ using stabilizer_subgroup assms group.lagrange group_hom group_hom.axioms(1) by blast
+ ultimately show ?thesis by auto
+qed
+
+
+
+subsection \<open>The Burnside's Lemma\<close>
+
+subsubsection \<open>Sums and Cardinals\<close>
+
+lemma card_as_sums:
+ assumes "A = {x \<in> B. P x}" "finite B"
+ shows "card A = (\<Sum>x\<in>B. (if P x then 1 else 0))"
+proof -
+ have "A \<subseteq> B" using assms(1) by blast
+ have "card A = (\<Sum>x\<in>A. 1)" by simp
+ also have " ... = (\<Sum>x\<in>A. (if P x then 1 else 0))"
+ by (simp add: assms(1))
+ also have " ... = (\<Sum>x\<in>A. (if P x then 1 else 0)) + (\<Sum>x\<in>(B - A). (if P x then 1 else 0))"
+ using assms(1) by auto
+ finally show "card A = (\<Sum>x\<in>B. (if P x then 1 else 0))"
+ using \<open>A \<subseteq> B\<close> add.commute assms(2) sum.subset_diff by metis
+qed
+
+lemma sum_invertion:
+ "\<lbrakk> finite A; finite B \<rbrakk> \<Longrightarrow> (\<Sum>x\<in>A. \<Sum>y\<in>B. f x y) = (\<Sum>y\<in>B. \<Sum>x\<in>A. f x y)"
+proof (induct set: finite)
+ case empty thus ?case by simp
+next
+ case (insert x A')
+ have "(\<Sum>x\<in>insert x A'. \<Sum>y\<in>B. f x y) = (\<Sum>y\<in>B. f x y) + (\<Sum>x\<in>A'. \<Sum>y\<in>B. f x y)"
+ by (simp add: insert.hyps)
+ also have " ... = (\<Sum>y\<in>B. f x y) + (\<Sum>y\<in>B. \<Sum>x\<in>A'. f x y)"
+ using insert.hyps by (simp add: insert.prems)
+ also have " ... = (\<Sum>y\<in>B. (f x y) + (\<Sum>x\<in>A'. f x y))"
+ by (simp add: sum.distrib)
+ finally have "(\<Sum>x\<in>insert x A'. \<Sum>y\<in>B. f x y) = (\<Sum>y\<in>B. \<Sum>x\<in>insert x A'. f x y)"
+ using sum.swap by blast
+ thus ?case by simp
+qed
+
+lemma (in group_action) card_stablizer_sum:
+ assumes "finite (carrier G)" "orb \<in> (orbits G E \<phi>)"
+ shows "(\<Sum>x \<in> orb. card (stabilizer G \<phi> x)) = order G"
+proof -
+ obtain x where x:"x \<in> E" and orb:"orb = orbit G \<phi> x"
+ using assms(2) unfolding orbits_def by blast
+ have "\<And>y. y \<in> orb \<Longrightarrow> card (stabilizer G \<phi> x) = card (stabilizer G \<phi> y)"
+ proof -
+ fix y assume "y \<in> orb"
+ hence y: "y \<in> E \<and> y \<in> orbit G \<phi> x"
+ using x orb assms(2) orbits_coverture by auto
+ hence same_orbit: "(orbit G \<phi> x) = (orbit G \<phi> y)"
+ using disjoint_union[of "orbit G \<phi> x" "orbit G \<phi> y"] orbit_refl x
+ unfolding orbits_def by auto
+ have "card (orbit G \<phi> x) * card (stabilizer G \<phi> x) =
+ card (orbit G \<phi> y) * card (stabilizer G \<phi> y)"
+ using y assms(1) x orbit_stabilizer_theorem by simp
+ hence "card (orbit G \<phi> x) * card (stabilizer G \<phi> x) =
+ card (orbit G \<phi> x) * card (stabilizer G \<phi> y)" using same_orbit by simp
+ moreover have "orbit G \<phi> x \<noteq> {} \<and> finite (orbit G \<phi> x)"
+ using y orbit_def[of G \<phi> x] assms(1) by auto
+ hence "card (orbit G \<phi> x) > 0"
+ by (simp add: card_gt_0_iff)
+ ultimately show "card (stabilizer G \<phi> x) = card (stabilizer G \<phi> y)" by auto
+ qed
+ hence "(\<Sum>x \<in> orb. card (stabilizer G \<phi> x)) = (\<Sum>y \<in> orb. card (stabilizer G \<phi> x))" by auto
+ also have " ... = card (stabilizer G \<phi> x) * (\<Sum>y \<in> orb. 1)" by simp
+ also have " ... = card (stabilizer G \<phi> x) * card (orbit G \<phi> x)"
+ using orb by auto
+ finally show "(\<Sum>x \<in> orb. card (stabilizer G \<phi> x)) = order G"
+ by (metis mult.commute orbit_stabilizer_theorem x)
+qed
+
+
+subsubsection \<open>The Lemma\<close>
+
+theorem (in group_action) burnside:
+ assumes "finite (carrier G)" "finite E"
+ shows "card (orbits G E \<phi>) * order G = (\<Sum>g \<in> carrier G. card(invariants E \<phi> g))"
+proof -
+ have "(\<Sum>g \<in> carrier G. card(invariants E \<phi> g)) =
+ (\<Sum>g \<in> carrier G. \<Sum>x \<in> E. (if (\<phi> g) x = x then 1 else 0))"
+ by (simp add: assms(2) card_as_sums invariants_def)
+ also have " ... = (\<Sum>x \<in> E. \<Sum>g \<in> carrier G. (if (\<phi> g) x = x then 1 else 0))"
+ using sum_invertion[where ?f = "\<lambda> g x. (if (\<phi> g) x = x then 1 else 0)"] assms by auto
+ also have " ... = (\<Sum>x \<in> E. card (stabilizer G \<phi> x))"
+ by (simp add: assms(1) card_as_sums stabilizer_def)
+ also have " ... = (\<Sum>orbit \<in> (orbits G E \<phi>). \<Sum>x \<in> orbit. card (stabilizer G \<phi> x))"
+ using disjoint_sum orbits_coverture assms(2) by metis
+ also have " ... = (\<Sum>orbit \<in> (orbits G E \<phi>). order G)"
+ by (simp add: assms(1) card_stablizer_sum)
+ finally have "(\<Sum>g \<in> carrier G. card(invariants E \<phi> g)) = card (orbits G E \<phi>) * order G" by simp
+ thus ?thesis by simp
+qed
+
+
+
+subsection \<open>Action by Conjugation\<close>
+
+
+subsubsection \<open>Action Over Itself\<close>
+
+text \<open>A Group Acts by Conjugation Over Itself\<close>
+
+lemma (in group) conjugation_is_inj:
+ assumes "g \<in> carrier G" "h1 \<in> carrier G" "h2 \<in> carrier G"
+ and "g \<otimes> h1 \<otimes> (inv g) = g \<otimes> h2 \<otimes> (inv g)"
+ shows "h1 = h2"
+ using assms by auto
+
+lemma (in group) conjugation_is_surj:
+ assumes "g \<in> carrier G" "h \<in> carrier G"
+ shows "g \<otimes> ((inv g) \<otimes> h \<otimes> g) \<otimes> (inv g) = h"
+ using assms m_assoc inv_closed inv_inv m_closed monoid_axioms r_inv r_one
+ by metis
+
+lemma (in group) conjugation_is_bij:
+ assumes "g \<in> carrier G"
+ shows "bij_betw (\<lambda>h \<in> carrier G. g \<otimes> h \<otimes> (inv g)) (carrier G) (carrier G)"
+ (is "bij_betw ?\<phi> (carrier G) (carrier G)")
+ unfolding bij_betw_def
+proof
+ show "inj_on ?\<phi> (carrier G)"
+ using conjugation_is_inj by (simp add: assms inj_on_def)
+next
+ have S: "\<And> h. h \<in> carrier G \<Longrightarrow> (inv g) \<otimes> h \<otimes> g \<in> carrier G"
+ using assms by blast
+ have "\<And> h. h \<in> carrier G \<Longrightarrow> ?\<phi> ((inv g) \<otimes> h \<otimes> g) = h"
+ using assms by (simp add: conjugation_is_surj)
+ hence "carrier G \<subseteq> ?\<phi> ` carrier G"
+ using S image_iff by fastforce
+ moreover have "\<And> h. h \<in> carrier G \<Longrightarrow> ?\<phi> h \<in> carrier G"
+ using assms by simp
+ hence "?\<phi> ` carrier G \<subseteq> carrier G" by blast
+ ultimately show "?\<phi> ` carrier G = carrier G" by blast
+qed
+
+lemma(in group) conjugation_is_hom:
+ "(\<lambda>g. \<lambda>h \<in> carrier G. g \<otimes> h \<otimes> inv g) \<in> hom G (BijGroup (carrier G))"
+ unfolding hom_def
+proof -
+ let ?\<psi> = "\<lambda>g. \<lambda>h. g \<otimes> h \<otimes> inv g"
+ let ?\<phi> = "\<lambda>g. restrict (?\<psi> g) (carrier G)"
+
+ (* First, we prove that ?\<phi>: G \<rightarrow> Bij(G) is well defined *)
+ have Step0: "\<And> g. g \<in> carrier G \<Longrightarrow> (?\<phi> g) \<in> Bij (carrier G)"
+ using Bij_def conjugation_is_bij by fastforce
+ hence Step1: "?\<phi>: carrier G \<rightarrow> carrier (BijGroup (carrier G))"
+ unfolding BijGroup_def by simp
+
+ (* Second, we prove that ?\<phi> is a homomorphism *)
+ have "\<And> g1 g2. \<lbrakk> g1 \<in> carrier G; g2 \<in> carrier G \<rbrakk> \<Longrightarrow>
+ (\<And> h. h \<in> carrier G \<Longrightarrow> ?\<psi> (g1 \<otimes> g2) h = (?\<phi> g1) ((?\<phi> g2) h))"
+ proof -
+ fix g1 g2 h assume g1: "g1 \<in> carrier G" and g2: "g2 \<in> carrier G" and h: "h \<in> carrier G"
+ have "inv (g1 \<otimes> g2) = (inv g2) \<otimes> (inv g1)"
+ using g1 g2 by (simp add: inv_mult_group)
+ thus "?\<psi> (g1 \<otimes> g2) h = (?\<phi> g1) ((?\<phi> g2) h)"
+ by (simp add: g1 g2 h m_assoc)
+ qed
+ hence "\<And> g1 g2. \<lbrakk> g1 \<in> carrier G; g2 \<in> carrier G \<rbrakk> \<Longrightarrow>
+ (\<lambda> h \<in> carrier G. ?\<psi> (g1 \<otimes> g2) h) = (\<lambda> h \<in> carrier G. (?\<phi> g1) ((?\<phi> g2) h))" by auto
+ hence Step2: "\<And> g1 g2. \<lbrakk> g1 \<in> carrier G; g2 \<in> carrier G \<rbrakk> \<Longrightarrow>
+ ?\<phi> (g1 \<otimes> g2) = (?\<phi> g1) \<otimes>\<^bsub>BijGroup (carrier G)\<^esub> (?\<phi> g2)"
+ unfolding BijGroup_def by (simp add: Step0 compose_def)
+
+ (* Finally, we combine both results to prove the lemma *)
+ thus "?\<phi> \<in> {h: carrier G \<rightarrow> carrier (BijGroup (carrier G)).
+ (\<forall>x \<in> carrier G. \<forall>y \<in> carrier G. h (x \<otimes> y) = h x \<otimes>\<^bsub>BijGroup (carrier G)\<^esub> h y)}"
+ using Step1 Step2 by auto
+qed
+
+theorem (in group) action_by_conjugation:
+ "group_action G (carrier G) (\<lambda>g. (\<lambda>h \<in> carrier G. g \<otimes> h \<otimes> (inv g)))"
+ unfolding group_action_def group_hom_def using conjugation_is_hom
+ by (simp add: group_BijGroup group_hom_axioms.intro is_group)
+
+
+subsubsection \<open>Action Over The Set of Subgroups\<close>
+
+text \<open>A Group Acts by Conjugation Over The Set of Subgroups\<close>
+
+lemma (in group) subgroup_conjugation_is_inj_aux:
+ assumes "g \<in> carrier G" "H1 \<subseteq> carrier G" "H2 \<subseteq> carrier G"
+ and "g <# H1 #> (inv g) = g <# H2 #> (inv g)"
+ shows "H1 \<subseteq> H2"
+proof
+ fix h1 assume h1: "h1 \<in> H1"
+ hence "g \<otimes> h1 \<otimes> (inv g) \<in> g <# H1 #> (inv g)"
+ unfolding l_coset_def r_coset_def using assms by blast
+ hence "g \<otimes> h1 \<otimes> (inv g) \<in> g <# H2 #> (inv g)"
+ using assms by auto
+ hence "\<exists>h2 \<in> H2. g \<otimes> h1 \<otimes> (inv g) = g \<otimes> h2 \<otimes> (inv g)"
+ unfolding l_coset_def r_coset_def by blast
+ then obtain h2 where "h2 \<in> H2 \<and> g \<otimes> h1 \<otimes> (inv g) = g \<otimes> h2 \<otimes> (inv g)" by blast
+ thus "h1 \<in> H2"
+ using assms conjugation_is_inj h1 by blast
+qed
+
+lemma (in group) subgroup_conjugation_is_inj:
+ assumes "g \<in> carrier G" "H1 \<subseteq> carrier G" "H2 \<subseteq> carrier G"
+ and "g <# H1 #> (inv g) = g <# H2 #> (inv g)"
+ shows "H1 = H2"
+ using subgroup_conjugation_is_inj_aux assms set_eq_subset by metis
+
+lemma (in group) subgroup_conjugation_is_surj0:
+ assumes "g \<in> carrier G" "H \<subseteq> carrier G"
+ shows "g <# ((inv g) <# H #> g) #> (inv g) = H"
+ using coset_assoc assms coset_mult_assoc l_coset_subset_G lcos_m_assoc
+ by (simp add: lcos_mult_one)
+
+lemma (in group) subgroup_conjugation_is_surj1:
+ assumes "g \<in> carrier G" "subgroup H G"
+ shows "subgroup ((inv g) <# H #> g) G"
+proof
+ show "\<one> \<in> inv g <# H #> g"
+ proof -
+ have "\<one> \<in> H" by (simp add: assms(2) subgroup.one_closed)
+ hence "inv g \<otimes> \<one> \<otimes> g \<in> inv g <# H #> g"
+ unfolding l_coset_def r_coset_def by blast
+ thus "\<one> \<in> inv g <# H #> g" using assms by simp
+ qed
+next
+ show "inv g <# H #> g \<subseteq> carrier G"
+ proof
+ fix x assume "x \<in> inv g <# H #> g"
+ hence "\<exists>h \<in> H. x = (inv g) \<otimes> h \<otimes> g"
+ unfolding r_coset_def l_coset_def by blast
+ hence "\<exists>h \<in> (carrier G). x = (inv g) \<otimes> h \<otimes> g"
+ by (meson assms subgroup.mem_carrier)
+ thus "x \<in> carrier G" using assms by blast
+ qed
+next
+ show " \<And> x y. \<lbrakk> x \<in> inv g <# H #> g; y \<in> inv g <# H #> g \<rbrakk> \<Longrightarrow> x \<otimes> y \<in> inv g <# H #> g"
+ proof -
+ fix x y assume "x \<in> inv g <# H #> g" "y \<in> inv g <# H #> g"
+ hence "\<exists> h1 \<in> H. \<exists> h2 \<in> H. x = (inv g) \<otimes> h1 \<otimes> g \<and> y = (inv g) \<otimes> h2 \<otimes> g"
+ unfolding l_coset_def r_coset_def by blast
+ hence "\<exists> h1 \<in> H. \<exists> h2 \<in> H. x \<otimes> y = ((inv g) \<otimes> h1 \<otimes> g) \<otimes> ((inv g) \<otimes> h2 \<otimes> g)" by blast
+ hence "\<exists> h1 \<in> H. \<exists> h2 \<in> H. x \<otimes> y = ((inv g) \<otimes> (h1 \<otimes> h2) \<otimes> g)"
+ using assms is_group inv_closed l_one m_assoc m_closed
+ monoid_axioms r_inv subgroup.mem_carrier by smt
+ hence "\<exists> h \<in> H. x \<otimes> y = (inv g) \<otimes> h \<otimes> g"
+ by (meson assms(2) subgroup_def)
+ thus "x \<otimes> y \<in> inv g <# H #> g"
+ unfolding l_coset_def r_coset_def by blast
+ qed
+next
+ show "\<And>x. x \<in> inv g <# H #> g \<Longrightarrow> inv x \<in> inv g <# H #> g"
+ proof -
+ fix x assume "x \<in> inv g <# H #> g"
+ hence "\<exists>h \<in> H. x = (inv g) \<otimes> h \<otimes> g"
+ unfolding r_coset_def l_coset_def by blast
+ then obtain h where h: "h \<in> H \<and> x = (inv g) \<otimes> h \<otimes> g" by blast
+ hence "x \<otimes> (inv g) \<otimes> (inv h) \<otimes> g = \<one>"
+ using assms inv_closed m_assoc m_closed monoid_axioms
+ r_inv r_one subgroup.mem_carrier by smt
+ hence "inv x = (inv g) \<otimes> (inv h) \<otimes> g"
+ using assms h inv_closed inv_inv inv_mult_group m_assoc
+ m_closed monoid_axioms subgroup.mem_carrier by smt
+ moreover have "inv h \<in> H"
+ by (simp add: assms h subgroup.m_inv_closed)
+ ultimately show "inv x \<in> inv g <# H #> g" unfolding r_coset_def l_coset_def by blast
+ qed
+qed
+
+lemma (in group) subgroup_conjugation_is_surj2:
+ assumes "g \<in> carrier G" "subgroup H G"
+ shows "subgroup (g <# H #> (inv g)) G"
+ using subgroup_conjugation_is_surj1 by (metis assms inv_closed inv_inv)
+
+lemma (in group) subgroup_conjugation_is_bij:
+ assumes "g \<in> carrier G"
+ shows "bij_betw (\<lambda>H \<in> {H. subgroup H G}. g <# H #> (inv g)) {H. subgroup H G} {H. subgroup H G}"
+ (is "bij_betw ?\<phi> {H. subgroup H G} {H. subgroup H G}")
+ unfolding bij_betw_def
+proof
+ show "inj_on ?\<phi> {H. subgroup H G}"
+ using subgroup_conjugation_is_inj assms inj_on_def subgroup.subset
+ by (metis (mono_tags, lifting) inj_on_restrict_eq mem_Collect_eq)
+next
+ have "\<And>H. H \<in> {H. subgroup H G} \<Longrightarrow> ?\<phi> ((inv g) <# H #> g) = H"
+ by (simp add: assms subgroup.subset subgroup_conjugation_is_surj0
+ subgroup_conjugation_is_surj1 is_group)
+ hence "\<And>H. H \<in> {H. subgroup H G} \<Longrightarrow> \<exists>H' \<in> {H. subgroup H G}. ?\<phi> H' = H"
+ using assms subgroup_conjugation_is_surj1 by fastforce
+ thus "?\<phi> ` {H. subgroup H G} = {H. subgroup H G}"
+ using subgroup_conjugation_is_surj2 assms by auto
+qed
+
+lemma (in group) subgroup_conjugation_is_hom:
+ "(\<lambda>g. \<lambda>H \<in> {H. subgroup H G}. g <# H #> (inv g)) \<in> hom G (BijGroup {H. subgroup H G})"
+ unfolding hom_def
+proof -
+ (* We follow the exact same structure of conjugation_is_hom's proof *)
+ let ?\<psi> = "\<lambda>g. \<lambda>H. g <# H #> (inv g)"
+ let ?\<phi> = "\<lambda>g. restrict (?\<psi> g) {H. subgroup H G}"
+
+ have Step0: "\<And> g. g \<in> carrier G \<Longrightarrow> (?\<phi> g) \<in> Bij {H. subgroup H G}"
+ using Bij_def subgroup_conjugation_is_bij by fastforce
+ hence Step1: "?\<phi>: carrier G \<rightarrow> carrier (BijGroup {H. subgroup H G})"
+ unfolding BijGroup_def by simp
+
+ have "\<And> g1 g2. \<lbrakk> g1 \<in> carrier G; g2 \<in> carrier G \<rbrakk> \<Longrightarrow>
+ (\<And> H. H \<in> {H. subgroup H G} \<Longrightarrow> ?\<psi> (g1 \<otimes> g2) H = (?\<phi> g1) ((?\<phi> g2) H))"
+ proof -
+ fix g1 g2 H assume g1: "g1 \<in> carrier G" and g2: "g2 \<in> carrier G" and H': "H \<in> {H. subgroup H G}"
+ hence H: "subgroup H G" by simp
+ have "(?\<phi> g1) ((?\<phi> g2) H) = g1 <# (g2 <# H #> (inv g2)) #> (inv g1)"
+ by (simp add: H g2 subgroup_conjugation_is_surj2)
+ also have " ... = g1 <# (g2 <# H) #> ((inv g2) \<otimes> (inv g1))"
+ by (simp add: H coset_mult_assoc g1 g2 group.coset_assoc
+ is_group l_coset_subset_G subgroup.subset)
+ also have " ... = g1 <# (g2 <# H) #> inv (g1 \<otimes> g2)"
+ using g1 g2 by (simp add: inv_mult_group)
+ finally have "(?\<phi> g1) ((?\<phi> g2) H) = ?\<psi> (g1 \<otimes> g2) H"
+ by (simp add: H g1 g2 lcos_m_assoc subgroup.subset)
+ thus "?\<psi> (g1 \<otimes> g2) H = (?\<phi> g1) ((?\<phi> g2) H)" by auto
+ qed
+ hence "\<And> g1 g2. \<lbrakk> g1 \<in> carrier G; g2 \<in> carrier G \<rbrakk> \<Longrightarrow>
+ (\<lambda>H \<in> {H. subgroup H G}. ?\<psi> (g1 \<otimes> g2) H) = (\<lambda>H \<in> {H. subgroup H G}. (?\<phi> g1) ((?\<phi> g2) H))"
+ by (meson restrict_ext)
+ hence Step2: "\<And> g1 g2. \<lbrakk> g1 \<in> carrier G; g2 \<in> carrier G \<rbrakk> \<Longrightarrow>
+ ?\<phi> (g1 \<otimes> g2) = (?\<phi> g1) \<otimes>\<^bsub>BijGroup {H. subgroup H G}\<^esub> (?\<phi> g2)"
+ unfolding BijGroup_def by (simp add: Step0 compose_def)
+
+ show "?\<phi> \<in> {h: carrier G \<rightarrow> carrier (BijGroup {H. subgroup H G}).
+ \<forall>x\<in>carrier G. \<forall>y\<in>carrier G. h (x \<otimes> y) = h x \<otimes>\<^bsub>BijGroup {H. subgroup H G}\<^esub> h y}"
+ using Step1 Step2 by auto
+qed
+
+theorem (in group) action_by_conjugation_on_subgroups_set:
+ "group_action G {H. subgroup H G} (\<lambda>g. \<lambda>H \<in> {H. subgroup H G}. g <# H #> (inv g))"
+ unfolding group_action_def group_hom_def using subgroup_conjugation_is_hom
+ by (simp add: group_BijGroup group_hom_axioms.intro is_group)
+
+
+subsubsection \<open>Action Over The Power Set\<close>
+
+text \<open>A Group Acts by Conjugation Over The Power Set\<close>
+
+lemma (in group) subset_conjugation_is_bij:
+ assumes "g \<in> carrier G"
+ shows "bij_betw (\<lambda>H \<in> {H. H \<subseteq> carrier G}. g <# H #> (inv g)) {H. H \<subseteq> carrier G} {H. H \<subseteq> carrier G}"
+ (is "bij_betw ?\<phi> {H. H \<subseteq> carrier G} {H. H \<subseteq> carrier G}")
+ unfolding bij_betw_def
+proof
+ show "inj_on ?\<phi> {H. H \<subseteq> carrier G}"
+ using subgroup_conjugation_is_inj assms inj_on_def
+ by (metis (mono_tags, lifting) inj_on_restrict_eq mem_Collect_eq)
+next
+ have "\<And>H. H \<in> {H. H \<subseteq> carrier G} \<Longrightarrow> ?\<phi> ((inv g) <# H #> g) = H"
+ by (simp add: assms l_coset_subset_G r_coset_subset_G subgroup_conjugation_is_surj0)
+ hence "\<And>H. H \<in> {H. H \<subseteq> carrier G} \<Longrightarrow> \<exists>H' \<in> {H. H \<subseteq> carrier G}. ?\<phi> H' = H"
+ by (metis assms l_coset_subset_G mem_Collect_eq r_coset_subset_G subgroup_def subgroup_self)
+ hence "{H. H \<subseteq> carrier G} \<subseteq> ?\<phi> ` {H. H \<subseteq> carrier G}" by blast
+ moreover have "?\<phi> ` {H. H \<subseteq> carrier G} \<subseteq> {H. H \<subseteq> carrier G}"
+ by (smt assms image_subsetI inv_closed l_coset_subset_G
+ mem_Collect_eq r_coset_subset_G restrict_apply')
+ ultimately show "?\<phi> ` {H. H \<subseteq> carrier G} = {H. H \<subseteq> carrier G}" by simp
+qed
+
+lemma (in group) subset_conjugation_is_hom:
+ "(\<lambda>g. \<lambda>H \<in> {H. H \<subseteq> carrier G}. g <# H #> (inv g)) \<in> hom G (BijGroup {H. H \<subseteq> carrier G})"
+ unfolding hom_def
+proof -
+ (* We follow the exact same structure of conjugation_is_hom's proof *)
+ let ?\<psi> = "\<lambda>g. \<lambda>H. g <# H #> (inv g)"
+ let ?\<phi> = "\<lambda>g. restrict (?\<psi> g) {H. H \<subseteq> carrier G}"
+
+ have Step0: "\<And> g. g \<in> carrier G \<Longrightarrow> (?\<phi> g) \<in> Bij {H. H \<subseteq> carrier G}"
+ using Bij_def subset_conjugation_is_bij by fastforce
+ hence Step1: "?\<phi>: carrier G \<rightarrow> carrier (BijGroup {H. H \<subseteq> carrier G})"
+ unfolding BijGroup_def by simp
+
+ have "\<And> g1 g2. \<lbrakk> g1 \<in> carrier G; g2 \<in> carrier G \<rbrakk> \<Longrightarrow>
+ (\<And> H. H \<in> {H. H \<subseteq> carrier G} \<Longrightarrow> ?\<psi> (g1 \<otimes> g2) H = (?\<phi> g1) ((?\<phi> g2) H))"
+ proof -
+ fix g1 g2 H assume g1: "g1 \<in> carrier G" and g2: "g2 \<in> carrier G" and H: "H \<in> {H. H \<subseteq> carrier G}"
+ hence "(?\<phi> g1) ((?\<phi> g2) H) = g1 <# (g2 <# H #> (inv g2)) #> (inv g1)"
+ using l_coset_subset_G r_coset_subset_G by auto
+ also have " ... = g1 <# (g2 <# H) #> ((inv g2) \<otimes> (inv g1))"
+ using H coset_assoc coset_mult_assoc g1 g2 l_coset_subset_G by auto
+ also have " ... = g1 <# (g2 <# H) #> inv (g1 \<otimes> g2)"
+ using g1 g2 by (simp add: inv_mult_group)
+ finally have "(?\<phi> g1) ((?\<phi> g2) H) = ?\<psi> (g1 \<otimes> g2) H"
+ using H g1 g2 lcos_m_assoc by force
+ thus "?\<psi> (g1 \<otimes> g2) H = (?\<phi> g1) ((?\<phi> g2) H)" by auto
+ qed
+ hence "\<And> g1 g2. \<lbrakk> g1 \<in> carrier G; g2 \<in> carrier G \<rbrakk> \<Longrightarrow>
+ (\<lambda>H \<in> {H. H \<subseteq> carrier G}. ?\<psi> (g1 \<otimes> g2) H) = (\<lambda>H \<in> {H. H \<subseteq> carrier G}. (?\<phi> g1) ((?\<phi> g2) H))"
+ by (meson restrict_ext)
+ hence Step2: "\<And> g1 g2. \<lbrakk> g1 \<in> carrier G; g2 \<in> carrier G \<rbrakk> \<Longrightarrow>
+ ?\<phi> (g1 \<otimes> g2) = (?\<phi> g1) \<otimes>\<^bsub>BijGroup {H. H \<subseteq> carrier G}\<^esub> (?\<phi> g2)"
+ unfolding BijGroup_def by (simp add: Step0 compose_def)
+
+ show "?\<phi> \<in> {h: carrier G \<rightarrow> carrier (BijGroup {H. H \<subseteq> carrier G}).
+ \<forall>x\<in>carrier G. \<forall>y\<in>carrier G. h (x \<otimes> y) = h x \<otimes>\<^bsub>BijGroup {H. H \<subseteq> carrier G}\<^esub> h y}"
+ using Step1 Step2 by auto
+qed
+
+theorem (in group) action_by_conjugation_on_power_set:
+ "group_action G {H. H \<subseteq> carrier G} (\<lambda>g. \<lambda>H \<in> {H. H \<subseteq> carrier G}. g <# H #> (inv g))"
+ unfolding group_action_def group_hom_def using subset_conjugation_is_hom
+ by (simp add: group_BijGroup group_hom_axioms.intro is_group)
+
+corollary (in group) normalizer_imp_subgroup:
+ assumes "H \<subseteq> carrier G"
+ shows "subgroup (normalizer G H) G"
+ unfolding normalizer_def
+ using group_action.stabilizer_subgroup[OF action_by_conjugation_on_power_set] assms by auto
+
+
+subsection \<open>Subgroup of an Acting Group\<close>
+
+text \<open>A Subgroup of an Acting Group Induces an Action\<close>
+
+lemma (in group_action) induced_homomorphism:
+ assumes "subgroup H G"
+ shows "\<phi> \<in> hom (G \<lparr>carrier := H\<rparr>) (BijGroup E)"
+ unfolding hom_def apply simp
+proof -
+ have S0: "H \<subseteq> carrier G" by (meson assms subgroup_def)
+ hence "\<phi>: H \<rightarrow> carrier (BijGroup E)"
+ by (simp add: BijGroup_def bij_prop0 subset_eq)
+ thus "\<phi>: H \<rightarrow> carrier (BijGroup E) \<and> (\<forall>x \<in> H. \<forall>y \<in> H. \<phi> (x \<otimes> y) = \<phi> x \<otimes>\<^bsub>BijGroup E\<^esub> \<phi> y)"
+ by (simp add: S0 group_hom group_hom.hom_mult set_rev_mp)
+qed
+
+theorem (in group_action) induced_action:
+ assumes "subgroup H G"
+ shows "group_action (G \<lparr>carrier := H\<rparr>) E \<phi>"
+ unfolding group_action_def group_hom_def
+ using induced_homomorphism assms group.subgroup_imp_group group_BijGroup
+ group_hom group_hom.axioms(1) group_hom_axioms_def by blast
+
+end
\ No newline at end of file
--- a/src/HOL/Algebra/Ideal.thy Thu Jun 14 10:51:12 2018 +0200
+++ b/src/HOL/Algebra/Ideal.thy Thu Jun 14 14:23:48 2018 +0100
@@ -14,7 +14,7 @@
locale ideal = additive_subgroup I R + ring R for I and R (structure) +
assumes I_l_closed: "\<lbrakk>a \<in> I; x \<in> carrier R\<rbrakk> \<Longrightarrow> x \<otimes> a \<in> I"
- and I_r_closed: "\<lbrakk>a \<in> I; x \<in> carrier R\<rbrakk> \<Longrightarrow> a \<otimes> x \<in> I"
+ and I_r_closed: "\<lbrakk>a \<in> I; x \<in> carrier R\<rbrakk> \<Longrightarrow> a \<otimes> x \<in> I"
sublocale ideal \<subseteq> abelian_subgroup I R
apply (intro abelian_subgroupI3 abelian_group.intro)
@@ -29,7 +29,7 @@
lemma idealI:
fixes R (structure)
assumes "ring R"
- assumes a_subgroup: "subgroup I \<lparr>carrier = carrier R, mult = add R, one = zero R\<rparr>"
+ assumes a_subgroup: "subgroup I (add_monoid R)"
and I_l_closed: "\<And>a x. \<lbrakk>a \<in> I; x \<in> carrier R\<rbrakk> \<Longrightarrow> x \<otimes> a \<in> I"
and I_r_closed: "\<And>a x. \<lbrakk>a \<in> I; x \<in> carrier R\<rbrakk> \<Longrightarrow> a \<otimes> x \<in> I"
shows "ideal I R"
@@ -74,7 +74,7 @@
locale maximalideal = ideal +
assumes I_notcarr: "carrier R \<noteq> I"
- and I_maximal: "\<lbrakk>ideal J R; I \<subseteq> J; J \<subseteq> carrier R\<rbrakk> \<Longrightarrow> J = I \<or> J = carrier R"
+ and I_maximal: "\<lbrakk>ideal J R; I \<subseteq> J; J \<subseteq> carrier R\<rbrakk> \<Longrightarrow> (J = I) \<or> (J = carrier R)"
lemma (in maximalideal) is_maximalideal: "maximalideal I R"
by (rule maximalideal_axioms)
@@ -83,7 +83,7 @@
fixes R
assumes "ideal I R"
and I_notcarr: "carrier R \<noteq> I"
- and I_maximal: "\<And>J. \<lbrakk>ideal J R; I \<subseteq> J; J \<subseteq> carrier R\<rbrakk> \<Longrightarrow> J = I \<or> J = carrier R"
+ and I_maximal: "\<And>J. \<lbrakk>ideal J R; I \<subseteq> J; J \<subseteq> carrier R\<rbrakk> \<Longrightarrow> (J = I) \<or> (J = carrier R)"
shows "maximalideal I R"
proof -
interpret ideal I R by fact
@@ -143,26 +143,17 @@
subsection \<open>Special Ideals\<close>
lemma (in ring) zeroideal: "ideal {\<zero>} R"
- apply (intro idealI subgroup.intro)
- apply (rule is_ring)
- apply simp+
- apply (fold a_inv_def, simp)
- apply simp+
- done
+ by (intro idealI subgroup.intro) (simp_all add: is_ring)
lemma (in ring) oneideal: "ideal (carrier R) R"
by (rule idealI) (auto intro: is_ring add.subgroupI)
lemma (in "domain") zeroprimeideal: "primeideal {\<zero>} R"
- apply (intro primeidealI)
- apply (rule zeroideal)
- apply (rule domain.axioms, rule domain_axioms)
- defer 1
- apply (simp add: integral)
-proof (rule ccontr, simp)
- assume "carrier R = {\<zero>}"
- then have "\<one> = \<zero>" by (rule one_zeroI)
- with one_not_zero show False by simp
+proof -
+ have "carrier R \<noteq> {\<zero>}"
+ by (simp add: carrier_one_not_zero)
+ then show ?thesis
+ by (metis (no_types, lifting) domain_axioms domain_def integral primeidealI singleton_iff zeroideal)
qed
@@ -651,6 +642,46 @@
qed
+(* Next lemma contributed by Paulo EmÃlio de Vilhena. *)
+
+text \<open>This next lemma would be trivial if placed in a theory that imports QuotRing,
+ but it makes more sense to have it here (easier to find and coherent with the
+ previous developments).\<close>
+
+lemma (in cring) cgenideal_prod:
+ assumes "a \<in> carrier R" "b \<in> carrier R"
+ shows "(PIdl a) <#> (PIdl b) = PIdl (a \<otimes> b)"
+proof -
+ have "(carrier R #> a) <#> (carrier R #> b) = carrier R #> (a \<otimes> b)"
+ proof
+ show "(carrier R #> a) <#> (carrier R #> b) \<subseteq> carrier R #> a \<otimes> b"
+ proof
+ fix x assume "x \<in> (carrier R #> a) <#> (carrier R #> b)"
+ then obtain r1 r2 where r1: "r1 \<in> carrier R" and r2: "r2 \<in> carrier R"
+ and "x = (r1 \<otimes> a) \<otimes> (r2 \<otimes> b)"
+ unfolding set_mult_def r_coset_def by blast
+ hence "x = (r1 \<otimes> r2) \<otimes> (a \<otimes> b)"
+ by (simp add: assms local.ring_axioms m_lcomm ring.ring_simprules(11))
+ thus "x \<in> carrier R #> a \<otimes> b"
+ unfolding r_coset_def using r1 r2 assms by blast
+ qed
+ next
+ show "carrier R #> a \<otimes> b \<subseteq> (carrier R #> a) <#> (carrier R #> b)"
+ proof
+ fix x assume "x \<in> carrier R #> a \<otimes> b"
+ then obtain r where r: "r \<in> carrier R" "x = r \<otimes> (a \<otimes> b)"
+ unfolding r_coset_def by blast
+ hence "x = (r \<otimes> a) \<otimes> (\<one> \<otimes> b)"
+ using assms by (simp add: m_assoc)
+ thus "x \<in> (carrier R #> a) <#> (carrier R #> b)"
+ unfolding set_mult_def r_coset_def using assms r by blast
+ qed
+ qed
+ thus ?thesis
+ using cgenideal_eq_rcos[of a] cgenideal_eq_rcos[of b] cgenideal_eq_rcos[of "a \<otimes> b"] by simp
+qed
+
+
subsection \<open>Prime Ideals\<close>
lemma (in ideal) primeidealCD:
@@ -708,10 +739,7 @@
qed
corollary (in cring) domain_eq_zeroprimeideal: "domain R = primeideal {\<zero>} R"
- apply rule
- apply (erule domain.zeroprimeideal)
- apply (erule zeroprimeideal_domainI)
- done
+ using domain.zeroprimeideal zeroprimeideal_domainI by blast
subsection \<open>Maximal Ideals\<close>
@@ -921,7 +949,7 @@
qed
qed (simp add: zeroideal oneideal)
-\<comment> \<open>Jacobson Theorem 2.2\<close>
+\<comment>\<open>"Jacobson Theorem 2.2"\<close>
lemma (in cring) trivialideals_eq_field:
assumes carrnzero: "carrier R \<noteq> {\<zero>}"
shows "({I. ideal I R} = {{\<zero>}, carrier R}) = field R"
@@ -963,9 +991,6 @@
qed
lemma (in cring) zeromaximalideal_eq_field: "maximalideal {\<zero>} R = field R"
- apply rule
- apply (erule zeromaximalideal_fieldI)
- apply (erule field.zeromaximalideal)
- done
+ using field.zeromaximalideal zeromaximalideal_fieldI by blast
end
--- a/src/HOL/Algebra/More_Finite_Product.thy Thu Jun 14 10:51:12 2018 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,102 +0,0 @@
-(* Title: HOL/Algebra/More_Finite_Product.thy
- Author: Jeremy Avigad
-*)
-
-section \<open>More on finite products\<close>
-
-theory More_Finite_Product
- imports More_Group
-begin
-
-lemma (in comm_monoid) finprod_UN_disjoint:
- "finite I \<Longrightarrow> (\<forall>i\<in>I. finite (A i)) \<longrightarrow> (\<forall>i\<in>I. \<forall>j\<in>I. i \<noteq> j \<longrightarrow> A i \<inter> A j = {}) \<longrightarrow>
- (\<forall>i\<in>I. \<forall>x \<in> A i. g x \<in> carrier G) \<longrightarrow>
- finprod G g (UNION I A) = finprod G (\<lambda>i. finprod G g (A i)) I"
- apply (induct set: finite)
- apply force
- apply clarsimp
- apply (subst finprod_Un_disjoint)
- apply blast
- apply (erule finite_UN_I)
- apply blast
- apply (fastforce)
- apply (auto intro!: funcsetI finprod_closed)
- done
-
-lemma (in comm_monoid) finprod_Union_disjoint:
- "finite C \<Longrightarrow>
- \<forall>A\<in>C. finite A \<and> (\<forall>x\<in>A. f x \<in> carrier G) \<Longrightarrow>
- \<forall>A\<in>C. \<forall>B\<in>C. A \<noteq> B \<longrightarrow> A \<inter> B = {} \<Longrightarrow>
- finprod G f (\<Union>C) = finprod G (finprod G f) C"
- apply (frule finprod_UN_disjoint [of C id f])
- apply auto
- done
-
-lemma (in comm_monoid) finprod_one: "finite A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x = \<one>) \<Longrightarrow> finprod G f A = \<one>"
- by (induct set: finite) auto
-
-
-(* need better simplification rules for rings *)
-(* the next one holds more generally for abelian groups *)
-
-lemma (in cring) sum_zero_eq_neg: "x \<in> carrier R \<Longrightarrow> y \<in> carrier R \<Longrightarrow> x \<oplus> y = \<zero> \<Longrightarrow> x = \<ominus> y"
- by (metis minus_equality)
-
-lemma (in domain) square_eq_one:
- fixes x
- assumes [simp]: "x \<in> carrier R"
- and "x \<otimes> x = \<one>"
- shows "x = \<one> \<or> x = \<ominus>\<one>"
-proof -
- have "(x \<oplus> \<one>) \<otimes> (x \<oplus> \<ominus> \<one>) = x \<otimes> x \<oplus> \<ominus> \<one>"
- by (simp add: ring_simprules)
- also from \<open>x \<otimes> x = \<one>\<close> have "\<dots> = \<zero>"
- by (simp add: ring_simprules)
- finally have "(x \<oplus> \<one>) \<otimes> (x \<oplus> \<ominus> \<one>) = \<zero>" .
- then have "(x \<oplus> \<one>) = \<zero> \<or> (x \<oplus> \<ominus> \<one>) = \<zero>"
- by (intro integral) auto
- then show ?thesis
- apply auto
- apply (erule notE)
- apply (rule sum_zero_eq_neg)
- apply auto
- apply (subgoal_tac "x = \<ominus> (\<ominus> \<one>)")
- apply (simp add: ring_simprules)
- apply (rule sum_zero_eq_neg)
- apply auto
- done
-qed
-
-lemma (in domain) inv_eq_self: "x \<in> Units R \<Longrightarrow> x = inv x \<Longrightarrow> x = \<one> \<or> x = \<ominus>\<one>"
- by (metis Units_closed Units_l_inv square_eq_one)
-
-
-text \<open>
- The following translates theorems about groups to the facts about
- the units of a ring. (The list should be expanded as more things are
- needed.)
-\<close>
-
-lemma (in ring) finite_ring_finite_units [intro]: "finite (carrier R) \<Longrightarrow> finite (Units R)"
- by (rule finite_subset) auto
-
-lemma (in monoid) units_of_pow:
- fixes n :: nat
- shows "x \<in> Units G \<Longrightarrow> x [^]\<^bsub>units_of G\<^esub> n = x [^]\<^bsub>G\<^esub> n"
- apply (induct n)
- apply (auto simp add: units_group group.is_monoid
- monoid.nat_pow_0 monoid.nat_pow_Suc units_of_one units_of_mult)
- done
-
-lemma (in cring) units_power_order_eq_one:
- "finite (Units R) \<Longrightarrow> a \<in> Units R \<Longrightarrow> a [^] card(Units R) = \<one>"
- apply (subst units_of_carrier [symmetric])
- apply (subst units_of_one [symmetric])
- apply (subst units_of_pow [symmetric])
- apply assumption
- apply (rule comm_group.power_order_eq_one)
- apply (rule units_comm_group)
- apply (unfold units_of_def, auto)
- done
-
-end
--- a/src/HOL/Algebra/More_Group.thy Thu Jun 14 10:51:12 2018 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,115 +0,0 @@
-(* Title: HOL/Algebra/More_Group.thy
- Author: Jeremy Avigad
-*)
-
-section \<open>More on groups\<close>
-
-theory More_Group
- imports Ring
-begin
-
-text \<open>
- Show that the units in any monoid give rise to a group.
-
- The file Residues.thy provides some infrastructure to use
- facts about the unit group within the ring locale.
-\<close>
-
-definition units_of :: "('a, 'b) monoid_scheme \<Rightarrow> 'a monoid"
- where "units_of G =
- \<lparr>carrier = Units G, Group.monoid.mult = Group.monoid.mult G, one = one G\<rparr>"
-
-lemma (in monoid) units_group: "group (units_of G)"
- apply (unfold units_of_def)
- apply (rule groupI)
- apply auto
- apply (subst m_assoc)
- apply auto
- apply (rule_tac x = "inv x" in bexI)
- apply auto
- done
-
-lemma (in comm_monoid) units_comm_group: "comm_group (units_of G)"
- apply (rule group.group_comm_groupI)
- apply (rule units_group)
- apply (insert comm_monoid_axioms)
- apply (unfold units_of_def Units_def comm_monoid_def comm_monoid_axioms_def)
- apply auto
- done
-
-lemma units_of_carrier: "carrier (units_of G) = Units G"
- by (auto simp: units_of_def)
-
-lemma units_of_mult: "mult (units_of G) = mult G"
- by (auto simp: units_of_def)
-
-lemma units_of_one: "one (units_of G) = one G"
- by (auto simp: units_of_def)
-
-lemma (in monoid) units_of_inv: "x \<in> Units G \<Longrightarrow> m_inv (units_of G) x = m_inv G x"
- apply (rule sym)
- apply (subst m_inv_def)
- apply (rule the1_equality)
- apply (rule ex_ex1I)
- apply (subst (asm) Units_def)
- apply auto
- apply (erule inv_unique)
- apply auto
- apply (rule Units_closed)
- apply (simp_all only: units_of_carrier [symmetric])
- apply (insert units_group)
- apply auto
- apply (subst units_of_mult [symmetric])
- apply (subst units_of_one [symmetric])
- apply (erule group.r_inv, assumption)
- apply (subst units_of_mult [symmetric])
- apply (subst units_of_one [symmetric])
- apply (erule group.l_inv, assumption)
- done
-
-lemma (in group) inj_on_const_mult: "a \<in> carrier G \<Longrightarrow> inj_on (\<lambda>x. a \<otimes> x) (carrier G)"
- unfolding inj_on_def by auto
-
-lemma (in group) surj_const_mult: "a \<in> carrier G \<Longrightarrow> (\<lambda>x. a \<otimes> x) ` carrier G = carrier G"
- apply (auto simp add: image_def)
- apply (rule_tac x = "(m_inv G a) \<otimes> x" in bexI)
- apply auto
-(* auto should get this. I suppose we need "comm_monoid_simprules"
- for ac_simps rewriting. *)
- apply (subst m_assoc [symmetric])
- apply auto
- done
-
-lemma (in group) l_cancel_one [simp]: "x \<in> carrier G \<Longrightarrow> a \<in> carrier G \<Longrightarrow> x \<otimes> a = x \<longleftrightarrow> a = one G"
- by (metis Units_eq Units_l_cancel monoid.r_one monoid_axioms one_closed)
-
-lemma (in group) r_cancel_one [simp]: "x \<in> carrier G \<Longrightarrow> a \<in> carrier G \<Longrightarrow> a \<otimes> x = x \<longleftrightarrow> a = one G"
- by (metis monoid.l_one monoid_axioms one_closed right_cancel)
-
-lemma (in group) l_cancel_one' [simp]: "x \<in> carrier G \<Longrightarrow> a \<in> carrier G \<Longrightarrow> x = x \<otimes> a \<longleftrightarrow> a = one G"
- using l_cancel_one by fastforce
-
-lemma (in group) r_cancel_one' [simp]: "x \<in> carrier G \<Longrightarrow> a \<in> carrier G \<Longrightarrow> x = a \<otimes> x \<longleftrightarrow> a = one G"
- using r_cancel_one by fastforce
-
-(* This should be generalized to arbitrary groups, not just commutative
- ones, using Lagrange's theorem. *)
-
-lemma (in comm_group) power_order_eq_one:
- assumes fin [simp]: "finite (carrier G)"
- and a [simp]: "a \<in> carrier G"
- shows "a [^] card(carrier G) = one G"
-proof -
- have "(\<Otimes>x\<in>carrier G. x) = (\<Otimes>x\<in>carrier G. a \<otimes> x)"
- by (subst (2) finprod_reindex [symmetric],
- auto simp add: Pi_def inj_on_const_mult surj_const_mult)
- also have "\<dots> = (\<Otimes>x\<in>carrier G. a) \<otimes> (\<Otimes>x\<in>carrier G. x)"
- by (auto simp add: finprod_multf Pi_def)
- also have "(\<Otimes>x\<in>carrier G. a) = a [^] card(carrier G)"
- by (auto simp add: finprod_const)
- finally show ?thesis
-(* uses the preceeding lemma *)
- by auto
-qed
-
-end
--- a/src/HOL/Algebra/More_Ring.thy Thu Jun 14 10:51:12 2018 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,74 +0,0 @@
-(* Title: HOL/Algebra/More_Ring.thy
- Author: Jeremy Avigad
-*)
-
-section \<open>More on rings etc.\<close>
-
-theory More_Ring
- imports Ring
-begin
-
-lemma (in cring) field_intro2: "\<zero>\<^bsub>R\<^esub> \<noteq> \<one>\<^bsub>R\<^esub> \<Longrightarrow> \<forall>x \<in> carrier R - {\<zero>\<^bsub>R\<^esub>}. x \<in> Units R \<Longrightarrow> field R"
- apply (unfold_locales)
- apply (use cring_axioms in auto)
- apply (rule trans)
- apply (subgoal_tac "a = (a \<otimes> b) \<otimes> inv b")
- apply assumption
- apply (subst m_assoc)
- apply auto
- apply (unfold Units_def)
- apply auto
- done
-
-lemma (in monoid) inv_char:
- "x \<in> carrier G \<Longrightarrow> y \<in> carrier G \<Longrightarrow> x \<otimes> y = \<one> \<Longrightarrow> y \<otimes> x = \<one> \<Longrightarrow> inv x = y"
- apply (subgoal_tac "x \<in> Units G")
- apply (subgoal_tac "y = inv x \<otimes> \<one>")
- apply simp
- apply (erule subst)
- apply (subst m_assoc [symmetric])
- apply auto
- apply (unfold Units_def)
- apply auto
- done
-
-lemma (in comm_monoid) comm_inv_char: "x \<in> carrier G \<Longrightarrow> y \<in> carrier G \<Longrightarrow> x \<otimes> y = \<one> \<Longrightarrow> inv x = y"
- apply (rule inv_char)
- apply auto
- apply (subst m_comm, auto)
- done
-
-lemma (in ring) inv_neg_one [simp]: "inv (\<ominus> \<one>) = \<ominus> \<one>"
- apply (rule inv_char)
- apply (auto simp add: l_minus r_minus)
- done
-
-lemma (in monoid) inv_eq_imp_eq: "x \<in> Units G \<Longrightarrow> y \<in> Units G \<Longrightarrow> inv x = inv y \<Longrightarrow> x = y"
- apply (subgoal_tac "inv (inv x) = inv (inv y)")
- apply (subst (asm) Units_inv_inv)+
- apply auto
- done
-
-lemma (in ring) Units_minus_one_closed [intro]: "\<ominus> \<one> \<in> Units R"
- apply (unfold Units_def)
- apply auto
- apply (rule_tac x = "\<ominus> \<one>" in bexI)
- apply auto
- apply (simp add: l_minus r_minus)
- done
-
-lemma (in monoid) inv_one [simp]: "inv \<one> = \<one>"
- apply (rule inv_char)
- apply auto
- done
-
-lemma (in ring) inv_eq_neg_one_eq: "x \<in> Units R \<Longrightarrow> inv x = \<ominus> \<one> \<longleftrightarrow> x = \<ominus> \<one>"
- apply auto
- apply (subst Units_inv_inv [symmetric])
- apply auto
- done
-
-lemma (in monoid) inv_eq_one_eq: "x \<in> Units G \<Longrightarrow> inv x = \<one> \<longleftrightarrow> x = \<one>"
- by (metis Units_inv_inv inv_one)
-
-end
--- a/src/HOL/Algebra/Multiplicative_Group.thy Thu Jun 14 10:51:12 2018 +0200
+++ b/src/HOL/Algebra/Multiplicative_Group.thy Thu Jun 14 14:23:48 2018 +0100
@@ -7,8 +7,6 @@
imports
Complex_Main
Group
- More_Group
- More_Finite_Product
Coset
UnivPoly
begin
--- a/src/HOL/Algebra/Ring.thy Thu Jun 14 10:51:12 2018 +0200
+++ b/src/HOL/Algebra/Ring.thy Thu Jun 14 14:23:48 2018 +0100
@@ -13,79 +13,100 @@
record 'a ring = "'a monoid" +
zero :: 'a ("\<zero>\<index>")
- add :: "['a, 'a] => 'a" (infixl "\<oplus>\<index>" 65)
+ add :: "['a, 'a] \<Rightarrow> 'a" (infixl "\<oplus>\<index>" 65)
+
+abbreviation
+ add_monoid :: "('a, 'm) ring_scheme \<Rightarrow> ('a, 'm) monoid_scheme"
+ where "add_monoid R \<equiv> \<lparr> carrier = carrier R, mult = add R, one = zero R, \<dots> = (undefined :: 'm) \<rparr>"
text \<open>Derived operations.\<close>
definition
- a_inv :: "[('a, 'm) ring_scheme, 'a ] => 'a" ("\<ominus>\<index> _" [81] 80)
- where "a_inv R = m_inv \<lparr>carrier = carrier R, mult = add R, one = zero R\<rparr>"
+ a_inv :: "[('a, 'm) ring_scheme, 'a ] \<Rightarrow> 'a" ("\<ominus>\<index> _" [81] 80)
+ where "a_inv R = m_inv (add_monoid R)"
definition
a_minus :: "[('a, 'm) ring_scheme, 'a, 'a] => 'a" ("(_ \<ominus>\<index> _)" [65,66] 65)
- where "[| x \<in> carrier R; y \<in> carrier R |] ==> x \<ominus>\<^bsub>R\<^esub> y = x \<oplus>\<^bsub>R\<^esub> (\<ominus>\<^bsub>R\<^esub> y)"
+ where "x \<ominus>\<^bsub>R\<^esub> y = x \<oplus>\<^bsub>R\<^esub> (\<ominus>\<^bsub>R\<^esub> y)"
+
+definition
+ add_pow :: "[_, ('b :: semiring_1), 'a] \<Rightarrow> 'a" ("[_] \<cdot>\<index> _" [81, 81] 80)
+ where "add_pow R k a = pow (add_monoid R) a k"
locale abelian_monoid =
fixes G (structure)
assumes a_comm_monoid:
- "comm_monoid \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
+ "comm_monoid (add_monoid G)"
definition
- finsum :: "[('b, 'm) ring_scheme, 'a => 'b, 'a set] => 'b" where
- "finsum G = finprod \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
+ finsum :: "[('b, 'm) ring_scheme, 'a \<Rightarrow> 'b, 'a set] \<Rightarrow> 'b" where
+ "finsum G = finprod (add_monoid G)"
syntax
- "_finsum" :: "index => idt => 'a set => 'b => 'b"
+ "_finsum" :: "index \<Rightarrow> idt \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"
("(3\<Oplus>__\<in>_. _)" [1000, 0, 51, 10] 10)
translations
- "\<Oplus>\<^bsub>G\<^esub>i\<in>A. b" \<rightleftharpoons> "CONST finsum G (%i. b) A"
+ "\<Oplus>\<^bsub>G\<^esub>i\<in>A. b" \<rightleftharpoons> "CONST finsum G (\<lambda>i. b) A"
\<comment> \<open>Beware of argument permutation!\<close>
locale abelian_group = abelian_monoid +
assumes a_comm_group:
- "comm_group \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
+ "comm_group (add_monoid G)"
subsection \<open>Basic Properties\<close>
lemma abelian_monoidI:
fixes R (structure)
- assumes a_closed:
- "!!x y. [| x \<in> carrier R; y \<in> carrier R |] ==> x \<oplus> y \<in> carrier R"
- and zero_closed: "\<zero> \<in> carrier R"
- and a_assoc:
- "!!x y z. [| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |] ==>
- (x \<oplus> y) \<oplus> z = x \<oplus> (y \<oplus> z)"
- and l_zero: "!!x. x \<in> carrier R ==> \<zero> \<oplus> x = x"
- and a_comm:
- "!!x y. [| x \<in> carrier R; y \<in> carrier R |] ==> x \<oplus> y = y \<oplus> x"
+ assumes "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> x \<oplus> y \<in> carrier R"
+ and "\<zero> \<in> carrier R"
+ and "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> (x \<oplus> y) \<oplus> z = x \<oplus> (y \<oplus> z)"
+ and "\<And>x. x \<in> carrier R \<Longrightarrow> \<zero> \<oplus> x = x"
+ and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> x \<oplus> y = y \<oplus> x"
shows "abelian_monoid R"
by (auto intro!: abelian_monoid.intro comm_monoidI intro: assms)
+lemma abelian_monoidE:
+ fixes R (structure)
+ assumes "abelian_monoid R"
+ shows "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> x \<oplus> y \<in> carrier R"
+ and "\<zero> \<in> carrier R"
+ and "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> (x \<oplus> y) \<oplus> z = x \<oplus> (y \<oplus> z)"
+ and "\<And>x. x \<in> carrier R \<Longrightarrow> \<zero> \<oplus> x = x"
+ and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> x \<oplus> y = y \<oplus> x"
+ using assms unfolding abelian_monoid_def comm_monoid_def comm_monoid_axioms_def monoid_def by auto
+
lemma abelian_groupI:
fixes R (structure)
- assumes a_closed:
- "!!x y. [| x \<in> carrier R; y \<in> carrier R |] ==> x \<oplus> y \<in> carrier R"
- and zero_closed: "zero R \<in> carrier R"
- and a_assoc:
- "!!x y z. [| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |] ==>
- (x \<oplus> y) \<oplus> z = x \<oplus> (y \<oplus> z)"
- and a_comm:
- "!!x y. [| x \<in> carrier R; y \<in> carrier R |] ==> x \<oplus> y = y \<oplus> x"
- and l_zero: "!!x. x \<in> carrier R ==> \<zero> \<oplus> x = x"
- and l_inv_ex: "\<And>x. x \<in> carrier R \<Longrightarrow> \<exists>y \<in> carrier R. y \<oplus> x = \<zero>"
+ assumes "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> x \<oplus> y \<in> carrier R"
+ and "\<zero> \<in> carrier R"
+ and "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> (x \<oplus> y) \<oplus> z = x \<oplus> (y \<oplus> z)"
+ and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> x \<oplus> y = y \<oplus> x"
+ and "\<And>x. x \<in> carrier R \<Longrightarrow> \<zero> \<oplus> x = x"
+ and "\<And>x. x \<in> carrier R \<Longrightarrow> \<exists>y \<in> carrier R. y \<oplus> x = \<zero>"
shows "abelian_group R"
by (auto intro!: abelian_group.intro abelian_monoidI
abelian_group_axioms.intro comm_monoidI comm_groupI
intro: assms)
+lemma abelian_groupE:
+ fixes R (structure)
+ assumes "abelian_group R"
+ shows "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> x \<oplus> y \<in> carrier R"
+ and "\<zero> \<in> carrier R"
+ and "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> (x \<oplus> y) \<oplus> z = x \<oplus> (y \<oplus> z)"
+ and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> x \<oplus> y = y \<oplus> x"
+ and "\<And>x. x \<in> carrier R \<Longrightarrow> \<zero> \<oplus> x = x"
+ and "\<And>x. x \<in> carrier R \<Longrightarrow> \<exists>y \<in> carrier R. y \<oplus> x = \<zero>"
+ using abelian_group.a_comm_group assms comm_groupE by fastforce+
+
lemma (in abelian_monoid) a_monoid:
- "monoid \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
+ "monoid (add_monoid G)"
by (rule comm_monoid.axioms, rule a_comm_monoid)
lemma (in abelian_group) a_group:
- "group \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
+ "group (add_monoid G)"
by (simp add: group_def a_monoid)
(simp add: comm_group.axioms group.axioms a_comm_group)
@@ -94,13 +115,15 @@
text \<open>Transfer facts from multiplicative structures via interpretation.\<close>
sublocale abelian_monoid <
- add: monoid "\<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
- rewrites "carrier \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = carrier G"
- and "mult \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = add G"
- and "one \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = zero G"
- by (rule a_monoid) auto
+ add: monoid "(add_monoid G)"
+ rewrites "carrier (add_monoid G) = carrier G"
+ and "mult (add_monoid G) = add G"
+ and "one (add_monoid G) = zero G"
+ and "(\<lambda>a k. pow (add_monoid G) a k) = (\<lambda>a k. add_pow G k a)"
+ by (rule a_monoid) (auto simp add: add_pow_def)
-context abelian_monoid begin
+context abelian_monoid
+begin
lemmas a_closed = add.m_closed
lemmas zero_closed = add.one_closed
@@ -112,12 +135,13 @@
end
sublocale abelian_monoid <
- add: comm_monoid "\<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
- rewrites "carrier \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = carrier G"
- and "mult \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = add G"
- and "one \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = zero G"
- and "finprod \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = finsum G"
- by (rule a_comm_monoid) (auto simp: finsum_def)
+ add: comm_monoid "(add_monoid G)"
+ rewrites "carrier (add_monoid G) = carrier G"
+ and "mult (add_monoid G) = add G"
+ and "one (add_monoid G) = zero G"
+ and "finprod (add_monoid G) = finsum G"
+ and "pow (add_monoid G) = (\<lambda>a k. add_pow G k a)"
+ by (rule a_comm_monoid) (auto simp: finsum_def add_pow_def)
context abelian_monoid begin
@@ -168,12 +192,13 @@
end
sublocale abelian_group <
- add: group "\<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
- rewrites "carrier \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = carrier G"
- and "mult \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = add G"
- and "one \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = zero G"
- and "m_inv \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = a_inv G"
- by (rule a_group) (auto simp: m_inv_def a_inv_def)
+ add: group "(add_monoid G)"
+ rewrites "carrier (add_monoid G) = carrier G"
+ and "mult (add_monoid G) = add G"
+ and "one (add_monoid G) = zero G"
+ and "m_inv (add_monoid G) = a_inv G"
+ and "pow (add_monoid G) = (\<lambda>a k. add_pow G k a)"
+ by (rule a_group) (auto simp: m_inv_def a_inv_def add_pow_def)
context abelian_group
begin
@@ -186,7 +211,6 @@
lemmas l_neg = add.l_inv [simp del]
lemmas r_neg = add.r_inv [simp del]
-lemmas minus_zero = add.inv_one
lemmas minus_minus = add.inv_inv
lemmas a_inv_inj = add.inv_inj
lemmas minus_equality = add.inv_equality
@@ -194,13 +218,14 @@
end
sublocale abelian_group <
- add: comm_group "\<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
- rewrites "carrier \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = carrier G"
- and "mult \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = add G"
- and "one \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = zero G"
- and "m_inv \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = a_inv G"
- and "finprod \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = finsum G"
- by (rule a_comm_group) (auto simp: m_inv_def a_inv_def finsum_def)
+ add: comm_group "(add_monoid G)"
+ rewrites "carrier (add_monoid G) = carrier G"
+ and "mult (add_monoid G) = add G"
+ and "one (add_monoid G) = zero G"
+ and "m_inv (add_monoid G) = a_inv G"
+ and "finprod (add_monoid G) = finsum G"
+ and "pow (add_monoid G) = (\<lambda>a k. add_pow G k a)"
+ by (rule a_comm_group) (auto simp: m_inv_def a_inv_def finsum_def add_pow_def)
lemmas (in abelian_group) minus_add = add.inv_mult
@@ -208,10 +233,10 @@
lemma comm_group_abelian_groupI:
fixes G (structure)
- assumes cg: "comm_group \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
+ assumes cg: "comm_group (add_monoid G)"
shows "abelian_group G"
proof -
- interpret comm_group "\<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
+ interpret comm_group "(add_monoid G)"
by (rule cg)
show "abelian_group G" ..
qed
@@ -219,26 +244,21 @@
subsection \<open>Rings: Basic Definitions\<close>
-locale semiring = abelian_monoid R + monoid R for R (structure) +
- assumes l_distr: "[| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |]
- ==> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z"
- and r_distr: "[| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |]
- ==> z \<otimes> (x \<oplus> y) = z \<otimes> x \<oplus> z \<otimes> y"
- and l_null[simp]: "x \<in> carrier R ==> \<zero> \<otimes> x = \<zero>"
- and r_null[simp]: "x \<in> carrier R ==> x \<otimes> \<zero> = \<zero>"
+locale semiring = abelian_monoid (* for add *) R + monoid (* for mult *) R for R (structure) +
+ assumes l_distr: "\<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z"
+ and r_distr: "\<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> z \<otimes> (x \<oplus> y) = z \<otimes> x \<oplus> z \<otimes> y"
+ and l_null[simp]: "x \<in> carrier R \<Longrightarrow> \<zero> \<otimes> x = \<zero>"
+ and r_null[simp]: "x \<in> carrier R \<Longrightarrow> x \<otimes> \<zero> = \<zero>"
-locale ring = abelian_group R + monoid R for R (structure) +
- assumes "[| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |]
- ==> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z"
- and "[| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |]
- ==> z \<otimes> (x \<oplus> y) = z \<otimes> x \<oplus> z \<otimes> y"
+locale ring = abelian_group (* for add *) R + monoid (* for mult *) R for R (structure) +
+ assumes "\<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z"
+ and "\<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> z \<otimes> (x \<oplus> y) = z \<otimes> x \<oplus> z \<otimes> y"
-locale cring = ring + comm_monoid R
+locale cring = ring + comm_monoid (* for mult *) R
locale "domain" = cring +
assumes one_not_zero [simp]: "\<one> \<noteq> \<zero>"
- and integral: "[| a \<otimes> b = \<zero>; a \<in> carrier R; b \<in> carrier R |] ==>
- a = \<zero> \<or> b = \<zero>"
+ and integral: "\<lbrakk> a \<otimes> b = \<zero>; a \<in> carrier R; b \<in> carrier R \<rbrakk> \<Longrightarrow> a = \<zero> \<or> b = \<zero>"
locale field = "domain" +
assumes field_Units: "Units R = carrier R - {\<zero>}"
@@ -248,16 +268,23 @@
lemma ringI:
fixes R (structure)
- assumes abelian_group: "abelian_group R"
- and monoid: "monoid R"
- and l_distr: "!!x y z. [| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |]
- ==> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z"
- and r_distr: "!!x y z. [| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |]
- ==> z \<otimes> (x \<oplus> y) = z \<otimes> x \<oplus> z \<otimes> y"
+ assumes "abelian_group R"
+ and "monoid R"
+ and "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z"
+ and "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> z \<otimes> (x \<oplus> y) = z \<otimes> x \<oplus> z \<otimes> y"
shows "ring R"
by (auto intro: ring.intro
abelian_group.axioms ring_axioms.intro assms)
+lemma ringE:
+ fixes R (structure)
+ assumes "ring R"
+ shows "abelian_group R"
+ and "monoid R"
+ and "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z"
+ and "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> z \<otimes> (x \<oplus> y) = z \<otimes> x \<oplus> z \<otimes> y"
+ using assms unfolding ring_def ring_axioms_def by auto
+
context ring begin
lemma is_abelian_group: "abelian_group R" ..
@@ -269,15 +296,15 @@
by (rule ring_axioms)
end
-
+thm monoid_record_simps
lemmas ring_record_simps = monoid_record_simps ring.simps
lemma cringI:
fixes R (structure)
assumes abelian_group: "abelian_group R"
and comm_monoid: "comm_monoid R"
- and l_distr: "!!x y z. [| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |]
- ==> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z"
+ and l_distr: "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow>
+ (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z"
shows "cring R"
proof (intro cring.intro ring.intro)
show "ring_axioms R"
@@ -300,20 +327,37 @@
qed (auto intro: cring.intro
abelian_group.axioms comm_monoid.axioms ring_axioms.intro assms)
+lemma cringE:
+ fixes R (structure)
+ assumes "cring R"
+ shows "comm_monoid R"
+ and "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z"
+ using assms cring_def apply auto by (simp add: assms cring.axioms(1) ringE(3))
+
(*
lemma (in cring) is_comm_monoid:
"comm_monoid R"
by (auto intro!: comm_monoidI m_assoc m_comm)
*)
-
lemma (in cring) is_cring:
"cring R" by (rule cring_axioms)
+lemma (in ring) minus_zero [simp]: "\<ominus> \<zero> = \<zero>"
+ by (simp add: a_inv_def)
subsubsection \<open>Normaliser for Rings\<close>
+lemma (in abelian_group) r_neg1:
+ "\<lbrakk> x \<in> carrier G; y \<in> carrier G \<rbrakk> \<Longrightarrow> (\<ominus> x) \<oplus> (x \<oplus> y) = y"
+proof -
+ assume G: "x \<in> carrier G" "y \<in> carrier G"
+ then have "(\<ominus> x \<oplus> x) \<oplus> y = y"
+ by (simp only: l_neg l_zero)
+ with G show ?thesis by (simp add: a_ac)
+qed
+
lemma (in abelian_group) r_neg2:
- "[| x \<in> carrier G; y \<in> carrier G |] ==> x \<oplus> (\<ominus> x \<oplus> y) = y"
+ "\<lbrakk> x \<in> carrier G; y \<in> carrier G \<rbrakk> \<Longrightarrow> x \<oplus> ((\<ominus> x) \<oplus> y) = y"
proof -
assume G: "x \<in> carrier G" "y \<in> carrier G"
then have "(x \<oplus> \<ominus> x) \<oplus> y = y"
@@ -322,15 +366,6 @@
by (simp add: a_ac)
qed
-lemma (in abelian_group) r_neg1:
- "[| x \<in> carrier G; y \<in> carrier G |] ==> \<ominus> x \<oplus> (x \<oplus> y) = y"
-proof -
- assume G: "x \<in> carrier G" "y \<in> carrier G"
- then have "(\<ominus> x \<oplus> x) \<oplus> y = y"
- by (simp only: l_neg l_zero)
- with G show ?thesis by (simp add: a_ac)
-qed
-
context ring begin
text \<open>
@@ -358,7 +393,7 @@
qed
lemma l_minus:
- "[| x \<in> carrier R; y \<in> carrier R |] ==> \<ominus> x \<otimes> y = \<ominus> (x \<otimes> y)"
+ "\<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> (\<ominus> x) \<otimes> y = \<ominus> (x \<otimes> y)"
proof -
assume R: "x \<in> carrier R" "y \<in> carrier R"
then have "(\<ominus> x) \<otimes> y \<oplus> x \<otimes> y = (\<ominus> x \<oplus> x) \<otimes> y" by (simp add: l_distr)
@@ -369,7 +404,7 @@
qed
lemma r_minus:
- "[| x \<in> carrier R; y \<in> carrier R |] ==> x \<otimes> \<ominus> y = \<ominus> (x \<otimes> y)"
+ "\<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> x \<otimes> (\<ominus> y) = \<ominus> (x \<otimes> y)"
proof -
assume R: "x \<in> carrier R" "y \<in> carrier R"
then have "x \<otimes> (\<ominus> y) \<oplus> x \<otimes> y = x \<otimes> (\<ominus> y \<oplus> y)" by (simp add: r_distr)
@@ -381,13 +416,13 @@
end
-lemma (in abelian_group) minus_eq:
- "[| x \<in> carrier G; y \<in> carrier G |] ==> x \<ominus> y = x \<oplus> \<ominus> y"
- by (simp only: a_minus_def)
+lemma (in abelian_group) minus_eq: "x \<ominus> y = x \<oplus> (\<ominus> y)"
+ by (rule a_minus_def)
text \<open>Setup algebra method:
compute distributive normal form in locale contexts\<close>
+
ML_file "ringsimp.ML"
attribute_setup algebra = \<open>
@@ -467,7 +502,7 @@
fixes R (structure) and S (structure)
assumes "ring R" "cring S"
assumes RS: "a \<in> carrier R" "b \<in> carrier R" "c \<in> carrier S" "d \<in> carrier S"
- shows "a \<oplus> \<ominus> (a \<oplus> \<ominus> b) = b \<and> c \<otimes>\<^bsub>S\<^esub> d = d \<otimes>\<^bsub>S\<^esub> c"
+ shows "a \<oplus> (\<ominus> (a \<oplus> (\<ominus> b))) = b \<and> c \<otimes>\<^bsub>S\<^esub> d = d \<otimes>\<^bsub>S\<^esub> c"
proof -
interpret ring R by fact
interpret cring S by fact
@@ -488,8 +523,8 @@
subsubsection \<open>Sums over Finite Sets\<close>
lemma (in semiring) finsum_ldistr:
- "[| finite A; a \<in> carrier R; f \<in> A \<rightarrow> carrier R |] ==>
- finsum R f A \<otimes> a = finsum R (%i. f i \<otimes> a) A"
+ "\<lbrakk> finite A; a \<in> carrier R; f: A \<rightarrow> carrier R \<rbrakk> \<Longrightarrow>
+ (\<Oplus> i \<in> A. (f i)) \<otimes> a = (\<Oplus> i \<in> A. ((f i) \<otimes> a))"
proof (induct set: finite)
case empty then show ?case by simp
next
@@ -497,25 +532,86 @@
qed
lemma (in semiring) finsum_rdistr:
- "[| finite A; a \<in> carrier R; f \<in> A \<rightarrow> carrier R |] ==>
- a \<otimes> finsum R f A = finsum R (%i. a \<otimes> f i) A"
+ "\<lbrakk> finite A; a \<in> carrier R; f: A \<rightarrow> carrier R \<rbrakk> \<Longrightarrow>
+ a \<otimes> (\<Oplus> i \<in> A. (f i)) = (\<Oplus> i \<in> A. (a \<otimes> (f i)))"
proof (induct set: finite)
case empty then show ?case by simp
next
case (insert x F) then show ?case by (simp add: Pi_def r_distr)
qed
+(* ************************************************************************** *)
+(* Contributed by Paulo E. de Vilhena. *)
+
+text \<open>A quick detour\<close>
+
+lemma add_pow_int_ge: "(k :: int) \<ge> 0 \<Longrightarrow> [ k ] \<cdot>\<^bsub>R\<^esub> a = [ nat k ] \<cdot>\<^bsub>R\<^esub> a"
+ by (simp add: add_pow_def int_pow_def nat_pow_def)
+
+lemma add_pow_int_lt: "(k :: int) < 0 \<Longrightarrow> [ k ] \<cdot>\<^bsub>R\<^esub> a = \<ominus>\<^bsub>R\<^esub> ([ nat (- k) ] \<cdot>\<^bsub>R\<^esub> a)"
+ by (simp add: int_pow_def nat_pow_def a_inv_def add_pow_def)
+
+corollary (in semiring) add_pow_ldistr:
+ assumes "a \<in> carrier R" "b \<in> carrier R"
+ shows "([(k :: nat)] \<cdot> a) \<otimes> b = [k] \<cdot> (a \<otimes> b)"
+proof -
+ have "([k] \<cdot> a) \<otimes> b = (\<Oplus> i \<in> {..< k}. a) \<otimes> b"
+ using add.finprod_const[OF assms(1), of "{..<k}"] by simp
+ also have " ... = (\<Oplus> i \<in> {..< k}. (a \<otimes> b))"
+ using finsum_ldistr[of "{..<k}" b "\<lambda>x. a"] assms by simp
+ also have " ... = [k] \<cdot> (a \<otimes> b)"
+ using add.finprod_const[of "a \<otimes> b" "{..<k}"] assms by simp
+ finally show ?thesis .
+qed
+
+corollary (in semiring) add_pow_rdistr:
+ assumes "a \<in> carrier R" "b \<in> carrier R"
+ shows "a \<otimes> ([(k :: nat)] \<cdot> b) = [k] \<cdot> (a \<otimes> b)"
+proof -
+ have "a \<otimes> ([k] \<cdot> b) = a \<otimes> (\<Oplus> i \<in> {..< k}. b)"
+ using add.finprod_const[OF assms(2), of "{..<k}"] by simp
+ also have " ... = (\<Oplus> i \<in> {..< k}. (a \<otimes> b))"
+ using finsum_rdistr[of "{..<k}" a "\<lambda>x. b"] assms by simp
+ also have " ... = [k] \<cdot> (a \<otimes> b)"
+ using add.finprod_const[of "a \<otimes> b" "{..<k}"] assms by simp
+ finally show ?thesis .
+qed
+
+(* For integers, we need the uniqueness of the additive inverse *)
+lemma (in ring) add_pow_ldistr_int:
+ assumes "a \<in> carrier R" "b \<in> carrier R"
+ shows "([(k :: int)] \<cdot> a) \<otimes> b = [k] \<cdot> (a \<otimes> b)"
+proof (cases "k \<ge> 0")
+ case True thus ?thesis
+ using add_pow_int_ge[of k R] add_pow_ldistr[OF assms] by auto
+next
+ case False thus ?thesis
+ using add_pow_int_lt[of k R a] add_pow_int_lt[of k R "a \<otimes> b"]
+ add_pow_ldistr[OF assms, of "nat (- k)"] assms l_minus by auto
+qed
+
+lemma (in ring) add_pow_rdistr_int:
+ assumes "a \<in> carrier R" "b \<in> carrier R"
+ shows "a \<otimes> ([(k :: int)] \<cdot> b) = [k] \<cdot> (a \<otimes> b)"
+proof (cases "k \<ge> 0")
+ case True thus ?thesis
+ using add_pow_int_ge[of k R] add_pow_rdistr[OF assms] by auto
+next
+ case False thus ?thesis
+ using add_pow_int_lt[of k R b] add_pow_int_lt[of k R "a \<otimes> b"]
+ add_pow_rdistr[OF assms, of "nat (- k)"] assms r_minus by auto
+qed
+
subsection \<open>Integral Domains\<close>
context "domain" begin
-lemma zero_not_one [simp]:
- "\<zero> \<noteq> \<one>"
+lemma zero_not_one [simp]: "\<zero> \<noteq> \<one>"
by (rule not_sym) simp
lemma integral_iff: (* not by default a simp rule! *)
- "[| a \<in> carrier R; b \<in> carrier R |] ==> (a \<otimes> b = \<zero>) = (a = \<zero> \<or> b = \<zero>)"
+ "\<lbrakk> a \<in> carrier R; b \<in> carrier R \<rbrakk> \<Longrightarrow> (a \<otimes> b = \<zero>) = (a = \<zero> \<or> b = \<zero>)"
proof
assume "a \<in> carrier R" "b \<in> carrier R" "a \<otimes> b = \<zero>"
then show "a = \<zero> \<or> b = \<zero>" by (simp add: integral)
@@ -533,7 +629,7 @@
with R have "a \<otimes> (b \<ominus> c) = \<zero>" by algebra
with R have "a = \<zero> \<or> (b \<ominus> c) = \<zero>" by (simp add: integral_iff)
with prem and R have "b \<ominus> c = \<zero>" by auto
- with R have "b = b \<ominus> (b \<ominus> c)" by algebra
+ with R have "b = b \<ominus> (b \<ominus> c)" by algebra
also from R have "b \<ominus> (b \<ominus> c) = c" by algebra
finally show "b = c" .
next
@@ -556,6 +652,7 @@
text \<open>Field would not need to be derived from domain, the properties
for domain follow from the assumptions of field\<close>
+
lemma (in cring) cring_fieldI:
assumes field_Units: "Units R = carrier R - {\<zero>}"
shows "field R"
@@ -614,49 +711,62 @@
lemma ring_hom_memI:
fixes R (structure) and S (structure)
- assumes hom_closed: "!!x. x \<in> carrier R ==> h x \<in> carrier S"
- and hom_mult: "!!x y. [| x \<in> carrier R; y \<in> carrier R |] ==>
- h (x \<otimes> y) = h x \<otimes>\<^bsub>S\<^esub> h y"
- and hom_add: "!!x y. [| x \<in> carrier R; y \<in> carrier R |] ==>
- h (x \<oplus> y) = h x \<oplus>\<^bsub>S\<^esub> h y"
- and hom_one: "h \<one> = \<one>\<^bsub>S\<^esub>"
+ assumes "\<And>x. x \<in> carrier R \<Longrightarrow> h x \<in> carrier S"
+ and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> h (x \<otimes> y) = h x \<otimes>\<^bsub>S\<^esub> h y"
+ and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> h (x \<oplus> y) = h x \<oplus>\<^bsub>S\<^esub> h y"
+ and "h \<one> = \<one>\<^bsub>S\<^esub>"
shows "h \<in> ring_hom R S"
by (auto simp add: ring_hom_def assms Pi_def)
+lemma ring_hom_memE:
+ fixes R (structure) and S (structure)
+ assumes "h \<in> ring_hom R S"
+ shows "\<And>x. x \<in> carrier R \<Longrightarrow> h x \<in> carrier S"
+ and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> h (x \<otimes> y) = h x \<otimes>\<^bsub>S\<^esub> h y"
+ and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> h (x \<oplus> y) = h x \<oplus>\<^bsub>S\<^esub> h y"
+ and "h \<one> = \<one>\<^bsub>S\<^esub>"
+ using assms unfolding ring_hom_def by auto
+
lemma ring_hom_closed:
- "[| h \<in> ring_hom R S; x \<in> carrier R |] ==> h x \<in> carrier S"
+ "\<lbrakk> h \<in> ring_hom R S; x \<in> carrier R \<rbrakk> \<Longrightarrow> h x \<in> carrier S"
by (auto simp add: ring_hom_def funcset_mem)
lemma ring_hom_mult:
fixes R (structure) and S (structure)
- shows
- "[| h \<in> ring_hom R S; x \<in> carrier R; y \<in> carrier R |] ==>
- h (x \<otimes> y) = h x \<otimes>\<^bsub>S\<^esub> h y"
+ shows "\<lbrakk> h \<in> ring_hom R S; x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> h (x \<otimes> y) = h x \<otimes>\<^bsub>S\<^esub> h y"
by (simp add: ring_hom_def)
lemma ring_hom_add:
fixes R (structure) and S (structure)
- shows
- "[| h \<in> ring_hom R S; x \<in> carrier R; y \<in> carrier R |] ==>
- h (x \<oplus> y) = h x \<oplus>\<^bsub>S\<^esub> h y"
+ shows "\<lbrakk> h \<in> ring_hom R S; x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> h (x \<oplus> y) = h x \<oplus>\<^bsub>S\<^esub> h y"
by (simp add: ring_hom_def)
lemma ring_hom_one:
fixes R (structure) and S (structure)
- shows "h \<in> ring_hom R S ==> h \<one> = \<one>\<^bsub>S\<^esub>"
+ shows "h \<in> ring_hom R S \<Longrightarrow> h \<one> = \<one>\<^bsub>S\<^esub>"
by (simp add: ring_hom_def)
-locale ring_hom_cring = R?: cring R + S?: cring S
- for R (structure) and S (structure) +
- fixes h
+lemma ring_hom_zero:
+ fixes R (structure) and S (structure)
+ assumes "h \<in> ring_hom R S" "ring R" "ring S"
+ shows "h \<zero> = \<zero>\<^bsub>S\<^esub>"
+proof -
+ have "h \<zero> = h \<zero> \<oplus>\<^bsub>S\<^esub> h \<zero>"
+ using ring_hom_add[OF assms(1), of \<zero> \<zero>] assms(2)
+ by (simp add: ring.ring_simprules(2) ring.ring_simprules(15))
+ thus ?thesis
+ by (metis abelian_group.l_neg assms ring.is_abelian_group ring.ring_simprules(18) ring.ring_simprules(2) ring_hom_closed)
+qed
+
+locale ring_hom_cring =
+ R?: cring R + S?: cring S for R (structure) and S (structure) + fixes h
assumes homh [simp, intro]: "h \<in> ring_hom R S"
notes hom_closed [simp, intro] = ring_hom_closed [OF homh]
and hom_mult [simp] = ring_hom_mult [OF homh]
and hom_add [simp] = ring_hom_add [OF homh]
and hom_one [simp] = ring_hom_one [OF homh]
-lemma (in ring_hom_cring) hom_zero [simp]:
- "h \<zero> = \<zero>\<^bsub>S\<^esub>"
+lemma (in ring_hom_cring) hom_zero [simp]: "h \<zero> = \<zero>\<^bsub>S\<^esub>"
proof -
have "h \<zero> \<oplus>\<^bsub>S\<^esub> h \<zero> = h \<zero> \<oplus>\<^bsub>S\<^esub> \<zero>\<^bsub>S\<^esub>"
by (simp add: hom_add [symmetric] del: hom_add)
@@ -664,7 +774,7 @@
qed
lemma (in ring_hom_cring) hom_a_inv [simp]:
- "x \<in> carrier R ==> h (\<ominus> x) = \<ominus>\<^bsub>S\<^esub> h x"
+ "x \<in> carrier R \<Longrightarrow> h (\<ominus> x) = \<ominus>\<^bsub>S\<^esub> h x"
proof -
assume R: "x \<in> carrier R"
then have "h x \<oplus>\<^bsub>S\<^esub> h (\<ominus> x) = h x \<oplus>\<^bsub>S\<^esub> (\<ominus>\<^bsub>S\<^esub> h x)"
@@ -673,19 +783,131 @@
qed
lemma (in ring_hom_cring) hom_finsum [simp]:
- "f \<in> A \<rightarrow> carrier R \<Longrightarrow>
- h (finsum R f A) = finsum S (h \<circ> f) A"
- by (induct A rule: infinite_finite_induct, auto simp: Pi_def)
+ assumes "f: A \<rightarrow> carrier R"
+ shows "h (\<Oplus> i \<in> A. f i) = (\<Oplus>\<^bsub>S\<^esub> i \<in> A. (h o f) i)"
+ using assms by (induct A rule: infinite_finite_induct, auto simp: Pi_def)
lemma (in ring_hom_cring) hom_finprod:
- "f \<in> A \<rightarrow> carrier R \<Longrightarrow>
- h (finprod R f A) = finprod S (h \<circ> f) A"
- by (induct A rule: infinite_finite_induct, auto simp: Pi_def)
+ assumes "f: A \<rightarrow> carrier R"
+ shows "h (\<Otimes> i \<in> A. f i) = (\<Otimes>\<^bsub>S\<^esub> i \<in> A. (h o f) i)"
+ using assms by (induct A rule: infinite_finite_induct, auto simp: Pi_def)
declare ring_hom_cring.hom_finprod [simp]
-lemma id_ring_hom [simp]:
- "id \<in> ring_hom R R"
+lemma id_ring_hom [simp]: "id \<in> ring_hom R R"
by (auto intro!: ring_hom_memI)
+(* Next lemma contributed by Paulo EmÃlio de Vilhena. *)
+
+lemma ring_hom_trans:
+ "\<lbrakk> f \<in> ring_hom R S; g \<in> ring_hom S T \<rbrakk> \<Longrightarrow> g \<circ> f \<in> ring_hom R T"
+ by (rule ring_hom_memI) (auto simp add: ring_hom_closed ring_hom_mult ring_hom_add ring_hom_one)
+
+subsection\<open>Jeremy Avigad's @{text"More_Finite_Product"} material\<close>
+
+(* need better simplification rules for rings *)
+(* the next one holds more generally for abelian groups *)
+
+lemma (in cring) sum_zero_eq_neg: "x \<in> carrier R \<Longrightarrow> y \<in> carrier R \<Longrightarrow> x \<oplus> y = \<zero> \<Longrightarrow> x = \<ominus> y"
+ by (metis minus_equality)
+
+lemma (in domain) square_eq_one:
+ fixes x
+ assumes [simp]: "x \<in> carrier R"
+ and "x \<otimes> x = \<one>"
+ shows "x = \<one> \<or> x = \<ominus>\<one>"
+proof -
+ have "(x \<oplus> \<one>) \<otimes> (x \<oplus> \<ominus> \<one>) = x \<otimes> x \<oplus> \<ominus> \<one>"
+ by (simp add: ring_simprules)
+ also from \<open>x \<otimes> x = \<one>\<close> have "\<dots> = \<zero>"
+ by (simp add: ring_simprules)
+ finally have "(x \<oplus> \<one>) \<otimes> (x \<oplus> \<ominus> \<one>) = \<zero>" .
+ then have "(x \<oplus> \<one>) = \<zero> \<or> (x \<oplus> \<ominus> \<one>) = \<zero>"
+ by (intro integral) auto
+ then show ?thesis
+ by (metis add.inv_closed add.inv_solve_right assms(1) l_zero one_closed zero_closed)
+qed
+
+lemma (in domain) inv_eq_self: "x \<in> Units R \<Longrightarrow> x = inv x \<Longrightarrow> x = \<one> \<or> x = \<ominus>\<one>"
+ by (metis Units_closed Units_l_inv square_eq_one)
+
+
+text \<open>
+ The following translates theorems about groups to the facts about
+ the units of a ring. (The list should be expanded as more things are
+ needed.)
+\<close>
+
+lemma (in ring) finite_ring_finite_units [intro]: "finite (carrier R) \<Longrightarrow> finite (Units R)"
+ by (rule finite_subset) auto
+
+lemma (in monoid) units_of_pow:
+ fixes n :: nat
+ shows "x \<in> Units G \<Longrightarrow> x [^]\<^bsub>units_of G\<^esub> n = x [^]\<^bsub>G\<^esub> n"
+ apply (induct n)
+ apply (auto simp add: units_group group.is_monoid
+ monoid.nat_pow_0 monoid.nat_pow_Suc units_of_one units_of_mult)
+ done
+
+lemma (in cring) units_power_order_eq_one:
+ "finite (Units R) \<Longrightarrow> a \<in> Units R \<Longrightarrow> a [^] card(Units R) = \<one>"
+ by (metis comm_group.power_order_eq_one units_comm_group units_of_carrier units_of_one units_of_pow)
+
+subsection\<open>Jeremy Avigad's @{text"More_Ring"} material\<close>
+
+lemma (in cring) field_intro2: "\<zero>\<^bsub>R\<^esub> \<noteq> \<one>\<^bsub>R\<^esub> \<Longrightarrow> \<forall>x \<in> carrier R - {\<zero>\<^bsub>R\<^esub>}. x \<in> Units R \<Longrightarrow> field R"
+ apply (unfold_locales)
+ apply (use cring_axioms in auto)
+ apply (rule trans)
+ apply (subgoal_tac "a = (a \<otimes> b) \<otimes> inv b")
+ apply assumption
+ apply (subst m_assoc)
+ apply auto
+ apply (unfold Units_def)
+ apply auto
+ done
+
+lemma (in monoid) inv_char:
+ "x \<in> carrier G \<Longrightarrow> y \<in> carrier G \<Longrightarrow> x \<otimes> y = \<one> \<Longrightarrow> y \<otimes> x = \<one> \<Longrightarrow> inv x = y"
+ apply (subgoal_tac "x \<in> Units G")
+ apply (subgoal_tac "y = inv x \<otimes> \<one>")
+ apply simp
+ apply (erule subst)
+ apply (subst m_assoc [symmetric])
+ apply auto
+ apply (unfold Units_def)
+ apply auto
+ done
+
+lemma (in comm_monoid) comm_inv_char: "x \<in> carrier G \<Longrightarrow> y \<in> carrier G \<Longrightarrow> x \<otimes> y = \<one> \<Longrightarrow> inv x = y"
+ by (simp add: inv_char m_comm)
+
+lemma (in ring) inv_neg_one [simp]: "inv (\<ominus> \<one>) = \<ominus> \<one>"
+ apply (rule inv_char)
+ apply (auto simp add: l_minus r_minus)
+ done
+
+lemma (in monoid) inv_eq_imp_eq: "x \<in> Units G \<Longrightarrow> y \<in> Units G \<Longrightarrow> inv x = inv y \<Longrightarrow> x = y"
+ apply (subgoal_tac "inv (inv x) = inv (inv y)")
+ apply (subst (asm) Units_inv_inv)+
+ apply auto
+ done
+
+lemma (in ring) Units_minus_one_closed [intro]: "\<ominus> \<one> \<in> Units R"
+ apply (unfold Units_def)
+ apply auto
+ apply (rule_tac x = "\<ominus> \<one>" in bexI)
+ apply auto
+ apply (simp add: l_minus r_minus)
+ done
+
+lemma (in ring) inv_eq_neg_one_eq: "x \<in> Units R \<Longrightarrow> inv x = \<ominus> \<one> \<longleftrightarrow> x = \<ominus> \<one>"
+ apply auto
+ apply (subst Units_inv_inv [symmetric])
+ apply auto
+ done
+
+lemma (in monoid) inv_eq_one_eq: "x \<in> Units G \<Longrightarrow> inv x = \<one> \<longleftrightarrow> x = \<one>"
+ by (metis Units_inv_inv inv_one)
+
end
--- a/src/HOL/Algebra/Sylow.thy Thu Jun 14 10:51:12 2018 +0200
+++ b/src/HOL/Algebra/Sylow.thy Thu Jun 14 14:23:48 2018 +0100
@@ -189,7 +189,7 @@
using rcosetGM1g_subset_G finite_G M1_subset_G cosets_finite rcosetsI by blast
lemma M1_cardeq_rcosetGM1g: "g \<in> carrier G \<Longrightarrow> card (M1 #> g) = card M1"
- by (simp add: card_cosets_equal rcosetsI)
+ by (metis M1_subset_G card_rcosets_equal rcosetsI)
lemma M1_RelM_rcosetGM1g: "g \<in> carrier G \<Longrightarrow> (M1, M1 #> g) \<in> RelM"
apply (simp add: RelM_def calM_def card_M1)
@@ -223,7 +223,7 @@
lemma M_funcset_rcosets_H:
"(\<lambda>x\<in>M. H #> (SOME g. g \<in> carrier G \<and> M1 #> g = x)) \<in> M \<rightarrow> rcosets H"
- by (metis (lifting) H_is_subgroup M_elem_map_carrier rcosetsI restrictI subgroup_imp_subset)
+ by (metis (lifting) H_is_subgroup M_elem_map_carrier rcosetsI restrictI subgroup.subset)
lemma inj_M_GmodH: "\<exists>f \<in> M \<rightarrow> rcosets H. inj_on f M"
apply (rule bexI)
--- a/src/HOL/Algebra/UnivPoly.thy Thu Jun 14 10:51:12 2018 +0200
+++ b/src/HOL/Algebra/UnivPoly.thy Thu Jun 14 14:23:48 2018 +0100
@@ -116,14 +116,15 @@
proof
assume R: "p \<in> up R"
then obtain n where "bound \<zero> n p" by auto
- then have "bound \<zero> n (\<lambda>i. \<ominus> p i)" by auto
+ then have "bound \<zero> n (\<lambda>i. \<ominus> p i)"
+ by (simp add: bound_def minus_equality)
then show "\<exists>n. bound \<zero> n (\<lambda>i. \<ominus> p i)" by auto
qed auto
lemma up_minus_closed:
"[| p \<in> up R; q \<in> up R |] ==> (\<lambda>i. p i \<ominus> q i) \<in> up R"
- using mem_upD [of p R] mem_upD [of q R] up_add_closed up_a_inv_closed a_minus_def [of _ R]
- by auto
+ unfolding a_minus_def
+ using mem_upD [of p R] mem_upD [of q R] up_add_closed up_a_inv_closed by auto
lemma up_mult_closed:
"[| p \<in> up R; q \<in> up R |] ==>
@@ -695,7 +696,7 @@
lemma monom_a_inv [simp]:
"a \<in> carrier R ==> monom P (\<ominus> a) n = \<ominus>\<^bsub>P\<^esub> monom P a n"
- by (rule up_eqI) simp_all
+ by (rule up_eqI) auto
lemma monom_inj:
"inj_on (\<lambda>a. monom P a n) (carrier R)"
@@ -1462,9 +1463,9 @@
subsection\<open>The long division algorithm: some previous facts.\<close>
lemma coeff_minus [simp]:
- assumes p: "p \<in> carrier P" and q: "q \<in> carrier P" shows "coeff P (p \<ominus>\<^bsub>P\<^esub> q) n = coeff P p n \<ominus> coeff P q n"
- unfolding a_minus_def [OF p q] unfolding coeff_add [OF p a_inv_closed [OF q]] unfolding coeff_a_inv [OF q]
- using coeff_closed [OF p, of n] using coeff_closed [OF q, of n] by algebra
+ assumes p: "p \<in> carrier P" and q: "q \<in> carrier P"
+ shows "coeff P (p \<ominus>\<^bsub>P\<^esub> q) n = coeff P p n \<ominus> coeff P q n"
+ by (simp add: a_minus_def p q)
lemma lcoeff_closed [simp]: assumes p: "p \<in> carrier P" shows "lcoeff p \<in> carrier R"
using coeff_closed [OF p, of "deg R p"] by simp
@@ -1719,10 +1720,7 @@
and min_mon0_closed: "\<ominus>\<^bsub>P\<^esub> monom P a 0 \<in> carrier P"
using a R.a_inv_closed by auto
have "eval R R id a ?g = eval R R id a (monom P \<one> 1) \<ominus> eval R R id a (monom P a 0)"
- unfolding P.minus_eq [OF mon1_closed mon0_closed]
- unfolding hom_add [OF mon1_closed min_mon0_closed]
- unfolding hom_a_inv [OF mon0_closed]
- using R.minus_eq [symmetric] mon1_closed mon0_closed by auto
+ by (simp add: a_minus_def mon0_closed)
also have "\<dots> = a \<ominus> a"
using eval_monom [OF R.one_closed a, of 1] using eval_monom [OF a a, of 0] using a by simp
also have "\<dots> = \<zero>"
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Algebra/Zassenhaus.thy Thu Jun 14 14:23:48 2018 +0100
@@ -0,0 +1,633 @@
+theory Zassenhaus
+ imports Coset Group_Action
+begin
+
+
+subsubsection \<open>Lemmas about normalizer\<close>
+
+
+lemma (in group) subgroup_in_normalizer:
+ assumes "subgroup H G"
+ shows "normal H (G\<lparr>carrier:= (normalizer G H)\<rparr>)"
+proof(intro group.normal_invI)
+ show "Group.group (G\<lparr>carrier := normalizer G H\<rparr>)"
+ by (simp add: assms group.normalizer_imp_subgroup is_group subgroup_imp_group subgroup_imp_subset)
+ have K:"H \<subseteq> (normalizer G H)" unfolding normalizer_def
+ proof
+ fix x assume xH: "x \<in> H"
+ from xH have xG : "x \<in> carrier G" using subgroup_imp_subset assms by auto
+ have "x <# H = H"
+ by (metis \<open>x \<in> H\<close> assms group.lcos_mult_one is_group
+ l_repr_independence one_closed subgroup_imp_subset)
+ moreover have "H #> inv x = H"
+ by (simp add: xH assms is_group subgroup.rcos_const subgroup.m_inv_closed)
+ ultimately have "x <# H #> (inv x) = H" by simp
+ thus " x \<in> stabilizer G (\<lambda>g. \<lambda>H\<in>{H. H \<subseteq> carrier G}. g <# H #> inv g) H"
+ using assms xG subgroup_imp_subset unfolding stabilizer_def by auto
+ qed
+ thus "subgroup H (G\<lparr>carrier:= (normalizer G H)\<rparr>)"
+ using subgroup_incl normalizer_imp_subgroup assms by (simp add: subgroup_imp_subset)
+ show " \<And>x h. x \<in> carrier (G\<lparr>carrier := normalizer G H\<rparr>) \<Longrightarrow> h \<in> H \<Longrightarrow>
+ x \<otimes>\<^bsub>G\<lparr>carrier := normalizer G H\<rparr>\<^esub> h
+ \<otimes>\<^bsub>G\<lparr>carrier := normalizer G H\<rparr>\<^esub> inv\<^bsub>G\<lparr>carrier := normalizer G H\<rparr>\<^esub> x \<in> H"
+ proof-
+ fix x h assume xnorm : "x \<in> carrier (G\<lparr>carrier := normalizer G H\<rparr>)" and hH : "h \<in> H"
+ have xnormalizer:"x \<in> normalizer G H" using xnorm by simp
+ moreover have hnormalizer:"h \<in> normalizer G H" using hH K by auto
+ ultimately have "x \<otimes>\<^bsub>G\<lparr>carrier := normalizer G H\<rparr>\<^esub> h = x \<otimes> h" by simp
+ moreover have " inv\<^bsub>G\<lparr>carrier := normalizer G H\<rparr>\<^esub> x = inv x"
+ using xnormalizer
+ by (simp add: assms normalizer_imp_subgroup subgroup_imp_subset subgroup_inv_equality)
+ ultimately have xhxegal: "x \<otimes>\<^bsub>G\<lparr>carrier := normalizer G H\<rparr>\<^esub> h
+ \<otimes>\<^bsub>G\<lparr>carrier := normalizer G H\<rparr>\<^esub> inv\<^bsub>G\<lparr>carrier := normalizer G H\<rparr>\<^esub> x
+ = x \<otimes>h \<otimes> inv x"
+ using hnormalizer by simp
+ have "x \<otimes>h \<otimes> inv x \<in> (x <# H #> inv x)"
+ unfolding l_coset_def r_coset_def using hH by auto
+ moreover have "x <# H #> inv x = H"
+ using xnormalizer assms subgroup_imp_subset[OF assms]
+ unfolding normalizer_def stabilizer_def by auto
+ ultimately have "x \<otimes>h \<otimes> inv x \<in> H" by simp
+ thus " x \<otimes>\<^bsub>G\<lparr>carrier := normalizer G H\<rparr>\<^esub> h
+ \<otimes>\<^bsub>G\<lparr>carrier := normalizer G H\<rparr>\<^esub> inv\<^bsub>G\<lparr>carrier := normalizer G H\<rparr>\<^esub> x \<in> H"
+ using xhxegal hH xnorm by simp
+ qed
+qed
+
+
+lemma (in group) normal_imp_subgroup_normalizer:
+ assumes "subgroup H G"
+ and "N \<lhd> (G\<lparr>carrier := H\<rparr>)"
+ shows "subgroup H (G\<lparr>carrier := normalizer G N\<rparr>)"
+proof-
+ have N_carrierG : "N \<subseteq> carrier(G)"
+ using assms normal_imp_subgroup subgroup_imp_subset
+ by (smt monoid.cases_scheme order_trans partial_object.simps(1) partial_object.update_convs(1))
+ {have "H \<subseteq> normalizer G N" unfolding normalizer_def stabilizer_def
+ proof
+ fix x assume xH : "x \<in> H"
+ hence xcarrierG : "x \<in> carrier(G)" using assms subgroup_imp_subset by auto
+ have " N #> x = x <# N" using assms xH
+ unfolding r_coset_def l_coset_def normal_def normal_axioms_def subgroup_imp_group by auto
+ hence "x <# N #> inv x =(N #> x) #> inv x"
+ by simp
+ also have "... = N #> \<one>"
+ using assms r_inv xcarrierG coset_mult_assoc[OF N_carrierG] by simp
+ finally have "x <# N #> inv x = N" by (simp add: N_carrierG)
+ thus "x \<in> {g \<in> carrier G. (\<lambda>H\<in>{H. H \<subseteq> carrier G}. g <# H #> inv g) N = N}"
+ using xcarrierG by (simp add : N_carrierG)
+ qed}
+ thus "subgroup H (G\<lparr>carrier := normalizer G N\<rparr>)"
+ using subgroup_incl[OF assms(1) normalizer_imp_subgroup]
+ assms normal_imp_subgroup subgroup_imp_subset
+ by (metis group.incl_subgroup is_group)
+qed
+
+
+subsection \<open>Second Isomorphism Theorem\<close>
+
+lemma (in group) mult_norm_subgroup:
+ assumes "normal N G"
+ and "subgroup H G"
+ shows "subgroup (N<#>H) G" unfolding subgroup_def
+proof-
+ have A :"N <#> H \<subseteq> carrier G"
+ using assms setmult_subset_G by (simp add: normal_imp_subgroup subgroup_imp_subset)
+
+ have B :"\<And> x y. \<lbrakk>x \<in> (N <#> H); y \<in> (N <#> H)\<rbrakk> \<Longrightarrow> (x \<otimes> y) \<in> (N<#>H)"
+ proof-
+ fix x y assume B1a: "x \<in> (N <#> H)" and B1b: "y \<in> (N <#> H)"
+ obtain n1 h1 where B2:"n1 \<in> N \<and> h1 \<in> H \<and> n1\<otimes>h1 = x"
+ using set_mult_def B1a by (metis (no_types, lifting) UN_E singletonD)
+ obtain n2 h2 where B3:"n2 \<in> N \<and> h2 \<in> H \<and> n2\<otimes>h2 = y"
+ using set_mult_def B1b by (metis (no_types, lifting) UN_E singletonD)
+ have "N #> h1 = h1 <# N"
+ using normalI B2 assms normal.coset_eq subgroup_imp_subset by blast
+ hence "h1\<otimes>n2 \<in> N #> h1"
+ using B2 B3 assms l_coset_def by fastforce
+ from this obtain y2 where y2_def:"y2 \<in> N" and y2_prop:"y2\<otimes>h1 = h1\<otimes>n2"
+ using singletonD by (metis (no_types, lifting) UN_E r_coset_def)
+ have " x\<otimes>y = n1 \<otimes> y2 \<otimes> h1 \<otimes> h2" using y2_def B2 B3
+ by (smt assms y2_prop m_assoc m_closed normal_imp_subgroup subgroup.mem_carrier)
+ moreover have B4 :"n1 \<otimes> y2 \<in>N"
+ using B2 y2_def assms normal_imp_subgroup by (metis subgroup_def)
+ moreover have "h1 \<otimes> h2 \<in>H" using B2 B3 assms by (simp add: subgroup.m_closed)
+ hence "(n1 \<otimes> y2) \<otimes> (h1 \<otimes> h2) \<in>(N<#>H) "
+ using B4 unfolding set_mult_def by auto
+ hence "n1 \<otimes> y2 \<otimes> h1 \<otimes> h2 \<in>(N<#>H)"
+ using m_assoc B2 B3 assms normal_imp_subgroup by (metis B4 subgroup.mem_carrier)
+ ultimately show "x \<otimes> y \<in> N <#> H" by auto
+ qed
+ have C :"\<And> x. x\<in>(N<#>H) \<Longrightarrow> (inv x)\<in>(N<#>H)"
+
+ proof-
+ fix x assume C1 : "x \<in> (N<#>H)"
+ obtain n h where C2:"n \<in> N \<and> h \<in> H \<and> n\<otimes>h = x"
+ using set_mult_def C1 by (metis (no_types, lifting) UN_E singletonD)
+ have C3 :"inv(n\<otimes>h) = inv(h)\<otimes>inv(n)"
+ by (meson C2 assms inv_mult_group normal_imp_subgroup subgroup.mem_carrier)
+ hence "... \<otimes>h \<in> N"
+ using assms C2
+ by (meson normal.inv_op_closed1 normal_def subgroup.m_inv_closed subgroup.mem_carrier)
+ hence C4:"(inv h \<otimes> inv n \<otimes> h) \<otimes> inv h \<in> (N<#>H)"
+ using C2 assms subgroup.m_inv_closed[of H G h] unfolding set_mult_def by auto
+ have "inv h \<otimes> inv n \<otimes> h \<otimes> inv h = inv h \<otimes> inv n"
+ using subgroup_imp_subset[OF assms(2)]
+ by (metis A C1 C2 C3 inv_closed inv_solve_right m_closed subsetCE)
+ thus "inv(x)\<in>N<#>H" using C4 C2 C3 by simp
+ qed
+
+ have D : "\<one> \<in> N <#> H"
+ proof-
+ have D1 : "\<one> \<in> N"
+ using assms by (simp add: normal_def subgroup.one_closed)
+ have D2 :"\<one> \<in> H"
+ using assms by (simp add: subgroup.one_closed)
+ thus "\<one> \<in> (N <#> H)"
+ using set_mult_def D1 assms by fastforce
+ qed
+ thus "(N <#> H \<subseteq> carrier G \<and> (\<forall>x y. x \<in> N <#> H \<longrightarrow> y \<in> N <#> H \<longrightarrow> x \<otimes> y \<in> N <#> H)) \<and>
+ \<one> \<in> N <#> H \<and> (\<forall>x. x \<in> N <#> H \<longrightarrow> inv x \<in> N <#> H)" using A B C D assms by blast
+qed
+
+
+lemma (in group) mult_norm_sub_in_sub:
+ assumes "normal N (G\<lparr>carrier:=K\<rparr>)"
+ assumes "subgroup H (G\<lparr>carrier:=K\<rparr>)"
+ assumes "subgroup K G"
+ shows "subgroup (N<#>H) (G\<lparr>carrier:=K\<rparr>)"
+proof-
+ have Hyp:"subgroup (N <#>\<^bsub>G\<lparr>carrier := K\<rparr>\<^esub> H) (G\<lparr>carrier := K\<rparr>)"
+ using group.mult_norm_subgroup[where ?G = "G\<lparr>carrier := K\<rparr>"] assms subgroup_imp_group by auto
+ have "H \<subseteq> carrier(G\<lparr>carrier := K\<rparr>)" using assms subgroup_imp_subset by blast
+ also have "... \<subseteq> K" by simp
+ finally have Incl1:"H \<subseteq> K" by simp
+ have "N \<subseteq> carrier(G\<lparr>carrier := K\<rparr>)" using assms normal_imp_subgroup subgroup_imp_subset by blast
+ also have "... \<subseteq> K" by simp
+ finally have Incl2:"N \<subseteq> K" by simp
+ have "(N <#>\<^bsub>G\<lparr>carrier := K\<rparr>\<^esub> H) = (N <#> H)"
+ using subgroup_set_mult_equality[of K] assms Incl1 Incl2 by simp
+ thus "subgroup (N<#>H) (G\<lparr>carrier:=K\<rparr>)" using Hyp by auto
+qed
+
+
+lemma (in group) subgroup_of_normal_set_mult:
+ assumes "normal N G"
+and "subgroup H G"
+shows "subgroup H (G\<lparr>carrier := N <#> H\<rparr>)"
+proof-
+ have "\<one> \<in> N" using normal_imp_subgroup assms(1) subgroup_def by blast
+ hence "\<one> <# H \<subseteq> N <#> H" unfolding set_mult_def l_coset_def by blast
+ hence H_incl : "H \<subseteq> N <#> H"
+ by (metis assms(2) lcos_mult_one subgroup_def)
+ show "subgroup H (G\<lparr>carrier := N <#> H\<rparr>)"
+ using subgroup_incl[OF assms(2) mult_norm_subgroup[OF assms(1) assms(2)] H_incl] .
+qed
+
+
+lemma (in group) normal_in_normal_set_mult:
+ assumes "normal N G"
+and "subgroup H G"
+shows "normal N (G\<lparr>carrier := N <#> H\<rparr>)"
+proof-
+ have "\<one> \<in> H" using assms(2) subgroup_def by blast
+ hence "N #> \<one> \<subseteq> N <#> H" unfolding set_mult_def r_coset_def by blast
+ hence N_incl : "N \<subseteq> N <#> H"
+ by (metis assms(1) normal_imp_subgroup coset_mult_one subgroup_def)
+ thus "normal N (G\<lparr>carrier := N <#> H\<rparr>)"
+ using normal_inter_subgroup[OF mult_norm_subgroup[OF assms] assms(1)]
+ by (simp add : inf_absorb1)
+qed
+
+
+proposition (in group) weak_snd_iso_thme:
+ assumes "subgroup H G"
+ and "N\<lhd>G"
+ shows "(G\<lparr>carrier := N<#>H\<rparr> Mod N \<cong> G\<lparr>carrier:=H\<rparr> Mod (N\<inter>H))"
+proof-
+ define f where "f = (#>) N"
+ have GroupNH : "Group.group (G\<lparr>carrier := N<#>H\<rparr>)"
+ using subgroup_imp_group assms mult_norm_subgroup by simp
+ have HcarrierNH :"H \<subseteq> carrier(G\<lparr>carrier := N<#>H\<rparr>)"
+ using assms subgroup_of_normal_set_mult subgroup_imp_subset by blast
+ hence HNH :"H \<subseteq> N<#>H" by simp
+ have op_hom : "f \<in> hom (G\<lparr>carrier := H\<rparr>) (G\<lparr>carrier := N <#> H\<rparr> Mod N)" unfolding hom_def
+ proof
+ have "\<And>x . x \<in> carrier (G\<lparr>carrier :=H\<rparr>) \<Longrightarrow>
+ (#>\<^bsub>G\<lparr>carrier := N <#> H\<rparr>\<^esub>) N x \<in> carrier (G\<lparr>carrier := N <#> H\<rparr> Mod N)"
+ proof-
+ fix x assume "x \<in> carrier (G\<lparr>carrier :=H\<rparr>)"
+ hence xH : "x \<in> H" by simp
+ hence "(#>\<^bsub>G\<lparr>carrier := N <#> H\<rparr>\<^esub>) N x \<in> rcosets\<^bsub>G\<lparr>carrier := N <#> H\<rparr>\<^esub> N"
+ using HcarrierNH RCOSETS_def[where ?G = "G\<lparr>carrier := N <#> H\<rparr>"] by blast
+ thus "(#>\<^bsub>G\<lparr>carrier := N <#> H\<rparr>\<^esub>) N x \<in> carrier (G\<lparr>carrier := N <#> H\<rparr> Mod N)"
+ unfolding FactGroup_def by simp
+ qed
+ hence "(#>\<^bsub>G\<lparr>carrier := N <#> H\<rparr>\<^esub>) N \<in> carrier (G\<lparr>carrier :=H\<rparr>) \<rightarrow>
+ carrier (G\<lparr>carrier := N <#> H\<rparr> Mod N)" by auto
+ hence "f \<in> carrier (G\<lparr>carrier :=H\<rparr>) \<rightarrow> carrier (G\<lparr>carrier := N <#> H\<rparr> Mod N)"
+ unfolding r_coset_def f_def by simp
+ moreover have "\<And>x y. x\<in>carrier (G\<lparr>carrier := H\<rparr>) \<Longrightarrow> y\<in>carrier (G\<lparr>carrier := H\<rparr>) \<Longrightarrow>
+ f (x \<otimes>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> y) = f(x) \<otimes>\<^bsub>G\<lparr>carrier := N <#> H\<rparr> Mod N\<^esub> f(y)"
+ proof-
+ fix x y assume "x\<in>carrier (G\<lparr>carrier := H\<rparr>)" "y\<in>carrier (G\<lparr>carrier := H\<rparr>)"
+ hence xHyH :"x \<in> H" "y \<in> H" by auto
+ have Nxeq :"N #>\<^bsub>G\<lparr>carrier := N<#>H\<rparr>\<^esub> x = N #>x" unfolding r_coset_def by simp
+ have Nyeq :"N #>\<^bsub>G\<lparr>carrier := N<#>H\<rparr>\<^esub> y = N #>y" unfolding r_coset_def by simp
+
+ have "x \<otimes>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> y =x \<otimes>\<^bsub>G\<lparr>carrier := N<#>H\<rparr>\<^esub> y" by simp
+ hence "N #>\<^bsub>G\<lparr>carrier := N<#>H\<rparr>\<^esub> x \<otimes>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> y
+ = N #>\<^bsub>G\<lparr>carrier := N<#>H\<rparr>\<^esub> x \<otimes>\<^bsub>G\<lparr>carrier := N<#>H\<rparr>\<^esub> y" by simp
+ also have "... = (N #>\<^bsub>G\<lparr>carrier := N<#>H\<rparr>\<^esub> x) <#>\<^bsub>G\<lparr>carrier := N<#>H\<rparr>\<^esub>
+ (N #>\<^bsub>G\<lparr>carrier := N<#>H\<rparr>\<^esub> y)"
+ using normal.rcos_sum[OF normal_in_normal_set_mult[OF assms(2) assms(1)], of x y]
+ xHyH assms HcarrierNH by auto
+ finally show "f (x \<otimes>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> y) = f(x) \<otimes>\<^bsub>G\<lparr>carrier := N <#> H\<rparr> Mod N\<^esub> f(y)"
+ unfolding FactGroup_def r_coset_def f_def using Nxeq Nyeq by auto
+ qed
+ hence "(\<forall>x\<in>carrier (G\<lparr>carrier := H\<rparr>). \<forall>y\<in>carrier (G\<lparr>carrier := H\<rparr>).
+ f (x \<otimes>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> y) = f(x) \<otimes>\<^bsub>G\<lparr>carrier := N <#> H\<rparr> Mod N\<^esub> f(y))" by blast
+ ultimately show " f \<in> carrier (G\<lparr>carrier := H\<rparr>) \<rightarrow> carrier (G\<lparr>carrier := N <#> H\<rparr> Mod N) \<and>
+ (\<forall>x\<in>carrier (G\<lparr>carrier := H\<rparr>). \<forall>y\<in>carrier (G\<lparr>carrier := H\<rparr>).
+ f (x \<otimes>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> y) = f(x) \<otimes>\<^bsub>G\<lparr>carrier := N <#> H\<rparr> Mod N\<^esub> f(y))"
+ by auto
+ qed
+ hence homomorphism : "group_hom (G\<lparr>carrier := H\<rparr>) (G\<lparr>carrier := N <#> H\<rparr> Mod N) f"
+ unfolding group_hom_def group_hom_axioms_def using subgroup_imp_group[OF assms(1)]
+ normal.factorgroup_is_group[OF normal_in_normal_set_mult[OF assms(2) assms(1)]] by auto
+ moreover have im_f : "(f ` carrier(G\<lparr>carrier:=H\<rparr>)) = carrier(G\<lparr>carrier := N <#> H\<rparr> Mod N)"
+ proof
+ show "f ` carrier (G\<lparr>carrier := H\<rparr>) \<subseteq> carrier (G\<lparr>carrier := N <#> H\<rparr> Mod N)"
+ using op_hom unfolding hom_def using funcset_image by blast
+ next
+ show "carrier (G\<lparr>carrier := N <#> H\<rparr> Mod N) \<subseteq> f ` carrier (G\<lparr>carrier := H\<rparr>)"
+ proof
+ fix x assume p : " x \<in> carrier (G\<lparr>carrier := N <#> H\<rparr> Mod N)"
+ hence "x \<in> \<Union>{y. \<exists>x\<in>carrier (G\<lparr>carrier := N <#> H\<rparr>). y = {N #>\<^bsub>G\<lparr>carrier := N <#> H\<rparr>\<^esub> x}}"
+ unfolding FactGroup_def RCOSETS_def by auto
+ hence hyp :"\<exists>y. \<exists>h\<in>carrier (G\<lparr>carrier := N <#> H\<rparr>). y = {N #>\<^bsub>G\<lparr>carrier := N <#> H\<rparr>\<^esub> h} \<and> x \<in> y"
+ using Union_iff by blast
+ from hyp obtain nh where nhNH:"nh \<in>carrier (G\<lparr>carrier := N <#> H\<rparr>)"
+ and "x \<in> {N #>\<^bsub>G\<lparr>carrier := N <#> H\<rparr>\<^esub> nh}"
+ by blast
+ hence K: "x = (#>\<^bsub>G\<lparr>carrier := N <#> H\<rparr>\<^esub>) N nh" by simp
+ have "nh \<in> N <#> H" using nhNH by simp
+ from this obtain n h where nN : "n \<in> N" and hH : " h \<in> H" and nhnh: "n \<otimes> h = nh"
+ unfolding set_mult_def by blast
+ have "x = (#>\<^bsub>G\<lparr>carrier := N <#> H\<rparr>\<^esub>) N (n \<otimes> h)" using K nhnh by simp
+ hence "x = (#>) N (n \<otimes> h)" using K nhnh unfolding r_coset_def by auto
+ also have "... = (N #> n) #>h"
+ using coset_mult_assoc hH nN assms subgroup_imp_subset normal_imp_subgroup
+ by (metis subgroup.mem_carrier)
+ finally have "x = (#>) N h"
+ using coset_join2[of n N] nN assms by (simp add: normal_imp_subgroup subgroup.mem_carrier)
+ thus "x \<in> f ` carrier (G\<lparr>carrier := H\<rparr>)" using hH unfolding f_def by simp
+ qed
+ qed
+ moreover have ker_f :"kernel (G\<lparr>carrier := H\<rparr>) (G\<lparr>carrier := N<#>H\<rparr> Mod N) f = N\<inter>H"
+ unfolding kernel_def f_def
+ proof-
+ have "{x \<in> carrier (G\<lparr>carrier := H\<rparr>). N #> x = \<one>\<^bsub>G\<lparr>carrier := N <#> H\<rparr> Mod N\<^esub>} =
+ {x \<in> carrier (G\<lparr>carrier := H\<rparr>). N #> x = N}" unfolding FactGroup_def by simp
+ also have "... = {x \<in> carrier (G\<lparr>carrier := H\<rparr>). x \<in> N}"
+ using coset_join1
+ by (metis (no_types, lifting) assms group.subgroup_self incl_subgroup is_group
+ normal_imp_subgroup subgroup.mem_carrier subgroup.rcos_const subgroup_imp_group)
+ also have "... =N \<inter> (carrier(G\<lparr>carrier := H\<rparr>))" by auto
+ finally show "{x \<in> carrier (G\<lparr>carrier := H\<rparr>). N#>x = \<one>\<^bsub>G\<lparr>carrier := N <#> H\<rparr> Mod N\<^esub>} = N \<inter> H"
+ by simp
+ qed
+ ultimately have "(G\<lparr>carrier := H\<rparr> Mod N \<inter> H) \<cong> (G\<lparr>carrier := N <#> H\<rparr> Mod N)"
+ using group_hom.FactGroup_iso[OF homomorphism im_f] by auto
+ hence "G\<lparr>carrier := N <#> H\<rparr> Mod N \<cong> G\<lparr>carrier := H\<rparr> Mod N \<inter> H"
+ by (simp add: group.iso_sym assms normal.factorgroup_is_group normal_inter_subgroup)
+ thus "G\<lparr>carrier := N <#> H\<rparr> Mod N \<cong> G\<lparr>carrier := H\<rparr> Mod N \<inter> H" by auto
+qed
+
+
+theorem (in group) snd_iso_thme:
+ assumes "subgroup H G"
+ and "subgroup N G"
+ and "subgroup H (G\<lparr>carrier:= (normalizer G N)\<rparr>)"
+ shows "(G\<lparr>carrier:= N<#>H\<rparr> Mod N) \<cong> (G\<lparr>carrier:= H\<rparr> Mod (H\<inter>N))"
+proof-
+ have "G\<lparr>carrier := normalizer G N, carrier := H\<rparr>
+ = G\<lparr>carrier := H\<rparr>" by simp
+ hence "G\<lparr>carrier := normalizer G N, carrier := H\<rparr> Mod N \<inter> H =
+ G\<lparr>carrier := H\<rparr> Mod N \<inter> H" by auto
+ moreover have "G\<lparr>carrier := normalizer G N,
+ carrier := N <#>\<^bsub>G\<lparr>carrier := normalizer G N\<rparr>\<^esub> H\<rparr> =
+ G\<lparr>carrier := N <#>\<^bsub>G\<lparr>carrier := normalizer G N\<rparr>\<^esub> H\<rparr>" by simp
+ hence "G\<lparr>carrier := normalizer G N,
+ carrier := N <#>\<^bsub>G\<lparr>carrier := normalizer G N\<rparr>\<^esub> H\<rparr> Mod N =
+ G\<lparr>carrier := N <#>\<^bsub>G\<lparr>carrier := normalizer G N\<rparr>\<^esub> H\<rparr> Mod N" by auto
+ hence "G\<lparr>carrier := normalizer G N,
+ carrier := N <#>\<^bsub>G\<lparr>carrier := normalizer G N\<rparr>\<^esub> H\<rparr> Mod N \<cong>
+ G\<lparr>carrier := normalizer G N, carrier := H\<rparr> Mod N \<inter> H =
+ (G\<lparr>carrier:= N<#>H\<rparr> Mod N) \<cong>
+ G\<lparr>carrier := normalizer G N, carrier := H\<rparr> Mod N \<inter> H"
+ using subgroup_set_mult_equality[OF normalizer_imp_subgroup[OF subgroup_imp_subset[OF assms(2)]], of N H]
+ subgroup_imp_subset[OF assms(3)]
+ subgroup_imp_subset[OF normal_imp_subgroup[OF subgroup_in_normalizer[OF assms(2)]]]
+ by simp
+ ultimately have "G\<lparr>carrier := normalizer G N,
+ carrier := N <#>\<^bsub>G\<lparr>carrier := normalizer G N\<rparr>\<^esub> H\<rparr> Mod N \<cong>
+ G\<lparr>carrier := normalizer G N, carrier := H\<rparr> Mod N \<inter> H =
+ (G\<lparr>carrier:= N<#>H\<rparr> Mod N) \<cong> G\<lparr>carrier := H\<rparr> Mod N \<inter> H" by auto
+ moreover have "G\<lparr>carrier := normalizer G N,
+ carrier := N <#>\<^bsub>G\<lparr>carrier := normalizer G N\<rparr>\<^esub> H\<rparr> Mod N \<cong>
+ G\<lparr>carrier := normalizer G N, carrier := H\<rparr> Mod N \<inter> H"
+ using group.weak_snd_iso_thme[OF subgroup_imp_group[OF normalizer_imp_subgroup[OF
+ subgroup_imp_subset[OF assms(2)]]] assms(3) subgroup_in_normalizer[OF assms(2)]]
+ by simp
+ moreover have "H\<inter>N = N\<inter>H" using assms by auto
+ ultimately show "(G\<lparr>carrier:= N<#>H\<rparr> Mod N) \<cong> G\<lparr>carrier := H\<rparr> Mod H \<inter> N" by auto
+qed
+
+
+corollary (in group) snd_iso_thme_recip :
+ assumes "subgroup H G"
+ and "subgroup N G"
+ and "subgroup H (G\<lparr>carrier:= (normalizer G N)\<rparr>)"
+ shows "(G\<lparr>carrier:= H<#>N\<rparr> Mod N) \<cong> (G\<lparr>carrier:= H\<rparr> Mod (H\<inter>N))"
+ by (metis assms commut_normal_subgroup group.subgroup_in_normalizer is_group subgroup_imp_subset
+ normalizer_imp_subgroup snd_iso_thme)
+
+
+subsection\<open>The Zassenhaus Lemma\<close>
+
+
+lemma (in group) distinc:
+ assumes "subgroup H G"
+ and "H1\<lhd>G\<lparr>carrier := H\<rparr>"
+ and "subgroup K G"
+ and "K1\<lhd>G\<lparr>carrier:=K\<rparr>"
+ shows "subgroup (H\<inter>K) (G\<lparr>carrier:=(normalizer G (H1<#>(H\<inter>K1))) \<rparr>)"
+proof (intro subgroup_incl[OF subgroups_Inter_pair[OF assms(1) assms(3)]])
+ show "subgroup (normalizer G (H1 <#> H \<inter> K1)) G"
+ using normalizer_imp_subgroup assms normal_imp_subgroup subgroup_imp_subset
+ by (metis group.incl_subgroup is_group setmult_subset_G subgroups_Inter_pair)
+next
+ show "H \<inter> K \<subseteq> normalizer G (H1 <#> H \<inter> K1)" unfolding normalizer_def stabilizer_def
+ proof
+ fix x assume xHK : "x \<in> H \<inter> K"
+ hence xG : "{x} \<subseteq> carrier G" "{inv x} \<subseteq> carrier G"
+ using subgroup_imp_subset assms inv_closed xHK by auto
+ have allG : "H \<subseteq> carrier G" "K \<subseteq> carrier G" "H1 \<subseteq> carrier G" "K1 \<subseteq> carrier G"
+ using assms subgroup_imp_subset normal_imp_subgroup incl_subgroup apply blast+ .
+ have HK1_normal: "H\<inter>K1 \<lhd> (G\<lparr>carrier := H \<inter> K\<rparr>)" using normal_inter[OF assms(3)assms(1)assms(4)]
+ by (simp add : inf_commute)
+ have "H \<inter> K \<subseteq> normalizer G (H \<inter> K1)"
+ using subgroup_imp_subset[OF normal_imp_subgroup_normalizer[OF subgroups_Inter_pair[OF
+ assms(1)assms(3)]HK1_normal]] by auto
+ hence "x <# (H \<inter> K1) #> inv x = (H \<inter> K1)"
+ using xHK subgroup_imp_subset[OF subgroups_Inter_pair[OF assms(1) incl_subgroup[OF assms(3)
+ normal_imp_subgroup[OF assms(4)]]]]
+ unfolding normalizer_def stabilizer_def by auto
+ moreover have "H \<subseteq> normalizer G H1"
+ using subgroup_imp_subset[OF normal_imp_subgroup_normalizer[OF assms(1)assms(2)]] by auto
+ hence "x <# H1 #> inv x = H1"
+ using xHK subgroup_imp_subset[OF incl_subgroup[OF assms(1) normal_imp_subgroup[OF assms(2)]]]
+ unfolding normalizer_def stabilizer_def by auto
+ ultimately have "H1 <#> H \<inter> K1 = (x <# H1 #> inv x) <#> (x <# H \<inter> K1 #> inv x)" by auto
+ also have "... = ({x} <#> H1) <#> {inv x} <#> ({x} <#> H \<inter> K1 <#> {inv x})"
+ by (simp add : r_coset_eq_set_mult l_coset_eq_set_mult)
+ also have "... = ({x} <#> H1 <#> {inv x} <#> {x}) <#> (H \<inter> K1 <#> {inv x})"
+ by (smt Int_lower1 allG xG set_mult_assoc subset_trans setmult_subset_G)
+ also have "... = ({x} <#> H1 <#> {\<one>}) <#> (H \<inter> K1 <#> {inv x})"
+ using allG xG coset_mult_assoc by (simp add: r_coset_eq_set_mult setmult_subset_G)
+ also have "... =({x} <#> H1) <#> (H \<inter> K1 <#> {inv x})"
+ using coset_mult_one r_coset_eq_set_mult[of G H1 \<one>] set_mult_assoc[OF xG(1) allG(3)] allG
+ by auto
+ also have "... = {x} <#> (H1 <#> H \<inter> K1) <#> {inv x}"
+ using allG xG set_mult_assoc setmult_subset_G by (metis inf.coboundedI2)
+ finally have "H1 <#> H \<inter> K1 = x <# (H1 <#> H \<inter> K1) #> inv x"
+ using xG setmult_subset_G allG by (simp add: l_coset_eq_set_mult r_coset_eq_set_mult)
+ thus "x \<in> {g \<in> carrier G. (\<lambda>H\<in>{H. H \<subseteq> carrier G}. g <# H #> inv g) (H1 <#> H \<inter> K1)
+ = H1 <#> H \<inter> K1}"
+ using xG allG setmult_subset_G[OF allG(3), where ?K = "H\<inter>K1"] xHK
+ by auto
+ qed
+qed
+
+lemma (in group) preliminary1:
+ assumes "subgroup H G"
+ and "H1\<lhd>G\<lparr>carrier := H\<rparr>"
+ and "subgroup K G"
+ and "K1\<lhd>G\<lparr>carrier:=K\<rparr>"
+ shows " (H\<inter>K) \<inter> (H1<#>(H\<inter>K1)) = (H1\<inter>K)<#>(H\<inter>K1)"
+proof
+ have all_inclG : "H \<subseteq> carrier G" "H1 \<subseteq> carrier G" "K \<subseteq> carrier G" "K1 \<subseteq> carrier G"
+ using assms subgroup_imp_subset normal_imp_subgroup incl_subgroup apply blast+.
+ show "H \<inter> K \<inter> (H1 <#> H \<inter> K1) \<subseteq> H1 \<inter> K <#> H \<inter> K1"
+ proof
+ fix x assume x_def : "x \<in> (H \<inter> K) \<inter> (H1 <#> (H \<inter> K1))"
+ from x_def have x_incl : "x \<in> H" "x \<in> K" "x \<in> (H1 <#> (H \<inter> K1))" by auto
+ then obtain h1 hk1 where h1hk1_def : "h1 \<in> H1" "hk1 \<in> H \<inter> K1" "h1 \<otimes> hk1 = x"
+ using assms unfolding set_mult_def by blast
+ hence "hk1 \<in> H \<inter> K" using subgroup_imp_subset[OF normal_imp_subgroup[OF assms(4)]] by auto
+ hence "inv hk1 \<in> H \<inter> K" using subgroup.m_inv_closed[OF subgroups_Inter_pair] assms by auto
+ moreover have "h1 \<otimes> hk1 \<in> H \<inter> K" using x_incl h1hk1_def by auto
+ ultimately have "h1 \<otimes> hk1 \<otimes> inv hk1 \<in> H \<inter> K"
+ using subgroup.m_closed[OF subgroups_Inter_pair] assms by auto
+ hence "h1 \<in> H \<inter> K" using h1hk1_def assms subgroup_imp_subset incl_subgroup normal_imp_subgroup
+ by (metis Int_iff contra_subsetD inv_solve_right m_closed)
+ hence "h1 \<in> H1 \<inter> H \<inter> K" using h1hk1_def by auto
+ hence "h1 \<in> H1 \<inter> K" using subgroup_imp_subset[OF normal_imp_subgroup[OF assms(2)]] by auto
+ hence "h1 \<otimes> hk1 \<in> (H1\<inter>K)<#>(H\<inter>K1)"
+ using h1hk1_def unfolding set_mult_def by auto
+ thus " x \<in> (H1\<inter>K)<#>(H\<inter>K1)" using h1hk1_def x_def by auto
+ qed
+ show "H1 \<inter> K <#> H \<inter> K1 \<subseteq> H \<inter> K \<inter> (H1 <#> H \<inter> K1)"
+ proof-
+ have "H1 \<inter> K \<subseteq> H \<inter> K" using subgroup_imp_subset[OF normal_imp_subgroup[OF assms(2)]] by auto
+ moreover have "H \<inter> K1 \<subseteq> H \<inter> K"
+ using subgroup_imp_subset[OF normal_imp_subgroup[OF assms(4)]] by auto
+ ultimately have "H1 \<inter> K <#> H \<inter> K1 \<subseteq> H \<inter> K" unfolding set_mult_def
+ using subgroup.m_closed[OF subgroups_Inter_pair [OF assms(1)assms(3)]] by blast
+ moreover have "H1 \<inter> K \<subseteq> H1" by auto
+ hence "H1 \<inter> K <#> H \<inter> K1 \<subseteq> (H1 <#> H \<inter> K1)" unfolding set_mult_def by auto
+ ultimately show "H1 \<inter> K <#> H \<inter> K1 \<subseteq> H \<inter> K \<inter> (H1 <#> H \<inter> K1)" by auto
+ qed
+qed
+
+lemma (in group) preliminary2:
+ assumes "subgroup H G"
+ and "H1\<lhd>G\<lparr>carrier := H\<rparr>"
+ and "subgroup K G"
+ and "K1\<lhd>G\<lparr>carrier:=K\<rparr>"
+ shows "(H1<#>(H\<inter>K1)) \<lhd> G\<lparr>carrier:=(H1<#>(H\<inter>K))\<rparr>"
+proof-
+ have all_inclG : "H \<subseteq> carrier G" "H1 \<subseteq> carrier G" "K \<subseteq> carrier G" "K1 \<subseteq> carrier G"
+ using assms subgroup_imp_subset normal_imp_subgroup incl_subgroup apply blast+.
+ have subH1:"subgroup (H1 <#> H \<inter> K) (G\<lparr>carrier := H\<rparr>)"
+ using mult_norm_sub_in_sub[OF assms(2)subgroup_incl[OF subgroups_Inter_pair[OF assms(1)assms(3)]
+ assms(1)]] assms by auto
+ have "Group.group (G\<lparr>carrier:=(H1<#>(H\<inter>K))\<rparr>)"
+ using subgroup_imp_group[OF incl_subgroup[OF assms(1) subH1]].
+ moreover have subH2 : "subgroup (H1 <#> H \<inter> K1) (G\<lparr>carrier := H\<rparr>)"
+ using mult_norm_sub_in_sub[OF assms(2) subgroup_incl[OF subgroups_Inter_pair[OF
+ assms(1) incl_subgroup[OF assms(3)normal_imp_subgroup[OF assms(4)]]]]] assms by auto
+ hence "(H\<inter>K1) \<subseteq> (H\<inter>K)"
+ using assms subgroup_imp_subset normal_imp_subgroup monoid.cases_scheme
+ by (metis inf.mono partial_object.simps(1) partial_object.update_convs(1) subset_refl)
+ hence incl:"(H1<#>(H\<inter>K1)) \<subseteq> H1<#>(H\<inter>K)" using assms subgroup_imp_subset normal_imp_subgroup
+ unfolding set_mult_def by blast
+ hence "subgroup (H1 <#> H \<inter> K1) (G\<lparr>carrier := (H1<#>(H\<inter>K))\<rparr>)"
+ using assms subgroup_incl[OF incl_subgroup[OF assms(1)subH2]incl_subgroup[OF assms(1)
+ subH1]] normal_imp_subgroup subgroup_imp_subset unfolding set_mult_def by blast
+ moreover have " (\<And> x. x\<in>carrier (G\<lparr>carrier := H1 <#> H \<inter> K\<rparr>) \<Longrightarrow>
+ H1 <#> H\<inter>K1 #>\<^bsub>G\<lparr>carrier := H1 <#> H\<inter>K\<rparr>\<^esub> x = x <#\<^bsub>G\<lparr>carrier := H1 <#> H\<inter>K\<rparr>\<^esub> (H1 <#> H\<inter>K1))"
+ proof-
+ fix x assume "x \<in>carrier (G\<lparr>carrier := H1 <#> H \<inter> K\<rparr>)"
+ hence x_def : "x \<in> H1 <#> H \<inter> K" by simp
+ from this obtain h1 hk where h1hk_def :"h1 \<in> H1" "hk \<in> H \<inter> K" "h1 \<otimes> hk = x"
+ unfolding set_mult_def by blast
+ have xH : "x \<in> H" using subgroup_imp_subset[OF subH1] using x_def by auto
+ hence allG : "h1 \<in> carrier G" "hk \<in> carrier G" "x \<in> carrier G"
+ using assms subgroup_imp_subset h1hk_def normal_imp_subgroup incl_subgroup apply blast+.
+ hence "x <#\<^bsub>G\<lparr>carrier := H1 <#> H\<inter>K\<rparr>\<^esub> (H1 <#> H\<inter>K1) =h1 \<otimes> hk <# (H1 <#> H\<inter>K1)"
+ using subgroup_set_mult_equality subgroup_imp_subset xH h1hk_def by (simp add: l_coset_def)
+ also have "... = h1 <# (hk <# (H1 <#> H\<inter>K1))"
+ using lcos_m_assoc[OF subgroup_imp_subset[OF incl_subgroup[OF assms(1) subH1]]allG(1)allG(2)]
+ by (metis allG(1) allG(2) assms(1) incl_subgroup lcos_m_assoc subH2 subgroup_imp_subset)
+ also have "... = h1 <# (hk <# H1 <#> H\<inter>K1)"
+ using set_mult_assoc all_inclG allG by (simp add: l_coset_eq_set_mult inf.coboundedI1)
+ also have "... = h1 <# (hk <# H1 #> \<one> <#> H\<inter>K1 #> \<one>)"
+ using coset_mult_one allG all_inclG l_coset_subset_G
+ by (smt inf_le2 setmult_subset_G subset_trans)
+ also have "... = h1 <# (hk <# H1 #> inv hk #> hk <#> H\<inter>K1 #> inv hk #> hk)"
+ using all_inclG allG coset_mult_assoc l_coset_subset_G
+ by (simp add: inf.coboundedI1 setmult_subset_G)
+ finally have "x <#\<^bsub>G\<lparr>carrier := H1 <#> H \<inter> K\<rparr>\<^esub> (H1 <#> H \<inter> K1) =
+ h1 <# ((hk <# H1 #> inv hk) <#> (hk <# H\<inter>K1 #> inv hk) #> hk)"
+ using rcos_assoc_lcos allG all_inclG
+ by (smt inf_le1 inv_closed l_coset_subset_G r_coset_subset_G setmult_rcos_assoc subset_trans)
+ moreover have "H \<subseteq> normalizer G H1"
+ using assms h1hk_def subgroup_imp_subset[OF normal_imp_subgroup_normalizer] by simp
+ hence "\<And>g. g \<in> H \<Longrightarrow> g \<in> {g \<in> carrier G. (\<lambda>H\<in>{H. H \<subseteq> carrier G}. g <# H #> inv g) H1 = H1}"
+ using all_inclG assms unfolding normalizer_def stabilizer_def by auto
+ hence "\<And>g. g \<in> H \<Longrightarrow> g <# H1 #> inv g = H1" using all_inclG by simp
+ hence "(hk <# H1 #> inv hk) = H1" using h1hk_def all_inclG by simp
+ moreover have "H\<inter>K \<subseteq> normalizer G (H\<inter>K1)"
+ using normal_inter[OF assms(3)assms(1)assms(4)] assms subgroups_Inter_pair
+ subgroup_imp_subset[OF normal_imp_subgroup_normalizer] by (simp add: inf_commute)
+ hence "\<And>g. g\<in>H\<inter>K \<Longrightarrow> g\<in>{g\<in>carrier G. (\<lambda>H\<in>{H. H \<subseteq> carrier G}. g <# H #> inv g) (H\<inter>K1) = H\<inter>K1}"
+ using all_inclG assms unfolding normalizer_def stabilizer_def by auto
+ hence "\<And>g. g \<in> H\<inter>K \<Longrightarrow> g <# (H\<inter>K1) #> inv g = H\<inter>K1"
+ using subgroup_imp_subset[OF subgroups_Inter_pair[OF assms(1) incl_subgroup[OF
+ assms(3)normal_imp_subgroup[OF assms(4)]]]] by auto
+ hence "(hk <# H\<inter>K1 #> inv hk) = H\<inter>K1" using h1hk_def by simp
+ ultimately have "x <#\<^bsub>G\<lparr>carrier := H1 <#> H \<inter> K\<rparr>\<^esub> (H1 <#> H \<inter> K1) = h1 <#(H1 <#> (H \<inter> K1)#> hk)"
+ by auto
+ also have "... = h1 <# H1 <#> ((H \<inter> K1)#> hk)"
+ using set_mult_assoc[where ?M = "{h1}" and ?H = "H1" and ?K = "(H \<inter> K1)#> hk"] allG all_inclG
+ by (simp add: l_coset_eq_set_mult inf.coboundedI2 r_coset_subset_G setmult_rcos_assoc)
+ also have "... = H1 <#> ((H \<inter> K1)#> hk)"
+ using coset_join3 allG incl_subgroup[OF assms(1)normal_imp_subgroup[OF assms(2)]] h1hk_def
+ by auto
+ finally have eq1 : "x <#\<^bsub>G\<lparr>carrier := H1 <#> H \<inter> K\<rparr>\<^esub> (H1 <#> H \<inter> K1) = H1 <#> (H \<inter> K1) #> hk"
+ by (simp add: allG(2) all_inclG inf.coboundedI2 setmult_rcos_assoc)
+ have "H1 <#> H \<inter> K1 #>\<^bsub>G\<lparr>carrier := H1 <#> H \<inter> K\<rparr>\<^esub> x = H1 <#> H \<inter> K1 #> (h1 \<otimes> hk)"
+ using subgroup_set_mult_equality subgroup_imp_subset xH h1hk_def by (simp add: r_coset_def)
+ also have "... = H1 <#> H \<inter> K1 #> h1 #> hk"
+ using coset_mult_assoc by (simp add: allG all_inclG inf.coboundedI2 setmult_subset_G)
+ also have"... = H \<inter> K1 <#> H1 #> h1 #> hk"
+ using commut_normal_subgroup[OF assms(1)assms(2)subgroup_incl[OF subgroups_Inter_pair[OF
+ assms(1)incl_subgroup[OF assms(3)normal_imp_subgroup[OF assms(4)]]]assms(1)]] by simp
+ also have "... = H \<inter> K1 <#> H1 #> hk"
+ using coset_join2[OF allG(1)incl_subgroup[OF assms(1)normal_imp_subgroup]
+ h1hk_def(1)] all_inclG allG assms by (metis inf.coboundedI2 setmult_rcos_assoc)
+ finally have "H1 <#> H \<inter> K1 #>\<^bsub>G\<lparr>carrier := H1 <#> H \<inter> K\<rparr>\<^esub> x =H1 <#> H \<inter> K1 #> hk"
+ using commut_normal_subgroup[OF assms(1)assms(2)subgroup_incl[OF subgroups_Inter_pair[OF
+ assms(1)incl_subgroup[OF assms(3)normal_imp_subgroup[OF assms(4)]]]assms(1)]] by simp
+ thus " H1 <#> H \<inter> K1 #>\<^bsub>G\<lparr>carrier := H1 <#> H \<inter> K\<rparr>\<^esub> x =
+ x <#\<^bsub>G\<lparr>carrier := H1 <#> H \<inter> K\<rparr>\<^esub> (H1 <#> H \<inter> K1)" using eq1 by simp
+ qed
+ ultimately show "H1 <#> H \<inter> K1 \<lhd> G\<lparr>carrier := H1 <#> H \<inter> K\<rparr>"
+ unfolding normal_def normal_axioms_def by auto
+qed
+
+
+proposition (in group) Zassenhaus_1:
+ assumes "subgroup H G"
+ and "H1\<lhd>G\<lparr>carrier := H\<rparr>"
+ and "subgroup K G"
+ and "K1\<lhd>G\<lparr>carrier:=K\<rparr>"
+ shows "(G\<lparr>carrier:= H1 <#> (H\<inter>K)\<rparr> Mod (H1<#>H\<inter>K1)) \<cong> (G\<lparr>carrier:= (H\<inter>K)\<rparr> Mod ((H1\<inter>K)<#>(H\<inter>K1)))"
+proof-
+ define N and N1 where "N = (H\<inter>K)" and "N1 =H1<#>(H\<inter>K1)"
+ have normal_N_N1 : "subgroup N (G\<lparr>carrier:=(normalizer G N1)\<rparr>)"
+ by (simp add: N1_def N_def assms distinc normal_imp_subgroup)
+ have Hp:"(G\<lparr>carrier:= N<#>N1\<rparr> Mod N1) \<cong> (G\<lparr>carrier:= N\<rparr> Mod (N\<inter>N1))"
+ by (metis N1_def N_def assms incl_subgroup inf_le1 mult_norm_sub_in_sub
+ normal_N_N1 normal_imp_subgroup snd_iso_thme_recip subgroup_incl subgroups_Inter_pair)
+ have H_simp: "N<#>N1 = H1<#> (H\<inter>K)"
+ proof-
+ have H1_incl_G : "H1 \<subseteq> carrier G"
+ using assms normal_imp_subgroup incl_subgroup subgroup_imp_subset by blast
+ have K1_incl_G :"K1 \<subseteq> carrier G"
+ using assms normal_imp_subgroup incl_subgroup subgroup_imp_subset by blast
+ have "N<#>N1= (H\<inter>K)<#> (H1<#>(H\<inter>K1))" by (auto simp add: N_def N1_def)
+ also have "... = ((H\<inter>K)<#>H1) <#>(H\<inter>K1)"
+ using set_mult_assoc[where ?M = "H\<inter>K"] K1_incl_G H1_incl_G assms
+ by (simp add: inf.coboundedI2 subgroup_imp_subset)
+ also have "... = (H1<#>(H\<inter>K))<#>(H\<inter>K1)"
+ using commut_normal_subgroup assms subgroup_incl subgroups_Inter_pair by auto
+ also have "... = H1 <#> ((H\<inter>K)<#>(H\<inter>K1))"
+ using set_mult_assoc K1_incl_G H1_incl_G assms
+ by (simp add: inf.coboundedI2 subgroup_imp_subset)
+ also have " ((H\<inter>K)<#>(H\<inter>K1)) = (H\<inter>K)"
+ proof (intro set_mult_subgroup_idem[where ?H = "H\<inter>K" and ?N="H\<inter>K1",
+ OF subgroups_Inter_pair[OF assms(1) assms(3)]])
+ show "subgroup (H \<inter> K1) (G\<lparr>carrier := H \<inter> K\<rparr>)"
+ using subgroup_incl[where ?I = "H\<inter>K1" and ?J = "H\<inter>K",OF subgroups_Inter_pair[OF assms(1)
+ incl_subgroup[OF assms(3) normal_imp_subgroup]] subgroups_Inter_pair] assms
+ normal_imp_subgroup by (metis inf_commute normal_inter)
+ qed
+ hence " H1 <#> ((H\<inter>K)<#>(H\<inter>K1)) = H1 <#> ((H\<inter>K))"
+ by simp
+ thus "N <#> N1 = H1 <#> H \<inter> K"
+ by (simp add: calculation)
+ qed
+
+ have "N\<inter>N1 = (H1\<inter>K)<#>(H\<inter>K1)"
+ using preliminary1 assms N_def N1_def by simp
+ thus "(G\<lparr>carrier:= H1 <#> (H\<inter>K)\<rparr> Mod N1) \<cong> (G\<lparr>carrier:= N\<rparr> Mod ((H1\<inter>K)<#>(H\<inter>K1)))"
+ using H_simp Hp by auto
+qed
+
+
+theorem (in group) Zassenhaus:
+ assumes "subgroup H G"
+ and "H1\<lhd>G\<lparr>carrier := H\<rparr>"
+ and "subgroup K G"
+ and "K1\<lhd>G\<lparr>carrier:=K\<rparr>"
+ shows "(G\<lparr>carrier:= H1 <#> (H\<inter>K)\<rparr> Mod (H1<#>(H\<inter>K1))) \<cong>
+ (G\<lparr>carrier:= K1 <#> (H\<inter>K)\<rparr> Mod (K1<#>(K\<inter>H1)))"
+proof-
+ define Gmod1 Gmod2 Gmod3 Gmod4
+ where "Gmod1 = (G\<lparr>carrier:= H1 <#> (H\<inter>K)\<rparr> Mod (H1<#>(H\<inter>K1))) "
+ and "Gmod2 = (G\<lparr>carrier:= K1 <#> (K\<inter>H)\<rparr> Mod (K1<#>(K\<inter>H1)))"
+ and "Gmod3 = (G\<lparr>carrier:= (H\<inter>K)\<rparr> Mod ((H1\<inter>K)<#>(H\<inter>K1)))"
+ and "Gmod4 = (G\<lparr>carrier:= (K\<inter>H)\<rparr> Mod ((K1\<inter>H)<#>(K\<inter>H1)))"
+ have Hyp : "Gmod1 \<cong> Gmod3" "Gmod2 \<cong> Gmod4"
+ using Zassenhaus_1 assms Gmod1_def Gmod2_def Gmod3_def Gmod4_def by auto
+ have Hp : "Gmod3 = G\<lparr>carrier:= (K\<inter>H)\<rparr> Mod ((K\<inter>H1)<#>(K1\<inter>H))"
+ by (simp add: Gmod3_def inf_commute)
+ have "(K\<inter>H1)<#>(K1\<inter>H) = (K1\<inter>H)<#>(K\<inter>H1)"
+ proof (intro commut_normal_subgroup[OF subgroups_Inter_pair[OF assms(1)assms(3)]])
+ show "K1 \<inter> H \<lhd> G\<lparr>carrier := H \<inter> K\<rparr>"
+ using normal_inter[OF assms(3)assms(1)assms(4)] by (simp add: inf_commute)
+ next
+ show "subgroup (K \<inter> H1) (G\<lparr>carrier := H \<inter> K\<rparr>)"
+ using subgroup_incl by (simp add: assms inf_commute normal_imp_subgroup normal_inter)
+ qed
+ hence "Gmod3 = Gmod4" using Hp Gmod4_def by simp
+ hence "Gmod1 \<cong> Gmod2"
+ using group.iso_sym group.iso_trans Hyp normal.factorgroup_is_group
+ by (metis assms Gmod1_def Gmod2_def preliminary2)
+ thus ?thesis using Gmod1_def Gmod2_def by (simp add: inf_commute)
+qed
+
+end
--- a/src/HOL/ROOT Thu Jun 14 10:51:12 2018 +0200
+++ b/src/HOL/ROOT Thu Jun 14 14:23:48 2018 +0100
@@ -294,20 +294,16 @@
theories
(* Orders and Lattices *)
Galois_Connection (* Knaster-Tarski theorem and Galois connections *)
-
(* Groups *)
FiniteProduct (* Product operator for commutative groups *)
Sylow (* Sylow's theorem *)
Bij (* Automorphism Groups *)
- More_Group
- More_Finite_Product
Multiplicative_Group
-
+ Zassenhaus (* The Zassenhaus lemma *)
(* Rings *)
Divisibility (* Rings *)
IntRing (* Ideals and residue classes *)
UnivPoly (* Polynomials *)
- More_Ring
document_files "root.bib" "root.tex"
session "HOL-Auth" (timing) in Auth = "HOL-Library" +