(* Title: HOL/Algebra/Group.thy Id: $Id: Group.thy,v 1.38 2008/05/07 08:56:58 berghofe Exp $ Author: Clemens Ballarin, started 4 February 2003 Based on work by Florian Kammueller, L C Paulson and Markus Wenzel. *) theory Group imports FuncSet Lattice begin section {* Monoids and Groups *} subsection {* Definitions *} text {* Definitions follow \cite{Jacobson:1985}. *} record 'a monoid = "'a partial_object" + mult :: "['a, 'a] => 'a" (infixl "⊗\<index>" 70) one :: 'a ("\<one>\<index>") constdefs (structure G) m_inv :: "('a, 'b) monoid_scheme => 'a => 'a" ("inv\<index> _" [81] 80) "inv x == (THE y. y ∈ carrier G & x ⊗ y = \<one> & y ⊗ x = \<one>)" Units :: "_ => 'a set" --{*The set of invertible elements*} "Units G == {y. y ∈ carrier G & (∃x ∈ carrier G. x ⊗ y = \<one> & y ⊗ x = \<one>)}" consts pow :: "[('a, 'm) monoid_scheme, 'a, 'b::number] => 'a" (infixr "'(^')\<index>" 75) defs (overloaded) nat_pow_def: "pow G a n == nat_rec \<one>G (%u b. b ⊗G a) n" int_pow_def: "pow G a z == let p = nat_rec \<one>G (%u b. b ⊗G a) in if neg z then invG (p (nat (-z))) else p (nat z)" locale monoid = fixes G (structure) assumes m_closed [intro, simp]: "[|x ∈ carrier G; y ∈ carrier G|] ==> x ⊗ y ∈ carrier G" and m_assoc: "[|x ∈ carrier G; y ∈ carrier G; z ∈ carrier G|] ==> (x ⊗ y) ⊗ z = x ⊗ (y ⊗ z)" and one_closed [intro, simp]: "\<one> ∈ carrier G" and l_one [simp]: "x ∈ carrier G ==> \<one> ⊗ x = x" and r_one [simp]: "x ∈ carrier G ==> x ⊗ \<one> = x" lemma monoidI: fixes G (structure) assumes m_closed: "!!x y. [| x ∈ carrier G; y ∈ carrier G |] ==> x ⊗ y ∈ carrier G" and one_closed: "\<one> ∈ carrier G" and m_assoc: "!!x y z. [| x ∈ carrier G; y ∈ carrier G; z ∈ carrier G |] ==> (x ⊗ y) ⊗ z = x ⊗ (y ⊗ z)" and l_one: "!!x. x ∈ carrier G ==> \<one> ⊗ x = x" and r_one: "!!x. x ∈ carrier G ==> x ⊗ \<one> = x" shows "monoid G" by (fast intro!: monoid.intro intro: prems) lemma (in monoid) Units_closed [dest]: "x ∈ Units G ==> x ∈ carrier G" by (unfold Units_def) fast lemma (in monoid) inv_unique: assumes eq: "y ⊗ x = \<one>" "x ⊗ y' = \<one>" and G: "x ∈ carrier G" "y ∈ carrier G" "y' ∈ carrier G" shows "y = y'" proof - from G eq have "y = y ⊗ (x ⊗ y')" by simp also from G have "... = (y ⊗ x) ⊗ y'" by (simp add: m_assoc) also from G eq have "... = y'" by simp finally show ?thesis . qed lemma (in monoid) Units_one_closed [intro, simp]: "\<one> ∈ Units G" by (unfold Units_def) auto lemma (in monoid) Units_inv_closed [intro, simp]: "x ∈ Units G ==> inv x ∈ carrier G" apply (unfold Units_def m_inv_def, auto) apply (rule theI2, fast) apply (fast intro: inv_unique, fast) done lemma (in monoid) Units_l_inv_ex: "x ∈ Units G ==> ∃y ∈ carrier G. y ⊗ x = \<one>" by (unfold Units_def) auto lemma (in monoid) Units_r_inv_ex: "x ∈ Units G ==> ∃y ∈ carrier G. x ⊗ y = \<one>" by (unfold Units_def) auto lemma (in monoid) Units_l_inv: "x ∈ Units G ==> inv x ⊗ x = \<one>" apply (unfold Units_def m_inv_def, auto) apply (rule theI2, fast) apply (fast intro: inv_unique, fast) done lemma (in monoid) Units_r_inv: "x ∈ Units G ==> x ⊗ inv x = \<one>" apply (unfold Units_def m_inv_def, auto) apply (rule theI2, fast) apply (fast intro: inv_unique, fast) done lemma (in monoid) Units_inv_Units [intro, simp]: "x ∈ Units G ==> inv x ∈ Units G" proof - assume x: "x ∈ Units G" show "inv x ∈ Units G" by (auto simp add: Units_def intro: Units_l_inv Units_r_inv x Units_closed [OF x]) qed lemma (in monoid) Units_l_cancel [simp]: "[| x ∈ Units G; y ∈ carrier G; z ∈ carrier G |] ==> (x ⊗ y = x ⊗ z) = (y = z)" proof assume eq: "x ⊗ y = x ⊗ z" and G: "x ∈ Units G" "y ∈ carrier G" "z ∈ carrier G" then have "(inv x ⊗ x) ⊗ y = (inv x ⊗ x) ⊗ z" by (simp add: m_assoc Units_closed) with G show "y = z" by (simp add: Units_l_inv) next assume eq: "y = z" and G: "x ∈ Units G" "y ∈ carrier G" "z ∈ carrier G" then show "x ⊗ y = x ⊗ z" by simp qed lemma (in monoid) Units_inv_inv [simp]: "x ∈ Units G ==> inv (inv x) = x" proof - assume x: "x ∈ Units G" then have "inv x ⊗ inv (inv x) = inv x ⊗ x" by (simp add: Units_l_inv Units_r_inv) with x show ?thesis by (simp add: Units_closed) qed lemma (in monoid) inv_inj_on_Units: "inj_on (m_inv G) (Units G)" proof (rule inj_onI) fix x y assume G: "x ∈ Units G" "y ∈ Units G" and eq: "inv x = inv y" then have "inv (inv x) = inv (inv y)" by simp with G show "x = y" by simp qed lemma (in monoid) Units_inv_comm: assumes inv: "x ⊗ y = \<one>" and G: "x ∈ Units G" "y ∈ Units G" shows "y ⊗ x = \<one>" proof - from G have "x ⊗ y ⊗ x = x ⊗ \<one>" by (auto simp add: inv Units_closed) with G show ?thesis by (simp del: r_one add: m_assoc Units_closed) qed text {* Power *} lemma (in monoid) nat_pow_closed [intro, simp]: "x ∈ carrier G ==> x (^) (n::nat) ∈ carrier G" by (induct n) (simp_all add: nat_pow_def) lemma (in monoid) nat_pow_0 [simp]: "x (^) (0::nat) = \<one>" by (simp add: nat_pow_def) lemma (in monoid) nat_pow_Suc [simp]: "x (^) (Suc n) = x (^) n ⊗ x" by (simp add: nat_pow_def) lemma (in monoid) nat_pow_one [simp]: "\<one> (^) (n::nat) = \<one>" by (induct n) simp_all lemma (in monoid) nat_pow_mult: "x ∈ carrier G ==> x (^) (n::nat) ⊗ x (^) m = x (^) (n + m)" by (induct m) (simp_all add: m_assoc [THEN sym]) lemma (in monoid) nat_pow_pow: "x ∈ carrier G ==> (x (^) n) (^) m = x (^) (n * m::nat)" by (induct m) (simp, simp add: nat_pow_mult add_commute) text {* A group is a monoid all of whose elements are invertible. *} locale group = monoid + assumes Units: "carrier G <= Units G" lemma (in group) is_group: "group G" by (rule group_axioms) theorem groupI: fixes G (structure) assumes m_closed [simp]: "!!x y. [| x ∈ carrier G; y ∈ carrier G |] ==> x ⊗ y ∈ carrier G" and one_closed [simp]: "\<one> ∈ carrier G" and m_assoc: "!!x y z. [| x ∈ carrier G; y ∈ carrier G; z ∈ carrier G |] ==> (x ⊗ y) ⊗ z = x ⊗ (y ⊗ z)" and l_one [simp]: "!!x. x ∈ carrier G ==> \<one> ⊗ x = x" and l_inv_ex: "!!x. x ∈ carrier G ==> ∃y ∈ carrier G. y ⊗ x = \<one>" shows "group G" proof - have l_cancel [simp]: "!!x y z. [| x ∈ carrier G; y ∈ carrier G; z ∈ carrier G |] ==> (x ⊗ y = x ⊗ z) = (y = z)" proof fix x y z assume eq: "x ⊗ y = x ⊗ z" and G: "x ∈ carrier G" "y ∈ carrier G" "z ∈ carrier G" with l_inv_ex obtain x_inv where xG: "x_inv ∈ carrier G" and l_inv: "x_inv ⊗ x = \<one>" by fast from G eq xG have "(x_inv ⊗ x) ⊗ y = (x_inv ⊗ x) ⊗ z" by (simp add: m_assoc) with G show "y = z" by (simp add: l_inv) next fix x y z assume eq: "y = z" and G: "x ∈ carrier G" "y ∈ carrier G" "z ∈ carrier G" then show "x ⊗ y = x ⊗ z" by simp qed have r_one: "!!x. x ∈ carrier G ==> x ⊗ \<one> = x" proof - fix x assume x: "x ∈ carrier G" with l_inv_ex obtain x_inv where xG: "x_inv ∈ carrier G" and l_inv: "x_inv ⊗ x = \<one>" by fast from x xG have "x_inv ⊗ (x ⊗ \<one>) = x_inv ⊗ x" by (simp add: m_assoc [symmetric] l_inv) with x xG show "x ⊗ \<one> = x" by simp qed have inv_ex: "!!x. x ∈ carrier G ==> ∃y ∈ carrier G. y ⊗ x = \<one> & x ⊗ y = \<one>" proof - fix x assume x: "x ∈ carrier G" with l_inv_ex obtain y where y: "y ∈ carrier G" and l_inv: "y ⊗ x = \<one>" by fast from x y have "y ⊗ (x ⊗ y) = y ⊗ \<one>" by (simp add: m_assoc [symmetric] l_inv r_one) with x y have r_inv: "x ⊗ y = \<one>" by simp from x y show "∃y ∈ carrier G. y ⊗ x = \<one> & x ⊗ y = \<one>" by (fast intro: l_inv r_inv) qed then have carrier_subset_Units: "carrier G <= Units G" by (unfold Units_def) fast show ?thesis by (fast intro!: group.intro monoid.intro group_axioms.intro carrier_subset_Units intro: prems r_one) qed lemma (in monoid) monoid_groupI: assumes l_inv_ex: "!!x. x ∈ carrier G ==> ∃y ∈ carrier G. y ⊗ x = \<one>" shows "group G" by (rule groupI) (auto intro: m_assoc l_inv_ex) lemma (in group) Units_eq [simp]: "Units G = carrier G" proof show "Units G <= carrier G" by fast next show "carrier G <= Units G" by (rule Units) qed lemma (in group) inv_closed [intro, simp]: "x ∈ carrier G ==> inv x ∈ carrier G" using Units_inv_closed by simp lemma (in group) l_inv_ex [simp]: "x ∈ carrier G ==> ∃y ∈ carrier G. y ⊗ x = \<one>" using Units_l_inv_ex by simp lemma (in group) r_inv_ex [simp]: "x ∈ carrier G ==> ∃y ∈ carrier G. x ⊗ y = \<one>" using Units_r_inv_ex by simp lemma (in group) l_inv [simp]: "x ∈ carrier G ==> inv x ⊗ x = \<one>" using Units_l_inv by simp subsection {* Cancellation Laws and Basic Properties *} lemma (in group) l_cancel [simp]: "[| x ∈ carrier G; y ∈ carrier G; z ∈ carrier G |] ==> (x ⊗ y = x ⊗ z) = (y = z)" using Units_l_inv by simp lemma (in group) r_inv [simp]: "x ∈ carrier G ==> x ⊗ inv x = \<one>" proof - assume x: "x ∈ carrier G" then have "inv x ⊗ (x ⊗ inv x) = inv x ⊗ \<one>" by (simp add: m_assoc [symmetric] l_inv) with x show ?thesis by (simp del: r_one) qed lemma (in group) r_cancel [simp]: "[| x ∈ carrier G; y ∈ carrier G; z ∈ carrier G |] ==> (y ⊗ x = z ⊗ x) = (y = z)" proof assume eq: "y ⊗ x = z ⊗ x" and G: "x ∈ carrier G" "y ∈ carrier G" "z ∈ carrier G" then have "y ⊗ (x ⊗ inv x) = z ⊗ (x ⊗ inv x)" by (simp add: m_assoc [symmetric] del: r_inv) with G show "y = z" by simp next assume eq: "y = z" and G: "x ∈ carrier G" "y ∈ carrier G" "z ∈ carrier G" then show "y ⊗ x = z ⊗ x" by simp qed lemma (in group) inv_one [simp]: "inv \<one> = \<one>" proof - have "inv \<one> = \<one> ⊗ (inv \<one>)" by (simp del: r_inv) moreover have "... = \<one>" by simp finally show ?thesis . qed lemma (in group) inv_inv [simp]: "x ∈ carrier G ==> inv (inv x) = x" using Units_inv_inv by simp lemma (in group) inv_inj: "inj_on (m_inv G) (carrier G)" using inv_inj_on_Units by simp lemma (in group) inv_mult_group: "[| x ∈ carrier G; y ∈ carrier G |] ==> inv (x ⊗ y) = inv y ⊗ inv x" proof - assume G: "x ∈ carrier G" "y ∈ carrier G" then have "inv (x ⊗ y) ⊗ (x ⊗ y) = (inv y ⊗ inv x) ⊗ (x ⊗ y)" by (simp add: m_assoc l_inv) (simp add: m_assoc [symmetric]) with G show ?thesis by (simp del: l_inv) qed lemma (in group) inv_comm: "[| x ⊗ y = \<one>; x ∈ carrier G; y ∈ carrier G |] ==> y ⊗ x = \<one>" by (rule Units_inv_comm) auto lemma (in group) inv_equality: "[|y ⊗ x = \<one>; x ∈ carrier G; y ∈ carrier G|] ==> inv x = y" apply (simp add: m_inv_def) apply (rule the_equality) apply (simp add: inv_comm [of y x]) apply (rule r_cancel [THEN iffD1], auto) done text {* Power *} lemma (in group) int_pow_def2: "a (^) (z::int) = (if neg z then inv (a (^) (nat (-z))) else a (^) (nat z))" by (simp add: int_pow_def nat_pow_def Let_def) lemma (in group) int_pow_0 [simp]: "x (^) (0::int) = \<one>" by (simp add: int_pow_def2) lemma (in group) int_pow_one [simp]: "\<one> (^) (z::int) = \<one>" by (simp add: int_pow_def2) subsection {* Subgroups *} locale subgroup = fixes H and G (structure) assumes subset: "H ⊆ carrier G" and m_closed [intro, simp]: "[|x ∈ H; y ∈ H|] ==> x ⊗ y ∈ H" and one_closed [simp]: "\<one> ∈ H" and m_inv_closed [intro,simp]: "x ∈ H ==> inv x ∈ H" lemma (in subgroup) is_subgroup: "subgroup H G" by (rule subgroup_axioms) declare (in subgroup) group.intro [intro] lemma (in subgroup) mem_carrier [simp]: "x ∈ H ==> x ∈ carrier G" using subset by blast lemma subgroup_imp_subset: "subgroup H G ==> H ⊆ carrier G" by (rule subgroup.subset) lemma (in subgroup) subgroup_is_group [intro]: includes group G shows "group (G(|carrier := H|)),)" by (rule groupI) (auto intro: m_assoc l_inv mem_carrier) text {* Since @{term H} is nonempty, it contains some element @{term x}. Since it is closed under inverse, it contains @{text "inv x"}. Since it is closed under product, it contains @{text "x ⊗ inv x = \<one>"}. *} lemma (in group) one_in_subset: "[| H ⊆ carrier G; H ≠ {}; ∀a ∈ H. inv a ∈ H; ∀a∈H. ∀b∈H. a ⊗ b ∈ H |] ==> \<one> ∈ H" by (force simp add: l_inv) text {* A characterization of subgroups: closed, non-empty subset. *} lemma (in group) subgroupI: assumes subset: "H ⊆ carrier G" and non_empty: "H ≠ {}" and inv: "!!a. a ∈ H ==> inv a ∈ H" and mult: "!!a b. [|a ∈ H; b ∈ H|] ==> a ⊗ b ∈ H" shows "subgroup H G" proof (simp add: subgroup_def prems) show "\<one> ∈ H" by (rule one_in_subset) (auto simp only: prems) qed declare monoid.one_closed [iff] group.inv_closed [simp] monoid.l_one [simp] monoid.r_one [simp] group.inv_inv [simp] lemma subgroup_nonempty: "~ subgroup {} G" by (blast dest: subgroup.one_closed) lemma (in subgroup) finite_imp_card_positive: "finite (carrier G) ==> 0 < card H" proof (rule classical) assume "finite (carrier G)" "~ 0 < card H" then have "finite H" by (blast intro: finite_subset [OF subset]) with prems have "subgroup {} G" by simp with subgroup_nonempty show ?thesis by contradiction qed (* lemma (in monoid) Units_subgroup: "subgroup (Units G) G" *) subsection {* Direct Products *} constdefs DirProd :: "_ => _ => ('a × 'b) monoid" (infixr "××" 80) "G ×× H ≡ (|carrier = carrier G × carrier H, mult = (λ(g, h) (g', h'). (g ⊗G g', h ⊗H h')), one = (\<one>G, \<one>H)|))," lemma DirProd_monoid: includes monoid G + monoid H shows "monoid (G ×× H)" proof - from prems show ?thesis by (unfold monoid_def DirProd_def, auto) qed text{*Does not use the previous result because it's easier just to use auto.*} lemma DirProd_group: includes group G + group H shows "group (G ×× H)" by (rule groupI) (auto intro: G.m_assoc H.m_assoc G.l_inv H.l_inv simp add: DirProd_def) lemma carrier_DirProd [simp]: "carrier (G ×× H) = carrier G × carrier H" by (simp add: DirProd_def) lemma one_DirProd [simp]: "\<one>G ×× H = (\<one>G, \<one>H)" by (simp add: DirProd_def) lemma mult_DirProd [simp]: "(g, h) ⊗(G ×× H) (g', h') = (g ⊗G g', h ⊗H h')" by (simp add: DirProd_def) lemma inv_DirProd [simp]: includes group G + group H assumes g: "g ∈ carrier G" and h: "h ∈ carrier H" shows "m_inv (G ×× H) (g, h) = (invG g, invH h)" apply (rule group.inv_equality [OF DirProd_group]) apply (simp_all add: prems group.l_inv) done text{*This alternative proof of the previous result demonstrates interpret. It uses @{text Prod.inv_equality} (available after @{text interpret}) instead of @{text "group.inv_equality [OF DirProd_group]"}. *} lemma includes group G + group H assumes g: "g ∈ carrier G" and h: "h ∈ carrier H" shows "m_inv (G ×× H) (g, h) = (invG g, invH h)" proof - interpret Prod: group ["G ×× H"] by (auto intro: DirProd_group group.intro group.axioms prems) show ?thesis by (simp add: Prod.inv_equality g h) qed subsection {* Homomorphisms and Isomorphisms *} constdefs (structure G and H) hom :: "_ => _ => ('a => 'b) set" "hom G H == {h. h ∈ carrier G -> carrier H & (∀x ∈ carrier G. ∀y ∈ carrier G. h (x ⊗G y) = h x ⊗H h y)}" lemma hom_mult: "[| h ∈ hom G H; x ∈ carrier G; y ∈ carrier G |] ==> h (x ⊗G y) = h x ⊗H h y" by (simp add: hom_def) lemma hom_closed: "[| h ∈ hom G H; x ∈ carrier G |] ==> h x ∈ carrier H" by (auto simp add: hom_def funcset_mem) lemma (in group) hom_compose: "[|h ∈ hom G H; i ∈ hom H I|] ==> compose (carrier G) i h ∈ hom G I" apply (auto simp add: hom_def funcset_compose) apply (simp add: compose_def funcset_mem) done constdefs iso :: "_ => _ => ('a => 'b) set" (infixr "≅" 60) "G ≅ H == {h. h ∈ hom G H & bij_betw h (carrier G) (carrier H)}" lemma iso_refl: "(%x. x) ∈ G ≅ G" by (simp add: iso_def hom_def inj_on_def bij_betw_def Pi_def) lemma (in group) iso_sym: "h ∈ G ≅ H ==> Inv (carrier G) h ∈ H ≅ G" apply (simp add: iso_def bij_betw_Inv) apply (subgoal_tac "Inv (carrier G) h ∈ carrier H -> carrier G") prefer 2 apply (simp add: bij_betw_imp_funcset [OF bij_betw_Inv]) apply (simp add: hom_def bij_betw_def Inv_f_eq funcset_mem f_Inv_f) done lemma (in group) iso_trans: "[|h ∈ G ≅ H; i ∈ H ≅ I|] ==> (compose (carrier G) i h) ∈ G ≅ I" by (auto simp add: iso_def hom_compose bij_betw_compose) lemma DirProd_commute_iso: shows "(λ(x,y). (y,x)) ∈ (G ×× H) ≅ (H ×× G)" by (auto simp add: iso_def hom_def inj_on_def bij_betw_def Pi_def) lemma DirProd_assoc_iso: shows "(λ(x,y,z). (x,(y,z))) ∈ (G ×× H ×× I) ≅ (G ×× (H ×× I))" by (auto simp add: iso_def hom_def inj_on_def bij_betw_def Pi_def) text{*Basis for homomorphism proofs: we assume two groups @{term G} and @{term H}, with a homomorphism @{term h} between them*} locale group_hom = group G + group H + var h + assumes homh: "h ∈ hom G H" notes hom_mult [simp] = hom_mult [OF homh] and hom_closed [simp] = hom_closed [OF homh] lemma (in group_hom) one_closed [simp]: "h \<one> ∈ carrier H" by simp lemma (in group_hom) hom_one [simp]: "h \<one> = \<one>H" proof - have "h \<one> ⊗H \<one>H = h \<one> ⊗H h \<one>" by (simp add: hom_mult [symmetric] del: hom_mult) then show ?thesis by (simp del: r_one) qed lemma (in group_hom) inv_closed [simp]: "x ∈ carrier G ==> h (inv x) ∈ carrier H" by simp lemma (in group_hom) hom_inv [simp]: "x ∈ carrier G ==> h (inv x) = invH (h x)" proof - assume x: "x ∈ carrier G" then have "h x ⊗H h (inv x) = \<one>H" by (simp add: hom_mult [symmetric] del: hom_mult) also from x have "... = h x ⊗H invH (h x)" by (simp add: hom_mult [symmetric] del: hom_mult) finally have "h x ⊗H h (inv x) = h x ⊗H invH (h x)" . with x show ?thesis by (simp del: H.r_inv) qed subsection {* Commutative Structures *} text {* Naming convention: multiplicative structures that are commutative are called \emph{commutative}, additive structures are called \emph{Abelian}. *} locale comm_monoid = monoid + assumes m_comm: "[|x ∈ carrier G; y ∈ carrier G|] ==> x ⊗ y = y ⊗ x" lemma (in comm_monoid) m_lcomm: "[|x ∈ carrier G; y ∈ carrier G; z ∈ carrier G|] ==> x ⊗ (y ⊗ z) = y ⊗ (x ⊗ z)" proof - assume xyz: "x ∈ carrier G" "y ∈ carrier G" "z ∈ carrier G" from xyz have "x ⊗ (y ⊗ z) = (x ⊗ y) ⊗ z" by (simp add: m_assoc) also from xyz have "... = (y ⊗ x) ⊗ z" by (simp add: m_comm) also from xyz have "... = y ⊗ (x ⊗ z)" by (simp add: m_assoc) finally show ?thesis . qed lemmas (in comm_monoid) m_ac = m_assoc m_comm m_lcomm lemma comm_monoidI: fixes G (structure) assumes m_closed: "!!x y. [| x ∈ carrier G; y ∈ carrier G |] ==> x ⊗ y ∈ carrier G" and one_closed: "\<one> ∈ carrier G" and m_assoc: "!!x y z. [| x ∈ carrier G; y ∈ carrier G; z ∈ carrier G |] ==> (x ⊗ y) ⊗ z = x ⊗ (y ⊗ z)" and l_one: "!!x. x ∈ carrier G ==> \<one> ⊗ x = x" and m_comm: "!!x y. [| x ∈ carrier G; y ∈ carrier G |] ==> x ⊗ y = y ⊗ x" shows "comm_monoid G" using l_one by (auto intro!: comm_monoid.intro comm_monoid_axioms.intro monoid.intro intro: prems simp: m_closed one_closed m_comm) lemma (in monoid) monoid_comm_monoidI: assumes m_comm: "!!x y. [| x ∈ carrier G; y ∈ carrier G |] ==> x ⊗ y = y ⊗ x" shows "comm_monoid G" by (rule comm_monoidI) (auto intro: m_assoc m_comm) (*lemma (in comm_monoid) r_one [simp]: "x ∈ carrier G ==> x ⊗ \<one> = x" proof - assume G: "x ∈ carrier G" then have "x ⊗ \<one> = \<one> ⊗ x" by (simp add: m_comm) also from G have "... = x" by simp finally show ?thesis . qed*) lemma (in comm_monoid) nat_pow_distr: "[| x ∈ carrier G; y ∈ carrier G |] ==> (x ⊗ y) (^) (n::nat) = x (^) n ⊗ y (^) n" by (induct n) (simp, simp add: m_ac) locale comm_group = comm_monoid + group lemma (in group) group_comm_groupI: assumes m_comm: "!!x y. [| x ∈ carrier G; y ∈ carrier G |] ==> x ⊗ y = y ⊗ x" shows "comm_group G" by unfold_locales (simp_all add: m_comm) lemma comm_groupI: fixes G (structure) assumes m_closed: "!!x y. [| x ∈ carrier G; y ∈ carrier G |] ==> x ⊗ y ∈ carrier G" and one_closed: "\<one> ∈ carrier G" and m_assoc: "!!x y z. [| x ∈ carrier G; y ∈ carrier G; z ∈ carrier G |] ==> (x ⊗ y) ⊗ z = x ⊗ (y ⊗ z)" and m_comm: "!!x y. [| x ∈ carrier G; y ∈ carrier G |] ==> x ⊗ y = y ⊗ x" and l_one: "!!x. x ∈ carrier G ==> \<one> ⊗ x = x" and l_inv_ex: "!!x. x ∈ carrier G ==> ∃y ∈ carrier G. y ⊗ x = \<one>" shows "comm_group G" by (fast intro: group.group_comm_groupI groupI prems) lemma (in comm_group) inv_mult: "[| x ∈ carrier G; y ∈ carrier G |] ==> inv (x ⊗ y) = inv x ⊗ inv y" by (simp add: m_ac inv_mult_group) subsection {* The Lattice of Subgroups of a Group *} text_raw {* \label{sec:subgroup-lattice} *} theorem (in group) subgroups_partial_order: "partial_order (| carrier = {H. subgroup H G}, le = op ⊆ |)" by (rule partial_order.intro) simp_all lemma (in group) subgroup_self: "subgroup (carrier G) G" by (rule subgroupI) auto lemma (in group) subgroup_imp_group: "subgroup H G ==> group (G(| carrier := H |))" by (erule subgroup.subgroup_is_group) (rule group_axioms) lemma (in group) is_monoid [intro, simp]: "monoid G" by (auto intro: monoid.intro m_assoc) lemma (in group) subgroup_inv_equality: "[| subgroup H G; x ∈ H |] ==> m_inv (G (| carrier := H |)) 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 theorem (in group) subgroups_Inter: assumes subgr: "(!!H. H ∈ A ==> subgroup H G)" and not_empty: "A ~= {}" shows "subgroup (\<Inter>A) G" proof (rule subgroupI) from subgr [THEN subgroup.subset] and not_empty show "\<Inter>A ⊆ carrier G" by blast next from subgr [THEN subgroup.one_closed] show "\<Inter>A ~= {}" by blast next fix x assume "x ∈ \<Inter>A" with subgr [THEN subgroup.m_inv_closed] show "inv x ∈ \<Inter>A" by blast next fix x y assume "x ∈ \<Inter>A" "y ∈ \<Inter>A" with subgr [THEN subgroup.m_closed] show "x ⊗ y ∈ \<Inter>A" by blast qed theorem (in group) subgroups_complete_lattice: "complete_lattice (| carrier = {H. subgroup H G}, le = op ⊆ |)" (is "complete_lattice ?L") proof (rule partial_order.complete_lattice_criterion1) show "partial_order ?L" by (rule subgroups_partial_order) next show "∃G. greatest ?L G (carrier ?L)" proof show "greatest ?L (carrier G) (carrier ?L)" by (unfold greatest_def) (simp add: subgroup.subset subgroup_self) qed next fix A assume L: "A ⊆ carrier ?L" and non_empty: "A ~= {}" then have Int_subgroup: "subgroup (\<Inter>A) G" by (fastsimp intro: subgroups_Inter) show "∃I. greatest ?L I (Lower ?L A)" proof show "greatest ?L (\<Inter>A) (Lower ?L A)" (is "greatest _ ?Int _") proof (rule greatest_LowerI) fix H assume H: "H ∈ A" with L have subgroupH: "subgroup H G" by auto from subgroupH have groupH: "group (G (| carrier := H |))" (is "group ?H") by (rule subgroup_imp_group) from groupH have monoidH: "monoid ?H" by (rule group.is_monoid) from H have Int_subset: "?Int ⊆ H" by fastsimp then show "le ?L ?Int H" by simp next fix H assume H: "H ∈ Lower ?L A" with L Int_subgroup show "le ?L H ?Int" by (fastsimp simp: Lower_def intro: Inter_greatest) next show "A ⊆ carrier ?L" by (rule L) next show "?Int ∈ carrier ?L" by simp (rule Int_subgroup) qed qed qed end
lemma monoidI:
[| !!x y. [| x ∈ carrier G; y ∈ carrier G |] ==> x ⊗G y ∈ carrier G;
\<one>G ∈ carrier G;
!!x y z.
[| x ∈ carrier G; y ∈ carrier G; z ∈ carrier G |]
==> x ⊗G y ⊗G z = x ⊗G (y ⊗G z);
!!x. x ∈ carrier G ==> \<one>G ⊗G x = x;
!!x. x ∈ carrier G ==> x ⊗G \<one>G = x |]
==> monoid G
lemma Units_closed:
x ∈ Units G ==> x ∈ carrier G
lemma inv_unique:
[| y ⊗ x = \<one>; x ⊗ y' = \<one>; x ∈ carrier G; y ∈ carrier G;
y' ∈ carrier G |]
==> y = y'
lemma Units_one_closed:
\<one> ∈ Units G
lemma Units_inv_closed:
x ∈ Units G ==> inv x ∈ carrier G
lemma Units_l_inv_ex:
x ∈ Units G ==> ∃y∈carrier G. y ⊗ x = \<one>
lemma Units_r_inv_ex:
x ∈ Units G ==> ∃y∈carrier G. x ⊗ y = \<one>
lemma Units_l_inv:
x ∈ Units G ==> inv x ⊗ x = \<one>
lemma Units_r_inv:
x ∈ Units G ==> x ⊗ inv x = \<one>
lemma Units_inv_Units:
x ∈ Units G ==> inv x ∈ Units G
lemma Units_l_cancel:
[| x ∈ Units G; y ∈ carrier G; z ∈ carrier G |] ==> (x ⊗ y = x ⊗ z) = (y = z)
lemma Units_inv_inv:
x ∈ Units G ==> inv (inv x) = x
lemma inv_inj_on_Units:
inj_on (m_inv G) (Units G)
lemma Units_inv_comm:
[| x ⊗ y = \<one>; x ∈ Units G; y ∈ Units G |] ==> y ⊗ x = \<one>
lemma nat_pow_closed:
x ∈ carrier G ==> x (^) n ∈ carrier G
lemma nat_pow_0:
x (^) 0 = \<one>
lemma nat_pow_Suc:
x (^) Suc n = x (^) n ⊗ x
lemma nat_pow_one:
\<one> (^) n = \<one>
lemma nat_pow_mult:
x ∈ carrier G ==> x (^) n ⊗ x (^) m = x (^) (n + m)
lemma nat_pow_pow:
x ∈ carrier G ==> (x (^) n) (^) m = x (^) (n * m)
lemma is_group:
group G
theorem groupI:
[| !!x y. [| x ∈ carrier G; y ∈ carrier G |] ==> x ⊗G y ∈ carrier G;
\<one>G ∈ carrier G;
!!x y z.
[| x ∈ carrier G; y ∈ carrier G; z ∈ carrier G |]
==> x ⊗G y ⊗G z = x ⊗G (y ⊗G z);
!!x. x ∈ carrier G ==> \<one>G ⊗G x = x;
!!x. x ∈ carrier G ==> ∃y∈carrier G. y ⊗G x = \<one>G |]
==> group G
lemma monoid_groupI:
(!!x. x ∈ carrier G ==> ∃y∈carrier G. y ⊗ x = \<one>) ==> group G
lemma Units_eq:
Units G = carrier G
lemma inv_closed:
x ∈ carrier G ==> inv x ∈ carrier G
lemma l_inv_ex:
x ∈ carrier G ==> ∃y∈carrier G. y ⊗ x = \<one>
lemma r_inv_ex:
x ∈ carrier G ==> ∃y∈carrier G. x ⊗ y = \<one>
lemma l_inv:
x ∈ carrier G ==> inv x ⊗ x = \<one>
lemma l_cancel:
[| x ∈ carrier G; y ∈ carrier G; z ∈ carrier G |] ==> (x ⊗ y = x ⊗ z) = (y = z)
lemma r_inv:
x ∈ carrier G ==> x ⊗ inv x = \<one>
lemma r_cancel:
[| x ∈ carrier G; y ∈ carrier G; z ∈ carrier G |] ==> (y ⊗ x = z ⊗ x) = (y = z)
lemma inv_one:
inv \<one> = \<one>
lemma inv_inv:
x ∈ carrier G ==> inv (inv x) = x
lemma inv_inj:
inj_on (m_inv G) (carrier G)
lemma inv_mult_group:
[| x ∈ carrier G; y ∈ carrier G |] ==> inv (x ⊗ y) = inv y ⊗ inv x
lemma inv_comm:
[| x ⊗ y = \<one>; x ∈ carrier G; y ∈ carrier G |] ==> y ⊗ x = \<one>
lemma inv_equality:
[| y ⊗ x = \<one>; x ∈ carrier G; y ∈ carrier G |] ==> inv x = y
lemma int_pow_def2:
a (^) z = (if neg z then inv (a (^) nat (- z)) else a (^) nat z)
lemma int_pow_0:
x (^) 0 = \<one>
lemma int_pow_one:
\<one> (^) z = \<one>
lemma is_subgroup:
subgroup H G
lemma mem_carrier:
x ∈ H ==> x ∈ carrier G
lemma subgroup_imp_subset:
subgroup H G ==> H ⊆ carrier G
lemma subgroup_is_group:
group G ==> group (G(| carrier := H |))
lemma one_in_subset:
[| H ⊆ carrier G; H ≠ {}; ∀a∈H. inv a ∈ H; ∀a∈H. ∀b∈H. a ⊗ b ∈ H |]
==> \<one> ∈ H
lemma subgroupI:
[| H ⊆ carrier G; H ≠ {}; !!a. a ∈ H ==> inv a ∈ H;
!!a b. [| a ∈ H; b ∈ H |] ==> a ⊗ b ∈ H |]
==> subgroup H G
lemma subgroup_nonempty:
¬ subgroup {} G
lemma finite_imp_card_positive:
finite (carrier G) ==> 0 < card H
lemma DirProd_monoid:
[| monoid G; monoid H |] ==> monoid (G ×× H)
lemma DirProd_group:
[| group G; group H |] ==> group (G ×× H)
lemma carrier_DirProd:
carrier (G ×× H) = carrier G × carrier H
lemma one_DirProd:
\<one>G ×× H = (\<one>G, \<one>H)
lemma mult_DirProd:
(g, h) ⊗G ×× H (g', h') = (g ⊗G g', h ⊗H h')
lemma inv_DirProd:
[| group G; group H; g ∈ carrier G; h ∈ carrier H |]
==> invG ×× H (g, h) = (invG g, invH h)
lemma
[| group G; group H; g ∈ carrier G; h ∈ carrier H |]
==> invG ×× H (g, h) = (invG g, invH h)
lemma hom_mult:
[| h ∈ hom G H; x ∈ carrier G; y ∈ carrier G |] ==> h (x ⊗G y) = h x ⊗H h y
lemma hom_closed:
[| h ∈ hom G H; x ∈ carrier G |] ==> h x ∈ carrier H
lemma hom_compose:
[| h ∈ hom G H; i ∈ hom H I |] ==> compose (carrier G) i h ∈ hom G I
lemma iso_refl:
(λx. x) ∈ G ≅ G
lemma iso_sym:
h ∈ G ≅ H ==> Inv (carrier G) h ∈ H ≅ G
lemma iso_trans:
[| h ∈ G ≅ H; i ∈ H ≅ I |] ==> compose (carrier G) i h ∈ G ≅ I
lemma DirProd_commute_iso:
(λ(x, y). (y, x)) ∈ G ×× H ≅ H ×× G
lemma DirProd_assoc_iso:
(λ(x, y, z). (x, y, z)) ∈ G ×× H ×× I ≅ G ×× H ×× I
lemma one_closed:
h \<one> ∈ carrier H
lemma hom_one:
h \<one> = \<one>H
lemma inv_closed:
x ∈ carrier G ==> h (inv x) ∈ carrier H
lemma hom_inv:
x ∈ carrier G ==> h (inv x) = invH h x
lemma m_lcomm:
[| x ∈ carrier G; y ∈ carrier G; z ∈ carrier G |] ==> x ⊗ (y ⊗ z) = y ⊗ (x ⊗ z)
lemma m_ac:
[| x ∈ carrier G; y ∈ carrier G; z ∈ carrier G |] ==> x ⊗ y ⊗ z = x ⊗ (y ⊗ z)
[| x ∈ carrier G; y ∈ carrier G |] ==> x ⊗ y = y ⊗ x
[| x ∈ carrier G; y ∈ carrier G; z ∈ carrier G |] ==> x ⊗ (y ⊗ z) = y ⊗ (x ⊗ z)
lemma comm_monoidI:
[| !!x y. [| x ∈ carrier G; y ∈ carrier G |] ==> x ⊗G y ∈ carrier G;
\<one>G ∈ carrier G;
!!x y z.
[| x ∈ carrier G; y ∈ carrier G; z ∈ carrier G |]
==> x ⊗G y ⊗G z = x ⊗G (y ⊗G z);
!!x. x ∈ carrier G ==> \<one>G ⊗G x = x;
!!x y. [| x ∈ carrier G; y ∈ carrier G |] ==> x ⊗G y = y ⊗G x |]
==> comm_monoid G
lemma monoid_comm_monoidI:
(!!x y. [| x ∈ carrier G; y ∈ carrier G |] ==> x ⊗ y = y ⊗ x) ==> comm_monoid G
lemma nat_pow_distr:
[| x ∈ carrier G; y ∈ carrier G |] ==> (x ⊗ y) (^) n = x (^) n ⊗ y (^) n
lemma group_comm_groupI:
(!!x y. [| x ∈ carrier G; y ∈ carrier G |] ==> x ⊗ y = y ⊗ x) ==> comm_group G
lemma comm_groupI:
[| !!x y. [| x ∈ carrier G; y ∈ carrier G |] ==> x ⊗G y ∈ carrier G;
\<one>G ∈ carrier G;
!!x y z.
[| x ∈ carrier G; y ∈ carrier G; z ∈ carrier G |]
==> x ⊗G y ⊗G z = x ⊗G (y ⊗G z);
!!x y. [| x ∈ carrier G; y ∈ carrier G |] ==> x ⊗G y = y ⊗G x;
!!x. x ∈ carrier G ==> \<one>G ⊗G x = x;
!!x. x ∈ carrier G ==> ∃y∈carrier G. y ⊗G x = \<one>G |]
==> comm_group G
lemma inv_mult:
[| x ∈ carrier G; y ∈ carrier G |] ==> inv (x ⊗ y) = inv x ⊗ inv y
theorem subgroups_partial_order:
partial_order (| carrier = {H. subgroup H G}, le = op ⊆ |)
lemma subgroup_self:
subgroup (carrier G) G
lemma subgroup_imp_group:
subgroup H G ==> group (G(| carrier := H |))
lemma is_monoid:
monoid G
lemma subgroup_inv_equality:
[| subgroup H G; x ∈ H |] ==> invG(| carrier := H |) x = inv x
theorem subgroups_Inter:
[| !!H. H ∈ A ==> subgroup H G; A ≠ {} |] ==> subgroup (Inter A) G
theorem subgroups_complete_lattice:
Lattice.complete_lattice (| carrier = {H. subgroup H G}, le = op ⊆ |)