diff -r 8882d47e075f -r 2c01c0bdb385 src/HOL/Algebra/Coset.thy --- a/src/HOL/Algebra/Coset.thy Tue Jul 15 16:02:10 2008 +0200 +++ b/src/HOL/Algebra/Coset.thy Tue Jul 15 16:50:09 2008 +0200 @@ -109,23 +109,27 @@ text (in group) {* Opposite of @{thm [source] "repr_independence"} *} lemma (in group) repr_independenceD: - includes subgroup H G + assumes "subgroup H G" assumes ycarr: "y \ carrier G" and repr: "H #> x = H #> y" shows "y \ H #> x" - apply (subst repr) +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 text {* Elements of a right coset are in the carrier *} lemma (in subgroup) elemrcos_carrier: - includes group + assumes "group G" assumes acarr: "a \ carrier G" and a': "a' \ H #> a" shows "a' \ carrier G" proof - + interpret group [G] by fact from subset and acarr have "H #> a \ carrier G" by (rule r_coset_subset_G) from this and a' @@ -134,38 +138,42 @@ qed lemma (in subgroup) rcos_const: - includes group + assumes "group G" assumes hH: "h \ H" shows "H #> h = H" - apply (unfold r_coset_def) - apply rule - apply rule - apply clarsimp - apply (intro subgroup.m_closed) - apply (rule is_subgroup) +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' \ H" - note carr = hH[THEN mem_carrier] h'H[THEN mem_carrier] - from carr - have a: "h' = (h' \ inv h) \ h" by (simp add: m_assoc) - from h'H hH - have "h' \ inv h \ H" by simp - from this and a - show "\x\H. h' = x \ h" by fast + apply (rule hH) + apply rule + apply simp + proof - + fix h' + assume h'H: "h' \ H" + note carr = hH[THEN mem_carrier] h'H[THEN mem_carrier] + from carr + have a: "h' = (h' \ inv h) \ h" by (simp add: m_assoc) + from h'H hH + have "h' \ inv h \ H" by simp + from this and a + show "\x\H. h' = x \ h" by fast + qed qed text {* Step one for lemma @{text "rcos_module"} *} lemma (in subgroup) rcos_module_imp: - includes group + assumes "group G" assumes xcarr: "x \ carrier G" and x'cos: "x' \ H #> x" shows "(x' \ inv x) \ H" proof - + interpret group [G] by fact from xcarr x'cos have x'carr: "x' \ carrier G" by (rule elemrcos_carrier[OF is_group]) @@ -200,11 +208,12 @@ text {* Step two for lemma @{text "rcos_module"} *} lemma (in subgroup) rcos_module_rev: - includes group + assumes "group G" assumes carr: "x \ carrier G" "x' \ carrier G" and xixH: "(x' \ inv x) \ H" shows "x' \ H #> x" proof - + interpret group [G] by fact from xixH have "\h\H. x' \ (inv x) = h" by fast from this @@ -231,27 +240,30 @@ text {* Module property of right cosets *} lemma (in subgroup) rcos_module: - includes group + assumes "group G" assumes carr: "x \ carrier G" "x' \ carrier G" shows "(x' \ H #> x) = (x' \ inv x \ H)" -proof - assume "x' \ H #> x" - from this and carr - show "x' \ inv x \ H" +proof - + interpret group [G] by fact + show ?thesis proof assume "x' \ H #> x" + from this and carr + show "x' \ inv x \ H" by (intro rcos_module_imp[OF is_group]) -next - assume "x' \ inv x \ H" - from this and carr - show "x' \ H #> x" + next + assume "x' \ inv x \ H" + from this and carr + show "x' \ H #> x" by (intro rcos_module_rev[OF is_group]) + qed qed text {* Right cosets are subsets of the carrier. *} lemma (in subgroup) rcosets_carrier: - includes group + assumes "group G" assumes XH: "X \ rcosets H" shows "X \ carrier G" proof - + interpret group [G] by fact from XH have "\x\ carrier G. X = H #> x" unfolding RCOSETS_def by fast @@ -331,11 +343,12 @@ qed lemma (in subgroup) lcos_module_rev: - includes group + assumes "group G" assumes carr: "x \ carrier G" "x' \ carrier G" and xixH: "(inv x \ x') \ H" shows "x' \ x <# H" proof - + interpret group [G] by fact from xixH have "\h\H. (inv x) \ x' = h" by fast from this @@ -584,29 +597,33 @@ lemma (in subgroup) equiv_rcong: - includes group G + assumes "group G" shows "equiv (carrier G) (rcong H)" -proof (intro equiv.intro) - show "refl (carrier G) (rcong H)" - by (auto simp add: r_congruent_def refl_def) -next - show "sym (rcong H)" - proof (simp add: r_congruent_def sym_def, clarify) - fix x y - assume [simp]: "x \ carrier G" "y \ carrier G" - and "inv x \ y \ H" - hence "inv (inv x \ y) \ H" by (simp add: m_inv_closed) - thus "inv y \ x \ H" by (simp add: inv_mult_group) - qed -next - show "trans (rcong H)" - proof (simp add: r_congruent_def trans_def, clarify) - fix x y z - assume [simp]: "x \ carrier G" "y \ carrier G" "z \ carrier G" - and "inv x \ y \ H" and "inv y \ z \ H" - hence "(inv x \ y) \ (inv y \ z) \ H" by simp - hence "inv x \ (y \ inv y) \ z \ H" by (simp add: m_assoc del: r_inv) - thus "inv x \ z \ H" by simp +proof - + interpret group [G] by fact + show ?thesis + proof (intro equiv.intro) + show "refl (carrier G) (rcong H)" + by (auto simp add: r_congruent_def refl_def) + next + show "sym (rcong H)" + proof (simp add: r_congruent_def sym_def, clarify) + fix x y + assume [simp]: "x \ carrier G" "y \ carrier G" + and "inv x \ y \ H" + hence "inv (inv x \ y) \ H" by (simp add: m_inv_closed) + thus "inv y \ x \ H" by (simp add: inv_mult_group) + qed + next + show "trans (rcong H)" + proof (simp add: r_congruent_def trans_def, clarify) + fix x y z + assume [simp]: "x \ carrier G" "y \ carrier G" "z \ carrier G" + and "inv x \ y \ H" and "inv y \ z \ H" + hence "(inv x \ y) \ (inv y \ z) \ H" by simp + hence "inv x \ (y \ inv y) \ z \ H" by (simp add: m_assoc del: r_inv) + thus "inv x \ z \ H" by simp + qed qed qed @@ -625,31 +642,38 @@ *) lemma (in subgroup) l_coset_eq_rcong: - includes group G + assumes "group G" assumes a: "a \ carrier G" shows "a <# H = rcong H `` {a}" -by (force simp add: r_congruent_def l_coset_def m_assoc [symmetric] a ) - +proof - + interpret group [G] by fact + show ?thesis by (force simp add: r_congruent_def l_coset_def m_assoc [symmetric] a ) +qed subsubsection{*Two Distinct Right Cosets are Disjoint*} lemma (in group) rcos_equation: - includes subgroup H G - shows - "\ha \ a = h \ b; a \ carrier G; b \ carrier G; - h \ H; ha \ H; hb \ H\ - \ hb \ a \ (\h\H. {h \ b})" -apply (rule UN_I [of "hb \ ((inv ha) \ h)"]) -apply (simp add: ) -apply (simp add: m_assoc transpose_inv) -done + assumes "subgroup H G" + assumes p: "ha \ a = h \ b" "a \ carrier G" "b \ carrier G" "h \ H" "ha \ H" "hb \ H" + shows "hb \ a \ (\h\H. {h \ b})" +proof - + interpret subgroup [H G] by fact + from p show ?thesis apply (rule_tac UN_I [of "hb \ ((inv ha) \ h)"]) + apply (simp add: ) + apply (simp add: m_assoc transpose_inv) + done +qed lemma (in group) rcos_disjoint: - includes subgroup H G - shows "\a \ rcosets H; b \ rcosets H; a\b\ \ a \ b = {}" -apply (simp add: RCOSETS_def r_coset_def) -apply (blast intro: rcos_equation prems sym) -done + assumes "subgroup H G" + assumes p: "a \ rcosets H" "b \ rcosets H" "a\b" + shows "a \ b = {}" +proof - + interpret subgroup [H G] by fact + from p show ?thesis apply (simp add: RCOSETS_def r_coset_def) + apply (blast intro: rcos_equation prems sym) + done +qed subsection {* Further lemmas for @{text "r_congruent"} *} @@ -732,12 +756,16 @@ "order S \ card (carrier S)" lemma (in group) rcosets_part_G: - includes subgroup + assumes "subgroup H G" shows "\(rcosets H) = carrier G" -apply (rule equalityI) - apply (force simp add: RCOSETS_def r_coset_def) -apply (auto simp add: RCOSETS_def intro: rcos_self prems) -done +proof - + interpret subgroup [H G] by fact + show ?thesis + apply (rule equalityI) + apply (force simp add: RCOSETS_def r_coset_def) + apply (auto simp add: RCOSETS_def intro: rcos_self prems) + done +qed lemma (in group) cosets_finite: "\c \ rcosets H; H \ carrier G; finite (carrier G)\ \ finite c" @@ -815,9 +843,10 @@ by (auto simp add: RCOSETS_def rcos_sum m_assoc) lemma (in subgroup) subgroup_in_rcosets: - includes group G + assumes "group G" shows "H \ rcosets H" proof - + interpret group [G] by fact from _ subgroup_axioms have "H #> \ = H" by (rule coset_join2) auto then show ?thesis