# HG changeset patch # User ballarin # Date 1044866722 -3600 # Node ID 722593f2f06887bc53d541400a69407107aa5b00 # Parent 91713a1915ee8505869afe8c1466d8323bcaca6a New development of algebra: Groups. diff -r 91713a1915ee -r 722593f2f068 src/HOL/Algebra/Group.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Algebra/Group.thy Mon Feb 10 09:45:22 2003 +0100 @@ -0,0 +1,372 @@ +(* + Title: HOL/Algebra/Group.thy + Id: $Id$ + Author: Clemens Ballarin, started 4 February 2003 + +Based on work by Florian Kammueller, L C Paulson and Markus Wenzel. +*) + +header {* Algebraic Structures up to Abelian Groups *} + +theory Group = FuncSet: + +text {* + Definitions follow Jacobson, Basic Algebra I, Freeman, 1985; with + the exception of \emph{magma} which, following Bourbaki, is a set + together with a binary, closed operation. +*} + +section {* From Magmas to Groups *} + +subsection {* Definitions *} + +record 'a magma = + carrier :: "'a set" + mult :: "['a, 'a] => 'a" (infixl "\\" 70) + +record 'a group = "'a magma" + + one :: 'a ("\\") + m_inv :: "'a => 'a" ("inv\ _" [81] 80) + +locale magma = struct G + + assumes m_closed [intro, simp]: + "[| x \ carrier G; y \ carrier G |] ==> x \ y \ carrier G" + +locale semigroup = magma + + assumes m_assoc: + "[| x \ carrier G; y \ carrier G; z \ carrier G |] ==> + (x \ y) \ z = x \ (y \ z)" + +locale group = semigroup + + assumes one_closed [intro, simp]: "\ \ carrier G" + and inv_closed [intro, simp]: "x \ carrier G ==> inv x \ carrier G" + and l_one [simp]: "x \ carrier G ==> \ \ x = x" + and l_inv: "x \ carrier G ==> inv x \ x = \" + +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)" +proof + assume eq: "x \ y = x \ z" + and G: "x \ carrier G" "y \ carrier G" "z \ carrier G" + then have "(inv x \ x) \ y = (inv x \ x) \ z" by (simp add: m_assoc) + with G show "y = z" by (simp add: l_inv) +next + assume eq: "y = z" + and G: "x \ carrier G" "y \ carrier G" "z \ carrier G" + then show "x \ y = x \ z" by simp +qed + +lemma (in group) r_one [simp]: + "x \ carrier G ==> x \ \ = x" +proof - + assume x: "x \ carrier G" + then have "inv x \ (x \ \) = inv x \ x" + by (simp add: m_assoc [symmetric] l_inv) + with x show ?thesis by simp +qed + +lemma (in group) r_inv: + "x \ carrier G ==> x \ inv x = \" +proof - + assume x: "x \ carrier G" + then have "inv x \ (x \ inv x) = inv x \ \" + 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]) + with G show "y = z" by (simp add: r_inv) +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_inv [simp]: + "x \ carrier G ==> inv (inv x) = x" +proof - + assume x: "x \ carrier G" + then have "inv x \ inv (inv x) = inv x \ x" by (simp add: l_inv r_inv) + with x show ?thesis by simp +qed + +lemma (in group) inv_mult: + "[| 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] l_inv) + with G show ?thesis by simp +qed + +subsection {* Substructures *} + +locale submagma = var H + struct G + + assumes subset [intro, simp]: "H \ carrier G" + and m_closed [intro, simp]: "[| x \ H; y \ H |] ==> x \ y \ H" + +declare (in submagma) magma.intro [intro] semigroup.intro [intro] + +(* +alternative definition of submagma + +locale submagma = var H + struct G + + assumes subset [intro, simp]: "carrier H \ carrier G" + and m_equal [simp]: "mult H = mult G" + and m_closed [intro, simp]: + "[| x \ carrier H; y \ carrier H |] ==> x \ y \ carrier H" +*) + +lemma submagma_imp_subset: + "submagma H G ==> H \ carrier G" + by (rule submagma.subset) + +lemma (in submagma) subsetD [dest, simp]: + "x \ H ==> x \ carrier G" + using subset by blast + +lemma (in submagma) magmaI [intro]: + includes magma G + shows "magma (G(| carrier := H |))" + by rule simp + +lemma (in submagma) semigroup_axiomsI [intro]: + includes semigroup G + shows "semigroup_axioms (G(| carrier := H |))" + by rule (simp add: m_assoc) + +lemma (in submagma) semigroupI [intro]: + includes semigroup G + shows "semigroup (G(| carrier := H |))" + using prems by fast + +locale subgroup = submagma H G + + assumes one_closed [intro, simp]: "\ \ H" + and m_inv_closed [intro, simp]: "x \ H ==> inv x \ H" + +declare (in subgroup) group.intro [intro] + +lemma (in subgroup) group_axiomsI [intro]: + includes group G + shows "group_axioms (G(| carrier := H |))" + by rule (simp_all add: l_inv) + +lemma (in subgroup) groupI [intro]: + includes group G + shows "group (G(| carrier := H |))" + using prems by fast + +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 = \"}. +*} + +lemma (in group) one_in_subset: + "[| H \ carrier G; H \ {}; \a \ H. inv a \ H; \a\H. \b\H. a \ b \ H |] + ==> \ \ 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 + from subset and mult show "submagma H G" .. +next + have "\ \ H" by (rule one_in_subset) (auto simp only: prems) + with inv show "subgroup_axioms H G" + by (intro subgroup_axioms.intro) simp_all +qed + +text {* + Repeat facts of submagmas for subgroups. Necessary??? +*} + +lemma (in subgroup) subset: + "H \ carrier G" + .. + +lemma (in subgroup) m_closed: + "[| x \ H; y \ H |] ==> x \ y \ H" + .. + +declare magma.m_closed [simp] + +declare group.one_closed [iff] group.inv_closed [simp] + group.l_one [simp] group.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) + have sub: "subgroup H G" using prems .. + assume fin: "finite (carrier G)" + and zero: "~ 0 < card H" + then have "finite H" by (blast intro: finite_subset dest: subset) + with zero sub have "subgroup {} G" by simp + with subgroup_nonempty show ?thesis by contradiction +qed + +subsection {* Direct Products *} + +constdefs + DirProdMagma :: + "[('a, 'c) magma_scheme, ('b, 'd) magma_scheme] => ('a \ 'b) magma" + (infixr "\\<^sub>m" 80) + "G \\<^sub>m H == (| carrier = carrier G \ carrier H, + mult = (%(xg, xh) (yg, yh). (mult G xg yg, mult H xh yh)) |)" + + DirProdGroup :: + "[('a, 'c) group_scheme, ('b, 'd) group_scheme] => ('a \ 'b) group" + (infixr "\\<^sub>g" 80) + "G \\<^sub>g H == (| carrier = carrier (G \\<^sub>m H), + mult = mult (G \\<^sub>m H), + one = (one G, one H), + m_inv = (%(g, h). (m_inv G g, m_inv H h)) |)" + +lemma DirProdMagma_magma: + includes magma G + magma H + shows "magma (G \\<^sub>m H)" + by rule (auto simp add: DirProdMagma_def) + +lemma DirProdMagma_semigroup_axioms: + includes semigroup G + semigroup H + shows "semigroup_axioms (G \\<^sub>m H)" + by rule (auto simp add: DirProdMagma_def G.m_assoc H.m_assoc) + +lemma DirProdMagma_semigroup: + includes semigroup G + semigroup H + shows "semigroup (G \\<^sub>m H)" + using prems + by (fast intro: semigroup.intro + DirProdMagma_magma DirProdMagma_semigroup_axioms) + +lemma DirProdGroup_magma: + includes magma G + magma H + shows "magma (G \\<^sub>g H)" + by rule (auto simp add: DirProdGroup_def DirProdMagma_def) + +lemma DirProdGroup_semigroup_axioms: + includes semigroup G + semigroup H + shows "semigroup_axioms (G \\<^sub>g H)" + by rule + (auto simp add: DirProdGroup_def DirProdMagma_def G.m_assoc H.m_assoc) + +lemma DirProdGroup_semigroup: + includes semigroup G + semigroup H + shows "semigroup (G \\<^sub>g H)" + using prems + by (fast intro: semigroup.intro + DirProdGroup_magma DirProdGroup_semigroup_axioms) + +(* ... and further lemmas for group ... *) + +lemma + includes group G + group H + shows "group (G \\<^sub>g H)" +by rule + (auto intro: magma.intro semigroup_axioms.intro group_axioms.intro + simp add: DirProdGroup_def DirProdMagma_def + G.m_assoc H.m_assoc G.l_inv H.l_inv) + +subsection {* Homomorphisms *} + +constdefs + hom :: "[('a, 'c) magma_scheme, ('b, 'd) magma_scheme] => ('a => 'b)set" + "hom G H == + {h. h \ carrier G -> carrier H & + (\x \ carrier G. \y \ carrier G. h (mult G x y) = mult H (h x) (h y))}" + +lemma (in semigroup) hom: + includes semigroup G + shows "semigroup (| carrier = hom G G, mult = op o |)" +proof + show "magma (| carrier = hom G G, mult = op o |)" + by rule (simp add: Pi_def hom_def) +next + show "semigroup_axioms (| carrier = hom G G, mult = op o |)" + by rule (simp add: o_assoc) +qed + +lemma hom_mult: + "[| h \ hom G H; x \ carrier G; y \ carrier G |] + ==> h (mult G x y) = mult H (h x) (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) + +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 \ \ carrier H" + by simp + +lemma (in group_hom) hom_one [simp]: + "h \ = \\<^sub>2" +proof - + have "h \ \\<^sub>2 \\<^sub>2 = h \ \\<^sub>2 h \" + 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) = inv\<^sub>2 (h x)" +proof - + assume x: "x \ carrier G" + then have "h x \\<^sub>2 h (inv x) = \\<^sub>2" + by (simp add: hom_mult [symmetric] G.r_inv del: hom_mult) + also from x have "... = h x \\<^sub>2 inv\<^sub>2 (h x)" + by (simp add: hom_mult [symmetric] H.r_inv del: hom_mult) + finally have "h x \\<^sub>2 h (inv x) = h x \\<^sub>2 inv\<^sub>2 (h x)" . + with x show ?thesis by simp +qed + +section {* Abelian Structures *} + +subsection {* Definition *} + +locale abelian_semigroup = semigroup + + assumes m_comm: "[| x \ carrier G; y \ carrier G |] ==> x \ y = y \ x" + +lemma (in abelian_semigroup) 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 abelian_semigroup) ac = m_assoc m_comm m_lcomm + +locale abelian_group = abelian_semigroup + group + +end