# HG changeset patch # User paulson # Date 1548775603 0 # Node ID 10e48c47a549871e1c0f431825be4ce9561b7bdd # Parent 7aafd0472661b8f82727d5fc034ca6d8a854d8f8 some new results in group theory diff -r 7aafd0472661 -r 10e48c47a549 src/HOL/Algebra/Coset.thy --- a/src/HOL/Algebra/Coset.thy Mon Jan 28 18:36:50 2019 -0500 +++ b/src/HOL/Algebra/Coset.thy Tue Jan 29 15:26:43 2019 +0000 @@ -38,6 +38,9 @@ normal_rel :: "['a set, ('a, 'b) monoid_scheme] \ bool" (infixl "\" 60) where "H \ G \ normal H G" +lemma (in comm_group) subgroup_imp_normal: "subgroup A G \ A \ G" + by (simp add: normal_def normal_axioms_def is_group l_coset_def r_coset_def m_comm subgroup.mem_carrier) + (*Next two lemmas contributed by Martin Baillon.*) lemma l_coset_eq_set_mult: @@ -196,6 +199,17 @@ ultimately show ?thesis by blast qed +lemma (in group_hom) inj_on_one_iff: + "inj_on h (carrier G) \ (\x. x \ carrier G \ h x = one H \ x = one G)" +using G.solve_equation G.subgroup_self by (force simp: inj_on_def) + +lemma inj_on_one_iff': + "\h \ hom G H; group G; group H\ \ inj_on h (carrier G) \ (\x. x \ carrier G \ h x = one H \ x = one G)" + using group_hom.inj_on_one_iff group_hom.intro group_hom_axioms.intro by blast + +lemma (in group_hom) iso_iff: "h \ iso G H \ carrier H \ h ` carrier G \ (\x\carrier G. h x = \\<^bsub>H\<^esub> \ x = \)" + by (auto simp: iso_def bij_betw_def inj_on_one_iff) + lemma (in group) repr_independence: assumes "y \ H #> x" "x \ carrier G" "subgroup H G" shows "H #> x = H #> y" using assms diff -r 7aafd0472661 -r 10e48c47a549 src/HOL/Algebra/Generated_Groups.thy --- a/src/HOL/Algebra/Generated_Groups.thy Mon Jan 28 18:36:50 2019 -0500 +++ b/src/HOL/Algebra/Generated_Groups.thy Tue Jan 29 15:26:43 2019 +0000 @@ -137,17 +137,24 @@ show "generate G { a } \ { a [^] (k :: int) | k. k \ UNIV }" proof fix h assume "h \ generate G { a }" hence "\k :: int. h = a [^] k" - proof (induction, metis int_pow_0[of a], metis singletonD int_pow_1[OF assms]) + proof (induction) + case one + then show ?case + using int_pow_0 [of G] by metis + next + case (incl h) + then show ?case + by (metis assms int_pow_1 singletonD) + next case (inv h) - hence "inv h = a [^] ((- 1) :: int)" - using assms unfolding int_pow_def2 by simp - thus ?case - by blast + then show ?case + by (metis assms int_pow_1 int_pow_neg singletonD) next - case eng thus ?case + case (eng h1 h2) + then show ?case using assms by (metis int_pow_mult) qed - thus "h \ { a [^] (k :: int) | k. k \ UNIV }" + then show "h \ { a [^] (k :: int) | k. k \ UNIV }" by blast qed qed @@ -436,4 +443,124 @@ assumes "K \ carrier G" shows "(derived H ^^ n) (h ` K) = h ` ((derived G ^^ n) K)" using derived_img[OF G.exp_of_derived_in_carrier[OF assms]] by (induct n) (auto) + +subsection \Generated subgroup of a group\ + +definition subgroup_generated :: "('a, 'b) monoid_scheme \ 'a set \ ('a, 'b) monoid_scheme" + where "subgroup_generated G S = G\carrier := generate G (carrier G \ S)\" + +lemma carrier_subgroup_generated: "carrier (subgroup_generated G S) = generate G (carrier G \ S)" + by (auto simp: subgroup_generated_def) + +lemma (in group) group_subgroup_generated [simp]: "group (subgroup_generated G S)" + unfolding subgroup_generated_def + by (simp add: generate_is_subgroup subgroup_imp_group) + +lemma (in group) abelian_subgroup_generated: + assumes "comm_group G" + shows "comm_group (subgroup_generated G S)" (is "comm_group ?GS") +proof (rule group.group_comm_groupI) + show "Group.group ?GS" + by simp +next + fix x y + assume "x \ carrier ?GS" "y \ carrier ?GS" + with assms show "x \\<^bsub>?GS\<^esub> y = y \\<^bsub>?GS\<^esub> x" + apply (simp add: subgroup_generated_def) + by (meson Int_lower1 comm_groupE(4) generate_in_carrier) +qed + +lemma (in group) subgroup_of_subgroup_generated: + assumes "H \ B" "subgroup H G" + shows "subgroup H (subgroup_generated G B)" +proof unfold_locales + fix x + assume "x \ H" + with assms show "inv\<^bsub>subgroup_generated G B\<^esub> x \ H" + unfolding subgroup_generated_def + by (metis IntI Int_commute Int_lower2 generate.incl generate_is_subgroup m_inv_consistent subgroup_def subsetCE) +next + show "H \ carrier (subgroup_generated G B)" + using assms apply (auto simp: carrier_subgroup_generated) + by (metis Int_iff generate.incl inf.orderE subgroup.mem_carrier) +qed (use assms in \auto simp: subgroup_generated_def subgroup_def\) + +lemma carrier_subgroup_generated_alt: + assumes "Group.group G" "S \ carrier G" + shows "carrier (subgroup_generated G S) = \{H. subgroup H G \ carrier G \ S \ H}" + using assms by (auto simp: group.generate_minimal subgroup_generated_def) + +lemma one_subgroup_generated [simp]: "\\<^bsub>subgroup_generated G S\<^esub> = \\<^bsub>G\<^esub>" + by (auto simp: subgroup_generated_def) + +lemma mult_subgroup_generated [simp]: "mult (subgroup_generated G S) = mult G" + by (auto simp: subgroup_generated_def) + +lemma (in group) inv_subgroup_generated [simp]: + assumes "f \ carrier (subgroup_generated G S)" + shows "inv\<^bsub>subgroup_generated G S\<^esub> f = inv f" +proof (rule group.inv_equality) + show "Group.group (subgroup_generated G S)" + by simp + have [simp]: "f \ carrier G" + by (metis Int_lower1 assms carrier_subgroup_generated generate_in_carrier) + show "inv f \\<^bsub>subgroup_generated G S\<^esub> f = \\<^bsub>subgroup_generated G S\<^esub>" + by (simp add: assms) + show "f \ carrier (subgroup_generated G S)" + using assms by (simp add: generate.incl subgroup_generated_def) + show "inv f \ carrier (subgroup_generated G S)" + using assms by (simp add: subgroup_generated_def generate_m_inv_closed) +qed + +lemma subgroup_generated_restrict [simp]: + "subgroup_generated G (carrier G \ S) = subgroup_generated G S" + by (simp add: subgroup_generated_def) + +lemma (in subgroup) carrier_subgroup_generated_subgroup [simp]: + "carrier (subgroup_generated G H) = H" + by (auto simp: generate.incl carrier_subgroup_generated elim: generate.induct) + + +lemma (in group) subgroup_subgroup_generated_iff: + "subgroup H (subgroup_generated G B) \ subgroup H G \ H \ carrier(subgroup_generated G B)" + (is "?lhs = ?rhs") +proof + assume L: ?lhs + then have Hsub: "H \ generate G (carrier G \ B)" + by (simp add: subgroup_def subgroup_generated_def) + then have H: "H \ carrier G" "H \ carrier(subgroup_generated G B)" + unfolding carrier_subgroup_generated + using generate_incl by blast+ + with Hsub have "subgroup H G" + by (metis Int_commute Int_lower2 L carrier_subgroup_generated generate_consistent + generate_is_subgroup inf.orderE subgroup.carrier_subgroup_generated_subgroup subgroup_generated_def) + with H show ?rhs + by blast +next + assume ?rhs + then show ?lhs + by (simp add: generate_is_subgroup subgroup_generated_def subgroup_incl) +qed + + +lemma pow_subgroup_generated: + "pow (subgroup_generated G S) = (pow G :: 'a \ nat \ 'a)" +proof - + have "x [^]\<^bsub>subgroup_generated G S\<^esub> n = x [^]\<^bsub>G\<^esub> n" for x and n::nat + by (induction n) auto + then show ?thesis + by force +qed + +lemma (in group) int_pow_subgroup_generated: + fixes n::int + assumes "x \ carrier (subgroup_generated G S)" + shows "x [^]\<^bsub>subgroup_generated G S\<^esub> n = x [^]\<^bsub>G\<^esub> n" +proof - + have "x [^] nat (- n) \ carrier (subgroup_generated G S)" + by (metis assms group.is_monoid group_subgroup_generated monoid.nat_pow_closed pow_subgroup_generated) + then show ?thesis + by (simp add: int_pow_def2 pow_subgroup_generated) +qed + end \ No newline at end of file diff -r 7aafd0472661 -r 10e48c47a549 src/HOL/Algebra/Group.thy --- a/src/HOL/Algebra/Group.thy Mon Jan 28 18:36:50 2019 -0500 +++ b/src/HOL/Algebra/Group.thy Tue Jan 29 15:26:43 2019 +0000 @@ -30,24 +30,6 @@ \ \The set of invertible elements\ where "Units G = {y. y \ carrier G \ (\x \ carrier G. x \\<^bsub>G\<^esub> y = \\<^bsub>G\<^esub> \ y \\<^bsub>G\<^esub> x = \\<^bsub>G\<^esub>)}" -consts - pow :: "[('a, 'm) monoid_scheme, 'a, 'b::semiring_1] => 'a" (infixr "[^]\" 75) - -overloading nat_pow == "pow :: [_, 'a, nat] => 'a" -begin - definition "nat_pow G a n = rec_nat \\<^bsub>G\<^esub> (%u b. b \\<^bsub>G\<^esub> a) n" -end - -overloading int_pow == "pow :: [_, 'a, int] => 'a" -begin - definition "int_pow G a z = - (let p = rec_nat \\<^bsub>G\<^esub> (%u b. b \\<^bsub>G\<^esub> a) - in if z < 0 then inv\<^bsub>G\<^esub> (p (nat (-z))) else p (nat z))" -end - -lemma int_pow_int: "x [^]\<^bsub>G\<^esub> (int n) = x [^]\<^bsub>G\<^esub> n" -by(simp add: int_pow_def nat_pow_def) - locale monoid = fixes G (structure) assumes m_closed [intro, simp]: @@ -191,46 +173,6 @@ lemma (in monoid) carrier_not_empty: "carrier G \ {}" by auto -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) = \" - 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]: - "\ [^] (n::nat) = \" - 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_comm: - "x \ carrier G \ (x [^] (n::nat)) \ (x [^] (m :: nat)) = (x [^] m) \ (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 \ carrier G \ x [^] (Suc n) = x \ (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 \ 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 \ carrier := H \)\<^esub> n" - unfolding nat_pow_def by simp - - (* Jacobson defines submonoid here. *) (* Jacobson defines the order of a monoid here. *) @@ -340,6 +282,20 @@ subsection \Cancellation Laws and Basic Properties\ +lemma (in group) inv_eq_1_iff [simp]: + assumes "x \ carrier G" shows "inv\<^bsub>G\<^esub> x = \\<^bsub>G\<^esub> \ x = \\<^bsub>G\<^esub>" +proof - + have "x = \" if "inv x = \" + proof - + have "inv x \ x = \" + using assms l_inv by blast + then show "x = \" + using that assms by simp + qed + then show ?thesis + by auto +qed + lemma (in group) r_inv [simp]: "x \ carrier G ==> x \ inv x = \" by simp @@ -374,30 +330,114 @@ "[|y \ x = \; x \ carrier G; y \ carrier G|] ==> inv x = y" using inv_unique r_inv by blast -(* Contributed by Joachim Breitner *) lemma (in group) inv_solve_left: "\ a \ carrier G; b \ carrier G; c \ carrier G \ \ a = inv b \ c \ c = b \ a" by (metis inv_equality l_inv_ex l_one m_assoc r_inv) + +lemma (in group) inv_solve_left': + "\ a \ carrier G; b \ carrier G; c \ carrier G \ \ inv b \ c = a \ c = b \ a" + by (metis inv_equality l_inv_ex l_one m_assoc r_inv) + lemma (in group) inv_solve_right: "\ a \ carrier G; b \ carrier G; c \ carrier G \ \ a = b \ inv c \ b = a \ c" by (metis inv_equality l_inv_ex l_one m_assoc r_inv) -text \Power\ +lemma (in group) inv_solve_right': + "\a \ carrier G; b \ carrier G; c \ carrier G\ \ b \ inv c = a \ b = a \ c" + by (auto simp: m_assoc) + + +subsection \Power\ + +consts + pow :: "[('a, 'm) monoid_scheme, 'a, 'b::semiring_1] => 'a" (infixr "[^]\" 75) + +overloading nat_pow == "pow :: [_, 'a, nat] => 'a" +begin + definition "nat_pow G a n = rec_nat \\<^bsub>G\<^esub> (%u b. b \\<^bsub>G\<^esub> a) n" +end + +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) = \" + 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]: + "\ [^] (n::nat) = \" + 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_comm: + "x \ carrier G \ (x [^] (n::nat)) \ (x [^] (m :: nat)) = (x [^] m) \ (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 \ carrier G \ x [^] (Suc n) = x \ (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 group) int_pow_def2: - "a [^] (z::int) = (if z < 0 then inv (a [^] (nat (-z))) else a [^] (nat z))" - by (simp add: int_pow_def nat_pow_def Let_def) +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) + +lemma (in monoid) nat_pow_consistent: + "x [^] (n :: nat) = x [^]\<^bsub>(G \ carrier := H \)\<^esub> n" + unfolding nat_pow_def by simp + +lemma nat_pow_0 [simp]: "x [^]\<^bsub>G\<^esub> (0::nat) = \\<^bsub>G\<^esub>" + by (simp add: nat_pow_def) + +lemma nat_pow_Suc [simp]: "x [^]\<^bsub>G\<^esub> (Suc n) = (x [^]\<^bsub>G\<^esub> n)\\<^bsub>G\<^esub> x" + by (simp add: nat_pow_def) -lemma (in group) int_pow_0 [simp]: - "x [^] (0::int) = \" - by (simp add: int_pow_def2) +lemma (in group) nat_pow_inv: + assumes "x \ carrier G" shows "(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) \ inv x" + by simp + also have " ... = (inv (x [^] i)) \ inv x" + by (simp add: Suc.IH Suc.prems) + also have " ... = inv (x \ (x [^] i))" + by (simp add: assms inv_mult_group) + also have " ... = inv (x [^] (Suc i))" + using assms nat_pow_Suc2 by auto + finally show ?case . +qed + +overloading int_pow == "pow :: [_, 'a, int] => 'a" +begin + definition "int_pow G a z = + (let p = rec_nat \\<^bsub>G\<^esub> (%u b. b \\<^bsub>G\<^esub> a) + in if z < 0 then inv\<^bsub>G\<^esub> (p (nat (-z))) else p (nat z))" +end + +lemma int_pow_int: "x [^]\<^bsub>G\<^esub> (int n) = x [^]\<^bsub>G\<^esub> n" +by(simp add: int_pow_def nat_pow_def) + +lemma int_pow_0 [simp]: "x [^]\<^bsub>G\<^esub> (0::int) = \\<^bsub>G\<^esub>" + by (simp add: int_pow_def) + +lemma int_pow_def2: "a [^]\<^bsub>G\<^esub> z = + (if z < 0 then inv\<^bsub>G\<^esub> (a [^]\<^bsub>G\<^esub> (nat (-z))) else a [^]\<^bsub>G\<^esub> (nat z))" + by (simp add: int_pow_def nat_pow_def) lemma (in group) int_pow_one [simp]: "\ [^] (z::int) = \" by (simp add: int_pow_def2) -(* The following are contributed by Joachim Breitner *) - lemma (in group) int_pow_closed [intro, simp]: "x \ carrier G ==> x [^] (i::int) \ carrier G" by (simp add: int_pow_def2) @@ -418,23 +458,6 @@ by (auto simp add: assms int_pow_def2 inv_solve_left inv_solve_right nat_add_distrib [symmetric] nat_pow_mult ) qed -lemma (in group) nat_pow_inv: - assumes "x \ carrier G" shows "(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) \ inv x" - by simp - also have " ... = (inv (x [^] i)) \ inv x" - by (simp add: Suc.IH Suc.prems) - also have " ... = inv (x \ (x [^] i))" - by (simp add: assms inv_mult_group) - also have " ... = inv (x [^] (Suc i))" - using assms nat_pow_Suc2 by auto - finally show ?case . -qed - lemma (in group) int_pow_inv: "x \ carrier G \ (inv x) [^] (i :: int) = inv (x [^] i)" by (simp add: nat_pow_inv int_pow_def2) @@ -446,7 +469,7 @@ assume n_ge: "n \ 0" thus ?thesis proof (cases) assume m_ge: "m \ 0" thus ?thesis - using n_ge nat_pow_pow[OF assms, of "nat n" "nat m"] int_pow_def2 + using n_ge nat_pow_pow[OF assms, of "nat n" "nat m"] int_pow_def2 [where G=G] by (simp add: mult_less_0_iff nat_mult_distrib) next assume m_lt: "\ m \ 0" @@ -592,7 +615,7 @@ assumes "subgroup H G" shows "H \ Units (G \ carrier := H \)" using group.Units[OF subgroup.subgroup_is_group[OF assms group_axioms]] by simp -lemma (in group) m_inv_consistent: +lemma (in group) m_inv_consistent [simp]: assumes "subgroup H G" "x \ H" shows "inv\<^bsub>(G \ carrier := H \)\<^esub> x = inv x" using assms m_inv_monoid_consistent[OF _ subgroup_is_submonoid] subgroup_Units[of H] by auto @@ -603,16 +626,16 @@ proof (cases) assume ge: "n \ 0" hence "x [^] n = x [^] (nat n)" - using int_pow_def2 by auto + using int_pow_def2 [of G] by auto also have " ... = x [^]\<^bsub>(G \ carrier := H \)\<^esub> (nat n)" using nat_pow_consistent by simp also have " ... = x [^]\<^bsub>(G \ carrier := H \)\<^esub> n" - using group.int_pow_def2[OF subgroup.subgroup_is_group[OF assms(1) is_group]] ge by auto + by (metis ge int_nat_eq int_pow_int) finally show ?thesis . next assume "\ n \ 0" hence lt: "n < 0" by simp hence "x [^] n = inv (x [^] (nat (- n)))" - using int_pow_def2 by auto + using int_pow_def2 [of G] by auto also have " ... = (inv x) [^] (nat (- n))" by (metis assms nat_pow_inv subgroup.mem_carrier) also have " ... = (inv\<^bsub>(G \ carrier := H \)\<^esub> x) [^]\<^bsub>(G \ carrier := H \)\<^esub> (nat (- n))" @@ -620,7 +643,7 @@ also have " ... = inv\<^bsub>(G \ carrier := H \)\<^esub> (x [^]\<^bsub>(G \ carrier := H \)\<^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 \ carrier := H \)\<^esub> n" - using group.int_pow_def2[OF subgroup.subgroup_is_group[OF assms(1) is_group]] lt by auto + by (simp add: int_pow_def2 lt) finally show ?thesis . qed @@ -781,6 +804,17 @@ {h. h \ carrier G \ carrier H \ (\x \ carrier G. \y \ carrier G. h (x \\<^bsub>G\<^esub> y) = h x \\<^bsub>H\<^esub> h y)}" +lemma homI: + "\\x. x \ carrier G \ h x \ carrier H; + \x y. \x \ carrier G; y \ carrier G\ \ h (x \\<^bsub>G\<^esub> y) = h x \\<^bsub>H\<^esub> h y\ \ h \ hom G H" + by (auto simp: hom_def) + +lemma hom_carrier: "h \ hom G H \ h ` carrier G \ carrier H" + by (auto simp: hom_def) + +lemma hom_in_carrier: "\h \ hom G H; x \ carrier G\ \ h x \ carrier H" + by (auto simp: hom_def) + lemma hom_compose: "\ f \ hom G H; g \ hom H I \ \ g \ f \ hom G I" unfolding hom_def by (auto simp add: Pi_iff) @@ -793,20 +827,50 @@ "[|h \ hom G H; i \ hom H I|] ==> compose (carrier G) i h \ hom G I" by (fastforce simp add: hom_def compose_def) -definition - iso :: "_ => _ => ('a => 'b) set" +definition iso :: "_ => _ => ('a => 'b) set" where "iso G H = {h. h \ hom G H \ bij_betw h (carrier G) (carrier H)}" -definition - is_iso :: "_ \ _ \ bool" (infixr "\" 60) +definition is_iso :: "_ \ _ \ bool" (infixr "\" 60) where "G \ H = (iso G H \ {})" +definition mon where "mon G H = {f \ hom G H. inj_on f (carrier G)}" + +definition epi where "epi G H = {f \ hom G H. f ` (carrier G) = carrier H}" + +lemma isoI: + "\h \ hom G H; bij_betw h (carrier G) (carrier H)\ \ h \ iso G H" + by (auto simp: iso_def) + +lemma epi_iff_subset: + "f \ epi G G' \ f \ hom G G' \ carrier G' \ f ` carrier G" + by (auto simp: epi_def hom_def) + +lemma iso_iff_mon_epi: "f \ iso G H \ f \ mon G H \ f \ epi G H" + by (auto simp: iso_def mon_def epi_def bij_betw_def) + lemma iso_set_refl: "(\x. x) \ iso G G" by (simp add: iso_def hom_def inj_on_def bij_betw_def Pi_def) +lemma id_iso: "id \ iso G G" + by (simp add: iso_def hom_def inj_on_def bij_betw_def Pi_def) + corollary iso_refl : "G \ G" using iso_set_refl unfolding is_iso_def by auto +lemma trivial_hom: + "group H \ (\x. one H) \ hom G H" + by (auto simp: hom_def Group.group_def) + +lemma (in group) hom_eq: + assumes "f \ hom G H" "\x. x \ carrier G \ f' x = f x" + shows "f' \ hom G H" + using assms by (auto simp: hom_def) + +lemma (in group) iso_eq: + assumes "f \ iso G H" "\x. x \ carrier G \ f' x = f x" + shows "f' \ iso G H" + using assms by (fastforce simp: iso_def inj_on_def bij_betw_def hom_eq image_iff) + lemma (in group) iso_set_sym: assumes "h \ iso G H" shows "inv_into (carrier G) h \ iso H G" @@ -846,10 +910,8 @@ corollary (in group) iso_trans: "\G \ H ; H \ I\ \ G \ I" using iso_set_trans unfolding is_iso_def by blast -(* NEW ====================================================================== *) lemma iso_same_card: "G \ H \ card (carrier G) = card (carrier H)" using bij_betw_same_card unfolding is_iso_def iso_def by auto -(* ========================================================================== *) (* Next four lemmas contributed by Paulo. *) @@ -1015,12 +1077,13 @@ shows "G \\ H \ G2 \\ I" using DirProd_iso_set_trans assms unfolding is_iso_def by blast +subsection\The locale for a homomorphism between two groups\ text\Basis for homomorphism proofs: we assume two groups \<^term>\G\ and \<^term>\H\, with a homomorphism \<^term>\h\ between them\ locale group_hom = G?: group G + H?: group H for G (structure) and H (structure) + fixes h - assumes homh: "h \ hom G H" + assumes homh [simp]: "h \ hom G H" lemma (in group_hom) hom_mult [simp]: "[| x \ carrier G; y \ carrier G |] ==> h (x \\<^bsub>G\<^esub> y) = h x \\<^bsub>H\<^esub> h y" @@ -1036,7 +1099,7 @@ with homh [unfolded hom_def] show ?thesis by auto qed -lemma (in group_hom) one_closed [simp]: "h \ \ carrier H" +lemma (in group_hom) one_closed: "h \ \ carrier H" by simp lemma (in group_hom) hom_one [simp]: "h \ = \\<^bsub>H\<^esub>" @@ -1046,6 +1109,16 @@ then show ?thesis by (simp del: r_one) qed +lemma hom_one: + assumes "h \ hom G H" "group G" "group H" + shows "h (one G) = one H" + apply (rule group_hom.hom_one) + by (simp add: assms group_hom_axioms_def group_hom_def) + +lemma hom_mult: + "\h \ hom G H; x \ carrier G; y \ carrier G\ \ h (x \\<^bsub>G\<^esub> y) = h x \\<^bsub>H\<^esub> h y" + by (auto simp: hom_def) + lemma (in group_hom) inv_closed [simp]: "x \ carrier G ==> h (inv x) \ carrier H" by simp @@ -1110,7 +1183,7 @@ lemma (in group_hom) int_pow_hom: "x \ carrier G \ h (x [^] (n :: int)) = (h x) [^]\<^bsub>H\<^esub> n" - using int_pow_def2 nat_pow_hom by (simp add: G.int_pow_def2) + using nat_pow_hom by (simp add: int_pow_def2) subsection \Commutative Structures\ diff -r 7aafd0472661 -r 10e48c47a549 src/HOL/Algebra/Multiplicative_Group.thy --- a/src/HOL/Algebra/Multiplicative_Group.thy Mon Jan 28 18:36:50 2019 -0500 +++ b/src/HOL/Algebra/Multiplicative_Group.thy Tue Jan 29 15:26:43 2019 +0000 @@ -433,7 +433,7 @@ fix b assume "b \ { a [^] k | k. k \ (UNIV :: nat set) }" then obtain k :: nat where "b = a [^] k" by blast hence "b = a [^] (int k)" - using int_pow_def2 by auto + by (simp add: int_pow_int) thus "b \ generate G { a }" unfolding generate_pow[OF assms(2)] by blast qed @@ -447,12 +447,12 @@ proof (cases "k < 0") assume "\ k < 0" hence "b = a [^] (nat k)" - using k int_pow_def2 by auto + by (simp add: int_pow_def2 k) thus ?thesis by blast next assume "k < 0" hence b: "b = inv (a [^] (nat (- k)))" - using k int_pow_def2 by auto + using k int_pow_def2[of G] by auto obtain m where m: "ord a * m \ nat (- k)" by (metis assms mult.left_neutral mult_le_mono1 ord_ge_1)