# HG changeset patch # User haftmann # Date 1487501931 -3600 # Node ID b46fe5138cb008e16ffd03e1d59d34dc45968365 # Parent 1846c455115382097ab28698dad6a7d264904a46 backed out unintended effects of 8355a6e2df79 in src/HOL/Algebra diff -r 1846c4551153 -r b46fe5138cb0 src/HOL/Algebra/Coset.thy --- a/src/HOL/Algebra/Coset.thy Sat Feb 18 19:49:29 2017 +0100 +++ b/src/HOL/Algebra/Coset.thy Sun Feb 19 11:58:51 2017 +0100 @@ -15,16 +15,16 @@ where "H #>\<^bsub>G\<^esub> a = (\h\H. {h \\<^bsub>G\<^esub> a})" definition - l_coset :: "[_, 'a, 'a set] \ 'a set" (infixl "\#\" 60) - where "a \#\<^bsub>G\<^esub> H = (\h\H. {a \\<^bsub>G\<^esub> h})" + l_coset :: "[_, 'a, 'a set] \ 'a set" (infixl "<#\" 60) + where "a <#\<^bsub>G\<^esub> H = (\h\H. {a \\<^bsub>G\<^esub> h})" definition RCOSETS :: "[_, 'a set] \ ('a set)set" ("rcosets\ _" [81] 80) where "rcosets\<^bsub>G\<^esub> H = (\a\carrier G. {H #>\<^bsub>G\<^esub> a})" definition - set_mult :: "[_, 'a set ,'a set] \ 'a set" (infixl "\#>\" 60) - where "H \#>\<^bsub>G\<^esub> K = (\h\H. \k\K. {h \\<^bsub>G\<^esub> k})" + set_mult :: "[_, 'a set ,'a set] \ 'a set" (infixl "<#>\" 60) + where "H <#>\<^bsub>G\<^esub> K = (\h\H. \k\K. {h \\<^bsub>G\<^esub> k})" definition SET_INV :: "[_,'a set] \ 'a set" ("set'_inv\ _" [81] 80) @@ -32,7 +32,7 @@ locale normal = subgroup + group + - assumes coset_eq: "(\x \ carrier G. H #> x = x \# H)" + assumes coset_eq: "(\x \ carrier G. H #> x = x <# H)" abbreviation normal_rel :: "['a set, ('a, 'b) monoid_scheme] \ bool" (infixl "\" 60) where @@ -287,7 +287,7 @@ lemma (in monoid) set_mult_closed: assumes Acarr: "A \ carrier G" and Bcarr: "B \ carrier G" - shows "A \#> B \ carrier G" + shows "A <#> B \ carrier G" apply rule apply (simp add: set_mult_def, clarsimp) proof - fix a b @@ -306,7 +306,7 @@ lemma (in comm_group) mult_subgroups: assumes subH: "subgroup H G" and subK: "subgroup K G" - shows "subgroup (H \#> 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 @@ -351,7 +351,7 @@ assumes "group G" assumes carr: "x \ carrier G" "x' \ carrier G" and xixH: "(inv x \ x') \ H" - shows "x' \ x \# H" + shows "x' \ x <# H" proof - interpret group G by fact from xixH @@ -375,7 +375,7 @@ have "x \ h = x'" by simp from this[symmetric] and hH - show "x' \ x \# H" + show "x' \ x <# H" unfolding l_coset_def by fast qed @@ -387,7 +387,7 @@ by (simp add: normal_def subgroup_def) lemma (in group) normalI: - "subgroup H G \ (\x \ carrier G. H #> x = x \# H) \ H \ G" + "subgroup H G \ (\x \ carrier G. H #> x = x <# H) \ H \ G" by (simp add: normal_def normal_axioms_def is_group) lemma (in normal) inv_op_closed1: @@ -460,18 +460,18 @@ lemma (in group) lcos_m_assoc: "[| M \ carrier G; g \ carrier G; h \ carrier G |] - ==> g \# (h \# M) = (g \ h) \# M" + ==> g <# (h <# M) = (g \ h) <# M" by (force simp add: l_coset_def m_assoc) -lemma (in group) lcos_mult_one: "M \ carrier G ==> \ \# M = M" +lemma (in group) lcos_mult_one: "M \ carrier G ==> \ <# M = M" by (force simp add: l_coset_def) lemma (in group) l_coset_subset_G: - "[| H \ carrier G; x \ carrier G |] ==> x \# H \ carrier G" + "[| H \ carrier G; x \ carrier G |] ==> x <# H \ carrier G" by (auto simp add: l_coset_def subsetD) lemma (in group) l_coset_swap: - "\y \ x \# H; x \ carrier G; subgroup H G\ \ x \ y \# H" + "\y \ x <# H; x \ carrier G; subgroup H G\ \ x \ y <# H" proof (simp add: l_coset_def) assume "\h\H. y = x \ h" and x: "x \ carrier G" @@ -487,13 +487,13 @@ qed lemma (in group) l_coset_carrier: - "[| y \ x \# H; x \ carrier G; subgroup H G |] ==> y \ carrier G" + "[| y \ x <# H; x \ carrier G; subgroup H G |] ==> y \ 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 \ x \# H" and x: "x \ carrier G" and sb: "subgroup H G" - shows "y \# H \ x \# H" + assumes y: "y \ x <# H" and x: "x \ carrier G" and sb: "subgroup H G" + shows "y <# H \ x <# H" proof - from y obtain h' where "h' \ H" "x \ h' = y" by (auto simp add: l_coset_def) @@ -503,20 +503,20 @@ qed lemma (in group) l_repr_independence: - assumes y: "y \ x \# H" and x: "x \ carrier G" and sb: "subgroup H G" - shows "x \# H = y \# H" + assumes y: "y \ x <# H" and x: "x \ carrier G" and sb: "subgroup H G" + shows "x <# H = y <# H" proof - show "x \# H \ y \# H" + show "x <# H \ y <# H" by (rule l_repr_imp_subset, (blast intro: l_coset_swap l_coset_carrier y x sb)+) - show "y \# H \ x \# H" by (rule l_repr_imp_subset [OF y x sb]) + show "y <# H \ x <# H" by (rule l_repr_imp_subset [OF y x sb]) qed lemma (in group) setmult_subset_G: - "\H \ carrier G; K \ carrier G\ \ H \#> K \ carrier G" + "\H \ carrier G; K \ carrier G\ \ H <#> K \ carrier G" by (auto simp add: set_mult_def subsetD) -lemma (in group) subgroup_mult_id: "subgroup H G \ H \#> H = H" +lemma (in group) subgroup_mult_id: "subgroup H G \ 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 _ "\"]) @@ -549,41 +549,41 @@ qed -subsubsection \Theorems for \\#>\ with \#>\ or \\#\.\ +subsubsection \Theorems for \<#>\ with \#>\ or \<#\.\ lemma (in group) setmult_rcos_assoc: "\H \ carrier G; K \ carrier G; x \ carrier G\ - \ H \#> (K #> x) = (H \#> K) #> x" + \ H <#> (K #> x) = (H <#> K) #> x" by (force simp add: r_coset_def set_mult_def m_assoc) lemma (in group) rcos_assoc_lcos: "\H \ carrier G; K \ carrier G; x \ carrier G\ - \ (H #> x) \#> K = H \#> (x \# K)" + \ (H #> x) <#> K = H <#> (x <# K)" by (force simp add: r_coset_def l_coset_def set_mult_def m_assoc) lemma (in normal) rcos_mult_step1: "\x \ carrier G; y \ carrier G\ - \ (H #> x) \#> (H #> y) = (H \#> (x \# H)) #> y" + \ (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) lemma (in normal) rcos_mult_step2: "\x \ carrier G; y \ carrier G\ - \ (H \#> (x \# H)) #> y = (H \#> (H #> x)) #> y" + \ (H <#> (x <# H)) #> y = (H <#> (H #> x)) #> y" by (insert coset_eq, simp add: normal_def) lemma (in normal) rcos_mult_step3: "\x \ carrier G; y \ carrier G\ - \ (H \#> (H #> x)) #> y = H #> (x \ y)" + \ (H <#> (H #> x)) #> y = H #> (x \ y)" by (simp add: setmult_rcos_assoc coset_mult_assoc subgroup_mult_id normal.axioms subset normal_axioms) lemma (in normal) rcos_sum: "\x \ carrier G; y \ carrier G\ - \ (H #> x) \#> (H #> y) = H #> (x \ y)" + \ (H #> x) <#> (H #> y) = H #> (x \ y)" by (simp add: rcos_mult_step1 rcos_mult_step2 rcos_mult_step3) -lemma (in normal) rcosets_mult_eq: "M \ rcosets H \ H \#> M = M" +lemma (in normal) rcosets_mult_eq: "M \ rcosets H \ H <#> M = M" \ \generalizes \subgroup_mult_id\\ by (auto simp add: RCOSETS_def subset setmult_rcos_assoc subgroup_mult_id normal.axioms normal_axioms) @@ -645,7 +645,7 @@ lemma (in subgroup) l_coset_eq_rcong: assumes "group G" assumes a: "a \ 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 ) @@ -684,7 +684,7 @@ text \The relation is a congruence\ lemma (in normal) congruent_rcong: - shows "congruent2 (rcong H) (rcong H) (\a b. a \ b \# H)" + shows "congruent2 (rcong H) (rcong H) (\a b. a \ b <# H)" proof (intro congruent2I[of "carrier G" _ "carrier G" _] equiv_rcong is_group) fix a b c assume abrcong: "(a, b) \ rcong H" @@ -712,10 +712,10 @@ ultimately have "(inv (a \ c)) \ (b \ c) \ H" by simp from carr and this - have "(b \ c) \ (a \ c) \# H" + have "(b \ c) \ (a \ c) <# H" by (simp add: lcos_module_rev[OF is_group]) from carr and this and is_subgroup - show "(a \ c) \# H = (b \ c) \# H" by (intro l_repr_independence, simp+) + show "(a \ c) <# H = (b \ c) <# H" by (intro l_repr_independence, simp+) next fix a b c assume abrcong: "(a, b) \ rcong H" @@ -746,10 +746,10 @@ have "inv (c \ a) \ (c \ b) \ H" by simp from carr and this - have "(c \ b) \ (c \ a) \# H" + have "(c \ b) \ (c \ a) <# H" by (simp add: lcos_module_rev[OF is_group]) from carr and this and is_subgroup - show "(c \ a) \# H = (c \ b) \# H" by (intro l_repr_independence, simp+) + show "(c \ a) <# H = (c \ b) <# H" by (intro l_repr_independence, simp+) qed @@ -835,7 +835,7 @@ where "FactGroup G H = \carrier = rcosets\<^bsub>G\<^esub> H, mult = set_mult G, one = H\" lemma (in normal) setmult_closed: - "\K1 \ rcosets H; K2 \ rcosets H\ \ K1 \#> K2 \ rcosets H" + "\K1 \ rcosets H; K2 \ rcosets H\ \ K1 <#> K2 \ rcosets H" by (auto simp add: rcos_sum RCOSETS_def) lemma (in normal) setinv_closed: @@ -844,7 +844,7 @@ lemma (in normal) rcosets_assoc: "\M1 \ rcosets H; M2 \ rcosets H; M3 \ rcosets H\ - \ M1 \#> M2 \#> M3 = M1 \#> (M2 \#> M3)" + \ M1 <#> M2 <#> M3 = M1 <#> (M2 <#> M3)" by (auto simp add: RCOSETS_def rcos_sum m_assoc) lemma (in subgroup) subgroup_in_rcosets: @@ -859,7 +859,7 @@ qed lemma (in normal) rcosets_inv_mult_group_eq: - "M \ rcosets H \ set_inv M \#> M = H" + "M \ rcosets H \ set_inv M <#> M = H" by (auto simp add: RCOSETS_def rcos_inv rcos_sum subgroup.subset normal.axioms normal_axioms) theorem (in normal) factorgroup_is_group: @@ -874,7 +874,7 @@ apply (auto dest: rcosets_inv_mult_group_eq simp add: setinv_closed) done -lemma mult_FactGroup [simp]: "X \\<^bsub>(G Mod H)\<^esub> X' = X \#>\<^bsub>G\<^esub> X'" +lemma mult_FactGroup [simp]: "X \\<^bsub>(G Mod H)\<^esub> X' = X <#>\<^bsub>G\<^esub> X'" by (simp add: FactGroup_def) lemma (in normal) inv_FactGroup: @@ -951,11 +951,11 @@ hence all: "\x\X. h x = h g" "\x\X'. h x = h g'" and Xsub: "X \ carrier G" and X'sub: "X' \ carrier G" by (force simp add: kernel_def r_coset_def image_def)+ - hence "h ` (X \#> X') = {h g \\<^bsub>H\<^esub> h g'}" using X X' + hence "h ` (X <#> X') = {h g \\<^bsub>H\<^esub> h g'}" using X X' by (auto dest!: FactGroup_nonempty intro!: image_eqI simp add: set_mult_def subsetD [OF Xsub] subsetD [OF X'sub]) - then show "the_elem (h ` (X \#> X')) = the_elem (h ` X) \\<^bsub>H\<^esub> the_elem (h ` X')" + then show "the_elem (h ` (X <#> X')) = the_elem (h ` X) \\<^bsub>H\<^esub> the_elem (h ` X')" by (auto simp add: all FactGroup_nonempty X X' the_elem_image_unique) qed