# HG changeset patch # User paulson # Date 1538663147 -3600 # Node ID 1b5178abaf97c06db00a35f79610cf8379836c2e # Parent 12dce58bcd3f970563c6465d6c52684cfb0771cb updates to Algebra from Baillon and de Vilhena diff -r 12dce58bcd3f -r 1b5178abaf97 src/HOL/Algebra/Chinese_Remainder.thy --- a/src/HOL/Algebra/Chinese_Remainder.thy Thu Oct 04 11:18:39 2018 +0200 +++ b/src/HOL/Algebra/Chinese_Remainder.thy Thu Oct 04 15:25:47 2018 +0100 @@ -3,1140 +3,507 @@ *) theory Chinese_Remainder - imports QuotRing Ideal_Product + imports Weak_Morphisms Ideal_Product + begin -section \Chinese Remainder Theorem\ -subsection \Direct Product of Rings\ +section \Direct Product of Rings\ + +subsection \Definitions\ -definition - RDirProd :: "[ ('a, 'n) ring_scheme, ('b, 'm) ring_scheme ] \ ('a \ 'b) ring" - where "RDirProd R S = - \ carrier = carrier R \ carrier S, - mult = (\(r, s). \(r', s'). (r \\<^bsub>R\<^esub> r', s \\<^bsub>S\<^esub> s')), - one = (\\<^bsub>R\<^esub>, \\<^bsub>S\<^esub>), - zero = (\\<^bsub>R\<^esub>, \\<^bsub>S\<^esub>), - add = (\(r, s). \(r', s'). (r \\<^bsub>R\<^esub> r', s \\<^bsub>S\<^esub> s')) \" +definition RDirProd :: "('a, 'n) ring_scheme \ ('b, 'm) ring_scheme \ ('a \ 'b) ring" + where "RDirProd R S = monoid.extend (R \\ S) + \ zero = one ((add_monoid R) \\ (add_monoid S)), + add = mult ((add_monoid R) \\ (add_monoid S)) \ " + +abbreviation nil_ring :: "('a list) ring" + where "nil_ring \ monoid.extend nil_monoid \ zero = [], add = (\a b. []) \" + +definition RDirProd_list :: "(('a, 'n) ring_scheme) list \ ('a list) ring" + where "RDirProd_list Rs = foldr (\R S. image_ring (\(a, as). a # as) (RDirProd R S)) Rs nil_ring" + -lemma RDirProd_monoid: - assumes "ring R" and "ring S" - shows "monoid (RDirProd R S)" - by (rule monoidI) (auto simp add: RDirProd_def assms ring.ring_simprules ring.is_monoid) +subsection \Basic Properties\ + +lemma RDirProd_carrier: "carrier (RDirProd R S) = carrier R \ carrier S" + unfolding RDirProd_def DirProd_def by (simp add: monoid.defs) + +lemma RDirProd_add_monoid [simp]: "add_monoid (RDirProd R S) = (add_monoid R) \\ (add_monoid S)" + by (simp add: RDirProd_def monoid.defs) -lemma RDirProd_abelian_group: - assumes "ring R" and "ring S" - shows "abelian_group (RDirProd R S)" - by (auto intro!: abelian_groupI - simp add: RDirProd_def assms ring.ring_simprules) - (meson assms ring.ring_simprules(3,16))+ - -lemma RDirProd_group: - assumes "ring R" and "ring S" - shows "ring (RDirProd R S)" +lemma RDirProd_ring: + assumes "ring R" and "ring S" shows "ring (RDirProd R S)" proof - + have "monoid (RDirProd R S)" + using DirProd_monoid[OF assms[THEN ring.axioms(2)]] unfolding monoid_def + by (auto simp add: DirProd_def RDirProd_def monoid.defs) + then interpret Prod: group "add_monoid (RDirProd R S)" + monoid "RDirProd R S" + using DirProd_group[OF assms[THEN abelian_group.a_group[OF ring.is_abelian_group]]] + unfolding RDirProd_add_monoid by auto show ?thesis - apply (rule ringI) - apply (simp_all add: assms RDirProd_abelian_group RDirProd_monoid) - by (auto simp add: RDirProd_def assms ring.ring_simprules) + by (unfold_locales, auto simp add: RDirProd_def DirProd_def monoid.defs assms ring.ring_simprules) qed -lemma RDirProd_isomorphism1: +lemma RDirProd_iso1: "(\(x, y). (y, x)) \ ring_iso (RDirProd R S) (RDirProd S R)" - unfolding ring_iso_def ring_hom_def bij_betw_def inj_on_def RDirProd_def by auto + unfolding ring_iso_def ring_hom_def bij_betw_def inj_on_def + by (auto simp add: RDirProd_def DirProd_def monoid.defs) -lemma RDirProd_isomorphism2: +lemma RDirProd_iso2: "(\(x, (y, z)). ((x, y), z)) \ ring_iso (RDirProd R (RDirProd S T)) (RDirProd (RDirProd R S) T)" - unfolding ring_iso_def ring_hom_def bij_betw_def inj_on_def RDirProd_def - by (auto simp add: image_iff) + unfolding ring_iso_def ring_hom_def bij_betw_def inj_on_def + by (auto simp add: image_iff RDirProd_def DirProd_def monoid.defs) -lemma RDirProd_isomorphism3: +lemma RDirProd_iso3: "(\((x, y), z). (x, (y, z))) \ ring_iso (RDirProd (RDirProd R S) T) (RDirProd R (RDirProd S T))" - unfolding ring_iso_def ring_hom_def bij_betw_def inj_on_def RDirProd_def - by (auto simp add: image_iff) + unfolding ring_iso_def ring_hom_def bij_betw_def inj_on_def + by (auto simp add: image_iff RDirProd_def DirProd_def monoid.defs) -lemma RDirProd_isomorphism4: - assumes "f \ ring_iso R S" - shows "(\(r, t). (f r, t)) \ ring_iso (RDirProd R T) (RDirProd S T)" - using assms unfolding ring_iso_def ring_hom_def bij_betw_def inj_on_def RDirProd_def - by (auto simp add: image_iff)+ +lemma RDirProd_iso4: + assumes "f \ ring_iso R S" shows "(\(r, t). (f r, t)) \ ring_iso (RDirProd R T) (RDirProd S T)" + using assms unfolding ring_iso_def ring_hom_def bij_betw_def inj_on_def + by (auto simp add: image_iff RDirProd_def DirProd_def monoid.defs)+ -lemma RDirProd_isomorphism5: - assumes "f \ ring_iso S T" - shows "(\(r, s). (r, f s)) \ ring_iso (RDirProd R S) (RDirProd R T)" - using ring_iso_set_trans[OF ring_iso_set_trans[OF RDirProd_isomorphism1[of R S] - RDirProd_isomorphism4[OF assms, of R]] - RDirProd_isomorphism1[of T R]] +lemma RDirProd_iso5: + assumes "f \ ring_iso S T" shows "(\(r, s). (r, f s)) \ ring_iso (RDirProd R S) (RDirProd R T)" + using ring_iso_set_trans[OF ring_iso_set_trans[OF RDirProd_iso1 RDirProd_iso4[OF assms]] RDirProd_iso1] by (simp add: case_prod_unfold comp_def) -lemma RDirProd_isomorphism6: +lemma RDirProd_iso6: assumes "f \ ring_iso R R'" and "g \ ring_iso S S'" shows "(\(r, s). (f r, g s)) \ ring_iso (RDirProd R S) (RDirProd R' S')" - using ring_iso_set_trans[OF RDirProd_isomorphism4[OF assms(1)] RDirProd_isomorphism5[OF assms(2)]] + using ring_iso_set_trans[OF RDirProd_iso4[OF assms(1)] RDirProd_iso5[OF assms(2)]] + by (simp add: case_prod_beta' comp_def) + +lemma RDirProd_iso7: + shows "(\a. (a, [])) \ ring_iso R (RDirProd R nil_ring)" + unfolding ring_iso_def ring_hom_def bij_betw_def inj_on_def + by (auto simp add: RDirProd_def DirProd_def monoid.defs) + +lemma RDirProd_hom1: + shows "(\a. (a, a)) \ ring_hom R (RDirProd R R)" + by (auto simp add: ring_hom_def RDirProd_def DirProd_def monoid.defs) + +lemma RDirProd_hom2: + assumes "f \ ring_hom S T" + shows "(\(x, y). (x, f y)) \ ring_hom (RDirProd R S) (RDirProd R T)" + and "(\(x, y). (f x, y)) \ ring_hom (RDirProd S R) (RDirProd T R)" + using assms by (auto simp add: ring_hom_def RDirProd_def DirProd_def monoid.defs) + +lemma RDirProd_hom3: + assumes "f \ ring_hom R R'" and "g \ ring_hom S S'" + shows "(\(r, s). (f r, g s)) \ ring_hom (RDirProd R S) (RDirProd R' S')" + using ring_hom_trans[OF RDirProd_hom2(2)[OF assms(1)] RDirProd_hom2(1)[OF assms(2)]] by (simp add: case_prod_beta' comp_def) -subsection \Simple Version of The Theorem\ - -text \We start by proving a simpler version of the theorem. The rest of the theory is - dedicated to its generalization\ +subsection \Direct Product of a List of Rings\ -lemma (in ideal) set_add_zero: - assumes "i \ I" - shows "I +> i = I" - by (simp add: a_rcos_const assms) +lemma RDirProd_list_nil [simp]: "RDirProd_list [] = nil_ring" + unfolding RDirProd_list_def by simp -lemma (in ideal) set_add_zero_imp_mem: - assumes "i \ carrier R" "I +> i = I" - shows "i \ I" - using a_rcos_self assms(1-2) by auto +lemma nil_ring_simprules [simp]: + "carrier nil_ring = { [] }" and "one nil_ring = []" and "zero nil_ring = []" + by (auto simp add: monoid.defs) -lemma (in ring) canonical_proj_is_surj: - assumes "ideal I R" "ideal J R" "I <+> J = carrier R" - shows "\x y. \ x \ carrier R; y \ carrier R \ \ - \a \ carrier R. I +> a = I +> x \ J +> a = J +> y" -proof - - { fix I J i j assume A: "ideal I R" "ideal J R" "i \ I" "j \ J" "\ = i \ j" - have "I +> \ = I +> j" - proof - - have "I +> \ = I +> (i \ j)" using A(5) by simp - also have " ... = (I +> i) <+> (I +> j)" - by (metis abelian_subgroup.a_rcos_sum abelian_subgroupI3 A(1-4) - ideal.Icarr ideal.axioms(1) is_abelian_group) - also have " ... = (I +> \) <+> (I +> j)" - using ideal.set_add_zero[OF A(1) A(3)] - by (simp add: A(1) additive_subgroup.a_subset ideal.axioms(1)) - also have " ... = I +> (\ \ j)" - by (meson A abelian_subgroup.a_rcos_sum abelian_subgroupI3 - additive_subgroup.a_Hcarr ideal.axioms(1) is_abelian_group zero_closed) - finally show "I +> \ = I +> j" - using A(2) A(4) ideal.Icarr by fastforce - qed } note aux_lemma = this - - fix x y assume x: "x \ carrier R" and y: "y \ carrier R" - - have "\ \ I <+> J" using assms by simp - then obtain i j where i: "i \ I" and j: "j \ J" and ij: "\ = i \ j" - using set_add_def'[of R I J] by auto - have mod_I: "I +> j = I +> \" and mod_J: "J +> i = J +> \" - using aux_lemma[OF assms(1-2) i j ij] apply simp - using aux_lemma[OF assms(2) assms(1) j i] ij - by (metis add.m_comm assms(1) assms(2) i ideal.Icarr j) - - have "I +> ((j \ x) \ (i \ y)) = (I +> (j \ x)) <+> (I +> (i \ y))" - by (metis abelian_subgroup.a_rcos_sum abelian_subgroupI3 assms i ideal.Icarr - ideal.axioms(1) is_abelian_group j m_closed x y) - also have " ... = (I +> (j \ x)) <+> (I +> \)" - using ideal.set_add_zero[OF assms(1), of "i \ y"] i assms(1) - by (simp add: additive_subgroup.a_subset ideal.I_r_closed ideal.axioms(1) y) - also have " ... = I +> (j \ x)" - by (metis abelian_subgroup.a_rcos_sum abelian_subgroupI3 additive_subgroup.a_Hcarr assms(1-2) - ideal.axioms(1) is_abelian_group j m_closed r_zero x zero_closed) - finally have Ix: "I +> ((j \ x) \ (i \ y)) = I +> x" using mod_I - by (metis (full_types) assms ideal.Icarr ideal.rcoset_mult_add is_monoid j monoid.l_one one_closed x) - have "J +> ((j \ x) \ (i \ y)) = (J +> (j \ x)) <+> (J +> (i \ y))" - by (metis abelian_subgroup.a_rcos_sum abelian_subgroupI3 assms i ideal.Icarr - ideal.axioms(1) is_abelian_group j m_closed x y) - also have " ... = (J +> \) <+> (J +> (i \ y))" - using ideal.set_add_zero[OF assms(2), of "j \ x"] j assms(2) - by (simp add: additive_subgroup.a_subset ideal.I_r_closed ideal.axioms(1) x) - also have " ... = J +> (i \ y)" - by (metis a_coset_add_zero a_rcosetsI abelian_subgroup.rcosets_add_eq abelian_subgroupI3 - additive_subgroup.a_Hcarr additive_subgroup.a_subset assms i ideal.axioms(1) - is_abelian_group m_closed y) - finally have Jy: "J +> ((j \ x) \ (i \ y)) = J +> y" using mod_J - by (metis (full_types) assms i ideal.Icarr ideal.rcoset_mult_add local.semiring_axioms one_closed semiring.semiring_simprules(9) y) - have "(j \ x) \ (i \ y) \ carrier R" - by (meson x y i j assms add.m_closed additive_subgroup.a_Hcarr ideal.axioms(1) m_closed) - thus "\a \ carrier R. I +> a = I +> x \ J +> a = J +> y" using Ix Jy by blast +lemma RDirProd_list_truncate: + shows "monoid.truncate (RDirProd_list Rs) = DirProd_list Rs" +proof (induct Rs, simp add: RDirProd_list_def DirProd_list_def monoid.defs) + case (Cons R Rs) + have "monoid.truncate (RDirProd_list (R # Rs)) = + monoid.truncate (image_ring (\(a, as). a # as) (RDirProd R (RDirProd_list Rs)))" + unfolding RDirProd_list_def by simp + also have " ... = image_group (\(a, as). a # as) (monoid.truncate (RDirProd R (RDirProd_list Rs)))" + by (simp add: image_ring_def image_group_def monoid.defs) + also have " ... = image_group (\(a, as). a # as) (R \\ (monoid.truncate (RDirProd_list Rs)))" + by (simp add: RDirProd_def DirProd_def monoid.defs) + also have " ... = DirProd_list (R # Rs)" + unfolding Cons DirProd_list_def by simp + finally show ?case . qed -lemma (in ring) canonical_proj_is_hom: - assumes "ideal I R" "ideal J R" "I <+> J = carrier R" - shows "(\a. (I +> a, J +> a)) \ ring_hom R (RDirProd (R Quot I) (R Quot J))" -proof (rule ring_hom_memI) - fix x y assume x: "x \ carrier R" and y: "y \ carrier R" - show "(I +> x, J +> x) \ carrier (RDirProd (R Quot I) (R Quot J))" - using A_RCOSETS_def'[of R I] A_RCOSETS_def'[of R J] x - unfolding RDirProd_def FactRing_def by auto - show "(I +> x \ y, J +> x \ y) = - (I +> x, J +> x) \\<^bsub>RDirProd (R Quot I) (R Quot J)\<^esub> (I +> y, J +> y)" - unfolding RDirProd_def FactRing_def by (simp add: assms ideal.rcoset_mult_add x y) - show "(I +> x \ y, J +> x \ y) = - (I +> x, J +> x) \\<^bsub>RDirProd (R Quot I) (R Quot J)\<^esub> (I +> y, J +> y)" - unfolding RDirProd_def FactRing_def - by (simp add: abelian_subgroup.a_rcos_sum abelian_subgroupI3 assms ideal.axioms(1) is_abelian_group x y) -next - show "(I +> \, J +> \) = \\<^bsub>RDirProd (R Quot I) (R Quot J)\<^esub>" - unfolding RDirProd_def FactRing_def by simp +lemma RDirProd_list_carrier_def': + shows "carrier (RDirProd_list Rs) = carrier (DirProd_list Rs)" +proof - + have "carrier (RDirProd_list Rs) = carrier (monoid.truncate (RDirProd_list Rs))" + by (simp add: monoid.defs) + thus ?thesis + unfolding RDirProd_list_truncate . qed -theorem (in ring) chinese_remainder_simple: - assumes "ideal I R" "ideal J R" "I <+> J = carrier R" - shows "(R Quot (I \ J)) \ (RDirProd (R Quot I) (R Quot J))" -proof - - let ?\ = "\a. (I +> a, J +> a)" +lemma RDirProd_list_carrier: + shows "carrier (RDirProd_list (G # Gs)) = (\(x, xs). x # xs) ` (carrier G \ carrier (RDirProd_list Gs))" + unfolding RDirProd_list_carrier_def' using DirProd_list_carrier . + +lemma RDirProd_list_one: + shows "one (RDirProd_list Rs) = foldr (\R tl. (one R) # tl) Rs []" + unfolding RDirProd_list_def RDirProd_def image_ring_def image_group_def + by (induct Rs) (auto simp add: monoid.defs) - have phi_hom: "?\ \ ring_hom R (RDirProd (R Quot I) (R Quot J))" - using canonical_proj_is_hom[OF assms] . +lemma RDirProd_list_zero: + shows "zero (RDirProd_list Rs) = foldr (\R tl. (zero R) # tl) Rs []" + unfolding RDirProd_list_def RDirProd_def image_ring_def + by (induct Rs) (auto simp add: monoid.defs) + +lemma RDirProd_list_zero': + shows "zero (RDirProd_list (R # Rs)) = (zero R) # (zero (RDirProd_list Rs))" + unfolding RDirProd_list_zero by simp + +lemma RDirProd_list_carrier_mem: + assumes "as \ carrier (RDirProd_list Rs)" + shows "length as = length Rs" and "\i. i < length Rs \ (as ! i) \ carrier (Rs ! i)" + using assms DirProd_list_carrier_mem unfolding RDirProd_list_carrier_def' by auto - moreover have "?\ ` (carrier R) = carrier (RDirProd (R Quot I) (R Quot J))" - proof - show "carrier (RDirProd (R Quot I) (R Quot J)) \ ?\ ` (carrier R)" - proof - fix t assume "t \ carrier (RDirProd (R Quot I) (R Quot J))" - then obtain x y where x: "x \ carrier R" and y: "y \ carrier R" - and t: "t = (I +> x, J +> y)" - using A_RCOSETS_def'[of R I] A_RCOSETS_def'[of R J] - unfolding RDirProd_def FactRing_def by auto - then obtain a where "a \ carrier R" "I +> a = I +> x" "J +> a = J +> y" - using canonical_proj_is_surj[OF assms x y] by blast - hence "?\ a = t" using t by simp - thus "t \ (?\ ` carrier R)" using \a \ carrier R\ by blast - qed - next - show "?\ ` carrier R \ carrier (RDirProd (R Quot I) (R Quot J))" - using phi_hom unfolding ring_hom_def by blast - qed +lemma RDirProd_list_carrier_memI: + assumes "length as = length Rs" and "\i. i < length Rs \ (as ! i) \ carrier (Rs ! i)" + shows "as \ carrier (RDirProd_list Rs)" + using assms DirProd_list_carrier_memI unfolding RDirProd_list_carrier_def' by auto + +lemma inj_on_RDirProd_carrier: + shows "inj_on (\(a, as). a # as) (carrier (RDirProd R (RDirProd_list Rs)))" + unfolding RDirProd_def DirProd_def inj_on_def by auto - moreover have "a_kernel R (RDirProd (R Quot I) (R Quot J)) ?\ = I \ J" - proof - show "I \ J \ a_kernel R (RDirProd (R Quot I) (R Quot J)) ?\" - proof - fix s assume s: "s \ I \ J" hence "I +> s = I \ J +> s = J" - by (simp add: additive_subgroup.zero_closed assms ideal.axioms(1) ideal.set_add_zero) - thus "s \ a_kernel R (RDirProd (R Quot I) (R Quot J)) ?\" - unfolding FactRing_def RDirProd_def a_kernel_def kernel_def - using s additive_subgroup.a_Hcarr assms(1) ideal.axioms(1) by fastforce - qed - next - show "a_kernel R (RDirProd (R Quot I) (R Quot J)) ?\ \ I \ J" - unfolding FactRing_def RDirProd_def a_kernel_def kernel_def apply simp - using ideal.set_add_zero_imp_mem assms(1-2) by fastforce - qed - - moreover have "ring (RDirProd (R Quot I) (R Quot J))" - by (simp add: RDirProd_group assms(1) assms(2) ideal.quotient_is_ring) - - ultimately show ?thesis - using ring_hom_ring.FactRing_iso[of R "RDirProd (R Quot I) (R Quot J)" ?\] is_ring - by (simp add: ring_hom_ringI2) +lemma RDirProd_list_is_ring: + assumes "\i. i < length Rs \ ring (Rs ! i)" shows "ring (RDirProd_list Rs)" + using assms +proof (induct Rs) + case Nil thus ?case + unfolding RDirProd_list_def by (unfold_locales, auto simp add: monoid.defs Units_def) +next + case (Cons R Rs) + hence is_ring: "ring (RDirProd R (RDirProd_list Rs))" + using RDirProd_ring[of R "RDirProd_list Rs"] by force + show ?case + using ring.inj_imp_image_ring_is_ring[OF is_ring inj_on_RDirProd_carrier] + unfolding RDirProd_list_def by auto qed +lemma RDirProd_list_iso1: + "(\(a, as). a # as) \ ring_iso (RDirProd R (RDirProd_list Rs)) (RDirProd_list (R # Rs))" + using inj_imp_image_ring_iso[OF inj_on_RDirProd_carrier] unfolding RDirProd_list_def by auto -subsection \First Generalization - The Extended Canonical Projection is Surjective\ +lemma RDirProd_list_iso2: + "Hilbert_Choice.inv (\(a, as). a # as) \ ring_iso (RDirProd_list (R # Rs)) (RDirProd R (RDirProd_list Rs))" + unfolding RDirProd_list_def by (auto intro: inj_imp_image_ring_inv_iso simp add: inj_def) -lemma (in cring) canonical_proj_ext_is_surj: - fixes n::nat - assumes "\i. i \ n \ x i \ carrier R" - and "\i. i \ n \ ideal (I i) R" - and "\i j. \ i \ n; j \ n; i \ j \ \ I i <+> I j = carrier R" - shows "\ a \ carrier R. \ i \ n. (I i) +> a = (I i) +> (x i)" using assms -proof (induct n) - case 0 thus ?case by blast -next - case (Suc n) - then obtain a where a: "a \ carrier R" "\i. i \ n \ (I i) +> a = (I i) +> (x i)" - by force - - have inter_is_ideal: "ideal (\ i \ n. I i) R" - by (metis (mono_tags, lifting) Suc.prems(2) atMost_iff i_Intersect imageE image_is_empty le_SucI not_empty_eq_Iic_eq_empty) - have "(\ i \ n. I i) <+> I (Suc n) = carrier R" - using inter_plus_ideal_eq_carrier Suc by simp - then obtain b where b: "b \ carrier R" - and "(\ i \ n. I i) +> b = (\ i \ n. I i) +> \" - and S: "I (Suc n) +> b = I (Suc n) +> (x (Suc n) \ a)" - using canonical_proj_is_surj[OF inter_is_ideal, of "I (Suc n)" \ "x (Suc n) \ a"] Suc a by auto - hence b_inter: "b \ (\ i \ n. I i)" - using ideal.set_add_zero_imp_mem[OF inter_is_ideal b] - by (metis additive_subgroup.zero_closed ideal.axioms(1) ideal.set_add_zero inter_is_ideal) - hence eq_zero: "\i. i \ n \ (I i) +> b = (I i) +> \" - proof - - fix i assume i: "i \ n" - hence "b \ I i" using b_inter by blast - moreover have "ideal (I i) R" using Suc i by simp - ultimately show "(I i) +> b = (I i) +> \" - by (metis b ideal.I_r_closed ideal.set_add_zero r_null zero_closed) - qed - - have "(I i) +> (a \ b) = (I i) +> (x i)" if "i \ Suc n" for i - proof - - show "(I i) +> (a \ b) = (I i) +> (x i)" - using that - proof (cases) - assume 1: "i \ n" - hence "(I i) +> (a \ b) = ((I i) +> (x i)) <+> ((I i) +> b)" - by (metis Suc.prems(2) a abelian_subgroup.a_rcos_sum abelian_subgroupI3 b ideal_def le_SucI ring_def) - also have " ... = ((I i) +> (x i)) <+> ((I i) +> \)" - using eq_zero[OF 1] by simp - also have " ... = I i +> ((x i) \ \)" - by (meson Suc.prems abelian_subgroup.a_rcos_sum abelian_subgroupI3 atMost_iff that ideal_def ring_def zero_closed) - finally show "(I i) +> (a \ b) = (I i) +> (x i)" - using Suc.prems(1) that by auto - next - assume "\ i \ n" hence 2: "i = Suc n" using that by simp - hence "I i +> (a \ b) = I (Suc n) +> (a \ b)" by simp - also have " ... = (I (Suc n) +> a) <+> (I (Suc n) +> (x (Suc n) \ a))" - by (metis le_Suc_eq S a b Suc.prems(2)[of "Suc n"] 2 abelian_subgroup.a_rcos_sum - abelian_subgroupI3 ideal.axioms(1) is_abelian_group) - also have " ... = I (Suc n) +> (a \ (x (Suc n) \ a))" - by (simp add: Suc.prems(1-2) a(1) abelian_subgroup.a_rcos_sum - abelian_subgroupI3 ideal.axioms(1) is_abelian_group) - also have " ... = I (Suc n) +> (x (Suc n))" - using a(1) Suc.prems(1)[of "Suc n"] abelian_group.minus_eq - abelian_group.r_neg add.m_lcomm is_abelian_group by fastforce - finally show "I i +> (a \ b) = (I i) +> (x i)" using 2 by simp - qed - qed - thus ?case using a b by auto +lemma RDirProd_list_iso3: + "(\a. [ a ]) \ ring_iso R (RDirProd_list [ R ])" +proof - + have [simp]: "(\a. [ a ]) = (\(a, as). a # as) \ (\a. (a, []))" by auto + show ?thesis + using ring_iso_set_trans[OF RDirProd_iso7] RDirProd_list_iso1[of R "[]"] + unfolding RDirProd_list_def by simp +qed + +lemma RDirProd_list_hom1: + "(\(a, as). a # as) \ ring_hom (RDirProd R (RDirProd_list Rs)) (RDirProd_list (R # Rs))" + using RDirProd_list_iso1 unfolding ring_iso_def by auto + +lemma RDirProd_list_hom2: + assumes "f \ ring_hom R S" shows "(\a. [ f a ]) \ ring_hom R (RDirProd_list [ S ])" +proof - + have hom1: "(\a. (a, [])) \ ring_hom R (RDirProd R nil_ring)" + using RDirProd_iso7 unfolding ring_iso_def by auto + have hom2: "(\(a, as). a # as) \ ring_hom (RDirProd S nil_ring) (RDirProd_list [ S ])" + using RDirProd_list_hom1[of _ "[]"] unfolding RDirProd_list_def by auto + have [simp]: "(\(a, as). a # as) \ ((\(x, y). (f x, y)) \ (\a. (a, []))) = (\a. [ f a ])" by auto + show ?thesis + using ring_hom_trans[OF ring_hom_trans[OF hom1 RDirProd_hom2(2)[OF assms]] hom2] by simp qed -subsection \Direct Product of a List of Monoid Structures\ - -fun DirProd_list :: "(('a, 'b) monoid_scheme) list \ (('a list), 'b) monoid_scheme" - where - "DirProd_list [] = \ carrier = {[]}, mult = (\a b. []), one = [], \ = (undefined :: 'b) \" - | "DirProd_list (Cons R Rs) = - \ carrier = { r # rs | r rs. r \ carrier R \ rs \ carrier (DirProd_list Rs) }, - mult = (\r1 r2. ((hd r1) \\<^bsub>R\<^esub> (hd r2)) # ((tl r1) \\<^bsub>(DirProd_list Rs)\<^esub> (tl r2))), - one = (\\<^bsub>R\<^esub>) # (\\<^bsub>(DirProd_list Rs)\<^esub>), \ = (undefined :: 'b) \" - - -lemma DirProd_list_carrier_elts: - assumes "rs \ carrier (DirProd_list Rs)" - shows "length rs = length Rs" using assms -proof (induct Rs arbitrary: rs rule: DirProd_list.induct) - case 1 thus ?case by simp -next - case (2 R Rs) - then obtain r' rs' where "r' \ carrier R" "rs' \ carrier (DirProd_list Rs)" - and "rs = r' # rs'" by auto - thus ?case by (simp add: "2.hyps"(1)) -qed - -lemma DirProd_list_in_carrierI: - assumes "\i. i < length rs \ rs ! i \ carrier (Rs ! i)" - and "length rs = length Rs" - shows "rs \ carrier (DirProd_list Rs)" using assms -proof (induct Rs arbitrary: rs rule: DirProd_list.induct) - case 1 thus ?case by simp -next - case (2 R Rs) - then obtain r' rs' where rs: "r' \ carrier R" "rs = r' # rs'" - by (metis Suc_length_conv lessThan_iff nth_Cons_0 zero_less_Suc) - hence "rs' \ carrier (DirProd_list Rs)" - using "2.hyps"(1) "2.prems"(1) "2.prems"(2) by force - thus ?case by (simp add: rs) -qed - -lemma DirProd_list_in_carrierE: - assumes "rs \ carrier (DirProd_list Rs)" - shows "\i. i < length rs \ rs ! i \ carrier (Rs ! i)" using assms -proof (induct Rs arbitrary: rs rule: DirProd_list.induct) - case 1 thus ?case by simp -next - case (2 R Rs) - then obtain r' rs' where r': " r' \ carrier R" - and rs': "rs' \ carrier (DirProd_list Rs)" - and rs: "rs = r' # rs'" by auto - hence "\i. i \ {..<(length rs')} \ rs' ! i \ carrier (Rs ! i)" - using "2.hyps"(1) by blast - hence "\i. i \ {(Suc 0 :: nat)..<(length rs)} \ rs ! i \ carrier ((R # Rs) ! i)" - by (simp add: less_eq_Suc_le rs) - moreover have "i = 0 \ rs ! i \ carrier ((R # Rs) ! i)" - using r' rs r' by simp - ultimately show ?case - using "2.prems"(1) by fastforce -qed +section \Chinese Remainder Theorem\ -lemma DirProd_list_m_closed: - assumes "r1 \ carrier (DirProd_list Rs)" "r2 \ carrier (DirProd_list Rs)" - and "\i. i < length Rs \ monoid (Rs ! i)" - shows "r1 \\<^bsub>(DirProd_list Rs)\<^esub> r2 \ carrier (DirProd_list Rs)" using assms -proof (induct Rs arbitrary: r1 r2 rule: DirProd_list.induct) - case 1 thus ?case by simp -next - case (2 R Rs) - then obtain r1' rs1' r2' rs2' - where r12': "r1' \ carrier R" "r2' \ carrier R" - and "rs1' \ carrier (DirProd_list Rs)" - and "rs2' \ carrier (DirProd_list Rs)" - and r1: "r1 = r1' # rs1'" - and r2: "r2 = r2' # rs2'" by auto - moreover have "\i. i < length Rs \ monoid (Rs ! i)" - using "2.prems"(3) by force - ultimately have "rs1' \\<^bsub>(DirProd_list Rs)\<^esub> rs2' \ carrier (DirProd_list Rs)" - using "2.hyps"(1) by blast - moreover have "monoid R" - using "2.prems"(3) by force - hence "r1' \\<^bsub>R\<^esub> r2' \ carrier R" - by (simp add: r12' monoid.m_closed) - ultimately show ?case by (simp add: r1 r2) -qed - -lemma DirProd_list_m_output: - assumes "r1 \ carrier (DirProd_list Rs)" "r2 \ carrier (DirProd_list Rs)" - shows "\i. i < length Rs \ - (r1 \\<^bsub>(DirProd_list Rs)\<^esub> r2) ! i = (r1 ! i) \\<^bsub>(Rs ! i)\<^esub> (r2 ! i)" using assms -proof (induct Rs arbitrary: r1 r2 rule: DirProd_list.induct) - case 1 thus ?case by simp -next - case (2 R Rs) - hence "\i. i \ {(Suc 0)..<(length (R # Rs))} \ - (r1 \\<^bsub>(DirProd_list (R # Rs))\<^esub> r2) ! i = (r1 ! i) \\<^bsub>((R # Rs) ! i)\<^esub> (r2 ! i)" - using "2"(5) "2"(6) by auto - moreover have "(r1 \\<^bsub>(DirProd_list (R # Rs))\<^esub> r2) ! 0 = (r1 ! 0) \\<^bsub>R\<^esub> (r2 ! 0)" - using "2.prems"(2) "2.prems"(3) by auto - ultimately show ?case - by (metis "2.prems"(1) atLeastLessThan_iff le_0_eq not_less_eq_eq nth_Cons') -qed - -lemma DirProd_list_m_comm: - assumes "r1 \ carrier (DirProd_list Rs)" "r2 \ carrier (DirProd_list Rs)" - and "\i. i < length Rs \ comm_monoid (Rs ! i)" - shows "r1 \\<^bsub>(DirProd_list Rs)\<^esub> r2 = r2 \\<^bsub>(DirProd_list Rs)\<^esub> r1" -proof (rule nth_equalityI) - show "length (r1 \\<^bsub>DirProd_list Rs\<^esub> r2) = length (r2 \\<^bsub>DirProd_list Rs\<^esub> r1)" - by (metis DirProd_list_carrier_elts DirProd_list_m_closed Group.comm_monoid.axioms(1) assms) - - fix i assume "i < length (r1 \\<^bsub>DirProd_list Rs\<^esub> r2)" - hence i: "i < length Rs" - by (metis DirProd_list_carrier_elts DirProd_list_m_closed Group.comm_monoid.axioms(1) assms) - have "(r1 \\<^bsub>DirProd_list Rs\<^esub> r2) ! i = (r1 ! i) \\<^bsub>(Rs ! i)\<^esub> (r2 ! i)" - using i DirProd_list_m_output[OF assms(1-2)] by simp - also have " ... = (r2 ! i) \\<^bsub>(Rs ! i)\<^esub> (r1 ! i)" - by (metis DirProd_list_carrier_elts DirProd_list_in_carrierE assms comm_monoid.m_comm i) - also have " ... = (r2 \\<^bsub>DirProd_list Rs\<^esub> r1) ! i" - using i DirProd_list_m_output[OF assms(2) assms(1)] by simp - finally show "(r1 \\<^bsub>DirProd_list Rs\<^esub> r2) ! i = (r2 \\<^bsub>DirProd_list Rs\<^esub> r1) ! i" . -qed - -lemma DirProd_list_m_assoc: - assumes "r1 \ carrier (DirProd_list Rs)" - and "r2 \ carrier (DirProd_list Rs)" - and "r3 \ carrier (DirProd_list Rs)" - and "\i. i < length Rs \ monoid (Rs ! i)" - shows "(r1 \\<^bsub>(DirProd_list Rs)\<^esub> r2) \\<^bsub>(DirProd_list Rs)\<^esub> r3 = - r1 \\<^bsub>(DirProd_list Rs)\<^esub> (r2 \\<^bsub>(DirProd_list Rs)\<^esub> r3)" -proof (rule nth_equalityI) - show "length ((r1 \\<^bsub>DirProd_list Rs\<^esub> r2) \\<^bsub>DirProd_list Rs\<^esub> r3) = - length (r1 \\<^bsub>DirProd_list Rs\<^esub> (r2 \\<^bsub>DirProd_list Rs\<^esub> r3))" - by (metis DirProd_list_carrier_elts DirProd_list_m_closed assms) - - fix i assume "i < length (r1 \\<^bsub>DirProd_list Rs\<^esub> r2 \\<^bsub>DirProd_list Rs\<^esub> r3)" - hence i: "i < length Rs" - by (metis DirProd_list_carrier_elts DirProd_list_m_closed assms) - have "((r1 \\<^bsub>DirProd_list Rs\<^esub> r2) \\<^bsub>DirProd_list Rs\<^esub> r3) ! i = - ((r1 ! i) \\<^bsub>(Rs ! i)\<^esub> (r2 ! i)) \\<^bsub>(Rs ! i)\<^esub> (r3 ! i)" - by (metis DirProd_list_m_closed DirProd_list_m_output i assms) - also have " ... = (r1 ! i) \\<^bsub>(Rs ! i)\<^esub> ((r2 ! i) \\<^bsub>(Rs ! i)\<^esub> (r3 ! i))" - by (metis DirProd_list_carrier_elts DirProd_list_in_carrierE assms i monoid.m_assoc) - also have " ... = (r1 \\<^bsub>DirProd_list Rs\<^esub> (r2 \\<^bsub>DirProd_list Rs\<^esub> r3)) ! i" - by (metis DirProd_list_m_closed DirProd_list_m_output i assms) - finally show "((r1 \\<^bsub>DirProd_list Rs\<^esub> r2) \\<^bsub>DirProd_list Rs\<^esub> r3) ! i = - (r1 \\<^bsub>DirProd_list Rs\<^esub> (r2 \\<^bsub>DirProd_list Rs\<^esub> r3))! i" . -qed - -lemma DirProd_list_one: - "\i. i < length Rs \ (\\<^bsub>(DirProd_list Rs)\<^esub>) ! i = \\<^bsub>(Rs ! i)\<^esub>" - by (induct Rs rule: DirProd_list.induct) (simp_all add: nth_Cons') - -lemma DirProd_list_one_closed: - assumes "\i. i < length Rs \ monoid (Rs ! i)" - shows "\\<^bsub>(DirProd_list Rs)\<^esub> \ carrier (DirProd_list Rs)" -proof (rule DirProd_list_in_carrierI) - show eq_len: "length \\<^bsub>DirProd_list Rs\<^esub> = length Rs" - by (induct Rs rule: DirProd_list.induct) (simp_all) - show "\i. i < length \\<^bsub>DirProd_list Rs\<^esub> \ \\<^bsub>DirProd_list Rs\<^esub> ! i \ carrier (Rs ! i)" - using eq_len DirProd_list_one[where ?Rs = Rs] monoid.one_closed by (simp add: assms) -qed +subsection \Definitions\ -lemma DirProd_list_l_one: - assumes "r1 \ carrier (DirProd_list Rs)" - and "\i. i < length Rs \ monoid (Rs ! i)" - shows "\\<^bsub>(DirProd_list Rs)\<^esub> \\<^bsub>(DirProd_list Rs)\<^esub> r1 = r1" -proof (rule nth_equalityI) - show eq_len: "length (\\<^bsub>DirProd_list Rs\<^esub> \\<^bsub>DirProd_list Rs\<^esub> r1) = length r1" - using DirProd_list_carrier_elts[of "\\<^bsub>DirProd_list Rs\<^esub> \\<^bsub>DirProd_list Rs\<^esub> r1" Rs] - DirProd_list_carrier_elts[OF assms(1)] - DirProd_list_m_closed[OF DirProd_list_one_closed[OF assms(2)] assms(1)] - by (simp add: assms) - fix i assume "i < length (\\<^bsub>DirProd_list Rs\<^esub> \\<^bsub>DirProd_list Rs\<^esub> r1)" - hence i: "i < length Rs" using DirProd_list_carrier_elts[OF assms(1)] eq_len by simp - hence "(\\<^bsub>DirProd_list Rs\<^esub> \\<^bsub>DirProd_list Rs\<^esub> r1) ! i = - (\\<^bsub>DirProd_list Rs\<^esub> ! i) \\<^bsub>(Rs ! i)\<^esub> (r1 ! i)" - using DirProd_list_m_output DirProd_list_one_closed assms by blast - also have " ... = \\<^bsub>(Rs ! i)\<^esub> \\<^bsub>(Rs ! i)\<^esub> (r1 ! i)" - by (simp add: DirProd_list_one i) - also have " ... = (r1 ! i)" - using DirProd_list_carrier_elts DirProd_list_in_carrierE i assms by fastforce - finally show "(\\<^bsub>DirProd_list Rs\<^esub> \\<^bsub>DirProd_list Rs\<^esub> r1) ! i = (r1 ! i)" . -qed - -lemma DirProd_list_r_one: - assumes "r1 \ carrier (DirProd_list Rs)" - and "\i. i < length Rs \ monoid (Rs ! i)" - shows "r1 \\<^bsub>(DirProd_list Rs)\<^esub> \\<^bsub>(DirProd_list Rs)\<^esub> = r1" -proof - - have "r1 \\<^bsub>(DirProd_list Rs)\<^esub> \\<^bsub>(DirProd_list Rs)\<^esub> = - \\<^bsub>(DirProd_list Rs)\<^esub> \\<^bsub>(DirProd_list Rs)\<^esub> r1" - proof (rule nth_equalityI) - show " length (r1 \\<^bsub>DirProd_list Rs\<^esub> \\<^bsub>DirProd_list Rs\<^esub>) = - length (\\<^bsub>DirProd_list Rs\<^esub> \\<^bsub>DirProd_list Rs\<^esub> r1)" - by (metis DirProd_list_carrier_elts DirProd_list_m_closed DirProd_list_one_closed assms) - - fix i assume "i < length (r1 \\<^bsub>DirProd_list Rs\<^esub> \\<^bsub>DirProd_list Rs\<^esub>)" - hence i: "i < length Rs" - by (metis DirProd_list_carrier_elts DirProd_list_m_closed DirProd_list_one_closed assms) - hence "(r1 \\<^bsub>DirProd_list Rs\<^esub> \\<^bsub>DirProd_list Rs\<^esub>) ! i = (r1 ! i) \\<^bsub>(Rs ! i)\<^esub> \\<^bsub>(Rs ! i)\<^esub>" - by (metis DirProd_list_m_output DirProd_list_one DirProd_list_one_closed assms) - also have " ... = \\<^bsub>(Rs ! i)\<^esub> \\<^bsub>(Rs ! i)\<^esub> (r1 ! i)" - using DirProd_list_carrier_elts DirProd_list_in_carrierE assms i by fastforce - also have " ... = (\\<^bsub>DirProd_list Rs\<^esub> \\<^bsub>DirProd_list Rs\<^esub> r1) ! i" - by (metis DirProd_list_m_output DirProd_list_one DirProd_list_one_closed assms i) - finally show "(r1 \\<^bsub>DirProd_list Rs\<^esub> \\<^bsub>DirProd_list Rs\<^esub>) ! i = - (\\<^bsub>DirProd_list Rs\<^esub> \\<^bsub>DirProd_list Rs\<^esub> r1) ! i" . - qed - thus ?thesis using DirProd_list_l_one assms by auto -qed - -lemma DirProd_list_monoid: - assumes "\i. i < length Rs \ monoid (Rs ! i)" - shows "monoid (DirProd_list Rs)" - unfolding monoid_def -proof (intro conjI allI impI) - show "\\<^bsub>DirProd_list Rs\<^esub> \ carrier (DirProd_list Rs)" - using DirProd_list_one_closed[of Rs] assms by simp - - fix x y z - assume x: "x \ carrier (DirProd_list Rs)" - and y: "y \ carrier (DirProd_list Rs)" - and z: "z \ carrier (DirProd_list Rs)" - show "x \\<^bsub>DirProd_list Rs\<^esub> y \ carrier (DirProd_list Rs)" - using DirProd_list_m_closed[OF x y] assms by simp - show "x \\<^bsub>DirProd_list Rs\<^esub> y \\<^bsub>DirProd_list Rs\<^esub> z = - x \\<^bsub>DirProd_list Rs\<^esub> (y \\<^bsub>DirProd_list Rs\<^esub> z)" - using DirProd_list_m_assoc[OF x y z] assms by simp - show "\\<^bsub>DirProd_list Rs\<^esub> \\<^bsub>DirProd_list Rs\<^esub> x = x" - using DirProd_list_l_one[OF x] assms by simp - show "x \\<^bsub>DirProd_list Rs\<^esub> \\<^bsub>DirProd_list Rs\<^esub> = x" - using DirProd_list_r_one[OF x] assms by simp -qed - -lemma DirProd_list_comm_monoid: - assumes "\i. i < length Rs \ comm_monoid (Rs ! i)" - shows "comm_monoid (DirProd_list Rs)" - unfolding comm_monoid_def comm_monoid_axioms_def apply auto - using DirProd_list_monoid Group.comm_monoid.axioms(1) assms apply blast - using DirProd_list_m_comm assms by blast - -lemma DirProd_list_isomorphism1: - "(\(hd, tl). hd # tl) \ iso (R \\ (DirProd_list Rs)) (DirProd_list (R # Rs))" - unfolding iso_def hom_def bij_betw_def inj_on_def by auto +abbreviation (in ring) canonical_proj :: "'a set \ 'a set \ 'a \ 'a set \ 'a set" + where "canonical_proj I J \ (\a. (I +> a, J +> a))" -lemma DirProd_list_isomorphism2: - "(\r. (hd r, tl r)) \ iso (DirProd_list (R # Rs)) (R \\ (DirProd_list Rs))" (is "?\ \ ?A") - unfolding iso_def hom_def bij_betw_def inj_on_def apply auto -proof - - fix a b assume "a \ carrier R" "b \ carrier (DirProd_list Rs)" - hence "a # b \ {r # rs |r rs. r \ carrier R \ rs \ carrier (DirProd_list Rs)} \ ?\ (a # b) = (a, b)" - by simp - thus "(a, b) \ ?\ ` {r # rs |r rs. r \ carrier R \ rs \ carrier (DirProd_list Rs)}" - by (metis (no_types, lifting) image_iff) -qed - -lemma DirProd_list_group: - assumes "\i. i < length Rs \ group (Rs ! i)" - shows "group (DirProd_list Rs)" using assms -proof (induction Rs rule: DirProd_list.induct) - case 1 thus ?case - unfolding group_def group_axioms_def Units_def monoid_def by auto -next - case (2 R Rs) - hence "group (DirProd_list Rs)" by force - moreover have "group R" - using "2.prems" by fastforce - moreover have "monoid (DirProd_list (R # Rs))" - using DirProd_list_monoid 2 group.is_monoid by blast - moreover have "R \\ DirProd_list Rs \ DirProd_list (R # Rs)" - unfolding is_iso_def using DirProd_list_isomorphism1 by auto - ultimately show ?case - using group.iso_imp_group[of "R \\ (DirProd_list Rs)" "DirProd_list (R # Rs)"] - DirProd_group[of R "DirProd_list Rs"] by simp -qed - -lemma DirProd_list_comm_group: - assumes "\i. i < length Rs \ comm_group (Rs ! i)" - shows "comm_group (DirProd_list Rs)" - using assms unfolding comm_group_def - using DirProd_list_group DirProd_list_comm_monoid by auto - -lemma DirProd_list_group_hom: - assumes "\i. i \ {..<(length (R # Rs))} \ group ((R # Rs) ! i)" - shows "group_hom (R \\ DirProd_list Rs) (DirProd_list (R # Rs)) (\(hd, tl). hd # tl)" -proof - - have "group R" - using assms by force - moreover have "group (DirProd_list Rs)" - using DirProd_list_group assms by fastforce - ultimately - - have "group (R \\ DirProd_list Rs)" - using DirProd_group[of R "DirProd_list Rs"] by simp - moreover have "group (DirProd_list (R # Rs))" - using DirProd_list_group assms by blast - moreover have "(\(x, y). x # y) \ hom (R \\ DirProd_list Rs) (DirProd_list (R # Rs))" - using DirProd_list_isomorphism1[of R Rs] unfolding iso_def by simp - ultimately show ?thesis - unfolding group_hom_def group_hom_axioms_def by simp -qed - -lemma DirProd_list_m_inv: - assumes "r \ carrier (DirProd_list Rs)" - and "\i. i < length Rs \ group (Rs ! i)" - shows "\i. i < length Rs \ (inv\<^bsub>(DirProd_list Rs)\<^esub> r) ! i = inv\<^bsub>(Rs ! i)\<^esub> (r ! i)" - using assms -proof (induct Rs arbitrary: r rule: DirProd_list.induct) - case 1 - have "group (DirProd_list [])" - unfolding group_def group_axioms_def Units_def monoid_def by auto - thus ?case using "1.prems"(1) group.inv_equality by fastforce -next - case (2 R Rs) - then obtain r' rs' where r': "r' \ carrier R" and rs': "rs' \ carrier (DirProd_list Rs)" - and r: "r = r' # rs'" by auto - hence "(r', rs') \ carrier (R \\ DirProd_list Rs)" by simp - moreover have "group_hom (R \\ DirProd_list Rs) (DirProd_list (R # Rs)) (\(hd, tl). hd # tl)" - using DirProd_list_group_hom[of R Rs] 2 by auto - moreover have "inv\<^bsub>(R \\ DirProd_list Rs)\<^esub> (r', rs') = (inv\<^bsub>R\<^esub> r', inv\<^bsub>(DirProd_list Rs)\<^esub> rs')" - using inv_DirProd[of R "DirProd_list Rs" r' rs'] "2.prems"(3) DirProd_list_group r' rs' by force - ultimately have "inv\<^bsub>(DirProd_list (R # Rs))\<^esub> r = (inv\<^bsub>R\<^esub> r') # (inv\<^bsub>(DirProd_list Rs)\<^esub> rs')" - using group_hom.hom_inv[of "R \\ DirProd_list Rs" "DirProd_list (R # Rs)" - "\(hd, tl). hd # tl" "(r', rs')"] r by simp - thus ?case - using 2 by simp (metis (no_types, lifting) less_Suc_eq_0_disj list.sel(3) nth_Cons_0 nth_Cons_Suc r) -qed +definition (in ring) canonical_proj_ext :: "(nat \ 'a set) \ nat \ 'a \ ('a set) list" + where "canonical_proj_ext I n = (\a. map (\i. (I i) +> a) [0..< Suc n])" -subsection \Direct Product for of a List of Rings\ - -text \In order to state a more general version of the Chinese Remainder Theorem, we need a new - structure: the direct product of a variable number of rings. The construction of this - structure as well as its algebraic properties are the subject of this subsection and follow - the similar study that has already been done for monoids in the previous subsection.\ +subsection \Chinese Remainder Simple\ -fun RDirProd_list :: "('a ring) list \ ('a list) ring" - where "RDirProd_list Rs = - monoid.extend (monoid.truncate (DirProd_list Rs)) - \ zero = one (DirProd_list (map add_monoid Rs)), - add = mult (DirProd_list (map add_monoid Rs)) \" - -lemma RDirProd_list_add_monoid: - "add_monoid (RDirProd_list Rs) = DirProd_list (map add_monoid Rs)" -proof - - have "carrier (RDirProd_list Rs) = carrier (DirProd_list (map add_monoid Rs))" - by (induct Rs rule: DirProd_list.induct) (simp_all add: monoid.defs) - thus ?thesis by (simp add: monoid.defs) -qed - -lemma RDirProd_list_mult_monoid: - "monoid.truncate (RDirProd_list Rs) = monoid.truncate (DirProd_list Rs)" - by (simp add: monoid.defs) +lemma (in ring) canonical_proj_is_surj: + assumes "ideal I R" "ideal J R" and "I <+> J = carrier R" + shows "(canonical_proj I J) ` carrier R = carrier (RDirProd (R Quot I) (R Quot J))" + unfolding RDirProd_def DirProd_def FactRing_def A_RCOSETS_def' +proof (auto simp add: monoid.defs) + { fix I i assume "ideal I R" "i \ I" hence "I +> i = \\<^bsub>R Quot I\<^esub>" + using a_rcos_zero by (simp add: FactRing_def) + } note aux_lemma1 = this -lemma RDirProd_list_monoid: - assumes "\i. i < length Rs \ monoid (Rs ! i)" - shows "monoid (RDirProd_list Rs)" -proof - - have "monoid (DirProd_list Rs)" - using DirProd_list_monoid assms by blast - hence "monoid (monoid.truncate (DirProd_list Rs))" - unfolding monoid_def by (auto intro: monoid.intro simp add: monoid.defs) - hence "monoid (monoid.truncate (RDirProd_list Rs))" - by (simp add: monoid.defs) - thus ?thesis - unfolding monoid_def by (auto intro: monoid.intro simp add: monoid.defs) -qed + { fix I i j assume A: "ideal I R" "i \ I" "j \ carrier R" "i \ j = \" + have "(I +> i) \\<^bsub>R Quot I\<^esub> (I +> j) = I +> \" + using ring_hom_memE(3)[OF ideal.rcos_ring_hom ideal.Icarr[OF _ A(2)] A(3)] A(1,4) by simp + moreover have "I +> i = I" + using abelian_subgroupI3[OF ideal.axioms(1) is_abelian_group] + by (simp add: A(1-2) abelian_subgroup.a_rcos_const) + moreover have "I +> j \ carrier (R Quot I)" and "I = \\<^bsub>R Quot I\<^esub>" and "I +> \ = \\<^bsub>R Quot I\<^esub>" + by (auto simp add: FactRing_def A_RCOSETS_def' A(3)) + ultimately have "I +> j = \\<^bsub>R Quot I\<^esub>" + using ring.ring_simprules(8)[OF ideal.quotient_is_ring[OF A(1)]] by simp + } note aux_lemma2 = this -lemma RDirProd_list_comm_monoid: - assumes "\i. i < length Rs \ comm_monoid (Rs ! i)" - shows "comm_monoid (RDirProd_list Rs)" -proof - - have "comm_monoid (DirProd_list Rs)" - using DirProd_list_comm_monoid assms by blast - hence "comm_monoid (monoid.truncate (DirProd_list Rs))" - unfolding comm_monoid_def monoid_def comm_monoid_axioms_def - by (auto simp add: monoid.defs) - hence "comm_monoid (monoid.truncate (RDirProd_list Rs))" - by (simp add: monoid.defs) - thus ?thesis - unfolding comm_monoid_def monoid_def comm_monoid_axioms_def - by (auto simp add: monoid.defs) -qed - -lemma RDirProd_list_abelian_monoid: - assumes "\i. i < length Rs \ abelian_monoid (Rs ! i)" - shows "abelian_monoid (RDirProd_list Rs)" -proof - - have "\i. i < length Rs \ comm_monoid ((map add_monoid Rs) ! i)" - using assms unfolding abelian_monoid_def by simp - hence "comm_monoid (DirProd_list (map add_monoid Rs))" - by (metis (no_types, lifting) DirProd_list_comm_monoid length_map) - thus ?thesis - unfolding abelian_monoid_def by (metis RDirProd_list_add_monoid) -qed + interpret I: ring "R Quot I" + J: ring "R Quot J" + using assms(1-2)[THEN ideal.quotient_is_ring] by auto -lemma RDirProd_list_abelian_group: - assumes "\i. i < length Rs \ abelian_group (Rs ! i)" - shows "abelian_group (RDirProd_list Rs)" -proof - - have "\i. i < length Rs \ comm_group ((map add_monoid Rs) ! i)" - using assms unfolding abelian_group_def abelian_group_axioms_def by simp - hence "comm_group (DirProd_list (map add_monoid Rs))" - by (metis (no_types, lifting) DirProd_list_comm_group length_map) - thus ?thesis - unfolding abelian_group_def abelian_group_axioms_def - by (metis RDirProd_list_abelian_monoid RDirProd_list_add_monoid abelian_group_def assms) -qed - -lemma RDirProd_list_carrier_elts: - assumes "rs \ carrier (RDirProd_list Rs)" - shows "length rs = length Rs" - using assms by (simp add: DirProd_list_carrier_elts monoid.defs) + fix a b assume a: "a \ carrier R" and b: "b \ carrier R" + have "\ \ I <+> J" + using assms(3) by blast + then obtain i j where i: "i \ carrier R" "i \ I" and j: "j \ carrier R" "j \ J" and ij: "i \ j = \" + using assms(1-2)[THEN ideal.Icarr] unfolding set_add_def' by auto + hence rcos_j: "I +> j = \\<^bsub>R Quot I\<^esub>" and rcos_i: "J +> i = \\<^bsub>R Quot J\<^esub>" + using assms(1-2)[THEN aux_lemma2] a_comm by simp+ -lemma RDirProd_list_in_carrierE: - assumes "rs \ carrier (RDirProd_list Rs)" - shows "\i. i < length rs \ rs ! i \ carrier (Rs ! i)" - using assms by (simp add: DirProd_list_in_carrierE monoid.defs) - -lemma RDirProd_list_in_carrierI: - assumes "\i. i < length rs \ rs ! i \ carrier (Rs ! i)" - and "length rs = length Rs" - shows "rs \ carrier (RDirProd_list Rs)" - using DirProd_list_in_carrierI assms by (simp add: monoid.defs, blast) - -lemma RDirProd_list_one: - "\i. i < length Rs \ (\\<^bsub>(RDirProd_list Rs)\<^esub>) ! i = \\<^bsub>(Rs ! i)\<^esub>" - by (simp add: DirProd_list_one monoid.defs) - -lemma RDirProd_list_zero: - "\i. i < length Rs \ (\\<^bsub>(RDirProd_list Rs)\<^esub>) ! i = \\<^bsub>(Rs ! i)\<^esub>" - by (induct Rs rule: DirProd_list.induct) (simp_all add: monoid.defs nth_Cons') + define s where "s = (a \ j) \ (b \ i)" + hence "s \ carrier R" + using a b i j by simp -lemma RDirProd_list_m_output: - assumes "r1 \ carrier (RDirProd_list Rs)" "r2 \ carrier (RDirProd_list Rs)" - shows "\i. i < length Rs \ - (r1 \\<^bsub>(RDirProd_list Rs)\<^esub> r2) ! i = (r1 ! i) \\<^bsub>(Rs ! i)\<^esub> (r2 ! i)" - using assms by (simp add: DirProd_list_m_output monoid.defs) + have "I +> s = ((I +> a) \\<^bsub>R Quot I\<^esub> (I +> j)) \\<^bsub>R Quot I\<^esub> (I +> (b \ i))" + using ring_hom_memE(2-3)[OF ideal.rcos_ring_hom[OF assms(1)]] + by (simp add: a b i(1) j(1) s_def) + moreover have "I +> a \ carrier (R Quot I)" + by (auto simp add: FactRing_def A_RCOSETS_def' a) + ultimately have "I +> s = I +> a" + unfolding rcos_j aux_lemma1[OF assms(1) ideal.I_l_closed[OF assms(1) i(2) b]] by simp -lemma RDirProd_list_a_output: - fixes i - assumes "r1 \ carrier (RDirProd_list Rs)" "r2 \ carrier (RDirProd_list Rs)" "i < length Rs" - shows "(r1 \\<^bsub>(RDirProd_list Rs)\<^esub> r2) ! i = (r1 ! i) \\<^bsub>(Rs ! i)\<^esub> (r2 ! i)" - using RDirProd_list_add_monoid[of Rs] monoid.defs(3) -proof - - have "(\\<^bsub>DirProd_list (map add_monoid Rs)\<^esub>) = (\\<^bsub>RDirProd_list Rs\<^esub>)" - by (metis \add_monoid (RDirProd_list Rs) = DirProd_list (map add_monoid Rs)\ monoid.select_convs(1)) - moreover have "r1 \ carrier (DirProd_list (map add_monoid Rs))" - by (metis \add_monoid (RDirProd_list Rs) = DirProd_list (map add_monoid Rs)\ assms(1) partial_object.select_convs(1)) - moreover have "r2 \ carrier (DirProd_list (map add_monoid Rs))" - by (metis \add_monoid (RDirProd_list Rs) = DirProd_list (map add_monoid Rs)\ assms(2) partial_object.select_convs(1)) - ultimately show ?thesis - by (simp add: DirProd_list_m_output assms(3)) -qed + have "J +> s = (J +> (a \ j)) \\<^bsub>R Quot J\<^esub> ((J +> b) \\<^bsub>R Quot J\<^esub> (J +> i))" + using ring_hom_memE(2-3)[OF ideal.rcos_ring_hom[OF assms(2)]] + by (simp add: a b i(1) j(1) s_def) + moreover have "J +> b \ carrier (R Quot J)" + by (auto simp add: FactRing_def A_RCOSETS_def' b) + ultimately have "J +> s = J +> b" + unfolding rcos_i aux_lemma1[OF assms(2) ideal.I_l_closed[OF assms(2) j(2) a]] by simp -lemma RDirProd_list_a_inv: - fixes i - assumes "r \ carrier (RDirProd_list Rs)" - and "\i. i < length Rs \ abelian_group (Rs ! i)" - and i: "i < length Rs" - shows "(\\<^bsub>(RDirProd_list Rs)\<^esub> r) ! i = \\<^bsub>(Rs ! i)\<^esub> (r ! i)" -proof - - have "m_inv (DirProd_list (map add_monoid Rs)) = a_inv (RDirProd_list Rs)" - by (metis RDirProd_list_add_monoid a_inv_def) - moreover have "r \ carrier (DirProd_list (map add_monoid Rs))" - by (metis RDirProd_list_add_monoid assms(1) partial_object.select_convs(1)) - moreover have "a_inv (Rs ! i) = m_inv (map add_monoid Rs ! i)" - by (simp add: a_inv_def i) - ultimately show ?thesis - by (metis (no_types, lifting) DirProd_list_carrier_elts DirProd_list_m_inv RDirProd_list_carrier_elts - abelian_group.a_group assms list_update_same_conv map_update) + from \I +> s = I +> a\ and \J +> s = J +> b\ and \s \ carrier R\ + show "(I +> a, J +> b) \ (canonical_proj I J) ` carrier R" by blast qed -lemma RDirProd_list_l_distr: - assumes "x \ carrier (RDirProd_list Rs)" - and "y \ carrier (RDirProd_list Rs)" - and "z \ carrier (RDirProd_list Rs)" - and "\i. i < length Rs \ ring (Rs ! i)" - shows "(x \\<^bsub>(RDirProd_list Rs)\<^esub> y) \\<^bsub>(RDirProd_list Rs)\<^esub> z = - (x \\<^bsub>(RDirProd_list Rs)\<^esub> z) \\<^bsub>(RDirProd_list Rs)\<^esub> (y \\<^bsub>(RDirProd_list Rs)\<^esub> z)" -proof - - have "length ((x \\<^bsub>(RDirProd_list Rs)\<^esub> y) \\<^bsub>(RDirProd_list Rs)\<^esub> z) = - length ((x \\<^bsub>(RDirProd_list Rs)\<^esub> z) \\<^bsub>(RDirProd_list Rs)\<^esub> (y \\<^bsub>(RDirProd_list Rs)\<^esub> z))" - by (metis RDirProd_list_abelian_group RDirProd_list_carrier_elts RDirProd_list_monoid - abelian_groupE(1) assms monoid.m_closed ring.is_abelian_group ring.is_monoid) - - moreover - have "\i. i < length Rs \ - ((x \\<^bsub>(RDirProd_list Rs)\<^esub> y) \\<^bsub>(RDirProd_list Rs)\<^esub> z) ! i = - ((x \\<^bsub>(RDirProd_list Rs)\<^esub> z) \\<^bsub>(RDirProd_list Rs)\<^esub> (y \\<^bsub>(RDirProd_list Rs)\<^esub> z)) ! i" - proof - - fix i assume i: "i < length Rs" - hence "((x \\<^bsub>(RDirProd_list Rs)\<^esub> y) \\<^bsub>(RDirProd_list Rs)\<^esub> z) ! i = - ((x ! i) \\<^bsub>(Rs ! i)\<^esub> (y ! i)) \\<^bsub>(Rs ! i)\<^esub> (z ! i)" - using RDirProd_list_m_output RDirProd_list_a_output assms - by (metis RDirProd_list_abelian_group abelian_groupE(1) lessThan_iff ring.is_abelian_group) - also have " ... = ((x ! i) \\<^bsub>(Rs ! i)\<^esub> (z ! i)) \\<^bsub>(Rs ! i)\<^esub> ((y ! i) \\<^bsub>(Rs ! i)\<^esub> (z ! i))" - by (metis RDirProd_list_carrier_elts RDirProd_list_in_carrierE - i assms lessThan_iff ring.ring_simprules(13)) - also - have " ... = ((x \\<^bsub>(RDirProd_list Rs)\<^esub> z) \\<^bsub>(RDirProd_list Rs)\<^esub> (y \\<^bsub>(RDirProd_list Rs)\<^esub> z)) ! i" - using RDirProd_list_m_output RDirProd_list_a_output assms - by (metis RDirProd_list_monoid i lessThan_iff monoid.m_closed ring.is_monoid) - finally - show "((x \\<^bsub>(RDirProd_list Rs)\<^esub> y) \\<^bsub>(RDirProd_list Rs)\<^esub> z) ! i = - ((x \\<^bsub>(RDirProd_list Rs)\<^esub> z) \\<^bsub>(RDirProd_list Rs)\<^esub> (y \\<^bsub>(RDirProd_list Rs)\<^esub> z)) ! i" . - qed - - moreover have "length ((x \\<^bsub>(RDirProd_list Rs)\<^esub> y) \\<^bsub>(RDirProd_list Rs)\<^esub> z) = length Rs" - by (meson RDirProd_list_abelian_group RDirProd_list_carrier_elts RDirProd_list_monoid - abelian_groupE(1) assms monoid.m_closed ring.is_abelian_group ring.is_monoid) - - ultimately show ?thesis by (simp add: nth_equalityI) -qed - -lemma RDirProd_list_r_distr: - assumes "x \ carrier (RDirProd_list Rs)" - and "y \ carrier (RDirProd_list Rs)" - and "z \ carrier (RDirProd_list Rs)" - and "\i. i < length Rs \ ring (Rs ! i)" - shows "z \\<^bsub>(RDirProd_list Rs)\<^esub> (x \\<^bsub>(RDirProd_list Rs)\<^esub> y) = - (z \\<^bsub>(RDirProd_list Rs)\<^esub> x) \\<^bsub>(RDirProd_list Rs)\<^esub> (z \\<^bsub>(RDirProd_list Rs)\<^esub> y)" -proof - - have "length (z \\<^bsub>(RDirProd_list Rs)\<^esub> (x \\<^bsub>(RDirProd_list Rs)\<^esub> y)) = - length ((z \\<^bsub>(RDirProd_list Rs)\<^esub> x) \\<^bsub>(RDirProd_list Rs)\<^esub> (z \\<^bsub>(RDirProd_list Rs)\<^esub> y))" - by (metis RDirProd_list_abelian_group RDirProd_list_carrier_elts RDirProd_list_monoid - abelian_groupE(1) assms monoid.m_closed ring.is_abelian_group ring.is_monoid) - - moreover - have "\i. i < length Rs \ - (z \\<^bsub>(RDirProd_list Rs)\<^esub> (x \\<^bsub>(RDirProd_list Rs)\<^esub> y)) ! i = - ((z \\<^bsub>(RDirProd_list Rs)\<^esub> x) \\<^bsub>(RDirProd_list Rs)\<^esub> (z \\<^bsub>(RDirProd_list Rs)\<^esub> y)) ! i" - proof - - fix i assume i: "i < length Rs" - hence "(z \\<^bsub>(RDirProd_list Rs)\<^esub> (x \\<^bsub>(RDirProd_list Rs)\<^esub> y)) ! i = - (z ! i) \\<^bsub>(Rs ! i)\<^esub> ((x ! i) \\<^bsub>(Rs ! i)\<^esub> (y ! i))" - using RDirProd_list_m_output RDirProd_list_a_output assms - by (metis RDirProd_list_abelian_group abelian_groupE(1) lessThan_iff ring.is_abelian_group) - also have " ... = ((z ! i) \\<^bsub>(Rs ! i)\<^esub> (x ! i)) \\<^bsub>(Rs ! i)\<^esub> ((z ! i) \\<^bsub>(Rs ! i)\<^esub> (y ! i))" - by (metis RDirProd_list_carrier_elts RDirProd_list_in_carrierE - assms i lessThan_iff ring.ring_simprules(23)) - also - have " ... = ((z \\<^bsub>(RDirProd_list Rs)\<^esub> x) \\<^bsub>(RDirProd_list Rs)\<^esub> (z \\<^bsub>(RDirProd_list Rs)\<^esub> y)) ! i" - using RDirProd_list_m_output RDirProd_list_a_output assms - by (metis RDirProd_list_monoid i lessThan_iff monoid.m_closed ring.is_monoid) - finally - show "(z \\<^bsub>(RDirProd_list Rs)\<^esub> (x \\<^bsub>(RDirProd_list Rs)\<^esub> y)) ! i = - ((z \\<^bsub>(RDirProd_list Rs)\<^esub> x) \\<^bsub>(RDirProd_list Rs)\<^esub> (z \\<^bsub>(RDirProd_list Rs)\<^esub> y)) ! i" . - qed - - moreover have "length (z \\<^bsub>(RDirProd_list Rs)\<^esub> (x \\<^bsub>(RDirProd_list Rs)\<^esub> y)) = length Rs" - by (meson RDirProd_list_abelian_group RDirProd_list_carrier_elts RDirProd_list_monoid - abelian_groupE(1) assms monoid.m_closed ring.is_abelian_group ring.is_monoid) - - ultimately show ?thesis by (simp add: nth_equalityI) -qed - -theorem RDirProd_list_ring: - assumes "\i. i < length Rs \ ring (Rs ! i)" - shows "ring (RDirProd_list Rs)" - using assms unfolding ring_def ring_axioms_def using assms - by (meson RDirProd_list_abelian_group RDirProd_list_l_distr - RDirProd_list_monoid RDirProd_list_r_distr) - -theorem RDirProd_list_cring: - assumes "\i. i < length Rs \ cring (Rs ! i)" - shows "cring (RDirProd_list Rs)" - by (meson RDirProd_list_comm_monoid RDirProd_list_ring assms cring_def) - -corollary (in cring) RDirProd_list_of_quot_is_cring: - assumes "\i. i < n \ ideal (I i) R" - shows "cring (RDirProd_list (map (\i. R Quot (I i)) [0..< n]))" - (is "cring (RDirProd_list ?Rs)") -proof - - have "\i. i \ {..<(length ?Rs)} \ cring (?Rs ! i)" - by (simp add: assms ideal.quotient_is_cring is_cring) - thus ?thesis - using RDirProd_list_cring by blast -qed - -lemma length_RDirProd_list_0: - assumes "\i. i < n \ cring (F i)" - shows "length (\\<^bsub>(RDirProd_list (map F [0..< n]))\<^esub>) = n" - by (metis (no_types, lifting) add_cancel_right_left RDirProd_list_carrier_elts RDirProd_list_cring cring.cring_simprules(2) diff_zero length_map length_upt nth_map_upt assms) - -lemma RDirProd_list_isomorphism1: - "(\(hd, tl). hd # tl) \ ring_iso (RDirProd R (RDirProd_list Rs)) (RDirProd_list (R # Rs))" - unfolding ring_iso_def ring_hom_def bij_betw_def inj_on_def RDirProd_def - by (auto simp add: monoid.defs) - -lemma RDirProd_list_isomorphism1': - "(RDirProd R (RDirProd_list Rs)) \ (RDirProd_list (R # Rs))" - unfolding is_ring_iso_def using RDirProd_list_isomorphism1 by blast - -lemma RDirProd_list_isomorphism2: - "(\r. (hd r, tl r)) \ ring_iso (RDirProd_list (R # Rs)) (RDirProd R (RDirProd_list Rs))" - unfolding ring_iso_def ring_hom_def bij_betw_def inj_on_def RDirProd_def -proof (auto simp add: monoid.defs) - let ?\ = "\r. (hd r, tl r)" - fix a b assume "a \ carrier R" "b \ carrier (DirProd_list Rs)" - hence "a # b \ {r # rs |r rs. r \ carrier R \ rs \ carrier (DirProd_list Rs)} \ ?\ (a # b) = (a, b)" - by simp - thus "(a, b) \ ?\ ` {r # rs |r rs. r \ carrier R \ rs \ carrier (DirProd_list Rs)}" - by (metis (no_types, lifting) image_iff) -qed - -lemma RDirProd_list_isomorphism3: - "(\(r, l). r @ [l]) \ ring_iso (RDirProd (RDirProd_list Rs) S) (RDirProd_list (Rs @ [S]))" -proof (induction Rs rule: DirProd_list.induct) - case 1 thus ?case - unfolding ring_iso_def ring_hom_def bij_betw_def inj_on_def RDirProd_def - by (auto simp add: monoid.defs image_iff) +lemma (in ring) canonical_proj_ker: + assumes "ideal I R" and "ideal J R" + shows "a_kernel R (RDirProd (R Quot I) (R Quot J)) (canonical_proj I J) = I \ J" +proof + show "a_kernel R (RDirProd (R Quot I) (R Quot J)) (canonical_proj I J) \ I \ J" + unfolding FactRing_def RDirProd_def DirProd_def a_kernel_def' + by (auto simp add: assms[THEN ideal.rcos_const_imp_mem] monoid.defs) next - case (2 R Rs) - - { fix r1 r2 assume A0: "r1 \ carrier (RDirProd_list (R # Rs))" - and A1: "r2 \ carrier (RDirProd_list (R # Rs))" - have "length r1 \ 1" - and "length r2 \ 1" - and "length (r1 \\<^bsub>(RDirProd_list (R # Rs))\<^esub> r2) \ 1" - and "length (r1 \\<^bsub>(RDirProd_list (R # Rs))\<^esub> r2) \ 1" - and "length (\\<^bsub>(RDirProd_list (R # Rs))\<^esub>) \ 1" - proof - - show len_r1: "length r1 \ 1" - and len_r2: "length r2 \ 1" - by (metis RDirProd_list_carrier_elts A0 A1 length_Cons less_one nat.simps(3) not_less)+ - show "length (r1 \\<^bsub>(RDirProd_list (R # Rs))\<^esub> r2) \ 1" - and "length (r1 \\<^bsub>(RDirProd_list (R # Rs))\<^esub> r2) \ 1" - and "length (\\<^bsub>(RDirProd_list (R # Rs))\<^esub>) \ 1" - using len_r1 len_r2 by (simp_all add: monoid.defs) - qed } note aux_lemma = this - - moreover - have "(\(r, s). (hd r, (tl r, s))) \ - ring_iso (RDirProd (RDirProd_list (R # Rs)) S) - (RDirProd R (RDirProd (RDirProd_list Rs) S))" - using ring_iso_set_trans[OF RDirProd_isomorphism4[OF RDirProd_list_isomorphism2[of R Rs],of S] - RDirProd_isomorphism3[of R "RDirProd_list Rs" S]] - by (simp add: case_prod_beta' comp_def) - - moreover - have "(\(hd, (tl, s)). hd # (tl @ [s])) \ - ring_iso (RDirProd R (RDirProd (RDirProd_list Rs) S)) - (RDirProd_list (R # (Rs @ [S])))" - using ring_iso_set_trans[OF RDirProd_isomorphism5[OF 2(1), of R] - RDirProd_list_isomorphism1[of R "Rs @ [S]"]] - by (simp add: case_prod_beta' comp_def) - - moreover - have "(\(r, s). (hd r) # ((tl r) @ [s])) \ - ring_iso (RDirProd (RDirProd_list (R # Rs)) S) (RDirProd_list (R # (Rs @ [S])))" - using ring_iso_set_trans[OF calculation(6-7)] by (simp add: case_prod_beta' comp_def) - hence iso: "(\(r, s). (hd r # tl r) @ [s]) \ - ring_iso (RDirProd (RDirProd_list (R # Rs)) S) (RDirProd_list ((R # Rs) @ [S]))" by simp - - show ?case - proof (rule ring_iso_morphic_prop[OF iso, of "\r. length (fst r) \ 1" "\(r, s). r @ [s]"]) - show "\r. 1 \ length (fst r) \ - (case r of (r, s) \ (hd r # tl r) @ [s]) = (case r of (r, s) \ r @ [s])" - by (simp add: Suc_le_eq case_prod_beta') - show "morphic_prop (RDirProd (RDirProd_list (R # Rs)) S) (\r. 1 \ length (fst r))" - unfolding RDirProd_def by (rule morphic_propI) (auto simp add: monoid.defs aux_lemma) + show "I \ J \ a_kernel R (RDirProd (R Quot I) (R Quot J)) (canonical_proj I J)" + proof + fix s assume s: "s \ I \ J" then have "I +> s = I" and "J +> s = J" + using abelian_subgroupI3[OF ideal.axioms(1) is_abelian_group] + by (simp add: abelian_subgroup.a_rcos_const assms)+ + thus "s \ a_kernel R (RDirProd (R Quot I) (R Quot J)) (canonical_proj I J)" + unfolding FactRing_def RDirProd_def DirProd_def a_kernel_def' + using ideal.Icarr[OF assms(1)] s by (simp add: monoid.defs) qed qed - -subsection \Second Generalization - The Extended Canonical Projection is a Homomorphism and - Description of its Kernel\ - -theorem (in cring) canonical_proj_ext_is_hom: - fixes n::nat - assumes "\i. i < n \ ideal (I i) R" - and "\i j. \ i < n; j < n; i \ j \ \ I i <+> I j = carrier R" - shows "(\a. (map (\i. (I i) +> a) [0..< n])) \ - ring_hom R (RDirProd_list (map (\i. R Quot (I i)) [0..< n]))" (is "?\ \ ?ring_hom") -proof (rule ring_hom_memI) - { fix x assume x: "x \ carrier R" - have "?\ x \ carrier (RDirProd_list (map (\i. R Quot I i) [0.. ring_hom R (RDirProd (R Quot I) (R Quot J))" + unfolding RDirProd_def DirProd_def FactRing_def A_RCOSETS_def' + by (auto intro!: ring_hom_memI + simp add: assms[THEN ideal.rcoset_mult_add] + assms[THEN ideal.a_rcos_sum] monoid.defs) - fix x y assume x: "x \ carrier R" and y: "y \ carrier R" - show x': "?\ x \ carrier (RDirProd_list (map (\i. R Quot I i) [0.. x \ carrier (DirProd_list (map (\i. R Quot I i) [0.. y \ carrier (RDirProd_list (map (\i. R Quot I i) [0..i. I i +> y) [0.. carrier (DirProd_list (map (\i. R Quot I i) [0.. J = carrier R" + shows "R Quot (I \ J) \ RDirProd (R Quot I) (R Quot J)" + using ring_hom_ring.FactRing_iso[OF canonical_proj_ring_hom canonical_proj_is_surj] + by (simp add: canonical_proj_ker assms) - show "?\ (x \ y) = ?\ x \\<^bsub>RDirProd_list (map (\i. R Quot I i) [0.. ?\ y" - apply (rule nth_equalityI) - apply (metis RDirProd_list_carrier_elts RDirProd_list_of_quot_is_cring assms(1) - cring.cring_simprules(5) length_map x' y') - apply (simp add: monoid.defs) - using DirProd_list_m_output [of "?\ x" "(map (\i. R Quot I i) [0.. y"] x'' y'' - by (simp add: x'' y'' FactRing_def assms(1) ideal.rcoset_mult_add x y) + +subsection \Chinese Remainder Generalized\ + +lemma (in ring) canonical_proj_ext_zero [simp]: "(canonical_proj_ext I 0) = (\a. [ (I 0) +> a ])" + unfolding canonical_proj_ext_def by simp - show "?\ (x \ y) = ?\ x \\<^bsub>RDirProd_list (map (\i. R Quot I i) [0.. ?\ y" - proof - - have "length (?\ (x \ y)) = - length ((?\ x) \\<^bsub>RDirProd_list (map (\i. R Quot I i) [0.. (?\ y))" - by (metis RDirProd_list_carrier_elts RDirProd_list_of_quot_is_cring assms(1) - cring.cring_simprules(1) length_map x' y') +lemma (in ring) canonical_proj_ext_tl: + "(\a. canonical_proj_ext I (Suc n) a) = (\a. ((I 0) +> a) # (canonical_proj_ext (\i. I (Suc i)) n a))" + unfolding canonical_proj_ext_def by (induct n) (auto, metis (lifting) append.assoc append_Cons append_Nil) + +lemma (in ring) canonical_proj_ext_is_hom: + assumes "\i. i \ n \ ideal (I i) R" + shows "(canonical_proj_ext I n) \ ring_hom R (RDirProd_list (map (\i. R Quot (I i)) [0..< Suc n]))" + using assms +proof (induct n arbitrary: I) + case 0 thus ?case + using RDirProd_list_hom2[OF ideal.rcos_ring_hom[of _ R]] by (simp add: canonical_proj_ext_def) +next + let ?DirProd = "\I n. RDirProd_list (map (\i. R Quot (I i)) [0..j. j < n \ - (?\ (x \ y)) ! j = ((?\ x) \\<^bsub>RDirProd_list (map (\i. R Quot I i) [0.. (?\ y)) ! j" - proof - - fix j assume j: "j < n" - have "(?\ (x \ y)) ! j = I j +> x \ y" using j by simp - also have " ... = (I j +> x) \\<^bsub>(R Quot I j)\<^esub> (I j +> y)" - by (simp add: FactRing_def abelian_subgroup.a_rcos_sum abelian_subgroupI3 - assms(1) ideal.axioms(1) is_abelian_group j x y) - also have " ... = ((?\ x) \\<^bsub>RDirProd_list (map (\i. R Quot I i) [0.. (?\ y)) ! j" - proof - - have "R Quot I j = map (\n. R Quot I n) [0.. x = I ([0.. x" - "I j +> y = I ([0.. y" - by (simp_all add: j) - moreover have "\n ns f. n < length ns \ map f ns ! n = (f (ns ! n::nat)::'a set)" - by simp - moreover have "\B ps C n. \B \ carrier (RDirProd_list ps); C \ carrier (RDirProd_list ps); n < length ps\ - \ (B \\<^bsub>RDirProd_list ps\<^esub> C) ! n = (B ! n::'a set) \\<^bsub>ps ! n\<^esub> C ! n" - by (meson RDirProd_list_a_output) - ultimately show ?thesis - by (metis (mono_tags, lifting) diff_zero j length_map length_upt x' y') - qed - finally show "(?\ (x \ y)) ! j = - ((?\ x) \\<^bsub>RDirProd_list (map (\i. R Quot I i) [0.. (?\ y)) ! j" . - qed - ultimately show "?\ (x \ y) = ?\ x \\<^bsub>RDirProd_list (map (\i. R Quot I i) [0.. ?\ y" - by (simp add: list_eq_iff_nth_eq) - qed -next - show "(?\ \) = \\<^bsub>RDirProd_list (map (\i. R Quot I i) [0.." - apply (rule nth_equalityI) - apply (metis RDirProd_list_carrier_elts cring.cring_simprules(6) - RDirProd_list_of_quot_is_cring assms(1) length_map) - using DirProd_list_one[where ?Rs = "map (\i. R Quot I i) [0..i. I (Suc i)) n) \ ring_hom R (?DirProd (\i. I (Suc i)) n)" + using Suc(1)[of "\i. I (Suc i)"] Suc(2) by simp + have [simp]: + "(\(a, as). a # as) \ ((\(r, s). (I 0 +> r, ?proj (\i. I (Suc i)) n s)) \ (\a. (a, a))) = ?proj I (Suc n)" + unfolding canonical_proj_ext_tl by auto + moreover have + "(R Quot I 0) # (map (\i. R Quot I (Suc i)) [0..< Suc n]) = map (\i. R Quot (I i)) [0..< Suc (Suc n)]" + by (induct n) (auto) + moreover show ?case + using ring_hom_trans[OF ring_hom_trans[OF RDirProd_hom1 + RDirProd_hom3[OF ideal.rcos_ring_hom[OF I] hom]] RDirProd_list_hom1] + unfolding calculation(2) by simp qed -theorem (in cring) canonical_proj_ext_kernel: - fixes n::nat +lemma (in ring) RDirProd_Quot_list_is_ring: + assumes "\i. i \ n \ ideal (I i) R" shows "ring (RDirProd_list (map (\i. R Quot (I i)) [0..< Suc n]))" +proof - + have ring_list: "\i. i < Suc n \ ring ((map (\i. R Quot I i) [0..< Suc n]) ! i)" + using ideal.quotient_is_ring[OF assms] + by (metis add.left_neutral diff_zero le_simps(2) nth_map_upt) + show ?thesis + using RDirProd_list_is_ring[OF ring_list] by simp +qed + +lemma (in ring) canonical_proj_ext_ring_hom: assumes "\i. i \ n \ ideal (I i) R" - and "\i j. \ i \ n; j \ n; i \ j \ \ I i <+> I j = carrier R" - shows "(\i \ n. I i) = a_kernel R (RDirProd_list (map (\i. R Quot (I i)) [0..< Suc n])) - (\a. (map (\i. (I i) +> a) [0..< Suc n]))" + shows "ring_hom_ring R (RDirProd_list (map (\i. R Quot (I i)) [0..< Suc n])) (canonical_proj_ext I n)" proof - - let ?\ = "\a. (map (\i. (I i) +> a) [0..< Suc n])" + have ring: "ring (RDirProd_list (map (\i. R Quot (I i)) [0..< Suc n]))" + using RDirProd_Quot_list_is_ring[OF assms] by simp show ?thesis - proof - show "(\i \ n. I i) \ a_kernel R (RDirProd_list (map (\i. R Quot (I i)) [0..< Suc n])) ?\" - proof - fix s assume s: "s \ (\i \ n. I i)" - hence "\i. i \ n \ (I i) +> s = I i" - by (simp add: additive_subgroup.zero_closed assms ideal.axioms(1) ideal.set_add_zero) - hence "\i. i \ n \ (?\ s) ! i = I i" - by (metis add.left_neutral diff_zero le_imp_less_Suc nth_map_upt) - moreover have - "\i. i \ n \ (\\<^bsub>(RDirProd_list (map (\i. R Quot (I i)) [0..< Suc n]))\<^esub>) ! i = - \\<^bsub>(R Quot (I i))\<^esub>" - using RDirProd_list_zero[where ?Rs = "map (\i. R Quot I i) [0..i. i \ n \ (\\<^bsub>(RDirProd_list (map (\i. R Quot (I i)) [0..< Suc n]))\<^esub>) ! i = I i" - unfolding FactRing_def by simp - moreover - have "length (\\<^bsub>(RDirProd_list (map (\i. R Quot (I i)) [0..< Suc n]))\<^esub>) = Suc n" - by (subst length_RDirProd_list_0) (simp_all add: length_RDirProd_list_0 assms(1) ideal.quotient_is_cring is_cring) - moreover have "length (?\ s) = Suc n" by simp - ultimately have "?\ s = \\<^bsub>(RDirProd_list (map (\i. R Quot (I i)) [0..< Suc n]))\<^esub>" - by (simp add: list_eq_iff_nth_eq) - moreover have "s \ carrier R" - using additive_subgroup.a_Hcarr assms(1) ideal.axioms(1) s by fastforce - ultimately show "s \ a_kernel R (RDirProd_list (map (\i. R Quot (I i)) [0..< Suc n])) ?\" - using a_kernel_def'[of R "RDirProd_list (map (\i. R Quot (I i)) [0..< Suc n])"] by simp - qed + using canonical_proj_ext_is_hom assms ring_hom_ring.intro[OF ring_axioms ring] + unfolding ring_hom_ring_axioms_def by blast +qed + +lemma (in ring) canonical_proj_ext_ker: + assumes "\i. i \ (n :: nat) \ ideal (I i) R" + shows "a_kernel R (RDirProd_list (map (\i. R Quot (I i)) [0..< Suc n])) (canonical_proj_ext I n) = (\i \ n. I i)" +proof - + let ?map_Quot = "\I n. map (\i. R Quot (I i)) [0..< Suc n]" + let ?ker = "\I n. a_kernel R (RDirProd_list (?map_Quot I n)) (canonical_proj_ext I n)" + + from assms show ?thesis + proof (induct n arbitrary: I) + case 0 then have I: "ideal (I 0) R" by simp + show ?case + unfolding a_kernel_def' RDirProd_list_zero canonical_proj_ext_def FactRing_def + using ideal.rcos_const_imp_mem a_rcos_zero ideal.Icarr I by auto next - show "a_kernel R (RDirProd_list (map (\i. R Quot (I i)) [0..< Suc n])) ?\ \ (\i \ n. I i)" - proof - fix s assume s: "s \ a_kernel R (RDirProd_list (map (\i. R Quot (I i)) [0..< Suc n])) ?\" - hence "?\ s = \\<^bsub>(RDirProd_list (map (\i. R Quot (I i)) [0..< Suc n]))\<^esub>" - unfolding a_kernel_def kernel_def by (simp add: monoid.defs) - hence "(I i) +> s = \\<^bsub>(R Quot (I i))\<^esub>" if "i \ n" for i - using RDirProd_list_zero[where ?Rs = "map (\i. R Quot I i) [0..i. i \ n \ (I i) +> s = I i" - unfolding FactRing_def by simp - moreover have "s \ carrier R" - using s unfolding a_kernel_def kernel_def by simp - ultimately show "s \ (\i \ n. I i)" - using ideal.set_add_zero_imp_mem[where ?i = s and ?R = R] by (simp add: assms(1)) - qed + case (Suc n) + hence I: "ideal (I 0) R" by simp + have map_simp: "?map_Quot I (Suc n) = (R Quot I 0) # (?map_Quot (\i. I (Suc i)) n)" + by (induct n) (auto) + have ker_I0: "I 0 = a_kernel R (R Quot (I 0)) (\a. (I 0) +> a)" + using ideal.rcos_const_imp_mem[OF I] a_rcos_zero[OF I] ideal.Icarr[OF I] + unfolding a_kernel_def' FactRing_def by auto + hence "?ker I (Suc n) = (?ker (\i. I (Suc i)) n) \ (I 0)" + unfolding a_kernel_def' map_simp RDirProd_list_zero' canonical_proj_ext_tl by auto + moreover have "?ker (\i. I (Suc i)) n = (\i \ n. I (Suc i))" + using Suc(1)[of "\i. I (Suc i)"] Suc(2) by auto + ultimately show ?case + by (auto simp add: INT_extend_simps(10) atMost_atLeast0) + (metis atLeastAtMost_iff le_zero_eq not_less_eq_eq) qed qed +lemma (in cring) canonical_proj_ext_is_surj: + assumes "\i. i \ n \ ideal (I i) R" and "\i j. \ i \ n; j \ n \ \ i \ j \ I i <+> I j = carrier R" + shows "(canonical_proj_ext I n) ` carrier R = carrier (RDirProd_list (map (\i. R Quot (I i)) [0..< Suc n]))" + using assms +proof (induct n arbitrary: I) + case 0 show ?case + by (auto simp add: RDirProd_list_carrier FactRing_def A_RCOSETS_def') +next + { fix S :: "'c ring" and T :: "'d ring" and f g + assume A: "ring T" "f \ ring_hom R S" "g \ ring_hom R T" "f ` carrier R \ f ` (a_kernel R T g)" + have "(\a. (f a, g a)) ` carrier R = (f ` carrier R) \ (g ` carrier R)" + proof + show "(\a. (f a, g a)) ` carrier R \ (f ` carrier R) \ (g ` carrier R)" + by blast + next + show "(f ` carrier R) \ (g ` carrier R) \ (\a. (f a, g a)) ` carrier R" + proof + fix t assume "t \ (f ` carrier R) \ (g ` carrier R)" + then obtain a b where a: "a \ carrier R" "f a = fst t" and b: "b \ carrier R" "g b = snd t" + by auto + obtain c where c: "c \ a_kernel R T g" "f c = f (a \ b)" + using A(4) minus_closed[OF a(1) b (1)] by auto + have "f (c \ b) = f (a \ b) \\<^bsub>S\<^esub> f b" + using ring_hom_memE(3)[OF A(2)] b c unfolding a_kernel_def' by auto + hence "f (c \ b) = f a" + using ring_hom_memE(3)[OF A(2) minus_closed[of a b], of b] a b by algebra + moreover have "g (c \ b) = g b" + using ring_hom_memE(1,3)[OF A(3)] b(1) c ring.ring_simprules(8)[OF A(1)] + unfolding a_kernel_def' by auto + ultimately have "(\a. (f a, g a)) (c \ b) = t" and "c \ b \ carrier R" + using a b c unfolding a_kernel_def' by auto + thus "t \ (\a. (f a, g a)) ` carrier R" + by blast + qed + qed } note aux_lemma = this -subsection \Final Generalization\ + let ?map_Quot = "\I n. map (\i. R Quot (I i)) [0..< Suc n]" + let ?DirProd = "\I n. RDirProd_list (?map_Quot I n)" + let ?proj = "\I n. canonical_proj_ext I n" + + case (Suc n) + interpret I: ideal "I 0" R + Inter: ideal "\i \ n. I (Suc i)" R + using i_Intersect[of "(\i. I (Suc i)) ` {..n}"] Suc(2) by auto + + have map_simp: "?map_Quot I (Suc n) = (R Quot I 0) # (?map_Quot (\i. I (Suc i)) n)" + by (induct n) (auto) + + have IH: "(?proj (\i. I (Suc i)) n) ` carrier R = carrier (?DirProd (\i. I (Suc i)) n)" + and ring: "ring (?DirProd (\i. I (Suc i)) n)" + and hom: "?proj (\i. I (Suc i)) n \ ring_hom R (?DirProd (\i. I (Suc i)) n)" + using RDirProd_Quot_list_is_ring[of n "\i. I (Suc i)"] Suc(1)[of "\i. I (Suc i)"] + canonical_proj_ext_is_hom[of n "\i. I (Suc i)"] Suc(2-3) by auto + + have ker: "a_kernel R (?DirProd (\i. I (Suc i)) n) (?proj (\i. I (Suc i)) n) = (\i \ n. I (Suc i))" + using canonical_proj_ext_ker[of n "\i. I (Suc i)"] Suc(2) by auto + have carrier_Quot: "carrier (R Quot (I 0)) = (\a. (I 0) +> a) ` carrier R" + by (auto simp add: RDirProd_list_carrier FactRing_def A_RCOSETS_def') + have ring: "ring (?DirProd (\i. I (Suc i)) n)" + and hom: "?proj (\i. I (Suc i)) n \ ring_hom R (?DirProd (\i. I (Suc i)) n)" + using RDirProd_Quot_list_is_ring[of n "\i. I (Suc i)"] + canonical_proj_ext_is_hom[of n "\i. I (Suc i)"] Suc(2) by auto + have "carrier (R Quot (I 0)) \ (\a. (I 0) +> a) ` (\i \ n. I (Suc i))" + proof + have "(\i \ {Suc 0.. Suc n}. I i) <+> (I 0) = carrier R" + using inter_plus_ideal_eq_carrier_arbitrary[of n I 0] + by (simp add: Suc(2-3) atLeast1_atMost_eq_remove0) + hence eq_carrier: "(I 0) <+> (\i \ n. I (Suc i)) = carrier R" + using set_add_comm[OF I.a_subset Inter.a_subset] + by (metis INT_extend_simps(10) atMost_atLeast0 image_Suc_atLeastAtMost) + + fix b assume "b \ carrier (R Quot (I 0))" + hence "(b, (\i \ n. I (Suc i))) \ carrier (R Quot I 0) \ carrier (R Quot (\i\n. I (Suc i)))" + using ring.ring_simprules(2)[OF Inter.quotient_is_ring] by (simp add: FactRing_def) + then obtain s + where "s \ carrier R" "(canonical_proj (I 0) (\i \ n. I (Suc i))) s = (b, (\i \ n. I (Suc i)))" + using canonical_proj_is_surj[OF I.is_ideal Inter.is_ideal eq_carrier] + unfolding RDirProd_carrier by (metis (no_types, lifting) imageE) + hence "s \ (\i \ n. I (Suc i))" and "(\a. (I 0) +> a) s = b" + using Inter.rcos_const_imp_mem by auto + thus "b \ (\a. (I 0) +> a) ` (\i \ n. I (Suc i))" + by auto + qed + hence "(\a. ((I 0) +> a, ?proj (\i. I (Suc i)) n a)) ` carrier R = + carrier (R Quot (I 0)) \ carrier (?DirProd (\i. I (Suc i)) n)" + using aux_lemma[OF ring I.rcos_ring_hom hom] unfolding carrier_Quot ker IH by simp + moreover show ?case + unfolding map_simp RDirProd_list_carrier sym[OF calculation(1)] canonical_proj_ext_tl by auto +qed theorem (in cring) chinese_remainder: - fixes n::nat - assumes "\i. i \ n \ ideal (I i) R" - and "\i j. \ i \ n; j \ n; i \ j \ \ I i <+> I j = carrier R" - shows "R Quot (\i \ n. I i) \ RDirProd_list (map (\i. R Quot (I i)) [0..< Suc n])" - using assms -proof (induct n) - case 0 - have "(\r. (r, [])) \ ring_iso (R Quot (I 0)) (RDirProd (R Quot (I 0)) (RDirProd_list []))" - unfolding ring_iso_def ring_hom_def bij_betw_def inj_on_def RDirProd_def - by (auto simp add: monoid.defs) - hence "(R Quot (I 0)) \ (RDirProd (R Quot (I 0)) (RDirProd_list []))" - unfolding is_ring_iso_def by blast - moreover - have "RDirProd ((R Quot (I 0)) :: 'a set ring) - (RDirProd_list ([] :: (('a set) ring) list)) \ (RDirProd_list [ (R Quot (I 0)) ])" - using RDirProd_list_isomorphism1'[of "(R Quot (I 0)) :: 'a set ring" "[] :: (('a set) ring) list"] . - ultimately show ?case - using ring_iso_trans by simp -next - case (Suc n) - have inter_ideal: "ideal (\ i \ n. I i) R" - using Suc.prems(1) i_Intersect[of "I ` {..n}"] atMost_Suc atLeast1_atMost_eq_remove0 by auto - hence "R Quot (\ i \ Suc n. I i) \ RDirProd (R Quot (\ i \ n. I i)) (R Quot (I (Suc n)))" - using chinese_remainder_simple[of "\ i \ n. I i" "I (Suc n)"] inter_plus_ideal_eq_carrier[of n I] - by (simp add: Int_commute Suc.prems(1-2) atMost_Suc) - moreover have "R Quot (\ i \ n. I i) \ RDirProd_list (map (\i. R Quot (I i)) [0..< Suc n])" - using Suc.hyps Suc.prems(1-2) by auto - hence "RDirProd (R Quot (\ i \ n. I i)) (R Quot (I (Suc n))) \ - RDirProd (RDirProd_list (map (\i. R Quot (I i)) [0..< Suc n])) (R Quot (I (Suc n)))" - unfolding is_ring_iso_def using RDirProd_isomorphism4 by blast - ultimately - have "R Quot (\ i \ Suc n. I i) \ - RDirProd (RDirProd_list (map (\i. R Quot (I i)) [0..< Suc n])) (R Quot (I (Suc n)))" - using ring_iso_trans by simp - moreover - have "RDirProd (RDirProd_list (map (\i. R Quot (I i)) [0..< Suc n])) (R Quot (I (Suc n))) \ - RDirProd_list ((map (\i. R Quot (I i)) [0..< Suc n]) @ [ R Quot (I (Suc n)) ])" - using RDirProd_list_isomorphism3 unfolding is_ring_iso_def by blast - hence "RDirProd (RDirProd_list (map (\i. R Quot (I i)) [0..< Suc n])) (R Quot (I (Suc n))) \ - RDirProd_list (map (\i. R Quot (I i)) [0..< Suc (Suc n)])" by simp - ultimately show ?case using ring_iso_trans by simp -qed + assumes "\i. i \ n \ ideal (I i) R" and "\i j. \ i \ n; j \ n \ \ i \ j \ I i <+> I j = carrier R" + shows "R Quot (\i \ n. I i) \ RDirProd_list (map (\i. R Quot (I i)) [0..< Suc n])" + using ring_hom_ring.FactRing_iso[OF canonical_proj_ext_ring_hom, of n I] + canonical_proj_ext_is_surj[of n I] canonical_proj_ext_ker[of n I] assms + by auto -end +end \ No newline at end of file diff -r 12dce58bcd3f -r 1b5178abaf97 src/HOL/Algebra/Coset.thy --- a/src/HOL/Algebra/Coset.thy Thu Oct 04 11:18:39 2018 +0200 +++ b/src/HOL/Algebra/Coset.thy Thu Oct 04 15:25:47 2018 +0100 @@ -440,6 +440,13 @@ shows "N \ G" using assms normal_inv_iff by blast +corollary (in group) normal_invE: + assumes "N \ G" + shows "subgroup N G" and "\x h. \ x \ carrier G; h \ N \ \ x \ h \ inv x \ N" + using assms normal_inv_iff apply blast + by (simp add: assms normal.inv_op_closed2) + + lemma (in group) one_is_normal: "{\} \ G" proof(intro normal_invI) show "subgroup {\} G" @@ -1076,7 +1083,7 @@ (* Next two lemmas contributed by Paulo Emílio de Vilhena. *) lemma (in group_hom) trivial_hom_iff: - "(h ` (carrier G) = { \\<^bsub>H\<^esub> }) = (kernel G H h = carrier G)" + "h ` (carrier G) = { \\<^bsub>H\<^esub> } \ kernel G H h = carrier G" unfolding kernel_def using one_closed by force lemma (in group_hom) trivial_ker_imp_inj: @@ -1091,6 +1098,90 @@ using A G.inv_equality G.inv_inv by blast qed +(* NEW ========================================================================== *) +lemma (in group_hom) inj_iff_trivial_ker: + shows "inj_on h (carrier G) \ kernel G H h = { \ }" +proof + assume inj: "inj_on h (carrier G)" show "kernel G H h = { \ }" + unfolding kernel_def + proof (auto) + fix a assume "a \ carrier G" "h a = \\<^bsub>H\<^esub>" thus "a = \" + using inj hom_one unfolding inj_on_def by force + qed +next + show "kernel G H h = { \ } \ inj_on h (carrier G)" + using trivial_ker_imp_inj by simp +qed + +(* NEW ========================================================================== *) +lemma (in group_hom) induced_group_hom': + assumes "subgroup I G" shows "group_hom (G \ carrier := I \) H h" +proof - + have "h \ hom (G \ carrier := I \) H" + using homh subgroup.subset[OF assms] unfolding hom_def by (auto, meson hom_mult subsetCE) + thus ?thesis + using subgroup.subgroup_is_group[OF assms G.group_axioms] group_axioms + unfolding group_hom_def group_hom_axioms_def by auto +qed + +(* NEW ========================================================================== *) +lemma (in group_hom) inj_on_subgroup_iff_trivial_ker: + assumes "subgroup I G" + shows "inj_on h I \ kernel (G \ carrier := I \) H h = { \ }" + using group_hom.inj_iff_trivial_ker[OF induced_group_hom'[OF assms]] by simp + +(* NEW ========================================================================== *) +lemma set_mult_hom: + assumes "h \ hom G H" "I \ carrier G" and "J \ carrier G" + shows "h ` (I <#>\<^bsub>G\<^esub> J) = (h ` I) <#>\<^bsub>H\<^esub> (h ` J)" +proof + show "h ` (I <#>\<^bsub>G\<^esub> J) \ (h ` I) <#>\<^bsub>H\<^esub> (h ` J)" + proof + fix a assume "a \ h ` (I <#>\<^bsub>G\<^esub> J)" + then obtain i j where i: "i \ I" and j: "j \ J" and "a = h (i \\<^bsub>G\<^esub> j)" + unfolding set_mult_def by auto + hence "a = (h i) \\<^bsub>H\<^esub> (h j)" + using assms unfolding hom_def by blast + thus "a \ (h ` I) <#>\<^bsub>H\<^esub> (h ` J)" + using i and j unfolding set_mult_def by auto + qed +next + show "(h ` I) <#>\<^bsub>H\<^esub> (h ` J) \ h ` (I <#>\<^bsub>G\<^esub> J)" + proof + fix a assume "a \ (h ` I) <#>\<^bsub>H\<^esub> (h ` J)" + then obtain i j where i: "i \ I" and j: "j \ J" and "a = (h i) \\<^bsub>H\<^esub> (h j)" + unfolding set_mult_def by auto + hence "a = h (i \\<^bsub>G\<^esub> j)" + using assms unfolding hom_def by fastforce + thus "a \ h ` (I <#>\<^bsub>G\<^esub> J)" + using i and j unfolding set_mult_def by auto + qed +qed + +(* NEW ========================================================================== *) +corollary coset_hom: + assumes "h \ hom G H" "I \ carrier G" "a \ carrier G" + shows "h ` (a <#\<^bsub>G\<^esub> I) = h a <#\<^bsub>H\<^esub> (h ` I)" and "h ` (I #>\<^bsub>G\<^esub> a) = (h ` I) #>\<^bsub>H\<^esub> h a" + unfolding l_coset_eq_set_mult r_coset_eq_set_mult using assms set_mult_hom[OF assms(1)] by auto + +(* NEW ========================================================================== *) +corollary (in group_hom) set_mult_ker_hom: + assumes "I \ carrier G" + shows "h ` (I <#> (kernel G H h)) = h ` I" and "h ` ((kernel G H h) <#> I) = h ` I" +proof - + have ker_in_carrier: "kernel G H h \ carrier G" + unfolding kernel_def by auto + + have "h ` (kernel G H h) = { \\<^bsub>H\<^esub> }" + unfolding kernel_def by force + moreover have "h ` I \ carrier H" + using assms by auto + hence "(h ` I) <#>\<^bsub>H\<^esub> { \\<^bsub>H\<^esub> } = h ` I" and "{ \\<^bsub>H\<^esub> } <#>\<^bsub>H\<^esub> (h ` I) = h ` I" + unfolding set_mult_def by force+ + ultimately show "h ` (I <#> (kernel G H h)) = h ` I" and "h ` ((kernel G H h) <#> I) = h ` I" + using set_mult_hom[OF homh assms ker_in_carrier] set_mult_hom[OF homh ker_in_carrier assms] by simp+ +qed + (* Next subsection contributed by Martin Baillon. *) diff -r 12dce58bcd3f -r 1b5178abaf97 src/HOL/Algebra/Cycles.thy --- a/src/HOL/Algebra/Cycles.thy Thu Oct 04 11:18:39 2018 +0200 +++ b/src/HOL/Algebra/Cycles.thy Thu Oct 04 15:25:47 2018 +0100 @@ -8,8 +8,10 @@ section \Cycles\ -abbreviation cycle :: "'a list \ bool" where - "cycle cs \ distinct cs" +subsection \Definitions\ + +abbreviation cycle :: "'a list \ bool" + where "cycle cs \ distinct cs" fun cycle_of_list :: "'a list \ 'a \ 'a" where @@ -17,41 +19,26 @@ | "cycle_of_list cs = id" -subsection\Cycles as Rotations\ +subsection \Basic Properties\ -text\We start proving that the function derived from a cycle rotates its support list.\ +text \We start proving that the function derived from a cycle rotates its support list.\ lemma id_outside_supp: - assumes "cycle cs" and "x \ set cs" - shows "(cycle_of_list cs) x = x" using assms -proof (induction cs rule: cycle_of_list.induct) - case (1 i j cs) - have "cycle_of_list (i # j # cs) x = (Fun.swap i j id) (cycle_of_list (j # cs) x)" by simp - also have " ... = (Fun.swap i j id) x" - using "1.IH" "1.prems"(1) "1.prems"(2) by auto - also have " ... = x" using 1(3) by simp - finally show ?case . -next - case "2_1" thus ?case by simp -next - case "2_2" thus ?case by simp -qed + assumes "x \ set cs" shows "(cycle_of_list cs) x = x" + using assms by (induct cs rule: cycle_of_list.induct) (simp_all) -lemma cycle_permutes: - assumes "cycle cs" - shows "permutation (cycle_of_list cs)" -proof (induction rule: cycle_of_list.induct) - case (1 i j cs) thus ?case - by (metis cycle_of_list.simps(1) permutation_compose permutation_swap_id) -next - case "2_1" thus ?case by simp -next - case "2_2" thus ?case by simp -qed +lemma permutation_of_cycle: "permutation (cycle_of_list cs)" +proof (induct cs rule: cycle_of_list.induct) + case 1 thus ?case + using permutation_compose[OF permutation_swap_id] unfolding comp_apply by simp +qed simp_all + +lemma cycle_permutes: "(cycle_of_list cs) permutes (set cs)" + using permutation_bijective[OF permutation_of_cycle] id_outside_supp[of _ cs] + by (simp add: bij_iff permutes_def) theorem cyclic_rotation: - assumes "cycle cs" - shows "map ((cycle_of_list cs) ^^ n) cs = rotate n cs" + assumes "cycle cs" shows "map ((cycle_of_list cs) ^^ n) cs = rotate n cs" proof - { have "map (cycle_of_list cs) cs = rotate1 cs" using assms(1) proof (induction cs rule: cycle_of_list.induct) @@ -72,92 +59,52 @@ also have " ... = rotate1 (i # j # cs)" by simp finally show ?thesis . qed - next - case "2_1" thus ?case by simp - next - case "2_2" thus ?case by simp - qed } + qed simp_all } note cyclic_rotation' = this - from assms show ?thesis - proof (induction n) - case 0 thus ?case by simp - next - case (Suc n) - have "map ((cycle_of_list cs) ^^ (Suc n)) cs = - map (cycle_of_list cs) (map ((cycle_of_list cs) ^^ n) cs)" by simp - also have " ... = map (cycle_of_list cs) (rotate n cs)" - by (simp add: Suc.IH assms) - also have " ... = rotate (Suc n) cs" - using cyclic_rotation' - by (metis rotate1_rotate_swap rotate_Suc rotate_map) - finally show ?case . - qed + show ?thesis + using cyclic_rotation' by (induct n) (auto, metis map_map rotate1_rotate_swap rotate_map) qed corollary cycle_is_surj: - assumes "cycle cs" - shows "(cycle_of_list cs) ` (set cs) = (set cs)" - using cyclic_rotation[OF assms, of 1] by (simp add: image_set) + assumes "cycle cs" shows "(cycle_of_list cs) ` (set cs) = (set cs)" + using cyclic_rotation[OF assms, of "Suc 0"] by (simp add: image_set) corollary cycle_is_id_root: - assumes "cycle cs" - shows "(cycle_of_list cs) ^^ (length cs) = id" + assumes "cycle cs" shows "(cycle_of_list cs) ^^ (length cs) = id" proof - have "map ((cycle_of_list cs) ^^ (length cs)) cs = cs" - by (simp add: assms cyclic_rotation) - hence "\x. x \ (set cs) \ ((cycle_of_list cs) ^^ (length cs)) x = x" - using map_eq_conv by fastforce - moreover have "\n x. x \ (set cs) \ ((cycle_of_list cs) ^^ n) x = x" - proof - - fix n show "\x. x \ (set cs) \ ((cycle_of_list cs) ^^ n) x = x" - proof (induction n) - case 0 thus ?case by simp - next - case (Suc n) thus ?case using id_outside_supp[OF assms] by simp - qed - qed - hence "\x. x \ (set cs) \ ((cycle_of_list cs) ^^ (length cs)) x = x" by simp + unfolding cyclic_rotation[OF assms] by simp + hence "((cycle_of_list cs) ^^ (length cs)) i = i" if "i \ set cs" for i + using that map_eq_conv by fastforce + moreover have "((cycle_of_list cs) ^^ n) i = i" if "i \ set cs" for i n + using id_outside_supp[OF that] by (induct n) (simp_all) ultimately show ?thesis - by (meson eq_id_iff) + by fastforce qed -corollary - assumes "cycle cs" - shows "(cycle_of_list cs) = (cycle_of_list (rotate n cs))" +corollary cycle_of_list_rotate_independent: + assumes "cycle cs" shows "(cycle_of_list cs) = (cycle_of_list (rotate n cs))" proof - - { fix cs :: "'a list" assume A: "cycle cs" + { fix cs :: "'a list" assume cs: "cycle cs" have "(cycle_of_list cs) = (cycle_of_list (rotate1 cs))" - proof - have "\x. x \ set cs \ (cycle_of_list cs) x = (cycle_of_list (rotate1 cs)) x" - proof - - have "cycle (rotate1 cs)" using A by simp - hence "map (cycle_of_list (rotate1 cs)) (rotate1 cs) = (rotate 2 cs)" - using cyclic_rotation by (metis Suc_eq_plus1 add.left_neutral - funpow.simps(2) funpow_simps_right(1) o_id one_add_one rotate_Suc rotate_def) - also have " ... = map (cycle_of_list cs) (rotate1 cs)" - using cyclic_rotation[OF A] - by (metis One_nat_def Suc_1 funpow.simps(2) id_apply map_map rotate0 rotate_Suc) - finally have "map (cycle_of_list (rotate1 cs)) (rotate1 cs) = - map (cycle_of_list cs) (rotate1 cs)" . - moreover fix x assume "x \ set cs" - ultimately show "(cycle_of_list cs) x = (cycle_of_list (rotate1 cs)) x" by auto - qed - moreover have "\x. x \ set cs \ (cycle_of_list cs) x = (cycle_of_list (rotate1 cs)) x" - using A by (simp add: id_outside_supp) - ultimately show "\x. (cycle_of_list cs) x = (cycle_of_list (rotate1 cs)) x" by blast - qed } - note rotate1_lemma = this + proof - + from cs have rotate1_cs: "cycle (rotate1 cs)" by simp + hence "map (cycle_of_list (rotate1 cs)) (rotate1 cs) = (rotate 2 cs)" + using cyclic_rotation[OF rotate1_cs, of 1] by (simp add: numeral_2_eq_2) + moreover have "map (cycle_of_list cs) (rotate1 cs) = (rotate 2 cs)" + using cyclic_rotation[OF cs] + by (metis One_nat_def Suc_1 funpow.simps(2) id_apply map_map rotate0 rotate_Suc) + ultimately have "(cycle_of_list cs) i = (cycle_of_list (rotate1 cs)) i" if "i \ set cs" for i + using that map_eq_conv unfolding sym[OF set_rotate1[of cs]] by fastforce + moreover have "(cycle_of_list cs) i = (cycle_of_list (rotate1 cs)) i" if "i \ set cs" for i + using that by (simp add: id_outside_supp) + ultimately show "(cycle_of_list cs) = (cycle_of_list (rotate1 cs))" + by blast + qed } note rotate1_lemma = this - show "cycle_of_list cs = cycle_of_list (rotate n cs)" - proof (induction n) - case 0 thus ?case by simp - next - case (Suc n) - have "cycle (rotate n cs)" by (simp add: assms) - thus ?case using rotate1_lemma[of "rotate n cs"] - by (simp add: Suc.IH) - qed + show ?thesis + using rotate1_lemma[of "rotate n cs"] by (induct n) (auto, metis assms distinct_rotate rotate1_lemma) qed @@ -190,563 +137,421 @@ subsection\When Cycles Commute\ lemma cycles_commute: - assumes "cycle \1" "cycle \2" and "set \1 \ set \2 = {}" - shows "(cycle_of_list \1) \ (cycle_of_list \2) = (cycle_of_list \2) \ (cycle_of_list \1)" -proof - - { fix \1 :: "'a list" and \2 :: "'a list" and x :: "'a" - assume A: "cycle \1" "cycle \2" "set \1 \ set \2 = {}" "x \ set \1" "x \ set \2" - have "((cycle_of_list \1) \ (cycle_of_list \2)) x = - ((cycle_of_list \2) \ (cycle_of_list \1)) x" + assumes "cycle p" "cycle q" and "set p \ set q = {}" + shows "(cycle_of_list p) \ (cycle_of_list q) = (cycle_of_list q) \ (cycle_of_list p)" +proof + { fix p :: "'a list" and q :: "'a list" and i :: "'a" + assume A: "cycle p" "cycle q" "set p \ set q = {}" "i \ set p" "i \ set q" + have "((cycle_of_list p) \ (cycle_of_list q)) i = + ((cycle_of_list q) \ (cycle_of_list p)) i" proof - - have "((cycle_of_list \1) \ (cycle_of_list \2)) x = (cycle_of_list \1) x" - using id_outside_supp[OF A(2) A(5)] by simp - also have " ... = ((cycle_of_list \2) \ (cycle_of_list \1)) x" - using id_outside_supp[OF A(2), of "(cycle_of_list \1) x"] - cycle_is_surj[OF A(1)] A(3) A(4) by fastforce + have "((cycle_of_list p) \ (cycle_of_list q)) i = (cycle_of_list p) i" + using id_outside_supp[OF A(5)] by simp + also have " ... = ((cycle_of_list q) \ (cycle_of_list p)) i" + using id_outside_supp[of "(cycle_of_list p) i"] cycle_is_surj[OF A(1)] A(3,4) by fastforce finally show ?thesis . - qed } - note aux_lemma = this - - let ?\12 = "\x. ((cycle_of_list \1) \ (cycle_of_list \2)) x" - let ?\21 = "\x. ((cycle_of_list \2) \ (cycle_of_list \1)) x" + qed } note aui_lemma = this - show ?thesis - proof - fix x have "x \ set \1 \ set \2 \ x \ set \1 \ set \2" by blast - from this show "?\12 x = ?\21 x" - proof - assume "x \ set \1 \ set \2" - hence "(x \ set \1 \ x \ set \2) \ (x \ set \1 \ x \ set \2)" using assms(3) by blast - from this show "?\12 x = ?\21 x" - proof - assume "x \ set \1 \ x \ set \2" thus ?thesis - using aux_lemma[OF assms(1-3)] by simp - next - assume "x \ set \1 \ x \ set \2" thus ?thesis - using assms aux_lemma inf_commute by metis - qed - next - assume "x \ set \1 \ set \2" thus ?thesis using id_outside_supp assms(1-2) - by (metis UnCI comp_apply) - qed + fix i consider "i \ set p" "i \ set q" | "i \ set p" "i \ set q" | "i \ set p" "i \ set q" + using \set p \ set q = {}\ by blast + thus "((cycle_of_list p) \ (cycle_of_list q)) i = ((cycle_of_list q) \ (cycle_of_list p)) i" + proof cases + case 1 thus ?thesis + using aui_lemma[OF assms] by simp + next + case 2 thus ?thesis + using aui_lemma[OF assms(2,1)] assms(3) by (simp add: ac_simps(8)) + next + case 3 thus ?thesis + by (simp add: id_outside_supp) qed qed -subsection\Cycles from Permutations\ +subsection \Cycles from Permutations\ + +subsubsection \Exponentiation of permutations\ -subsubsection\Exponentiation of permutations\ +text \Some important properties of permutations before defining how to extract its cycles.\ -text\Some important properties of permutations before defining how to extract its cycles\ +lemma permutation_funpow: + assumes "permutation p" shows "permutation (p ^^ n)" + using assms by (induct n) (simp_all add: permutation_compose) -lemma exp_of_permutation1: - assumes "p permutes S" - shows "(p ^^ n) permutes S" using assms -proof (induction n) - case 0 thus ?case by (simp add: permutes_def) -next - case (Suc n) thus ?case by (metis funpow_Suc_right permutes_compose) -qed +lemma permutes_funpow: + assumes "p permutes S" shows "(p ^^ n) permutes S" + using assms by (induct n) (simp add: permutes_def, metis funpow_Suc_right permutes_compose) -lemma exp_of_permutation2: - assumes "p permutes S" - and "i < j" "(p ^^ j) = (p ^^ i)" - shows "(p ^^ (j - i)) = id" using assms +lemma funpow_diff: + assumes "inj p" and "i \ j" "(p ^^ i) a = (p ^^ j) a" shows "(p ^^ (j - i)) a = a" proof - - have "(p ^^ i) \ (p ^^ (j - i)) = (p ^^ j)" - by (metis add_diff_inverse_nat assms(2) funpow_add le_eq_less_or_eq not_le) - also have " ... = (p ^^ i)" using assms(3) by simp - finally have "(p ^^ i) \ (p ^^ (j - i)) = (p ^^ i)" . - moreover have "bij (p ^^ i)" using exp_of_permutation1[OF assms(1)] - using permutes_bij by auto - ultimately show ?thesis - by (metis (no_types, lifting) bij_is_inj comp_assoc fun.map_id inv_o_cancel) + have "(p ^^ i) ((p ^^ (j - i)) a) = (p ^^ i) a" + using assms(2-3) by (metis (no_types) add_diff_inverse_nat funpow_add not_le o_def) + thus ?thesis + unfolding inj_eq[OF inj_fn[OF assms(1)], of i] . qed -lemma exp_of_permutation3: - assumes "p permutes S" "finite S" - shows "\n. (p ^^ n) = id \ n > 0" -proof (rule ccontr) - assume "\n. (p ^^ n) = id \ 0 < n" - hence S: "\n. n > 0 \ (p ^^ n) \ id" by auto - hence "\i j. \ i \ 0; j \ 0 \ \ i \ j \ (p ^^ i) \ (p ^^ j)" - proof - - fix i :: "nat" and j :: "nat" assume "i \ 0" "j \ 0" and Ineq: "i \ j" - show "(p ^^ i) \ (p ^^ j)" - proof (rule ccontr) - assume "\ (p ^^ i) \ (p ^^ j)" hence Eq: "(p ^^ i) = (p ^^ j)" by simp - have "(p ^^ (j - i)) = id" if "j > i" - using Eq exp_of_permutation2[OF assms(1) that] by simp - moreover have "(p ^^ (i - j)) = id" if "i > j" - using Eq exp_of_permutation2[OF assms(1) that] by simp - ultimately show False using Ineq S - by (meson le_eq_less_or_eq not_le zero_less_diff) - qed +lemma permutation_is_nilpotent: + assumes "permutation p" obtains n where "(p ^^ n) = id" and "n > 0" +proof - + obtain S where "finite S" and "p permutes S" + using assms unfolding permutation_permutes by blast + hence "\n. (p ^^ n) = id \ n > 0" + proof (induct S arbitrary: p) + case empty thus ?case + using id_funpow[of 1] unfolding permutes_empty by blast + next + case (insert s S) + have "(\n. (p ^^ n) s) ` UNIV \ (insert s S)" + using permutes_in_image[OF permutes_funpow[OF insert(4)], of _ s] by auto + hence "\ inj_on (\n. (p ^^ n) s) UNIV" + using insert(1) infinite_iff_countable_subset unfolding sym[OF finite_insert, of S s] by metis + then obtain i j where ij: "i < j" "(p ^^ i) s = (p ^^ j) s" + unfolding inj_on_def by (metis nat_neq_iff) + hence "(p ^^ (j - i)) s = s" + using funpow_diff[OF permutes_inj[OF insert(4)]] le_eq_less_or_eq by blast + hence "p ^^ (j - i) permutes S" + using permutes_superset[OF permutes_funpow[OF insert(4), of "j - i"], of S] by auto + then obtain n where n: "((p ^^ (j - i)) ^^ n) = id" "n > 0" + using insert(3) by blast + thus ?case + using ij(1) nat_0_less_mult_iff zero_less_diff unfolding funpow_mult by metis qed - hence "bij_betw (\i. (p ^^ i)) {i :: nat . i \ 0} {(p ^^ i) | i :: nat . i \ 0}" - unfolding bij_betw_def inj_on_def by blast - hence "infinite {(p ^^ i) | i :: nat . i \ 0}" - using bij_betw_finite by auto - moreover have "{(p ^^ i) | i :: nat . i \ 0} \ {\. \ permutes S}" - using exp_of_permutation1[OF assms(1)] by blast - hence "finite {(p ^^ i) | i :: nat . i \ 0}" - by (simp add: assms(2) finite_permutations finite_subset) - ultimately show False .. + thus thesis + using that by blast qed -lemma power_prop: - assumes "(p ^^ k) x = x" - shows "(p ^^ (k * l)) x = x" -proof (induction l) - case 0 thus ?case by simp -next - case (Suc l) - hence "(p ^^ (k * (Suc l))) x = ((p ^^ (k * l)) \ (p ^^ k)) x" - by (metis funpow_Suc_right funpow_mult) - also have " ... = (p ^^ (k * l)) x" - by (simp add: assms) - also have " ... = x" - using Suc.IH Suc.prems assms by blast - finally show ?case . -qed - -lemma exp_of_permutation4: - assumes "p permutes S" "finite S" - shows "\n. (p ^^ n) = id \ n > m" +lemma permutation_is_nilpotent': + assumes "permutation p" obtains n where "(p ^^ n) = id" and "n > m" proof - - obtain k where "k > 0" "(p ^^ k) = id" - using exp_of_permutation3[OF assms] by blast - moreover obtain n where "n * k > m" - by (metis calculation(1) dividend_less_times_div mult.commute mult_Suc_right) - ultimately show ?thesis - by (metis (full_types) funpow_mult id_funpow mult.commute) + obtain n where "(p ^^ n) = id" and "n > 0" + using permutation_is_nilpotent[OF assms] by blast + then obtain k where "n * k > m" + by (metis dividend_less_times_div mult_Suc_right) + from \(p ^^ n) = id\ have "p ^^ (n * k) = id" + by (induct k) (simp, metis funpow_mult id_funpow) + with \n * k > m\ show thesis + using that by blast qed -subsubsection\Extraction of cycles from permutations\ +subsubsection \Extraction of cycles from permutations\ -definition - least_power :: "('a \ 'a) \ 'a \ nat" +definition least_power :: "('a \ 'a) \ 'a \ nat" where "least_power f x = (LEAST n. (f ^^ n) x = x \ n > 0)" -abbreviation - support :: "('a \ 'a) \ 'a \ 'a list" +abbreviation support :: "('a \ 'a) \ 'a \ 'a list" where "support p x \ map (\i. (p ^^ i) x) [0..< (least_power p x)]" -lemma least_power_wellfounded: - assumes "permutation p" - shows "(p ^^ (least_power p x)) x = x" -proof - - obtain S where "p permutes S" "finite S" - using assms permutation_permutes by blast - hence "\n. (p ^^ n) x = x \ n > 0" - using eq_id_iff exp_of_permutation3 by metis - thus ?thesis unfolding least_power_def - by (metis (mono_tags, lifting) LeastI_ex) -qed + +lemma least_powerI: + assumes "(f ^^ n) x = x" and "n > 0" + shows "(f ^^ (least_power f x)) x = x" and "least_power f x > 0" + using assms unfolding least_power_def by (metis (mono_tags, lifting) LeastI)+ -lemma least_power_gt_zero: - assumes "permutation p" - shows "least_power p x > 0" -proof (rule ccontr) - obtain S where "p permutes S" "finite S" - using assms permutation_permutes by blast - hence Ex: "\n. (p ^^ n) x = x \ n > 0" - using eq_id_iff exp_of_permutation3 by metis - assume "\ 0 < least_power p x" hence "least_power p x = 0" - using Ex unfolding least_power_def by (metis (mono_tags, lifting) LeastI) - thus False unfolding least_power_def - by (metis (mono_tags, lifting) Ex LeastI_ex less_irrefl) -qed +lemma least_power_le: + assumes "(f ^^ n) x = x" and "n > 0" shows "least_power f x \ n" + using assms unfolding least_power_def by (simp add: Least_le) + +lemma least_power_of_permutation: + assumes "permutation p" shows "(p ^^ (least_power p a)) a = a" and "least_power p a > 0" + using permutation_is_nilpotent[OF assms] least_powerI by (metis id_apply)+ lemma least_power_gt_one: - assumes "permutation p" "p x \ x" - shows "least_power p x > Suc 0" -proof (rule ccontr) - obtain S where "p permutes S" "finite S" - using assms permutation_permutes by blast - hence Ex: "\n. (p ^^ n) x = x \ n > 0" - using eq_id_iff exp_of_permutation3 by metis - assume "\ Suc 0 < least_power p x" hence "least_power p x = (Suc 0)" - using Ex unfolding least_power_def by (metis (mono_tags, lifting) LeastI Suc_lessI) - hence "p x = x" using least_power_wellfounded[OF assms(1), of x] by simp - from \p x = x\ and \p x \ x\ show False by simp + assumes "permutation p" and "p a \ a" shows "least_power p a > Suc 0" + using least_power_of_permutation[OF assms(1)] assms(2) + by (metis Suc_lessI funpow.simps(2) funpow_simps_right(1) o_id) + +lemma least_power_minimal: + assumes "(p ^^ n) a = a" shows "(least_power p a) dvd n" +proof (cases "n = 0", simp) + let ?lpow = "least_power p" + + assume "n \ 0" then have "n > 0" by simp + hence "(p ^^ (?lpow a)) a = a" and "least_power p a > 0" + using assms unfolding least_power_def by (metis (mono_tags, lifting) LeastI)+ + hence aux_lemma: "(p ^^ ((?lpow a) * k)) a = a" for k :: nat + by (induct k) (simp_all add: funpow_add) + + have "(p ^^ (n mod ?lpow a)) ((p ^^ (n - (n mod ?lpow a))) a) = (p ^^ n) a" + by (metis add_diff_inverse_nat funpow_add mod_less_eq_dividend not_less o_apply) + with \(p ^^ n) a = a\ have "(p ^^ (n mod ?lpow a)) a = a" + using aux_lemma by (simp add: minus_mod_eq_mult_div) + hence "?lpow a \ n mod ?lpow a" if "n mod ?lpow a > 0" + using least_power_le[OF _ that, of p a] by simp + with \least_power p a > 0\ show "(least_power p a) dvd n" + using mod_less_divisor not_le by blast qed -lemma least_power_bound: - assumes "permutation p" shows "\m > 0. (least_power p x) \ m" -proof - - obtain S where "p permutes S" "finite S" - using assms permutation_permutes by blast - hence "\n. (p ^^ n) x = x \ n > 0" - using eq_id_iff exp_of_permutation3 by metis - then obtain m :: "nat" where "m > 0" "m = (least_power p x)" - unfolding least_power_def by (metis (mono_tags, lifting) LeastI_ex) - thus ?thesis by blast +lemma least_power_dvd: + assumes "permutation p" shows "(least_power p a) dvd n \ (p ^^ n) a = a" +proof + show "(p ^^ n) a = a \ (least_power p a) dvd n" + using least_power_minimal[of _ p] by simp +next + have "(p ^^ ((least_power p a) * k)) a = a" for k :: nat + using least_power_of_permutation(1)[OF assms(1)] by (induct k) (simp_all add: funpow_add) + thus "(least_power p a) dvd n \ (p ^^ n) a = a" by blast qed -lemma lt_least_power: - assumes "Suc n = least_power p x" - and "0 < i" "i \ n" - shows "(p ^^ i) x \ x" -proof (rule ccontr) - assume "\ (p ^^ i) x \ x" hence "(p ^^ i) x = x" by simp - hence "i \ {n. (p ^^ n) x = x \ n > 0}" - using assms(2-3) by blast - moreover have "i < least_power p x" - using assms(3) assms(1) by linarith - ultimately show False unfolding least_power_def - using not_less_Least by auto +theorem cycle_of_permutation: + assumes "permutation p" shows "cycle (support p a)" +proof - + have "(least_power p a) dvd (j - i)" if "i \ j" "j < least_power p a" and "(p ^^ i) a = (p ^^ j) a" for i j + using funpow_diff[OF bij_is_inj that(1,3)] assms by (simp add: permutation least_power_dvd) + moreover have "i = j" if "i \ j" "j < least_power p a" and "(least_power p a) dvd (j - i)" for i j + using that le_eq_less_or_eq nat_dvd_not_less by auto + ultimately have "inj_on (\i. (p ^^ i) a) {..< (least_power p a)}" + unfolding inj_on_def by (metis le_cases lessThan_iff) + thus ?thesis + by (simp add: atLeast_upt distinct_map) qed -lemma least_power_welldefined: - assumes "permutation p" and "y \ {(p ^^ k) x | k. k \ 0}" - shows "least_power p x = least_power p y" -proof - - have aux_lemma: "\z. least_power p z = least_power p (p z)" - proof - - fix z - have "(p ^^ (least_power p z)) z = z" - by (metis assms(1) least_power_wellfounded) - hence "(p ^^ (least_power p z)) (p z) = (p z)" - by (metis funpow_swap1) - hence "least_power p z \ least_power p (p z)" - by (metis assms(1) inc_induct le_SucE least_power_gt_zero lt_least_power nat_le_linear) + +subsection \Decomposition on Cycles\ + +text \We show that a permutation can be decomposed on cycles\ + +subsubsection \Preliminaries\ - moreover have "(p ^^ (least_power p (p z))) (p z) = (p z)" - by (simp add: assms(1) least_power_wellfounded) - hence "(p ^^ (least_power p (p z))) z = z" - by (metis assms(1) funpow_swap1 permutation_permutes permutes_def) - hence "least_power p z \ least_power p (p z)" - by (metis assms(1) least_power_gt_zero less_imp_Suc_add lt_least_power not_less_eq_eq) - - ultimately show "least_power p z = least_power p (p z)" by simp - qed - - obtain k where "k \ 0" "y = (p ^^ k) x" - using assms(2) by blast - thus ?thesis - proof (induction k arbitrary: x y) - case 0 thus ?case by simp - next - case (Suc k) - have "least_power p ((p ^^ k) x) = least_power p x" - using Suc.IH by fastforce - thus ?case using aux_lemma - using Suc.prems(2) by auto +lemma support_set: + assumes "permutation p" shows "set (support p a) = range (\i. (p ^^ i) a)" +proof + show "set (support p a) \ range (\i. (p ^^ i) a)" + by auto +next + show "range (\i. (p ^^ i) a) \ set (support p a)" + proof (auto) + fix i + have "(p ^^ i) a = (p ^^ (i mod (least_power p a))) ((p ^^ (i - (i mod (least_power p a)))) a)" + by (metis add_diff_inverse_nat funpow_add mod_less_eq_dividend not_le o_apply) + also have " ... = (p ^^ (i mod (least_power p a))) a" + using least_power_dvd[OF assms] by (metis dvd_minus_mod) + also have " ... \ (\i. (p ^^ i) a) ` {0..< (least_power p a)}" + using least_power_of_permutation(2)[OF assms] by fastforce + finally show "(p ^^ i) a \ (\i. (p ^^ i) a) ` {0..< (least_power p a)}" . qed qed -theorem cycle_of_permutation: - assumes "permutation p" - shows "cycle (support p x)" -proof (rule ccontr) - assume "\ cycle (support p x)" - hence "\ i j. i \ {0.. j \ {0.. i \ j \ (p ^^ i) x = (p ^^ j) x" - using atLeast_upt by (simp add: distinct_conv_nth) - then obtain i j where ij: "0 \ i" "i < j" "j < least_power p x" - and "(p ^^ i) x = (p ^^ j) x" - by (metis atLeast_upt le0 le_eq_less_or_eq lessThan_iff not_less set_upt) - hence "(p ^^ i) x = (p ^^ i) ((p ^^ (j - i)) x)" - by (metis add_diff_inverse_nat funpow_add not_less_iff_gr_or_eq o_apply) - hence "(p ^^ (j - i)) x = x" - using exp_of_permutation1 assms by (metis bij_pointE permutation_permutes permutes_bij) - moreover have "0 \ j - i \ j - i < least_power p x" - by (simp add: ij(3) less_imp_diff_less) - hence "(p ^^ (j - i)) x \ x" using lt_least_power ij - by (metis diff_le_self lessE less_imp_diff_less less_imp_le zero_less_diff) - ultimately show False by simp +lemma disjoint_support: + assumes "permutation p" shows "disjoint (range (\a. set (support p a)))" (is "disjoint ?A") +proof (rule disjointI) + { fix i j a b + assume "set (support p a) \ set (support p b) \ {}" have "set (support p a) \ set (support p b)" + unfolding support_set[OF assms] + proof (auto) + from \set (support p a) \ set (support p b) \ {}\ + obtain i j where ij: "(p ^^ i) a = (p ^^ j) b" + by auto + + fix k + have "(p ^^ k) a = (p ^^ (k + (least_power p a) * l)) a" for l + using least_power_dvd[OF assms] by (induct l) (simp, metis dvd_triv_left funpow_add o_def) + then obtain m where "m \ i" and "(p ^^ m) a = (p ^^ k) a" + using least_power_of_permutation(2)[OF assms] + by (metis dividend_less_times_div le_eq_less_or_eq mult_Suc_right trans_less_add2) + hence "(p ^^ m) a = (p ^^ (m - i)) ((p ^^ i) a)" + by (metis Nat.le_imp_diff_is_add funpow_add o_apply) + with \(p ^^ m) a = (p ^^ k) a\ have "(p ^^ k) a = (p ^^ ((m - i) + j)) b" + unfolding ij by (simp add: funpow_add) + thus "(p ^^ k) a \ range (\i. (p ^^ i) b)" + by blast + qed } note aux_lemma = this + + fix supp_a supp_b + assume "supp_a \ ?A" and "supp_b \ ?A" + then obtain a b where a: "supp_a = set (support p a)" and b: "supp_b = set (support p b)" + by auto + assume "supp_a \ supp_b" thus "supp_a \ supp_b = {}" + using aux_lemma unfolding a b by blast qed - -subsection\Decomposition on Cycles\ - -text\We show that a permutation can be decomposed on cycles\ - -subsubsection\Preliminaries\ - -lemma support_set: +lemma disjoint_support': assumes "permutation p" - shows "set (support p x) = {(p ^^ k) x | k. k \ 0}" + shows "set (support p a) \ set (support p b) = {} \ a \ set (support p b)" proof - - have "{(p ^^ k) x | k. k \ 0} = {(p ^^ k) x | k. 0 \ k \ k < (least_power p x)}" (is "?A = ?B") + have "a \ set (support p a)" + using least_power_of_permutation(2)[OF assms] by force + show ?thesis proof - show "?B \ ?A" by blast - next - show "?A \ ?B" - proof - fix y assume "y \ ?A" - then obtain k :: "nat" where k: "k \ 0" "(p ^^ k) x = y" by blast - hence "k = (least_power p x) * (k div (least_power p x)) + (k mod (least_power p x))" by simp - hence "y = (p ^^ ((least_power p x) * (k div (least_power p x)) + (k mod (least_power p x)))) x" - using k by auto - hence " y = (p ^^ (k mod (least_power p x))) x" - using power_prop[OF least_power_wellfounded[OF assms, of x], of "k div (least_power p x)"] - by (metis add.commute funpow_add o_apply) - moreover have "k mod (least_power p x) < least_power p x" - using k mod_less_divisor[of "least_power p x" k, OF least_power_gt_zero[OF assms]] by simp - ultimately show "y \ ?B" - by blast - qed - qed - - moreover have "{(p ^^ k) x | k. 0 \ k \ k < (least_power p x)} = set (support p x)" (is "?B = ?C") - proof - show "?B \ ?C" - proof - fix y assume "y \ {(p ^^ k) x | k. 0 \ k \ k < (least_power p x)}" - then obtain k where "k \ 0" "k < (least_power p x)" "y = (p ^^ k) x" by blast - thus "y \ ?C" by auto - qed + assume "set (support p a) \ set (support p b) = {}" + with \a \ set (support p a)\ show "a \ set (support p b)" + by blast next - show "?C \ ?B" - proof - fix y assume "y \ ?C" - then obtain k where "k \ 0" "k < (least_power p x)" "(support p x) ! k = y" by auto - thus "y \ ?B" by auto + assume "a \ set (support p b)" show "set (support p a) \ set (support p b) = {}" + proof (rule ccontr) + assume "set (support p a) \ set (support p b) \ {}" + hence "set (support p a) = set (support p b)" + using disjoint_support[OF assms] by (meson UNIV_I disjoint_def image_iff) + with \a \ set (support p a)\ and \a \ set (support p b)\ show False + by simp qed qed - - ultimately show ?thesis by simp -qed - -lemma disjoint_support: - assumes "p permutes S" "finite S" - shows "disjoint {{(p ^^ k) x | k. k \ 0} | x. x \ S}" (is "disjoint ?A") -proof (rule disjointI) - { fix a b assume "a \ ?A" "b \ ?A" "a \ b \ {}" - then obtain x y where x: "x \ S" "a = {(p ^^ k) x | k. k \ 0}" - and y: "y \ S" "b = {(p ^^ k) y | k. k \ 0}" by blast - assume "a \ b \ {}" - then obtain z kx ky where z: "kx \ 0" "ky \ 0" "z = (p ^^ kx) x" "z = (p ^^ ky) y" - using x(2) y(2) by blast - have "a \ b" - proof - fix w assume "w \ a" - then obtain k where k: "k \ 0" "w = (p ^^ k) x" using x by blast - define l where "l = (kx div (least_power p w)) + 1" - hence l: "l * (least_power p w) > kx" - using least_power_gt_zero assms One_nat_def add.right_neutral add_Suc_right - mult.commute permutation_permutes - by (metis dividend_less_times_div mult_Suc_right) - - have "w = (p ^^ (l * (least_power p w))) w" - by (metis assms least_power_wellfounded mult.commute permutation_permutes power_prop) - also have "... = (p ^^ (l * (least_power p w) + k)) x" - using k by (simp add: funpow_add) - also have " ... = (p ^^ (l * (least_power p w) + k - kx + kx)) x" - using l by auto - also have " ... = (p ^^ (l * (least_power p w) + k - kx)) ((p ^^ kx) x)" - by (simp add: funpow_add) - also have " ... = (p ^^ (l * (least_power p w) + k - kx)) ((p ^^ ky) y)" using z - by simp - finally have "w = (p ^^ (l * (least_power p w) + k - kx + ky)) y" - by (simp add: funpow_add) - thus "w \ b" using y by blast - qed } note aux_lemma = this - - fix a b assume ab: "a \ ?A" "b \ ?A" "a \ b" - show "a \ b = {}" - proof (rule ccontr) - assume "a \ b \ {}" thus False using aux_lemma ab - by (metis (no_types, lifting) inf.absorb2 inf.orderE) - qed qed lemma support_coverture: - assumes "p permutes S" "finite S" - shows "\{{(p ^^ k) x | k. k \ 0} | x. x \ S} = S" + assumes "permutation p" shows "\ { set (support p a) | a. p a \ a } = { a. p a \ a }" proof - show "\{{(p ^^ k) x |k. 0 \ k} |x. x \ S} \ S" + show "{ a. p a \ a } \ \ { set (support p a) | a. p a \ a }" proof - fix y assume "y \ \{{(p ^^ k) x |k. 0 \ k} |x. x \ S}" - then obtain x k where x: "x \ S" and k: "k \ 0" and y: "y = (p ^^ k) x" by blast - have "(p ^^ k) x \ S" - proof (induction k) - case 0 thus ?case using x by simp - next - case (Suc k) thus ?case using assms - by (simp add: permutes_in_image) - qed - thus "y \ S" using y by simp + fix a assume "a \ { a. p a \ a }" + have "a \ set (support p a)" + using least_power_of_permutation(2)[OF assms, of a] by force + with \a \ { a. p a \ a }\ show "a \ \ { set (support p a) | a. p a \ a }" + by blast qed next - show "S \ \{{(p ^^ k) x |k. 0 \ k} |x. x \ S}" + show "\ { set (support p a) | a. p a \ a } \ { a. p a \ a }" proof - fix x assume x: "x \ S" - hence "x \ {(p ^^ k) x |k. 0 \ k}" - by (metis (mono_tags, lifting) CollectI funpow_0 le_numeral_extra(3)) - thus "x \ \{{(p ^^ k) x |k. 0 \ k} |x. x \ S}" using x by blast + fix b assume "b \ \ { set (support p a) | a. p a \ a }" + then obtain a i where "p a \ a" and "(p ^^ i) a = b" + by auto + have "p a = a" if "(p ^^ i) a = (p ^^ Suc i) a" + using funpow_diff[OF bij_is_inj _ that] assms unfolding permutation by simp + with \p a \ a\ and \(p ^^ i) a = b\ show "b \ { a. p a \ a }" + by auto qed qed theorem cycle_restrict: - assumes "permutation p" "y \ set (support p x)" - shows "p y = (cycle_of_list (support p x)) y" + assumes "permutation p" and "b \ set (support p a)" shows "p b = (cycle_of_list (support p a)) b" proof - - have "\ i. \ 0 \ i; i < length (support p x) - 1 \ \ - p ((support p x) ! i) = (support p x) ! (i + 1)" - proof - - fix i assume i: "0 \ i" "i < length (support p x) - 1" - hence "p ((support p x) ! i) = p ((p ^^ i) x)" by simp - also have " ... = (p ^^ (i + 1)) x" by simp - also have " ... = (support p x) ! (i + 1)" - using i by simp - finally show "p ((support p x) ! i) = (support p x) ! (i + 1)" . + note least_power_props [simp] = least_power_of_permutation[OF assms(1)] + + have "map (cycle_of_list (support p a)) (support p a) = rotate1 (support p a)" + using cyclic_rotation[OF cycle_of_permutation[OF assms(1)], of 1 a] by simp + hence "map (cycle_of_list (support p a)) (support p a) = tl (support p a) @ [ a ]" + by (simp add: hd_map rotate1_hd_tl) + also have " ... = map p (support p a)" + proof (rule nth_equalityI, auto) + fix i assume "i < least_power p a" show "(tl (support p a) @ [a]) ! i = p ((p ^^ i) a)" + proof (cases) + assume i: "i = least_power p a - 1" + hence "(tl (support p a) @ [ a ]) ! i = a" + by (metis (no_types, lifting) diff_zero length_map length_tl length_upt nth_append_length) + also have " ... = p ((p ^^ i) a)" + by (metis (mono_tags, hide_lams) least_power_props i Suc_diff_1 funpow_simps_right(2) funpow_swap1 o_apply) + finally show ?thesis . + next + assume "i \ least_power p a - 1" + with \i < least_power p a\ have "i < least_power p a - 1" + by simp + hence "(tl (support p a) @ [ a ]) ! i = (p ^^ (Suc i)) a" + by (metis One_nat_def Suc_eq_plus1 add.commute length_map length_upt map_tl nth_append nth_map_upt tl_upt) + thus ?thesis + by simp + qed qed - hence 1: "map p (butlast (support p x)) = tl (support p x)" - using nth_butlast [where 'a='a] nth_tl [where 'a='a] - nth_equalityI[of "map p (butlast (support p x))" "tl (support p x)"] by force - have "p ((support p x) ! (length (support p x) - 1)) = p ((p ^^ (length (support p x) - 1)) x)" + finally have "map (cycle_of_list (support p a)) (support p a) = map p (support p a)" . + thus ?thesis using assms(2) by auto - also have " ... = (p ^^ (length (support p x))) x" - by (metis (mono_tags, lifting) Suc_pred' assms(2) funpow_Suc_right - funpow_swap1 length_pos_if_in_set o_apply) - also have " ... = x" - by (simp add: assms(1) least_power_wellfounded) - also have " ... = (support p x) ! 0" - by (simp add: assms(1) least_power_gt_zero) - finally have "p ((support p x) ! (length (support p x) - 1)) = (support p x) ! 0" . - hence 2: "p (last (support p x)) = hd (support p x)" - by (metis assms(2) hd_conv_nth last_conv_nth length_greater_0_conv length_pos_if_in_set) - - have "map p (support p x) = (tl (support p x)) @ [hd (support p x)]" using 1 2 - by (metis (no_types, lifting) assms(2) last_map length_greater_0_conv - length_pos_if_in_set list.map_disc_iff map_butlast snoc_eq_iff_butlast) - hence "map p (support p x) = rotate1 (support p x)" - by (metis assms(2) length_greater_0_conv length_pos_if_in_set rotate1_hd_tl) - - moreover have "map (cycle_of_list (support p x)) (support p x) = rotate1 (support p x)" - using cyclic_rotation[OF cycle_of_permutation[OF assms(1)], of 1 x] by simp - - ultimately show ?thesis using assms(2) - using map_eq_conv by fastforce qed subsubsection\Decomposition\ -inductive cycle_decomp :: "'a set \ ('a \ 'a) \ bool" where -empty: "cycle_decomp {} id" | -comp: "\ cycle_decomp I p; cycle cs; set cs \ I = {} \ \ - cycle_decomp (set cs \ I) ((cycle_of_list cs) \ p)" +inductive cycle_decomp :: "'a set \ ('a \ 'a) \ bool" + where + empty: "cycle_decomp {} id" + | comp: "\ cycle_decomp I p; cycle cs; set cs \ I = {} \ \ + cycle_decomp (set cs \ I) ((cycle_of_list cs) \ p)" lemma semidecomposition: - assumes "p permutes S" "finite S" - shows "(\y. if y \ (S - set (support p x)) then p y else y) permutes (S - set (support p x))" - (is "?q permutes ?S'") -proof - - have "\y. y \ S \ p y = y" - by (meson assms permutes_not_in) - - moreover have cycle_surj: "(cycle_of_list (support p x)) ` set (support p x) = set (support p x)" - using cycle_is_surj cycle_of_permutation assms permutation_permutes by metis - hence "\y. y \ set (support p x) \ p y \ set (support p x)" - using cycle_restrict assms permutation_permutes by (metis imageI) - - ultimately - have 1: "\y. y \ ?S' \ p y \ ?S'" by auto - have 2: "\y. y \ ?S' \ p y \ ?S'" - proof - - fix y assume y: "y \ ?S'" show "p y \ ?S'" - proof (rule ccontr) - assume "p y \ ?S'" hence "p y \ set (support p x)" - using assms(1) permutes_in_image y by fastforce - then obtain y' where y': "y' \ set (support p x)" "(cycle_of_list (support p x)) y' = p y" - using cycle_surj by (metis (mono_tags, lifting) imageE) - hence "p y' = p y" - using cycle_restrict assms permutation_permutes by metis - hence "y = y'" by (metis assms(1) permutes_def) - thus False using y y' by blast - qed - qed - - have "p ` ?S' = ?S'" - proof - - have "\ y. y \ ?S' \ \!x. p x = y" - by (metis assms(1) permutes_def) - hence "\ y. y \ ?S' \ \x \ ?S'. p x = y" using 1 by metis - thus ?thesis using 2 by blast - qed - hence "bij_betw p ?S' ?S'" - by (metis DiffD1 assms(1) bij_betw_subset permutes_imp_bij subsetI) - hence "bij_betw ?q ?S' ?S'" - by (rule rev_iffD1 [OF _ bij_betw_cong]) auto - moreover have "\y. y \ ?S' \ ?q y = y" by auto - ultimately show ?thesis - using bij_imp_permutes by blast -qed - + assumes "p permutes S" and "finite S" + shows "(\y. if y \ (S - set (support p a)) then p y else y) permutes (S - set (support p a))" +proof (rule bij_imp_permutes) + show "(if b \ (S - set (support p a)) then p b else b) = b" if "b \ S - set (support p a)" for b + using that by auto +next + have is_permutation: "permutation p" + using assms unfolding permutation_permutes by blast -lemma cycle_decomposition_aux: - assumes "p permutes S" "finite S" "card S = k" - shows "cycle_decomp S p" using assms -proof(induct arbitrary: S p rule: less_induct) - case (less x) thus ?case - proof (cases "S = {}") - case True thus ?thesis - by (metis empty less.prems(1) permutes_empty) + let ?q = "\y. if y \ (S - set (support p a)) then p y else y" + show "bij_betw ?q (S - set (support p a)) (S - set (support p a))" + proof (rule bij_betw_imageI) + show "inj_on ?q (S - set (support p a))" + using permutes_inj[OF assms(1)] unfolding inj_on_def by auto next - case False - then obtain x where x: "x \ S" by blast - define S' :: "'a set" where S': "S' = S - set (support p x)" - define q :: "'a \ 'a" where q: "q = (\x. if x \ S' then p x else x)" - hence q_permutes: "q permutes S'" - using semidecomposition[OF less.prems(1-2), of x] S' q by blast - moreover have "x \ set (support p x)" - by (metis (no_types, lifting) add.left_neutral diff_zero funpow_0 in_set_conv_nth least_power_gt_zero - length_map length_upt less.prems(1) less.prems(2) nth_map_upt permutation_permutes) - hence "card S' < card S" - by (metis Diff_iff S' \x \ S\ card_seteq leI less.prems(2) subsetI) - ultimately have "cycle_decomp S' q" - using S' less.hyps less.prems(2) less.prems(3) by blast - - moreover have "p = (cycle_of_list (support p x)) \ q" - proof - fix y show "p y = ((cycle_of_list (support p x)) \ q) y" - proof (cases) - assume y: "y \ set (support p x)" hence "y \ S'" using S' by simp - hence "q y = y" using q by simp - thus ?thesis - using comp_apply cycle_restrict less.prems permutation_permutes y by fastforce - next - assume y: "y \ set (support p x)" thus ?thesis - proof (cases) - assume "y \ S'" - hence "q y \ S'" using q_permutes - by (simp add: permutes_in_image) - hence "q y \ set (support p x)" - using S' by blast - hence "(cycle_of_list (support p x)) (q y) = (q y)" - by (metis cycle_of_permutation id_outside_supp less.prems(1-2) permutation_permutes) - thus ?thesis by (simp add: \y \ S'\ q) - next - assume "y \ S'" hence "y \ S" using y S' by blast - hence "(cycle_of_list (support p x) \ q) y = (cycle_of_list (support p x)) y" - by (simp add: \y \ S'\ q) - also have " ... = y" - by (metis cycle_of_permutation id_outside_supp less.prems(1-2) permutation_permutes y) - also have " ... = p y" - by (metis \y \ S\ less.prems(1) permutes_def) - finally show ?thesis by simp - qed - qed + have aux_lemma: "set (support p s) \ (S - set (support p a))" if "s \ S - set (support p a)" for s + proof - + have "(p ^^ i) s \ S" for i + using that unfolding permutes_in_image[OF permutes_funpow[OF assms(1)]] by simp + thus ?thesis + using that disjoint_support'[OF is_permutation, of s a] by auto qed - moreover have "cycle (support p x)" - using cycle_of_permutation less.prems permutation_permutes by fastforce - moreover have "set (support p x) \ S' = {}" using S' by simp - moreover have "set (support p x) \ S" - using support_coverture[OF less.prems(1-2)] support_set[of p x] x - permutation_permutes[of p] less.prems(1-2) by blast - hence "S = set (support p x) \ S'" using S' by blast - ultimately show ?thesis using comp[of S' q "support p x"] by auto + have "(p ^^ 1) s \ set (support p s)" for s + unfolding support_set[OF is_permutation] by blast + hence "p s \ set (support p s)" for s + by simp + hence "p ` (S - set (support p a)) \ S - set (support p a)" + using aux_lemma by blast + moreover have "(p ^^ ((least_power p s) - 1)) s \ set (support p s)" for s + unfolding support_set[OF is_permutation] by blast + hence "\s' \ set (support p s). p s' = s" for s + using least_power_of_permutation[OF is_permutation] by (metis Suc_diff_1 funpow.simps(2) o_apply) + hence "S - set (support p a) \ p ` (S - set (support p a))" + using aux_lemma + by (clarsimp simp add: image_iff) (metis image_subset_iff) + ultimately show "?q ` (S - set (support p a)) = (S - set (support p a))" + by auto qed qed theorem cycle_decomposition: - assumes "p permutes S" "finite S" - shows "cycle_decomp S p" - using assms cycle_decomposition_aux by blast + assumes "p permutes S" and "finite S" shows "cycle_decomp S p" + using assms +proof(induct "card S" arbitrary: S p rule: less_induct) + case less show ?case + proof (cases) + assume "S = {}" thus ?thesis + using empty less(2) by auto + next + have is_permutation: "permutation p" + using less(2-3) unfolding permutation_permutes by blast + + assume "S \ {}" then obtain s where "s \ S" + by blast + define q where "q = (\y. if y \ (S - set (support p s)) then p y else y)" + have "(cycle_of_list (support p s) \ q) = p" + proof + fix a + consider "a \ S - set (support p s)" | "a \ set (support p s)" | "a \ S" "a \ set (support p s)" + by blast + thus "((cycle_of_list (support p s) \ q)) a = p a" + proof cases + case 1 + have "(p ^^ 1) a \ set (support p a)" + unfolding support_set[OF is_permutation] by blast + with \a \ S - set (support p s)\ have "p a \ set (support p s)" + using disjoint_support'[OF is_permutation, of a s] by auto + with \a \ S - set (support p s)\ show ?thesis + using id_outside_supp[of _ "support p s"] unfolding q_def by simp + next + case 2 thus ?thesis + using cycle_restrict[OF is_permutation] unfolding q_def by simp + next + case 3 thus ?thesis + using id_outside_supp[OF 3(2)] less(2) permutes_not_in unfolding q_def by fastforce + qed + qed + + moreover from \s \ S\ have "(p ^^ i) s \ S" for i + unfolding permutes_in_image[OF permutes_funpow[OF less(2)]] . + hence "set (support p s) \ (S - set (support p s)) = S" + by auto + + moreover have "s \ set (support p s)" + using least_power_of_permutation[OF is_permutation] by force + with \s \ S\ have "card (S - set (support p s)) < card S" + using less(3) by (metis DiffE card_seteq linorder_not_le subsetI) + hence "cycle_decomp (S - set (support p s)) q" + using less(1)[OF _ semidecomposition[OF less(2-3)], of s] less(3) unfolding q_def by blast + + moreover show ?thesis + using comp[OF calculation(3) cycle_of_permutation[OF is_permutation], of s] + unfolding calculation(1-2) by blast + qed +qed end diff -r 12dce58bcd3f -r 1b5178abaf97 src/HOL/Algebra/Embedded_Algebras.thy --- a/src/HOL/Algebra/Embedded_Algebras.thy Thu Oct 04 11:18:39 2018 +0200 +++ b/src/HOL/Algebra/Embedded_Algebras.thy Thu Oct 04 15:25:47 2018 +0100 @@ -412,9 +412,6 @@ lemma Span_eq_generate: assumes "set Us \ carrier R" shows "Span K Us = generate (add_monoid R) (K <#> (set Us))" proof (rule add.generateI) - show "K <#> set Us \ carrier R" - using subring_props(1) assms unfolding set_mult_def by blast -next show "subgroup (Span K Us) (add_monoid R)" using Span_is_add_subgroup[OF assms] . next @@ -474,8 +471,7 @@ corollary mono_Span_sublist: assumes "set Us \ set Vs" "set Vs \ carrier R" shows "Span K Us \ Span K Vs" - using add.mono_generate[OF mono_set_mult[OF _ assms(1), of K K R] - set_mult_closed[OF subring_props(1) assms(2)]] + using add.mono_generate[OF mono_set_mult[OF _ assms(1), of K K R]] Span_eq_generate[OF assms(2)] Span_eq_generate[of Us] assms by auto corollary mono_Span_append: diff -r 12dce58bcd3f -r 1b5178abaf97 src/HOL/Algebra/Generated_Groups.thy --- a/src/HOL/Algebra/Generated_Groups.thy Thu Oct 04 11:18:39 2018 +0200 +++ b/src/HOL/Algebra/Generated_Groups.thy Thu Oct 04 15:25:47 2018 +0100 @@ -4,12 +4,12 @@ theory Generated_Groups imports Group Coset + begin -section\Generated Groups\ +section \Generated Groups\ -inductive_set - generate :: "('a, 'b) monoid_scheme \ 'a set \ 'a set" +inductive_set generate :: "('a, 'b) monoid_scheme \ 'a set \ 'a set" for G and H where one: "\\<^bsub>G\<^esub> \ generate G H" | incl: "h \ H \ h \ generate G H" @@ -17,624 +17,423 @@ | eng: "h1 \ generate G H \ h2 \ generate G H \ h1 \\<^bsub>G\<^esub> h2 \ generate G H" -subsection\Basic Properties of Generated Groups - First Part\ - -lemma (in group) generate_in_carrier: - assumes "H \ carrier G" - shows "h \ generate G H \ h \ carrier G" - apply (induction rule: generate.induct) using assms by blast+ - -lemma (in group) generate_m_inv_closed: - assumes "H \ carrier G" - shows "h \ generate G H \ (inv h) \ generate G H" -proof (induction rule: generate.induct) - case one thus ?case by (simp add: generate.one) -next - case (incl h) thus ?case using generate.inv[OF incl(1), of G] by simp -next - case (inv h) thus ?case using assms generate.incl by fastforce -next - case (eng h1 h2) - hence "inv (h1 \ h2) = (inv h2) \ (inv h1)" - by (meson assms generate_in_carrier group.inv_mult_group is_group) - thus ?case using generate.eng[OF eng(4) eng(3)] by simp -qed - -lemma (in group) generate_is_subgroup: - assumes "H \ carrier G" - shows "subgroup (generate G H) G" -proof (intro subgroupI) - show "generate G H \ carrier G" using generate_in_carrier[OF assms] by blast - show "generate G H \ {}" using generate.one by auto - show "\h. h \ generate G H \ inv h \ generate G H" - using generate_m_inv_closed[OF assms] by blast - show "\h1 h2. \ h1 \ generate G H; h2 \ generate G H \ \ h1 \ h2 \ generate G H" - by (simp add: generate.eng) -qed - - -subsection\Characterisations of Generated Groups\ - -lemma (in group) generate_min_subgroup1: - assumes "H \ carrier G" - and "subgroup E G" "H \ E" - shows "generate G H \ E" -proof - fix h show "h \ generate G H \ h \ E" - proof (induct rule: generate.induct) - case one thus ?case using subgroup.one_closed[OF assms(2)] by simp - case incl thus ?case using assms(3) by blast - case inv thus ?case using subgroup.m_inv_closed[OF assms(2)] assms(3) by blast - next - case eng thus ?case using subgroup.m_closed[OF assms(2)] by simp - qed -qed - -lemma (in group) generateI: - assumes "H \ carrier G" - and "subgroup E G" "H \ E" - and "\K. \ subgroup K G; H \ K \ \ E \ K" - shows "E = generate G H" -proof - show "E \ generate G H" - using assms generate_is_subgroup generate.incl by (metis subset_iff) - show "generate G H \ E" - using generate_min_subgroup1[OF assms(1-3)] by simp -qed +subsection \Basic Properties\ -lemma (in group) generateE: - assumes "H \ carrier G" and "E = generate G H" - shows "subgroup E G" and "H \ E" and "\K. \ subgroup K G; H \ K \ \ E \ K" -proof - - show "subgroup E G" using assms generate_is_subgroup by simp - show "H \ E" using assms(2) by (simp add: generate.incl subsetI) - show "\K. subgroup K G \ H \ K \ E \ K" - using assms generate_min_subgroup1 by auto -qed - -lemma (in group) generate_min_subgroup2: - assumes "H \ carrier G" - shows "generate G H = \{K. subgroup K G \ H \ K}" +lemma (in group) generate_consistent: + assumes "K \ H" "subgroup H G" shows "generate (G \ carrier := H \) K = generate G K" proof - have "subgroup (generate G H) G \ H \ generate G H" - by (simp add: assms generateE(2) generate_is_subgroup) - thus "\{K. subgroup K G \ H \ K} \ generate G H" by blast + show "generate (G \ carrier := H \) K \ generate G K" + proof + fix h assume "h \ generate (G \ carrier := H \) K" thus "h \ generate G K" + proof (induction, simp add: one, simp_all add: incl[of _ K G] eng) + case inv thus ?case + using m_inv_consistent assms generate.inv[of _ K G] by auto + qed + qed next - have "\K. subgroup K G \ H \ K \ generate G H \ K" - by (simp add: assms generate_min_subgroup1) - thus "generate G H \ \{K. subgroup K G \ H \ K}" by blast -qed - - -subsection\Representation of Elements from a Generated Group\ - -text\We define a sort of syntax tree to allow induction arguments with elements of a generated group\ - -datatype 'a repr = - One | Inv "'a" | Leaf "'a" | Mult "'a repr" "'a repr" - -fun norm :: "('a, 'b) monoid_scheme \ 'a repr \ 'a" - where - "norm G (One) = \\<^bsub>G\<^esub>" - | "norm G (Inv h) = (inv\<^bsub>G\<^esub> h)" - | "norm G (Leaf h) = h" - | "norm G (Mult h1 h2) = (norm G h1) \\<^bsub>G\<^esub> (norm G h2)" - -fun elts :: "'a repr \ 'a set" - where - "elts (One) = {}" - | "elts (Inv h) = { h }" - | "elts (Leaf h) = { h }" - | "elts (Mult h1 h2) = (elts h1) \ (elts h2)" - -lemma (in group) generate_repr_iff: - assumes "H \ carrier G" - shows "(h \ generate G H) \ (\r. (elts r) \ H \ norm G r = h)" -proof - show "h \ generate G H \ \r. (elts r) \ H \ norm G r = h" - proof (induction rule: generate.induct) - case one thus ?case - using elts.simps(1) norm.simps(1)[of G] by fastforce - next - case (incl h) - hence "elts (Leaf h) \ H \ norm G (Leaf h) = h" by simp - thus ?case by blast - next - case (inv h) - hence "elts (Inv h) \ H \ norm G (Inv h) = inv h" by auto - thus ?case by blast - next - case (eng h1 h2) - then obtain r1 r2 where r1: "elts r1 \ H" "norm G r1 = h1" - and r2: "elts r2 \ H" "norm G r2 = h2" by blast - hence "elts (Mult r1 r2) \ H \ norm G (Mult r1 r2) = h1 \ h2" by simp - thus ?case by blast - qed - - show "\r. elts r \ H \ norm G r = h \ h \ generate G H" - proof - - assume "\r. elts r \ H \ norm G r = h" - then obtain r where "elts r \ H" "norm G r = h" by blast - thus "h \ generate G H" - proof (induction arbitrary: h rule: repr.induct) - case One thus ?case using generate.one by auto - next - case Inv thus ?case using generate.simps by force - next - case Leaf thus ?case using generate.simps by force - next - case Mult thus ?case using generate.eng by fastforce + show "generate G K \ generate (G \ carrier := H \) K" + proof + note gen_simps = one incl eng + fix h assume "h \ generate G K" thus "h \ generate (G \ carrier := H \) K" + using gen_simps[where ?G = "G \ carrier := H \"] + proof (induction, auto) + fix h assume "h \ K" thus "inv h \ generate (G \ carrier := H \) K" + using m_inv_consistent assms generate.inv[of h K "G \ carrier := H \"] by auto qed qed qed -corollary (in group) generate_repr_set: - assumes "H \ carrier G" - shows "generate G H = {norm G r | r. (elts r) \ H}" (is "?A = ?B") +lemma (in group) generate_in_carrier: + assumes "H \ carrier G" and "h \ generate G H" shows "h \ carrier G" + using assms(2,1) by (induct h rule: generate.induct) (auto) + +lemma (in group) generate_incl: + assumes "H \ carrier G" shows "generate G H \ carrier G" + using generate_in_carrier[OF assms(1)] by auto + +lemma (in group) generate_m_inv_closed: + assumes "H \ carrier G" and "h \ generate G H" shows "(inv h) \ generate G H" + using assms(2,1) +proof (induction rule: generate.induct, auto simp add: one inv incl) + fix h1 h2 + assume h1: "h1 \ generate G H" "inv h1 \ generate G H" + and h2: "h2 \ generate G H" "inv h2 \ generate G H" + hence "inv (h1 \ h2) = (inv h2) \ (inv h1)" + by (meson assms generate_in_carrier group.inv_mult_group is_group) + thus "inv (h1 \ h2) \ generate G H" + using generate.eng[OF h2(2) h1(2)] by simp +qed + +lemma (in group) generate_is_subgroup: + assumes "H \ carrier G" shows "subgroup (generate G H) G" + using subgroup.intro[OF generate_incl eng one generate_m_inv_closed] assms by auto + +lemma (in group) mono_generate: + assumes "K \ H" shows "generate G K \ generate G H" proof - show "?A \ ?B" - proof - fix h assume "h \ generate G H" thus "h \ {norm G r |r. elts r \ H}" - using generate_repr_iff[OF assms] by auto - qed -next - show "?B \ ?A" - proof - fix h assume "h \ {norm G r |r. elts r \ H}" thus "h \ generate G H" - using generate_repr_iff[OF assms] by auto + fix h assume "h \ generate G K" thus "h \ generate G H" + using assms by (induction) (auto simp add: one incl inv eng) +qed + +lemma (in group) generate_subgroup_incl: + assumes "K \ H" "subgroup H G" shows "generate G K \ H" + using group.generate_incl[OF subgroup_imp_group[OF assms(2)], of K] assms(1) + by (simp add: generate_consistent[OF assms]) + +lemma (in group) generate_minimal: + assumes "H \ carrier G" shows "generate G H = \ { H'. subgroup H' G \ H \ H' }" + using generate_subgroup_incl generate_is_subgroup[OF assms] incl[of _ H] by blast + +lemma (in group) generateI: + assumes "subgroup E G" "H \ E" and "\K. \ subgroup K G; H \ K \ \ E \ K" + shows "E = generate G H" +proof - + have subset: "H \ carrier G" + using subgroup.subset assms by auto + show ?thesis + using assms unfolding generate_minimal[OF subset] by blast +qed + +lemma (in group) normal_generateI: + assumes "H \ carrier G" and "\h g. \ h \ H; g \ carrier G \ \ g \ h \ (inv g) \ H" + shows "generate G H \ G" +proof (rule normal_invI[OF generate_is_subgroup[OF assms(1)]]) + fix g h assume g: "g \ carrier G" show "h \ generate G H \ g \ h \ (inv g) \ generate G H" + proof (induct h rule: generate.induct) + case one thus ?case + using g generate.one by auto + next + case incl show ?case + using generate.incl[OF assms(2)[OF incl g]] . + next + case (inv h) + hence h: "h \ carrier G" + using assms(1) by auto + hence "inv (g \ h \ (inv g)) = g \ (inv h) \ (inv g)" + using g by (simp add: inv_mult_group m_assoc) + thus ?case + using generate_m_inv_closed[OF assms(1) generate.incl[OF assms(2)[OF inv g]]] by simp + next + case (eng h1 h2) + note in_carrier = eng(1,3)[THEN generate_in_carrier[OF assms(1)]] + have "g \ (h1 \ h2) \ inv g = (g \ h1 \ inv g) \ (g \ h2 \ inv g)" + using in_carrier g by (simp add: inv_solve_left m_assoc) + thus ?case + using generate.eng[OF eng(2,4)] by simp qed qed -corollary (in group) mono_generate: - assumes "I \ J" and "J \ carrier G" - shows "generate G I \ generate G J" - using assms generate_repr_iff by fastforce +lemma (in group) subgroup_int_pow_closed: + assumes "subgroup H G" "h \ H" shows "h [^] (k :: int) \ H" + using group.int_pow_closed[OF subgroup_imp_group[OF assms(1)]] assms(2) + unfolding int_pow_consistent[OF assms] by simp -lemma (in group) subgroup_gen_equality: - assumes "subgroup H G" "K \ H" - shows "generate G K = generate (G \ carrier := H \) K" -proof - - have "generate G K \ H" - by (meson assms generate_min_subgroup1 order.trans subgroup.subset) - have mult_eq: "\k1 k2. \ k1 \ generate G K; k2 \ generate G K \ \ - k1 \\<^bsub>G \ carrier := H \\<^esub> k2 = k1 \ k2" - using \generate G K \ H\ subgroup_mult_equality by simp - - { fix r assume A: "elts r \ K" - hence "norm G r = norm (G \ carrier := H \) r" - proof (induction r rule: repr.induct) - case One thus ?case by simp - next - case (Inv k) hence "k \ K" using A by simp - thus ?case using m_inv_consistent[OF assms(1)] assms(2) by auto +lemma (in group) generate_pow: + assumes "a \ carrier G" shows "generate G { a } = { a [^] (k :: int) | k. k \ UNIV }" +proof + show "{ a [^] (k :: int) | k. k \ UNIV } \ generate G { a }" + using subgroup_int_pow_closed[OF generate_is_subgroup[of "{ a }"] incl[of a]] assms by auto +next + 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]) + case (inv h) + hence "inv h = a [^] ((- 1) :: int)" + using assms unfolding int_pow_def2 by simp + thus ?case + by blast next - case (Leaf k) hence "k \ K" using A by simp - thus ?case using m_inv_consistent[OF assms(1)] assms(2) by auto - next - case (Mult k1 k2) thus ?case using mult_eq by auto - qed } note aux_lemma = this + case eng thus ?case + using assms by (metis int_pow_mult) + qed + thus "h \ { a [^] (k :: int) | k. k \ UNIV }" + by blast + qed +qed + +corollary (in group) generate_one: "generate G { \ } = { \ }" + using generate_pow[of "\", OF one_closed] by simp + +corollary (in group) generate_empty: "generate G {} = { \ }" + using mono_generate[of "{}" "{ \ }"] generate.one unfolding generate_one by auto - show ?thesis +lemma (in group_hom) + "subgroup K G \ subgroup (h ` K) H" + using subgroup_img_is_subgroup by auto + +lemma (in group_hom) generate_img: + assumes "K \ carrier G" shows "generate H (h ` K) = h ` (generate G K)" +proof + have "h ` K \ h ` (generate G K)" + using incl[of _ K G] by auto + thus "generate H (h ` K) \ h ` (generate G K)" + using generate_subgroup_incl subgroup_img_is_subgroup[OF G.generate_is_subgroup[OF assms]] by auto +next + show "h ` (generate G K) \ generate H (h ` K)" proof - show "generate G K \ generate (G\carrier := H\) K" - proof - fix h assume "h \ generate G K" then obtain r where r: "elts r \ K" "h = norm G r" - using generate_repr_iff assms by (metis order.trans subgroup.subset) - hence "h = norm (G \ carrier := H \) r" using aux_lemma by simp - thus "h \ generate (G\carrier := H\) K" - using r assms group.generate_repr_iff [of "G \ carrier := H \" K] - subgroup.subgroup_is_group[OF assms(1) is_group] by auto - qed - show "generate (G\carrier := H\) K \ generate G K" - proof - fix h assume "h \ generate (G\carrier := H\) K" - then obtain r where r: "elts r \ K" "h = norm (G\carrier := H\) r" - using group.generate_repr_iff [of "G \ carrier := H \" K] - subgroup.subgroup_is_group[OF assms(1) is_group] assms by auto - hence "h = norm G r" using aux_lemma by simp - thus "h \ generate G K" - by (meson assms generate_repr_iff order.trans r(1) subgroup.subset) + fix a assume "a \ h ` (generate G K)" + then obtain k where "k \ generate G K" "a = h k" + by blast + show "a \ generate H (h ` K)" + using \k \ generate G K\ unfolding \a = h k\ + proof (induct k, auto simp add: generate.one[of H] generate.incl[of _ "h ` K" H]) + case (inv k) show ?case + using assms generate.inv[of "h k" "h ` K" H] inv by auto + next + case eng show ?case + using generate.eng[OF eng(2,4)] eng(1,3)[THEN G.generate_in_carrier[OF assms]] by auto qed qed qed -corollary (in group) gen_equality_betw_subgroups: - assumes "subgroup I G" "subgroup J G" "K \ (I \ J)" - shows "generate (G \ carrier := I \) K = generate (G \ carrier := J \) K" - by (metis Int_subset_iff assms subgroup_gen_equality) -lemma (in group) normal_generateI: - assumes "H \ carrier G" - and "\h g. \ h \ H; g \ carrier G \ \ g \ h \ (inv g) \ H" - shows "generate G H \ G" -proof (rule normal_invI) - show "subgroup (generate G H) G" - by (simp add: assms(1) generate_is_subgroup) -next - have "\r g. \ elts r \ H; g \ carrier G \ \ (g \ (norm G r) \ (inv g)) \ (generate G H)" - proof - - fix r g assume "elts r \ H" "g \ carrier G" - thus "(g \ (norm G r) \ (inv g)) \ (generate G H)" - proof (induction r rule: repr.induct) - case One thus ?case - by (simp add: generate.one) - next - case (Inv h) - hence "g \ h \ (inv g) \ H" using assms(2) by simp - moreover have "norm G (Inv (g \ h \ (inv g))) = g \ (inv h) \ (inv g)" - using Inv.prems(1) Inv.prems(2) assms(1) inv_mult_group m_assoc by auto - ultimately have "\r. elts r \ H \ norm G r = g \ (inv h) \ (inv g)" - by (metis elts.simps(2) empty_subsetI insert_subset) - thus ?case by (simp add: assms(1) generate_repr_iff) - next - case (Leaf h) - thus ?case using assms(2)[of h g] generate.incl[of "g \ h \ inv g" H] by simp - next - case (Mult h1 h2) - hence "elts h1 \ H \ elts h2 \ H" using Mult(3) by simp - hence in_gen: "norm G h1 \ generate G H \ norm G h2 \ generate G H" - using assms(1) generate_repr_iff by auto - - have "g \ norm G (Mult h1 h2) \ inv g = g \ (norm G h1 \ norm G h2) \ inv g" by simp - also have " ... = g \ (norm G h1 \ (inv g \ g) \ norm G h2) \ inv g" - using Mult(4) in_gen assms(1) generate_in_carrier by auto - also have " ... = (g \ norm G h1 \ inv g) \ (g \ norm G h2 \ inv g)" - using Mult.prems(2) assms(1) generate_in_carrier in_gen inv_closed m_assoc m_closed by presburger - finally have "g \ norm G (Mult h1 h2) \ inv g = - (g \ norm G h1 \ inv g) \ (g \ norm G h2 \ inv g)" . - - thus ?case - using generate.eng[of "g \ norm G h1 \ inv g" G H "g \ norm G h2 \ inv g"] - by (simp add: Mult.IH Mult.prems(2) \elts h1 \ H \ elts h2 \ H\) - qed - qed - thus "\x h. x \ carrier G \ h \ generate G H \ x \ h \ inv x \ generate G H" - using assms(1) generate_repr_iff by auto -qed - - -subsection\Basic Properties of Generated Groups - Second Part\ +section \Derived Subgroup\ -lemma (in group) generate_pow: - assumes "a \ carrier G" - shows "generate G { a } = range (\k::int. a [^] k)" (is "?lhs = ?rhs") -proof - show "?lhs \ ?rhs" - proof - fix h show "h \ generate G { a } \ h \ range (\k::int. a [^] k)" - proof (induction rule: generate.induct) - case one thus ?case - by (metis (full_types) int_pow_0 rangeI) - next - case (incl h) hence "h = a" by auto thus ?case - by (metis \h = a\ assms group.int_pow_1 is_group rangeI) - next - case (inv h) hence "h = a" by auto thus ?case - by (metis (mono_tags) rangeI assms group.int_pow_1 int_pow_neg is_group) - next - case (eng h1 h2) thus ?case - using assms is_group by (auto simp: image_iff simp flip: int_pow_mult) - qed - qed +subsection \Definitions\ - show "?rhs \ ?lhs" - proof - { fix k :: "nat" have "a [^] k \ generate G { a }" - proof (induction k) - case 0 thus ?case by (simp add: generate.one) - next - case (Suc k) thus ?case by (simp add: generate.eng generate.incl) - qed } note aux_lemma = this - - fix h assume "h \ ?rhs" - then obtain k :: "nat" where "h = a [^] k \ h = inv (a [^] k)" - by (auto simp: int_pow_def2) - thus "h \ generate G { a }" using aux_lemma - using assms generate_m_inv_closed by auto - qed -qed - -(* { a [^] k | k. k \ (UNIV :: int set) } *) - -corollary (in group) generate_one: "generate G { \ } = { \ }" - using generate_pow[of "\", OF one_closed] by auto - -corollary (in group) generate_empty: "generate G {} = { \ }" - using mono_generate[of "{}" "{ \ }"] generate_one generate.one one_closed by blast - -corollary (in group) - assumes "H \ carrier G" "h \ H" - shows "h [^] (k :: int) \ generate G H" - using mono_generate[of "{ h }" H] generate_pow[of h] assms by auto - - -subsection\Derived Subgroup\ - -abbreviation derived_set :: "('a, 'b) monoid_scheme \ 'a set \ 'a set" where - "derived_set G H \ - \h1 \ H. (\h2 \ H. { h1 \\<^bsub>G\<^esub> h2 \\<^bsub>G\<^esub> (inv\<^bsub>G\<^esub> h1) \\<^bsub>G\<^esub> (inv\<^bsub>G\<^esub> h2) })" +abbreviation derived_set :: "('a, 'b) monoid_scheme \ 'a set \ 'a set" + where "derived_set G H \ + \h1 \ H. (\h2 \ H. { h1 \\<^bsub>G\<^esub> h2 \\<^bsub>G\<^esub> (inv\<^bsub>G\<^esub> h1) \\<^bsub>G\<^esub> (inv\<^bsub>G\<^esub> h2) })" definition derived :: "('a, 'b) monoid_scheme \ 'a set \ 'a set" where "derived G H = generate G (derived_set G H)" + +subsection \Basic Properties\ + lemma (in group) derived_set_incl: - assumes "subgroup H G" - shows "derived_set G H \ H" - by (auto simp add: m_inv_consistent[OF assms] subgroupE[OF assms]) + assumes "K \ H" "subgroup H G" shows "derived_set G K \ H" + using assms(1) subgroupE(3-4)[OF assms(2)] by (auto simp add: subset_iff) lemma (in group) derived_incl: - assumes "subgroup H G" - shows "derived G H \ H" - unfolding derived_def using derived_set_incl[OF assms] assms - by (meson generate_min_subgroup1 order.trans subgroup.subset) - -lemma (in group) subgroup_derived_equality: - assumes "subgroup H G" "K \ H" - shows "derived (G \ carrier := H \) K = derived G K" -proof - - have "derived_set G K \ H" - proof - fix x assume "x \ derived_set G K" - then obtain k1 k2 - where k12: "k1 \ K" "k2 \ K" - and "x = k1 \ k2 \ inv k1 \ inv k2" by blast - thus "x \ H" using k12 assms by (meson subgroup_def subsetCE) - qed - - moreover have "derived_set (G \ carrier := H \) K = derived_set G K" - proof - show "derived_set G K \ derived_set (G\carrier := H\) K" - proof - fix x assume "x \ derived_set G K" - then obtain k1 k2 where k12: "k1 \ K" "k2 \ K" - and "x = k1 \ k2 \ inv k1 \ inv k2" by blast - hence "x = k1 \\<^bsub>G\carrier := H\\<^esub> k2 \\<^bsub>G\carrier := H\\<^esub> inv\<^bsub>G\carrier := H\\<^esub> k1 \\<^bsub>G\carrier := H\\<^esub> inv\<^bsub>G\carrier := H\\<^esub> k2" - using subgroup_mult_equality[OF assms(1)] m_inv_consistent[OF assms(1)] assms(2) k12 - by (simp add: subset_iff) - thus "x \ derived_set (G\carrier := H\) K" using k12 by blast - qed - next - show "derived_set (G \ carrier := H \) K \ derived_set G K" - proof - fix x assume "x \ derived_set (G \ carrier := H \) K" - then obtain k1 k2 - where k12: "k1 \ K" "k2 \ K" - and "x = k1 \\<^bsub>G\carrier := H\\<^esub> k2 \\<^bsub>G\carrier := H\\<^esub> inv\<^bsub>G\carrier := H\\<^esub> k1 \\<^bsub>G\carrier := H\\<^esub> inv\<^bsub>G\carrier := H\\<^esub> k2" - by blast - hence "x = k1 \ k2 \ inv k1 \ inv k2" - using subgroup_mult_equality[OF assms(1)] m_inv_consistent[OF assms(1)] assms(2) k12 - by (simp add: subset_iff) - thus "x \ derived_set G K" using k12 assms by blast - qed - qed - - ultimately show ?thesis unfolding derived_def - using subgroup_gen_equality[OF assms(1), of "derived_set (G\carrier := H\) K"] by simp -qed - -lemma (in comm_group) derived_set: - assumes "H \ carrier G" - shows "derived G H = { \ }" -proof - - have "derived_set G H = {} \ derived_set G H = { \ }" - proof (cases) - assume "H = {}" thus ?thesis by simp - next - assume "H \ {}" then obtain h' where h': "h' \ H" by blast - have "derived_set G H = { \ }" - proof - - { fix h assume A: "h \ derived_set G H" - have "h = \" - proof - - obtain h1 h2 where h1: "h1 \ H" and h2: "h2 \ H" and h: "h = h1 \ h2 \ inv h1 \ inv h2" - using A by blast - then have "h1 \ carrier G" "h2 \ carrier G" - using assms by auto - then have "\ = h" - by (metis \h1 \ carrier G\ \h2 \ carrier G\ h inv_closed inv_mult m_assoc m_closed r_inv) - then show ?thesis - using \h1 \ carrier G\ \h2 \ carrier G\ by force - qed } note aux_lemma = this - show ?thesis - proof - show "derived_set G H \ { \ }" using aux_lemma by blast - next - show "{ \ } \ derived_set G H" - proof - - have "h' \ h' \ inv h' \ inv h' \ derived_set G H" using h' by blast - thus ?thesis using aux_lemma by auto - qed - qed - qed - thus ?thesis by simp - qed - thus ?thesis unfolding derived_def using generate_empty generate_one by presburger -qed + assumes "K \ H" "subgroup H G" shows "derived G K \ H" + using generate_subgroup_incl[OF derived_set_incl] assms unfolding derived_def by auto lemma (in group) derived_set_in_carrier: - assumes "H \ carrier G" - shows "derived_set G H \ carrier G" -proof - fix h assume "h \ derived_set G H" - then obtain h1 h2 where "h1 \ H" "h2 \ H" "h = h1 \ h2 \ inv h1 \ inv h2" - by blast - thus "h \ carrier G" using assms by blast + assumes "H \ carrier G" shows "derived_set G H \ carrier G" + using derived_set_incl[OF assms subgroup_self] . + +lemma (in group) derived_in_carrier: + assumes "H \ carrier G" shows "derived G H \ carrier G" + using derived_incl[OF assms subgroup_self] . + +lemma (in group) exp_of_derived_in_carrier: + assumes "H \ carrier G" shows "(derived G ^^ n) H \ carrier G" + using assms derived_in_carrier by (induct n) (auto) + +lemma (in group) derived_is_subgroup: + assumes "H \ carrier G" shows "subgroup (derived G H) G" + unfolding derived_def using generate_is_subgroup[OF derived_set_in_carrier[OF assms]] . + +lemma (in group) exp_of_derived_is_subgroup: + assumes "subgroup H G" shows "subgroup ((derived G ^^ n) H) G" + using assms derived_is_subgroup subgroup.subset by (induct n) (auto) + +lemma (in group) exp_of_derived_is_subgroup': + assumes "H \ carrier G" shows "subgroup ((derived G ^^ (Suc n)) H) G" + using assms derived_is_subgroup[OF subgroup.subset] derived_is_subgroup by (induct n) (auto) + +lemma (in group) mono_derived_set: + assumes "K \ H" shows "derived_set G K \ derived_set G H" + using assms by auto + +lemma (in group) mono_derived: + assumes "K \ H" shows "derived G K \ derived G H" + unfolding derived_def using mono_generate[OF mono_derived_set[OF assms]] . + +lemma (in group) mono_exp_of_derived: + assumes "K \ H" shows "(derived G ^^ n) K \ (derived G ^^ n) H" + using assms mono_derived by (induct n) (auto) + +lemma (in group) derived_set_consistent: + assumes "K \ H" "subgroup H G" shows "derived_set (G \ carrier := H \) K = derived_set G K" + using m_inv_consistent[OF assms(2)] assms(1) by (auto simp add: subset_iff) + +lemma (in group) derived_consistent: + assumes "K \ H" "subgroup H G" shows "derived (G \ carrier := H \) K = derived G K" + using generate_consistent[OF derived_set_incl] derived_set_consistent assms by (simp add: derived_def) + +lemma (in comm_group) derived_eq_singleton: + assumes "H \ carrier G" shows "derived G H = { \ }" +proof (cases "derived_set G H = {}") + case True show ?thesis + using generate_empty unfolding derived_def True by simp +next + case False + have aux_lemma: "h \ derived_set G H \ h = \" for h + using assms by (auto simp add: subset_iff) + (metis (no_types, lifting) m_comm m_closed inv_closed inv_solve_right l_inv l_inv_ex) + have "derived_set G H = { \ }" + proof + show "derived_set G H \ { \ }" + using aux_lemma by auto + next + obtain h where h: "h \ derived_set G H" + using False by blast + thus "{ \ } \ derived_set G H" + using aux_lemma[OF h] by auto + qed + thus ?thesis + using generate_one unfolding derived_def by auto qed lemma (in group) derived_is_normal: - assumes "H \ G" - shows "derived G H \ G" unfolding derived_def -proof (rule normal_generateI) - show "(\h1 \ H. \h2 \ H. { h1 \ h2 \ inv h1 \ inv h2 }) \ carrier G" - using subgroup.subset assms normal_imp_subgroup by blast -next - show "\h g. \ h \ derived_set G H; g \ carrier G \ \ g \ h \ inv g \ derived_set G H" - proof - + assumes "H \ G" shows "derived G H \ G" +proof - + interpret H: normal H G + using assms . + + show ?thesis + unfolding derived_def + proof (rule normal_generateI[OF derived_set_in_carrier[OF H.subset]]) fix h g assume "h \ derived_set G H" and g: "g \ carrier G" - then obtain h1 h2 where h1: "h1 \ H" "h1 \ carrier G" - and h2: "h2 \ H" "h2 \ carrier G" - and h: "h = h1 \ h2 \ inv h1 \ inv h2" - using subgroup.subset assms normal_imp_subgroup by blast - hence "g \ h \ inv g = + then obtain h1 h2 where h: "h1 \ H" "h2 \ H" "h = h1 \ h2 \ inv h1 \ inv h2" + by auto + hence in_carrier: "h1 \ carrier G" "h2 \ carrier G" "g \ carrier G" + using H.subset g by auto + have "g \ h \ inv g = g \ h1 \ (inv g \ g) \ h2 \ (inv g \ g) \ inv h1 \ (inv g \ g) \ inv h2 \ inv g" - by (simp add: g m_assoc) - also - have " ... = + unfolding h(3) by (simp add: in_carrier m_assoc) + also have " ... = (g \ h1 \ inv g) \ (g \ h2 \ inv g) \ (g \ inv h1 \ inv g) \ (g \ inv h2 \ inv g)" - using g h1 h2 m_assoc inv_closed m_closed by presburger - finally - have "g \ h \ inv g = - (g \ h1 \ inv g) \ (g \ h2 \ inv g) \ inv (g \ h1 \ inv g) \ inv (g \ h2 \ inv g)" - by (simp add: g h1 h2 inv_mult_group m_assoc) - moreover have "g \ h1 \ inv g \ H" by (simp add: assms normal.inv_op_closed2 g h1) - moreover have "g \ h2 \ inv g \ H" by (simp add: assms normal.inv_op_closed2 g h2) - ultimately show "g \ h \ inv g \ derived_set G H" by blast + using in_carrier m_assoc inv_closed m_closed by presburger + finally have "g \ h \ inv g = + (g \ h1 \ inv g) \ (g \ h2 \ inv g) \ inv (g \ h1 \ inv g) \ inv (g \ h2 \ inv g)" + by (simp add: in_carrier inv_mult_group m_assoc) + thus "g \ h \ inv g \ derived_set G H" + using h(1-2)[THEN H.inv_op_closed2[OF g]] by auto qed qed +lemma (in group) normal_self: "carrier G \ G" + by (rule normal_invI[OF subgroup_self], simp) + corollary (in group) derived_self_is_normal: "derived G (carrier G) \ G" - by (simp add: group.derived_is_normal group.normal_invI is_group subgroup_self) + using derived_is_normal[OF normal_self] . corollary (in group) derived_subgroup_is_normal: - assumes "subgroup H G" - shows "derived G H \ G \ carrier := H \" -proof - - have "H \ G \ carrier := H \" - by (metis assms group.coset_join3 group.normalI group.subgroup_self - partial_object.cases_scheme partial_object.select_convs(1) partial_object.update_convs(1) - subgroup.rcos_const subgroup_imp_group) - hence "derived (G \ carrier := H \) H \ G \ carrier := H \" - using group.derived_is_normal[of "G \ carrier := H \" H] normal_def by blast - thus ?thesis using subgroup_derived_equality[OF assms] by simp -qed + assumes "subgroup H G" shows "derived G H \ G \ carrier := H \" + using group.derived_self_is_normal[OF subgroup_imp_group[OF assms]] + derived_consistent[OF _ assms] + by simp corollary (in group) derived_quot_is_group: "group (G Mod (derived G (carrier G)))" - using derived_self_is_normal normal.factorgroup_is_group by auto + using normal.factorgroup_is_group[OF derived_self_is_normal] by auto + +lemma (in group) derived_quot_is_comm_group: "comm_group (G Mod (derived G (carrier G)))" +proof (rule group.group_comm_groupI[OF derived_quot_is_group], simp add: FactGroup_def) + interpret DG: normal "derived G (carrier G)" G + using derived_self_is_normal . -lemma (in group) derived_quot_is_comm: - assumes "H \ carrier (G Mod (derived G (carrier G)))" - and "K \ carrier (G Mod (derived G (carrier G)))" - shows "H <#> K = K <#> H" -proof - - { fix H K assume A1: "H \ carrier (G Mod (derived G (carrier G)))" - and A2: "K \ carrier (G Mod (derived G (carrier G)))" - have "H <#> K \ K <#> H" - proof - - obtain h k where hk: "h \ carrier G" "k \ carrier G" - and H: "H = (derived G (carrier G)) #> h" - and K: "K = (derived G (carrier G)) #> k" - using A1 A2 unfolding FactGroup_def RCOSETS_def by auto - have "H <#> K = (h \ k) <# (derived G (carrier G))" - using hk H K derived_self_is_normal m_closed normal.coset_eq normal.rcos_sum - by (metis (no_types, lifting)) - moreover have "K <#> H = (k \ h) <# (derived G (carrier G))" - using hk H K derived_self_is_normal m_closed normal.coset_eq normal.rcos_sum - by (metis (no_types, lifting)) - moreover have "(h \ k) <# (derived G (carrier G)) \ (k \ h) <# (derived G (carrier G))" + fix H K assume "H \ rcosets derived G (carrier G)" and "K \ rcosets derived G (carrier G)" + then obtain g1 g2 + where g1: "g1 \ carrier G" "H = derived G (carrier G) #> g1" + and g2: "g2 \ carrier G" "K = derived G (carrier G) #> g2" + unfolding RCOSETS_def by auto + hence "H <#> K = derived G (carrier G) #> (g1 \ g2)" + by (simp add: DG.rcos_sum) + also have " ... = derived G (carrier G) #> (g2 \ g1)" + proof - + { fix g1 g2 assume g1: "g1 \ carrier G" and g2: "g2 \ carrier G" + have "derived G (carrier G) #> (g1 \ g2) \ derived G (carrier G) #> (g2 \ g1)" proof - fix x assume "x \ h \ k <# derived G (carrier G)" - then obtain d where d: "d \ derived G (carrier G)" "d \ carrier G" "x = h \ k \ d" - using generate_is_subgroup[of "derived_set G (carrier G)"] - subgroup.subset[of "derived G (carrier G)" G] - derived_set_in_carrier[of "carrier G"] - unfolding l_coset_def derived_def by auto - hence "x = (k \ (h \ inv h) \ inv k) \ h \ k \ d" - using hk by simp - also have " ... = (k \ h) \ (inv h \ inv k) \ h \ k \ d" - using d(2) hk m_assoc by (metis subgroup_def subgroup_self) - also have " ... = (k \ h) \ ((inv h \ inv k \ h \ k) \ d)" - using d(2) hk m_assoc by simp - finally have "x = (k \ h) \ ((inv h \ inv k \ h \ k) \ d)" . - - moreover have "inv h \ inv k \ inv (inv h) \ inv (inv k) \ derived_set G (carrier G)" - using inv_closed[OF hk(1)] inv_closed[OF hk(2)] by blast - hence "inv h \ inv k \ h \ k \ derived_set G (carrier G)" - by (simp add: hk(1) hk(2)) - hence "(inv h \ inv k \ h \ k) \ d \ derived G (carrier G)" - using d(1) unfolding derived_def by (simp add: generate.eng generate.incl) - - ultimately show "x \ (k \ h) <# (derived G (carrier G))" - unfolding l_coset_def by blast - qed - ultimately show ?thesis by simp - qed } - thus ?thesis using assms by auto + fix h assume "h \ derived G (carrier G) #> (g1 \ g2)" + then obtain g' where h: "g' \ carrier G" "g' \ derived G (carrier G)" "h = g' \ (g1 \ g2)" + using DG.subset unfolding r_coset_def by auto + hence "h = g' \ (g1 \ g2) \ (inv g1 \ inv g2 \ g2 \ g1)" + using g1 g2 by (simp add: m_assoc) + hence "h = (g' \ (g1 \ g2 \ inv g1 \ inv g2)) \ (g2 \ g1)" + using h(1) g1 g2 inv_closed m_assoc m_closed by presburger + moreover have "g1 \ g2 \ inv g1 \ inv g2 \ derived G (carrier G)" + using incl[of _ "derived_set G (carrier G)"] g1 g2 unfolding derived_def by blast + hence "g' \ (g1 \ g2 \ inv g1 \ inv g2) \ derived G (carrier G)" + using DG.m_closed[OF h(2)] by simp + ultimately show "h \ derived G (carrier G) #> (g2 \ g1)" + unfolding r_coset_def by blast + qed } + thus ?thesis + using g1(1) g2(1) by auto + qed + also have " ... = K <#> H" + by (simp add: g1 g2 DG.rcos_sum) + finally show "H <#> K = K <#> H" . qed -theorem (in group) derived_quot_is_comm_group: - "comm_group (G Mod (derived G (carrier G)))" - apply (intro group.group_comm_groupI[OF derived_quot_is_group]) - using derived_quot_is_comm by simp - corollary (in group) derived_quot_of_subgroup_is_comm_group: - assumes "subgroup H G" - shows "comm_group ((G \ carrier := H \) Mod (derived G H))" -proof - - have "group (G \ carrier := H \)" - using assms subgroup_imp_group by auto - thus ?thesis - using group.derived_quot_is_comm_group subgroup_derived_equality[OF assms] by fastforce -qed + assumes "subgroup H G" shows "comm_group ((G \ carrier := H \) Mod (derived G H))" + using group.derived_quot_is_comm_group[OF subgroup_imp_group[OF assms]] + derived_consistent[OF _ assms] + by simp -lemma (in group) mono_derived: - assumes "J \ carrier G" "I \ J" - shows "(derived G ^^ n) I \ (derived G ^^ n) J" +proposition (in group) derived_minimal: + assumes "H \ G" and "comm_group (G Mod H)" shows "derived G (carrier G) \ H" proof - - { fix I J assume A: "J \ carrier G" "I \ J" have "derived G I \ derived G J" - proof - - have "derived_set G I \ derived_set G J" using A by blast - thus ?thesis unfolding derived_def using mono_generate A by (simp add: derived_set_in_carrier) - qed } note aux_lemma1 = this - - { fix n I assume A: "I \ carrier G" have "(derived G ^^ n) I \ carrier G" - proof (induction n) - case 0 thus ?case using A by simp - next - case (Suc n) - with aux_lemma1 have "(derived G ^^ Suc n) I \ derived G (carrier G)" - by auto - also have "... \ carrier G" - by (simp add: derived_incl subgroup_self) - finally show ?case . - qed } note aux_lemma2 = this + interpret H: normal H G + using assms(1) . show ?thesis - proof (induction n) - case 0 thus ?case using assms by simp - next - case (Suc n) thus ?case using aux_lemma1 aux_lemma2 assms(1) by auto + unfolding derived_def + proof (rule generate_subgroup_incl[OF _ H.subgroup_axioms]) + show "derived_set G (carrier G) \ H" + proof + fix h assume "h \ derived_set G (carrier G)" + then obtain g1 g2 where h: "g1 \ carrier G" "g2 \ carrier G" "h = g1 \ g2 \ inv g1 \ inv g2" + by auto + have "H #> (g1 \ g2) = (H #> g1) <#> (H #> g2)" + by (simp add: h(1-2) H.rcos_sum) + also have " ... = (H #> g2) <#> (H #> g1)" + using comm_groupE(4)[OF assms(2)] h(1-2) unfolding FactGroup_def RCOSETS_def by auto + also have " ... = H #> (g2 \ g1)" + by (simp add: h(1-2) H.rcos_sum) + finally have "H #> (g1 \ g2) = H #> (g2 \ g1)" . + then obtain h' where "h' \ H" "\ \ (g1 \ g2) = h' \ (g2 \ g1)" + using H.one_closed unfolding r_coset_def by blast + thus "h \ H" + using h m_assoc by auto + qed qed qed -lemma (in group) exp_of_derived_in_carrier: - assumes "H \ carrier G" - shows "(derived G ^^ n) H \ carrier G" using assms -proof (induction n) - case 0 thus ?case by simp -next - case (Suc n) - have "(derived G ^^ Suc n) H = (derived G) ((derived G ^^ n) H)" by simp - also have " ... \ (derived G) (carrier G)" - using mono_derived[of "carrier G" "(derived G ^^ n) H" 1] Suc by simp - also have " ... \ carrier G" unfolding derived_def - by (simp add: derived_set_incl generate_min_subgroup1 subgroup_self) - finally show ?case . +proposition (in group) derived_of_subgroup_minimal: + assumes "K \ G \ carrier := H \" "subgroup H G" and "comm_group ((G \ carrier := H \) Mod K)" + shows "derived G H \ K" + using group.derived_minimal[OF subgroup_imp_group[OF assms(2)] assms(1,3)] + derived_consistent[OF _ assms(2)] + by simp + +lemma (in group_hom) derived_img: + assumes "K \ carrier G" shows "derived H (h ` K) = h ` (derived G K)" +proof - + have "derived_set H (h ` K) = h ` (derived_set G K)" + proof + show "derived_set H (h ` K) \ h ` derived_set G K" + proof + fix a assume "a \ derived_set H (h ` K)" + then obtain k1 k2 + where "k1 \ K" "k2 \ K" "a = (h k1) \\<^bsub>H\<^esub> (h k2) \\<^bsub>H\<^esub> inv\<^bsub>H\<^esub> (h k1) \\<^bsub>H\<^esub> inv\<^bsub>H\<^esub> (h k2)" + by auto + hence "a = h (k1 \ k2 \ inv k1 \ inv k2)" + using assms by (simp add: subset_iff) + from this \k1 \ K\ and \k2 \ K\ show "a \ h ` derived_set G K" by auto + qed + next + show "h ` (derived_set G K) \ derived_set H (h ` K)" + proof + fix a assume "a \ h ` (derived_set G K)" + then obtain k1 k2 where "k1 \ K" "k2 \ K" "a = h (k1 \ k2 \ inv k1 \ inv k2)" + by auto + hence "a = (h k1) \\<^bsub>H\<^esub> (h k2) \\<^bsub>H\<^esub> inv\<^bsub>H\<^esub> (h k1) \\<^bsub>H\<^esub> inv\<^bsub>H\<^esub> (h k2)" + using assms by (simp add: subset_iff) + from this \k1 \ K\ and \k2 \ K\ show "a \ derived_set H (h ` K)" by auto + qed + qed + thus ?thesis + unfolding derived_def using generate_img[OF G.derived_set_in_carrier[OF assms]] by simp qed -lemma (in group) exp_of_derived_is_subgroup: - assumes "subgroup H G" - shows "subgroup ((derived G ^^ n) H) G" using assms -proof (induction n) - case 0 thus ?case by simp -next - case (Suc n) - have "(derived G ^^ n) H \ carrier G" - by (simp add: Suc.IH assms subgroup.subset) - hence "subgroup ((derived G) ((derived G ^^ n) H)) G" - unfolding derived_def using derived_set_in_carrier generate_is_subgroup by auto - thus ?case by simp -qed +lemma (in group_hom) exp_of_derived_img: + 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) -hide_const (open) norm - -end +end \ No newline at end of file diff -r 12dce58bcd3f -r 1b5178abaf97 src/HOL/Algebra/Group.thy --- a/src/HOL/Algebra/Group.thy Thu Oct 04 11:18:39 2018 +0200 +++ b/src/HOL/Algebra/Group.thy Thu Oct 04 15:25:47 2018 +0100 @@ -781,6 +781,19 @@ {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)}" + +(* NEW ========================================================================== *) +lemma hom_trans: + "\ f \ hom G H; g \ hom H I \ \ g \ f \ hom G I" + unfolding hom_def by (auto simp add: Pi_iff) +(* ============================================================================== *) + +(* NEW ============================================================================ *) +lemma (in group) hom_restrict: + assumes "h \ hom G H" and "\g. g \ carrier G \ h g = t g" shows "t \ hom G H" + using assms unfolding hom_def by (auto simp add: Pi_iff) +(* ============================================================================== *) + lemma (in group) hom_compose: "[|h \ hom G H; i \ hom H I|] ==> compose (carrier G) i h \ hom G I" by (fastforce simp add: hom_def compose_def) @@ -838,6 +851,12 @@ 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. *) lemma (in monoid) hom_imp_img_monoid: diff -r 12dce58bcd3f -r 1b5178abaf97 src/HOL/Algebra/Ideal.thy --- a/src/HOL/Algebra/Ideal.thy Thu Oct 04 11:18:39 2018 +0200 +++ b/src/HOL/Algebra/Ideal.thy Thu Oct 04 15:25:47 2018 +0100 @@ -39,7 +39,7 @@ proof - interpret ring R by fact show ?thesis - by (auto simp: ideal.intro ideal_axioms.intro additive_subgroupI a_subgroup is_ring I_l_closed I_r_closed) + by (auto simp: ideal.intro ideal_axioms.intro additive_subgroupI a_subgroup ring_axioms I_l_closed I_r_closed) qed @@ -68,6 +68,46 @@ (rule is_ideal, rule generate) qed +(* NEW ====== *) +lemma (in ideal) rcos_const_imp_mem: + assumes "i \ carrier R" and "I +> i = I" shows "i \ I" + using additive_subgroup.zero_closed[OF ideal.axioms(1)[OF ideal_axioms]] assms + by (force simp add: a_r_coset_def') +(* ========== *) + +(* NEW ====== *) +lemma (in ring) a_rcos_zero: + assumes "ideal I R" "i \ I" shows "I +> i = I" + using abelian_subgroupI3[OF ideal.axioms(1) is_abelian_group] + by (simp add: abelian_subgroup.a_rcos_const assms) +(* ========== *) + +(* NEW ====== *) +lemma (in ring) ideal_is_normal: + assumes "ideal I R" shows "I \ (add_monoid R)" + using abelian_subgroup.a_normal[OF abelian_subgroupI3[OF ideal.axioms(1)]] + abelian_group_axioms assms + by auto +(* ========== *) + +(* NEW ====== *) +lemma (in ideal) a_rcos_sum: + assumes "a \ carrier R" and "b \ carrier R" shows "(I +> a) <+> (I +> b) = I +> (a \ b)" + using normal.rcos_sum[OF ideal_is_normal[OF ideal_axioms]] assms + unfolding set_add_def a_r_coset_def by simp +(* ========== *) + +(* NEW ====== *) +lemma (in ring) set_add_comm: + assumes "I \ carrier R" "J \ carrier R" shows "I <+> J = J <+> I" +proof - + { fix I J assume "I \ carrier R" "J \ carrier R" hence "I <+> J \ J <+> I" + using a_comm unfolding set_add_def' by (auto, blast) } + thus ?thesis + using assms by auto +qed +(* ========== *) + subsubsection \Maximal Ideals\ @@ -128,9 +168,10 @@ proof - interpret additive_subgroup I R by fact interpret cring R by fact - show ?thesis - apply intro_locales - apply (simp add: I_l_closed I_r_closed ideal_axioms_def) + show ?thesis apply intro_locales + apply (intro ideal_axioms.intro) + apply (erule (1) I_l_closed) + apply (erule (1) I_r_closed) by (simp add: I_notcarr I_prime primeideal_axioms.intro) qed @@ -138,10 +179,10 @@ subsection \Special Ideals\ lemma (in ring) zeroideal: "ideal {\} R" - by (intro idealI subgroup.intro) (simp_all add: is_ring) + by (intro idealI subgroup.intro) (simp_all add: ring_axioms) lemma (in ring) oneideal: "ideal (carrier R) R" - by (rule idealI) (auto intro: is_ring add.subgroupI) + by (rule idealI) (auto intro: ring_axioms add.subgroupI) lemma (in "domain") zeroprimeideal: "primeideal {\} R" proof - @@ -186,7 +227,7 @@ by (force simp: a_subset) show ?thesis apply (intro idealI subgroup.intro) - apply (simp_all add: IJ is_ring I_l_closed assms ideal.I_l_closed ideal.I_r_closed flip: a_inv_def) + apply (simp_all add: IJ ring_axioms I_l_closed assms ideal.I_l_closed ideal.I_r_closed flip: a_inv_def) done qed @@ -239,7 +280,7 @@ show "additive_subgroup (I <+> J) R" by (intro ideal.axioms[OF idealI] ideal.axioms[OF idealJ] add_additive_subgroups) show "ring R" - by (rule is_ring) + by (rule ring_axioms) show "ideal_axioms (I <+> J) R" proof - { fix x i j @@ -347,17 +388,14 @@ text \Generation of Principal Ideals in Commutative Rings\ definition cgenideal :: "_ \ 'a \ 'a set" ("PIdl\ _" [80] 79) - where "cgenideal R a \ {x \\<^bsub>R\<^esub> a | x. x \ carrier R}" - -lemma cginideal_def': "cgenideal R a = (\x. x \\<^bsub>R\<^esub> a) ` carrier R" - by (auto simp add: cgenideal_def) + where "cgenideal R a = {x \\<^bsub>R\<^esub> a | x. x \ carrier R}" text \genhideal (?) really generates an ideal\ lemma (in cring) cgenideal_ideal: assumes acarr: "a \ carrier R" shows "ideal (PIdl a) R" unfolding cgenideal_def -proof (intro subgroup.intro idealI[OF is_ring], simp_all) +proof (intro subgroup.intro idealI[OF ring_axioms], simp_all) show "{x \ a |x. x \ carrier R} \ carrier R" by (blast intro: acarr) show "\x y. \\u. x = u \ a \ u \ carrier R; \x. y = x \ a \ x \ carrier R\ @@ -431,7 +469,7 @@ shows "Idl (I \ J) = I <+> J" proof show "Idl (I \ J) \ I <+> J" - proof (rule ring.genideal_minimal [OF is_ring]) + proof (rule ring.genideal_minimal [OF ring_axioms]) show "ideal (I <+> J) R" by (rule add_ideals[OF idealI idealJ]) have "\x. x \ I \ \xa\I. \xb\J. x = xa \ xb" diff -r 12dce58bcd3f -r 1b5178abaf97 src/HOL/Algebra/QuotRing.thy --- a/src/HOL/Algebra/QuotRing.thy Thu Oct 04 11:18:39 2018 +0200 +++ b/src/HOL/Algebra/QuotRing.thy Thu Oct 04 15:25:47 2018 +0100 @@ -214,6 +214,39 @@ shows "inj_on h (carrier R)" using group_hom.trivial_ker_imp_inj[OF a_group_hom] assms a_kernel_def[of R S h] by simp +(* NEW ========================================================================== *) +lemma (in ring_hom_ring) inj_iff_trivial_ker: + shows "inj_on h (carrier R) \ a_kernel R S h = { \ }" + using group_hom.inj_iff_trivial_ker[OF a_group_hom] a_kernel_def[of R S h] by simp + +(* NEW ========================================================================== *) +corollary ring_hom_in_hom: + assumes "h \ ring_hom R S" shows "h \ hom R S" and "h \ hom (add_monoid R) (add_monoid S)" + using assms unfolding ring_hom_def hom_def by auto + +(* NEW ========================================================================== *) +corollary set_add_hom: + assumes "h \ ring_hom R S" "I \ carrier R" and "J \ carrier R" + shows "h ` (I <+>\<^bsub>R\<^esub> J) = h ` I <+>\<^bsub>S\<^esub> h ` J" + using set_mult_hom[OF ring_hom_in_hom(2)[OF assms(1)]] assms(2-3) + unfolding a_kernel_def[of R S h] set_add_def by simp + +(* NEW ========================================================================== *) +corollary a_coset_hom: + assumes "h \ ring_hom R S" "I \ carrier R" "a \ carrier R" + shows "h ` (a <+\<^bsub>R\<^esub> I) = h a <+\<^bsub>S\<^esub> (h ` I)" and "h ` (I +>\<^bsub>R\<^esub> a) = (h ` I) +>\<^bsub>S\<^esub> h a" + using assms coset_hom[OF ring_hom_in_hom(2)[OF assms(1)], of I a] + unfolding a_l_coset_def l_coset_eq_set_mult + a_r_coset_def r_coset_eq_set_mult + by simp_all + +(* NEW ========================================================================== *) +corollary (in ring_hom_ring) set_add_ker_hom: + assumes "I \ carrier R" + shows "h ` (I <+> (a_kernel R S h)) = h ` I" and "h ` ((a_kernel R S h) <+> I) = h ` I" + using group_hom.set_mult_ker_hom[OF a_group_hom] assms + unfolding a_kernel_def[of R S h] set_add_def by simp+ + lemma (in ring_hom_ring) non_trivial_field_hom_imp_inj: assumes "field R" shows "h ` (carrier R) \ { \\<^bsub>S\<^esub> } \ inj_on h (carrier R)" @@ -226,6 +259,23 @@ thus "inj_on h (carrier R)" using trivial_ker_imp_inj by blast qed +lemma "field R \ cring R" + using fieldE(1) by blast + +lemma non_trivial_field_hom_is_inj: + assumes "h \ ring_hom R S" and "field R" and "field S" shows "inj_on h (carrier R)" +proof - + interpret ring_hom_cring R S h + using assms(1) ring_hom_cring.intro[OF assms(2-3)[THEN fieldE(1)]] + unfolding symmetric[OF ring_hom_cring_axioms_def] by simp + + have "h \\<^bsub>R\<^esub> = \\<^bsub>S\<^esub>" and "\\<^bsub>S\<^esub> \ \\<^bsub>S\<^esub>" + using domain.one_not_zero[OF field.axioms(1)[OF assms(3)]] by auto + hence "h ` (carrier R) \ { \\<^bsub>S\<^esub> }" + using ring.kernel_zero ring.trivial_hom_iff by fastforce + thus ?thesis + using ring.non_trivial_field_hom_imp_inj[OF assms(2)] by simp +qed lemma (in ring_hom_ring) img_is_add_subgroup: assumes "subgroup H (add_monoid R)" @@ -538,24 +588,21 @@ and "\r1 r2. \ r1 \ carrier R; r2 \ carrier R \ \ P (r1 \ r2)" using assms unfolding morphic_prop_def by auto -lemma ring_iso_restrict: - assumes "f \ ring_iso R S" - and "\r. r \ carrier R \ f r = g r" - and "ring R" - shows "g \ ring_iso R S" -proof (rule ring_iso_memI) - show "bij_betw g (carrier R) (carrier S)" - using assms(1-2) bij_betw_cong ring_iso_memE(5) by blast - show "g \\<^bsub>R\<^esub> = \\<^bsub>S\<^esub>" - using assms ring.ring_simprules(6) ring_iso_memE(4) by force -next - fix x y assume x: "x \ carrier R" and y: "y \ carrier R" - show "g x \ carrier S" - using assms(1-2) ring_iso_memE(1) x by fastforce - show "g (x \\<^bsub>R\<^esub> y) = g x \\<^bsub>S\<^esub> g y" - by (metis assms ring.ring_simprules(5) ring_iso_memE(2) x y) - show "g (x \\<^bsub>R\<^esub> y) = g x \\<^bsub>S\<^esub> g y" - by (metis assms ring.ring_simprules(1) ring_iso_memE(3) x y) +(* NEW ============================================================================ *) +lemma (in ring) ring_hom_restrict: + assumes "f \ ring_hom R S" and "\r. r \ carrier R \ f r = g r" shows "g \ ring_hom R S" + using assms(2) ring_hom_memE[OF assms(1)] by (auto intro: ring_hom_memI) + +(* PROOF ========================================================================== *) +lemma (in ring) ring_iso_restrict: + assumes "f \ ring_iso R S" and "\r. r \ carrier R \ f r = g r" shows "g \ ring_iso R S" +proof - + have hom: "g \ ring_hom R S" + using ring_hom_restrict assms unfolding ring_iso_def by auto + have "bij_betw g (carrier R) (carrier S)" + using bij_betw_cong[of "carrier R" f g] ring_iso_memE(5)[OF assms(1)] assms(2) by simp + thus ?thesis + using ring_hom_memE[OF hom] by (auto intro!: ring_iso_memI) qed lemma ring_iso_morphic_prop: @@ -580,7 +627,7 @@ lemma (in ring) ring_hom_imp_img_ring: assumes "h \ ring_hom R S" - shows "ring (S \ carrier := h ` (carrier R), one := h \, zero := h \ \)" (is "ring ?h_img") + shows "ring (S \ carrier := h ` (carrier R), zero := h \ \)" (is "ring ?h_img") proof - have "h \ hom (add_monoid R) (add_monoid S)" using assms unfolding hom_def ring_hom_def by auto @@ -594,8 +641,7 @@ hence "monoid (S \ carrier := h ` (carrier R), one := h \ \)" using hom_imp_img_monoid[of h S] by simp hence monoid: "monoid ?h_img" - unfolding monoid_def by (simp add: monoid.defs) - + using ring_hom_memE(4)[OF assms] unfolding monoid_def by (simp add: monoid.defs) show ?thesis proof (rule ringI, simp_all add: comm_group_abelian_groupI[OF comm_group] monoid) fix x y z assume "x \ h ` carrier R" "y \ h ` carrier R" "z \ h ` carrier R" @@ -623,9 +669,9 @@ lemma (in ring) ring_iso_imp_img_ring: assumes "h \ ring_iso R S" - shows "ring (S \ one := h \, zero := h \ \)" + shows "ring (S \ zero := h \ \)" proof - - have "ring (S \ carrier := h ` (carrier R), one := h \, zero := h \ \)" + have "ring (S \ carrier := h ` (carrier R), zero := h \ \)" using ring_hom_imp_img_ring[of h S] assms unfolding ring_iso_def by auto moreover have "h ` (carrier R) = carrier S" using assms unfolding ring_iso_def bij_betw_def by auto @@ -634,7 +680,7 @@ lemma (in cring) ring_iso_imp_img_cring: assumes "h \ ring_iso R S" - shows "cring (S \ one := h \, zero := h \ \)" (is "cring ?h_img") + shows "cring (S \ zero := h \ \)" (is "cring ?h_img") proof - note m_comm interpret h_img?: ring ?h_img @@ -659,16 +705,20 @@ lemma (in domain) ring_iso_imp_img_domain: assumes "h \ ring_iso R S" - shows "domain (S \ one := h \, zero := h \ \)" (is "domain ?h_img") + shows "domain (S \ zero := h \ \)" (is "domain ?h_img") proof - note aux = m_closed integral one_not_zero one_closed zero_closed interpret h_img?: cring ?h_img using ring_iso_imp_img_cring[OF assms] . show ?thesis proof (unfold_locales) - show "\\<^bsub>?h_img\<^esub> \ \\<^bsub>?h_img\<^esub>" + have "\\<^bsub>?h_img\<^esub> = \\<^bsub>?h_img\<^esub> \ h \ = h \" + using ring_iso_memE(4)[OF assms] by simp + moreover have "h \ \ h \" using ring_iso_memE(5)[OF assms] aux(3-4) unfolding bij_betw_def inj_on_def by force + ultimately show "\\<^bsub>?h_img\<^esub> \ \\<^bsub>?h_img\<^esub>" + by auto next fix a b assume A: "a \\<^bsub>?h_img\<^esub> b = \\<^bsub>?h_img\<^esub>" "a \ carrier ?h_img" "b \ carrier ?h_img" @@ -693,21 +743,21 @@ lemma (in field) ring_iso_imp_img_field: assumes "h \ ring_iso R S" - shows "field (S \ one := h \, zero := h \ \)" (is "field ?h_img") + shows "field (S \ zero := h \ \)" (is "field ?h_img") proof - interpret h_img?: domain ?h_img using ring_iso_imp_img_domain[OF assms] . show ?thesis proof (unfold_locales, auto simp add: Units_def) interpret field R using field_axioms . - fix a assume a: "a \ carrier S" "a \\<^bsub>S\<^esub> h \ = h \" + fix a assume a: "a \ carrier S" "a \\<^bsub>S\<^esub> h \ = \\<^bsub>S\<^esub>" then obtain r where r: "r \ carrier R" "a = h r" using assms image_iff[where ?f = h and ?A = "carrier R"] unfolding ring_iso_def bij_betw_def by auto have "a \\<^bsub>S\<^esub> h \ = h (r \ \)" unfolding r(2) using ring_iso_memE(2)[OF assms r(1)] by simp hence "h \ = h \" - using r(1) a(2) by simp + using ring_iso_memE(4)[OF assms] r(1) a(2) by simp thus False using ring_iso_memE(5)[OF assms] unfolding bij_betw_def inj_on_def by force @@ -723,19 +773,14 @@ have "h (inv r) \\<^bsub>S\<^esub> h r = h \" and "h r \\<^bsub>S\<^esub> h (inv r) = h \" using ring_iso_memE(2)[OF assms inv_r(1) r(1)] inv_r(3-4) ring_iso_memE(2)[OF assms r(1) inv_r(1)] by auto - thus "\s' \ carrier S. s' \\<^bsub>S\<^esub> s = h \ \ s \\<^bsub>S\<^esub> s' = h \" - using ring_iso_memE(1)[OF assms inv_r(1)] r(2) by auto + thus "\s' \ carrier S. s' \\<^bsub>S\<^esub> s = \\<^bsub>S\<^esub> \ s \\<^bsub>S\<^esub> s' = \\<^bsub>S\<^esub>" + using ring_iso_memE(1,4)[OF assms] inv_r(1) r(2) by auto qed qed lemma ring_iso_same_card: "R \ S \ card (carrier R) = card (carrier S)" -proof - - assume "R \ S" - then obtain h where "bij_betw h (carrier R) (carrier S)" - unfolding is_ring_iso_def ring_iso_def by auto - thus "card (carrier R) = card (carrier S)" - using bij_betw_same_card[of h "carrier R" "carrier S"] by simp -qed + using bij_betw_same_card unfolding is_ring_iso_def ring_iso_def by auto +(* ========================================================================== *) lemma ring_iso_set_refl: "id \ ring_iso R R" by (rule ring_iso_memI) (auto) @@ -934,8 +979,7 @@ have FactRing_is_ring: "ring (R Quot (a_kernel R S h))" by (simp add: ideal.quotient_is_ring kernel_is_ideal) have "ring ((S \ carrier := ?the_elem ` (carrier (R Quot (a_kernel R S h))) \) - \ one := ?the_elem \\<^bsub>(R Quot (a_kernel R S h))\<^esub>, - zero := ?the_elem \\<^bsub>(R Quot (a_kernel R S h))\<^esub> \)" + \ zero := ?the_elem \\<^bsub>(R Quot (a_kernel R S h))\<^esub> \)" using ring.ring_iso_imp_img_ring[OF FactRing_is_ring, of ?the_elem "S \ carrier := ?the_elem ` (carrier (R Quot (a_kernel R S h))) \"] FactRing_iso_set_aux the_elem_surj by auto diff -r 12dce58bcd3f -r 1b5178abaf97 src/HOL/Algebra/Solvable_Groups.thy --- a/src/HOL/Algebra/Solvable_Groups.thy Thu Oct 04 11:18:39 2018 +0200 +++ b/src/HOL/Algebra/Solvable_Groups.thy Thu Oct 04 15:25:47 2018 +0100 @@ -3,162 +3,82 @@ *) theory Solvable_Groups - imports Group Coset Generated_Groups + imports Generated_Groups + begin -inductive solvable_seq :: "('a, 'b) monoid_scheme \ 'a set \ bool" for G where -unity: "solvable_seq G { \\<^bsub>G\<^esub> }" | -extension: "\ solvable_seq G K; K \ H; subgroup H G; K \ (G \ carrier := H \); - comm_group ((G \ carrier := H \) Mod K) \ \ solvable_seq G H" +section \Solvable Groups\ + +subsection \Definitions\ -definition - solvable :: "('a, 'b) monoid_scheme \ bool" +inductive solvable_seq :: "('a, 'b) monoid_scheme \ 'a set \ bool" + for G where + unity: "solvable_seq G { \\<^bsub>G\<^esub> }" + | extension: "\ solvable_seq G K; K \ (G \ carrier := H \); subgroup H G; + comm_group ((G \ carrier := H \) Mod K) \ \ solvable_seq G H" + +definition solvable :: "('a, 'b) monoid_scheme \ bool" where "solvable G \ solvable_seq G (carrier G)" subsection \Solvable Groups and Derived Subgroups\ text \We show that a group G is solvable iff the subgroup (derived G ^^ n) (carrier G) - is trivial for a sufficiently large n\ + is trivial for a sufficiently large n. \ lemma (in group) solvable_imp_subgroup: - assumes "solvable_seq G H" - shows "subgroup H G" using assms -proof (induction) - case unity thus ?case - using generate_empty generate_is_subgroup by force -next - case extension thus ?case by simp -qed + assumes "solvable_seq G H" shows "subgroup H G" + using assms normal.axioms(1)[OF one_is_normal] by (induction) (auto) lemma (in group) augment_solvable_seq: - assumes "subgroup H G" - and "solvable_seq G (derived G H)" - shows "solvable_seq G H" -proof (cases) - assume "derived G H = H" thus ?thesis - unfolding solvable_def using assms by simp -next - assume "derived G H \ H" - thus ?thesis unfolding solvable_def - using solvable_seq.extension[OF assms(2), of H] assms(1) - derived_quot_of_subgroup_is_comm_group[of H, OF assms(1)] - derived_incl[OF assms(1)] derived_subgroup_is_normal[OF assms(1)] by simp -qed + assumes "subgroup H G" and "solvable_seq G (derived G H)" shows "solvable_seq G H" + using extension[OF _ derived_subgroup_is_normal _ derived_quot_of_subgroup_is_comm_group] assms by simp theorem (in group) trivial_derived_seq_imp_solvable: - assumes "subgroup H G" and "((derived G) ^^ n) H = { \ }" - shows "solvable_seq G H" using assms -proof (induction n arbitrary: H) - case 0 hence "H = { \ }" by simp thus ?case by (simp add: unity) -next - case (Suc n) - hence "(derived G ^^ n) (derived G H) = { \ }" - by (simp add: funpow_swap1) - moreover have "subgroup (derived G H) G" unfolding derived_def - using Suc.prems(1) derived_set_incl generate_is_subgroup order.trans subgroup.subset - by (metis (no_types, lifting)) - ultimately have "solvable_seq G (derived G H)" by (simp add: Suc.IH) - thus ?case by (simp add: Suc.prems(1) augment_solvable_seq) + assumes "subgroup H G" and "((derived G) ^^ n) H = { \ }" shows "solvable_seq G H" + using assms +proof (induct n arbitrary: H, simp add: unity[of G]) + case (Suc n) thus ?case + using augment_solvable_seq derived_is_subgroup[OF subgroup.subset] by (simp add: funpow_swap1) qed theorem (in group) solvable_imp_trivial_derived_seq: - assumes "solvable_seq G H" and "subgroup H G" - shows "\n. (derived G ^^ n) H = { \ }" -proof - - { fix K H assume A: "K \ H" "K \ (G \ carrier := H \)" "subgroup K G" "subgroup H G" - "comm_group ((G \ carrier := H \) Mod K)" - have "derived G H \ K" - proof - - have Hcarr: "\a. a \ H \ a \ carrier G" - by (meson A(4) subgroup.mem_carrier) - have "derived_set G H \ K" - proof - fix h assume "h \ derived_set G H" - then obtain h1 h2 where h12: "h1 \ H" "h2 \ H" "h = h1 \ h2 \ inv h1 \ inv h2" by blast - - hence K_h12: "(K #> (h1 \ h2)) \ carrier ((G \ carrier := H \) Mod K)" - unfolding FactGroup_def RCOSETS_def r_coset_def apply simp by (metis A(4) subgroup_def) - have K_h1: "K #>\<^bsub>G\carrier := H\\<^esub> h1 \ carrier ((G \ carrier := H \) Mod K)" - unfolding FactGroup_def RCOSETS_def r_coset_def apply simp using h12(1) by blast - have K_h2: "K #>\<^bsub>G\carrier := H\\<^esub> h2 \ carrier ((G \ carrier := H \) Mod K)" - unfolding FactGroup_def RCOSETS_def r_coset_def apply simp using h12(2) by blast - - hence "K #>\<^bsub>G\carrier := H\\<^esub> (h1 \ h2) = - (K #>\<^bsub>G\carrier := H\\<^esub> h1) <#>\<^bsub>G\carrier := H\\<^esub> (K #>\<^bsub>G\carrier := H\\<^esub> h2)" - using normal.rcos_sum[OF A(2),of h1 h2] h12(1-2) by simp - also have " ... = - (K #>\<^bsub>G\carrier := H\\<^esub> h2) <#>\<^bsub>G\carrier := H\\<^esub> (K #>\<^bsub>G\carrier := H\\<^esub> h1)" - using comm_groupE(4)[OF A(5) K_h1 K_h2] by simp - finally have "K #>\<^bsub>G\carrier := H\\<^esub> (h1 \ h2) = K #>\<^bsub>G\carrier := H\\<^esub> (h2 \ h1)" - using normal.rcos_sum[OF A(2),of h2 h1] h12(1-2) by simp - - moreover have h12H: "h1 \ h2 \ H" and "h2 \ h1 \ H" - using h12 subgroupE(4)[OF A(4)] by auto - ultimately have "K #> (h1 \ h2) = K #> (h2 \ h1)" by auto - - then obtain k where k: "k \ K" "\ \ (h1 \ h2) = k \ (h2 \ h1)" - using subgroup.one_closed[OF A(3)] unfolding r_coset_def by blast - hence "(h1 \ h2) \ (inv h1 \ inv h2) = k" - proof - - have "k \ carrier G" - by (meson A(3) k(1) subgroup.mem_carrier) - with Hcarr h12 show ?thesis - by (metis h12H inv_mult_group inv_solve_right k(2) r_cancel_one' subgroup_def subgroup_self) - qed - hence "h = k" using h12 - by (meson A(4) \h1 \ h2 \ H\ inv_closed m_assoc subgroup.mem_carrier) - thus "h \ K" using k(1) by simp - qed - thus ?thesis unfolding derived_def by (meson A(3) generateE(3) order.trans subgroupE(1)) - qed } note aux_lemma = this - - show "\n. (derived G ^^ n) H = { \ }" using assms - proof (induct H rule: solvable_seq.induct) - case unity have "(derived G ^^ 0) { \ } = { \ }" by simp thus ?case by blast - next - case (extension K H) - then obtain n where n: "(derived G ^^ n) K = { \ }" - using solvable_imp_subgroup extension by blast - have "derived G H \ K" using solvable_imp_subgroup extension aux_lemma by blast - hence "(derived G ^^ n) (derived G H) \ (derived G ^^ n) K" - using mono_derived solvable_imp_subgroup extension.hyps(4) - by (simp add: extension.hyps(1) subgroup.subset) - hence "(derived G ^^ (Suc n)) H \ { \ }" - by (metis funpow_simps_right(2) n o_apply) - moreover have "\ \ derived G ((derived G ^^ n) H)" - unfolding derived_def using generate.one by auto - hence "{ \ } \ (derived G ^^ (Suc n)) H" by simp - ultimately show ?case by blast - qed + assumes "solvable_seq G H" shows "\n. (derived G ^^ n) H = { \ }" + using assms +proof (induction) + case unity + have "(derived G ^^ 0) { \ } = { \ }" + by simp + thus ?case by blast +next + case (extension K H) + obtain n where "(derived G ^^ n) K = { \ }" + using solvable_imp_subgroup extension(1,5) by auto + hence "(derived G ^^ (Suc n)) H \ { \ }" + using mono_exp_of_derived[OF derived_of_subgroup_minimal[OF extension(2-4)], of n] by (simp add: funpow_swap1) + moreover have "{ \ } \ (derived G ^^ (Suc n)) H" + using subgroup.one_closed[OF exp_of_derived_is_subgroup[OF extension(3)], of "Suc n"] by auto + ultimately show ?case + by blast qed theorem (in group) solvable_iff_trivial_derived_seq: "solvable G \ (\n. (derived G ^^ n) (carrier G) = { \ })" - unfolding solvable_def - using solvable_imp_trivial_derived_seq subgroup_self - trivial_derived_seq_imp_solvable by blast + using solvable_imp_trivial_derived_seq subgroup_self trivial_derived_seq_imp_solvable + by (auto simp add: solvable_def) corollary (in group) solvable_subgroup: - assumes "subgroup H G" - shows "solvable G \ solvable_seq G H" + assumes "subgroup H G" and "solvable G" shows "solvable_seq G H" proof - - { fix I J assume A: "subgroup I G" "subgroup J G" "I \ J" "solvable_seq G J" - have "solvable_seq G I" - proof - - obtain n where "(derived G ^^ n) J = { \ }" - using solvable_imp_trivial_derived_seq[OF A(4) A(2)] by auto - hence "(derived G ^^ n) I \ { \ }" - using mono_derived[OF subgroup.subset[OF A(2)] A(3)] by auto - hence "(derived G ^^ n) I = { \ }" - using subgroup.one_closed[OF exp_of_derived_is_subgroup[OF A(1), of n]] by auto - thus ?thesis - using trivial_derived_seq_imp_solvable[OF A(1), of n] by auto - qed } note aux_lemma = this - assume "solvable G" - thus ?thesis - using aux_lemma[OF assms subgroup_self subgroup.subset[OF assms]] - unfolding solvable_def by simp + obtain n where n: "(derived G ^^ n) (carrier G) = { \ }" + using assms(2) solvable_imp_trivial_derived_seq by (auto simp add: solvable_def) + show ?thesis + proof (rule trivial_derived_seq_imp_solvable[OF assms(1), of n]) + show "(derived G ^^ n) H = { \ }" + using subgroup.one_closed[OF exp_of_derived_is_subgroup[OF assms(1)], of n] + mono_exp_of_derived[OF subgroup.subset[OF assms(1)], of n] n + by auto + qed qed @@ -168,275 +88,78 @@ injective homomorphism from a group H to a group G, if H isn't solvable the group G isn't neither. \ -lemma (in group_hom) generate_of_img: - assumes "K \ carrier G" - shows "generate H (h ` K) = h ` (generate G K)" -proof - have img_in_carrier: "h ` K \ carrier H" - by (meson assms group_hom.hom_closed group_hom_axioms image_subsetI subsetCE) - - show "generate H (h ` K) \ h ` generate G K" - proof - fix x assume "x \ generate H (h ` K)" - then obtain r where r: "elts r \ (h ` K)" "Generated_Groups.norm H r = x" - using H.generate_repr_iff img_in_carrier by auto - from \elts r \ (h ` K)\ have "Generated_Groups.norm H r \ h ` generate G K" - proof (induct r rule: repr.induct) - case One - have "\\<^bsub>G\<^esub> \ generate G K" using generate.one[of G] by simp - hence "h \\<^bsub>G\<^esub> \ h ` generate G K" by blast - thus ?case by simp - next - case (Inv x) hence "x \ h ` K" by simp - then obtain k where k: "k \ K" "x = h k" by blast - hence "inv\<^bsub>H\<^esub> x = h (inv k)" using assms by auto - thus ?case using k by (simp add: generate.inv) - next - case (Leaf x) hence "x \ h ` K" by simp - then obtain k where "k \ K" "x = h k" by blast - thus ?case by (simp add: generate.incl) - next - case (Mult x1 x2) hence A: "elts x1 \ elts x2 \ h ` K" by simp - have "Generated_Groups.norm H x1 \ h ` (generate G K)" using A Mult by simp - moreover have "Generated_Groups.norm H x2 \ h ` (generate G K)" using A Mult by simp - ultimately obtain k1 k2 where k1: "k1 \ generate G K" "Generated_Groups.norm H x1 = h k1" - and k2: "k2 \ generate G K" "Generated_Groups.norm H x2 = h k2" by blast - hence "Generated_Groups.norm H (Mult x1 x2) = h (k1 \ k2)" - using G.generate_in_carrier assms by auto - thus ?case using k1 k2 by (simp add: generate.eng) - qed - thus "x \ h ` generate G K" using r by simp - qed - -next - show "h ` generate G K \ generate H (h ` K)" - proof - fix x assume "x \ h ` generate G K" - then obtain r where r: "elts r \ K" "x = h (Generated_Groups.norm G r)" using G.generate_repr_iff assms by auto - from \elts r \ K\ have "h (Generated_Groups.norm G r) \ generate H (h ` K)" - proof (induct r rule: repr.induct) - case One thus ?case by (simp add: generate.one) - next - case (Inv x) hence A: "x \ K" by simp - hence "h (Generated_Groups.norm G (Inv x)) = inv\<^bsub>H\<^esub> (h x)" using assms by auto - moreover have "h x \ generate H (h ` K)" using A by (simp add: generate.incl) - ultimately show ?case by (simp add: A generate.inv) - next - case (Leaf x) thus ?case by (simp add: generate.incl) - next - case (Mult x1 x2) hence A: "elts x1 \ elts x2 \ K" by simp - have "Generated_Groups.norm G x1 \ carrier G" - by (meson A G.generateE(1) G.generate_repr_iff Un_subset_iff assms subgroup.mem_carrier) - moreover have "Generated_Groups.norm G x2 \ carrier G" - by (meson A G.generateE(1) G.generate_repr_iff Un_subset_iff assms subgroup.mem_carrier) - ultimately have "h (Generated_Groups.norm G (Mult x1 x2)) = h (Generated_Groups.norm G x1) \\<^bsub>H\<^esub> h (Generated_Groups.norm G x2)" by simp - thus ?case using Mult A by (simp add: generate.eng) - qed - thus "x \ generate H (h ` K)" using r by simp - qed -qed - -lemma (in group_hom) derived_of_img: - assumes "K \ carrier G" - shows "(derived H ^^ n) (h ` K) = h ` ((derived G ^^ n) K)" +theorem (in group_hom) solvable_img_imp_solvable: + assumes "subgroup K G" and "inj_on h K" and "solvable_seq H (h ` K)" shows "solvable_seq G K" proof - - { fix K assume A: "K \ carrier G" - have "derived H (h ` K) = h ` (derived G K)" - proof - - have Kcarr: "\a. a \ K \ a \ carrier G" - by (metis (no_types) A subsetCE) - have "derived_set H (h ` K) = h ` (derived_set G K)" - proof - show "derived_set H (h ` K) \ h ` derived_set G K" - proof - fix x assume "x \ derived_set H (h ` K)" - then obtain k1 k2 - where k12: "k1 \ K" "k2 \ K" - and xeq: "x = (h k1) \\<^bsub>H\<^esub> (h k2) \\<^bsub>H\<^esub> inv\<^bsub>H\<^esub> (h k1) \\<^bsub>H\<^esub> inv\<^bsub>H\<^esub>(h k2)" by blast - hence "x = h (k1 \ k2 \ inv k1 \ inv k2)" - proof - - have "k1 \ carrier G" "k2 \ carrier G" - using A \k1 \ K\ \k2 \ K\ by blast+ - then show ?thesis - using G.inv_closed G.m_closed xeq hom_inv hom_mult by presburger - qed - thus "x \ h ` (derived_set G K)" using k12 by blast - qed - next - show "h ` derived_set G K \ derived_set H (h ` K)" - proof - fix x assume " x \ h ` derived_set G K" - then obtain k1 k2 where k12: "k1 \ K" "k2 \ K" - and "x = h (k1 \ k2 \ inv k1 \ inv k2)" by blast - hence "x = (h k1) \\<^bsub>H\<^esub> (h k2) \\<^bsub>H\<^esub> inv\<^bsub>H\<^esub> (h k1) \\<^bsub>H\<^esub> inv\<^bsub>H\<^esub>(h k2)" - by (metis (no_types) Kcarr G.inv_closed G.m_closed hom_inv hom_mult) - thus "x \ derived_set H (h ` K)" using k12 by blast - qed - qed - thus ?thesis unfolding derived_def using generate_of_img - by (simp add: G.derived_set_in_carrier A) - qed } note aux_lemma = this - - from \K \ carrier G\ show ?thesis - proof (induction n) - case 0 thus ?case by simp - next - case (Suc n) - have "(derived H ^^ Suc n) (h ` K) = (derived H) ((derived H ^^ n) (h ` K))" by simp - also have " ... = (derived H) (h ` ((derived G ^^ n) K))" using Suc by simp - also have " ... = h ` ((derived G) ((derived G ^^ n) K))" - using aux_lemma[of "(derived G ^^ n) K"] G.exp_of_derived_in_carrier[OF Suc(2),of n] by linarith - also have " ... = h ` ((derived G ^^ (Suc n)) K)" by simp - finally show ?case . - qed -qed - -theorem (in group_hom) solvable_img_imp_solvable: - assumes "subgroup I G" - and "inj_on h I" - and "solvable_seq H (h ` I)" - shows "solvable_seq G I" -proof - - { fix n I assume A: "subgroup I G" "inj_on h I" - hence "inj_on h ((derived G ^^ n) I)" - proof - - have "(derived G ^^ n) I \ I" - proof (induction n) - case 0 thus ?case by simp - next - case (Suc n) - hence "(derived G) ((derived G ^^ n) I) \ (derived G) I" - using G.mono_derived[of I "(derived G ^^ n) I" 1, - OF subgroup.subset[OF A(1)] Suc] by simp - thus ?case using A(1) G.derived_incl by auto - qed - thus ?thesis using A(2) inj_on_subset by blast - qed } note aux_lemma = this - - obtain n where "(derived H ^^ n) (h ` I) = { \\<^bsub>H\<^esub> }" - using H.solvable_imp_subgroup H.solvable_imp_trivial_derived_seq assms(3) by blast - hence "h ` ((derived G ^^ n) I) = { \\<^bsub>H\<^esub> }" - by (metis derived_of_img assms(1) subgroup.subset) - moreover have "inj_on h ((derived G ^^ n) I)" - using aux_lemma[OF assms(1-2), of n] by simp - hence "\x. \ x \ ((derived G ^^ n) I); h x = \\<^bsub>H\<^esub> \ \ x = \" - by (metis G.exp_of_derived_is_subgroup assms(1) hom_one inj_on_eq_iff subgroup_def) - ultimately have "(derived G ^^ n) I = { \ }" by blast - thus ?thesis - using G.trivial_derived_seq_imp_solvable[OF assms(1), of n] by simp -qed - -corollary (in group_hom) not_solvable: - assumes "inj_on h (carrier G)" - and "\ solvable G" - shows "\ solvable H" -proof - - { fix I J assume A: "subgroup I H" "subgroup J H" "I \ J" "solvable_seq H J" - have "solvable_seq H I" - proof - - obtain n where n: "(derived H ^^ n) J = { \\<^bsub>H\<^esub> }" - using A(4) H.solvable_imp_subgroup H.solvable_imp_trivial_derived_seq by blast - have "(derived H ^^ n) I \ (derived H ^^ n) J" - using A by (simp add: H.mono_derived subgroupE(1)) - hence "(derived H ^^ n) I \ { \\<^bsub>H\<^esub> }" using n by simp - hence "(derived H ^^ n) I = { \\<^bsub>H\<^esub> }" - by (simp add: A(1) subgroupE(2)[OF H.exp_of_derived_is_subgroup] subset_singleton_iff) - thus ?thesis - using A(1) H.trivial_derived_seq_imp_solvable by blast - qed } note aux_lemma = this - - show ?thesis - proof (rule ccontr) - assume "\ \ solvable H" - hence "solvable_seq H (carrier H)" unfolding solvable_def by simp - hence "solvable_seq H (h ` (carrier G))" - using aux_lemma[of "h ` (carrier G)" "carrier H"] - by (metis G.generateI G.subgroupE(1) G.subgroup_self H.generateE(1) - H.subgroup_self generate_of_img hom_closed image_subsetI) - hence "solvable_seq G (carrier G)" - using G.subgroup_self assms(1) solvable_img_imp_solvable by blast - hence "solvable G" unfolding solvable_def by simp - thus False using assms(2) by simp - qed + obtain n where "(derived H ^^ n) (h ` K) = { \\<^bsub>H\<^esub> }" + using solvable_imp_trivial_derived_seq assms(1,3) by auto + hence "h ` ((derived G ^^ n) K) = { \\<^bsub>H\<^esub> }" + unfolding exp_of_derived_img[OF subgroup.subset[OF assms(1)]] . + moreover have "(derived G ^^ n) K \ K" + using G.mono_derived[of _ K] G.derived_incl[OF _ assms(1)] by (induct n) (auto) + hence "inj_on h ((derived G ^^ n) K)" + using inj_on_subset[OF assms(2)] by blast + moreover have "{ \ } \ (derived G ^^ n) K" + using subgroup.one_closed[OF G.exp_of_derived_is_subgroup[OF assms(1)]] by blast + ultimately show ?thesis + using G.trivial_derived_seq_imp_solvable[OF assms(1), of n] + by (metis (no_types, lifting) hom_one image_empty image_insert inj_on_image_eq_iff order_refl) qed corollary (in group_hom) inj_hom_imp_solvable: - assumes "inj_on h (carrier G)" - shows "solvable H \ solvable G" - using not_solvable[OF assms] by auto + assumes "inj_on h (carrier G)" and "solvable H" shows "solvable G" + using solvable_img_imp_solvable[OF _ assms(1)] G.subgroup_self + solvable_subgroup[OF subgroup_img_is_subgroup assms(2)] + unfolding solvable_def + by simp theorem (in group_hom) solvable_imp_solvable_img: - assumes "subgroup I G" - and "solvable_seq G I" - shows "solvable_seq H (h ` I)" + assumes "solvable_seq G K" shows "solvable_seq H (h ` K)" proof - - obtain n where "(derived G ^^ n) I = { \\<^bsub>G\<^esub> }" - using G.solvable_imp_trivial_derived_seq[OF assms(2) assms(1)] .. - hence "(derived H ^^ n) (h ` I) = { \\<^bsub>H\<^esub> }" - using derived_of_img[OF G.subgroupE(1)[OF assms(1)], of n] by simp + obtain n where "(derived G ^^ n) K = { \ }" + using G.solvable_imp_trivial_derived_seq[OF assms] by blast thus ?thesis - using H.trivial_derived_seq_imp_solvable[OF subgroup_img_is_subgroup[OF assms(1)]] by simp + using trivial_derived_seq_imp_solvable[OF subgroup_img_is_subgroup, of _ n] + exp_of_derived_img[OF subgroup.subset, of _ n] G.solvable_imp_subgroup[OF assms] + by auto qed corollary (in group_hom) surj_hom_imp_solvable: - assumes "h ` (carrier G) = (carrier H)" - shows "solvable G \ solvable H" - using solvable_imp_solvable_img[OF G.subgroup_self] assms unfolding solvable_def by auto + assumes "h ` carrier G = carrier H" and "solvable G" shows "solvable H" + using assms solvable_imp_solvable_img[of "carrier G"] unfolding solvable_def by simp lemma solvable_seq_condition: - assumes "group_hom G1 G2 h" "group_hom G2 G3 f" - and "subgroup I G1" "subgroup J G2" - and "h ` I \ J" - and "\g. \ g \ carrier G2; f g = \\<^bsub>G3\<^esub> \ \ g \ h ` I" - shows "\ solvable_seq G1 I; solvable_seq G3 (f ` J) \ \ solvable_seq G2 J" + assumes "group_hom G H f" "group_hom H K g" and "f ` I \ J" and "kernel H K g \ f ` I" + and "subgroup J H" and "solvable_seq G I" "solvable_seq K (g ` J)" + shows "solvable_seq H J" proof - - have G1: "group G1" and G2: "group G2" and G3: "group G3" - using assms(1-2) unfolding group_hom_def by auto + interpret G: group G + H: group H + K: group K + J: subgroup J H + I: subgroup I G + using assms(1-2,5) group.solvable_imp_subgroup[OF _ assms(6)] unfolding group_hom_def by auto - assume "solvable_seq G1 I" "solvable_seq G3 (f ` J)" - then obtain n m :: nat - where n: "(derived G1 ^^ n) I = { \\<^bsub>G1\<^esub> }" - and m: "(derived G3 ^^ m) (f ` J) = { \\<^bsub>G3\<^esub> }" - using group.solvable_imp_trivial_derived_seq[OF G1, of I] - group.solvable_imp_trivial_derived_seq[OF G3, of "f ` J"] - group_hom.subgroup_img_is_subgroup[OF assms(2) assms(4)] assms(2-4) by auto - have "f ` ((derived G2 ^^ m) J) = (derived G3 ^^ m) (f ` J)" - using group_hom.derived_of_img[OF assms(2), of J m] subgroup.subset[OF assms(4)] by simp - hence "f ` ((derived G2 ^^ m) J) = { \\<^bsub>G3\<^esub> }" - using m by simp - hence "((derived G2 ^^ m) J) \ h ` I" - using assms(6) group.exp_of_derived_in_carrier[OF G2 subgroup.subset[OF assms(4)], of m] - by blast - hence "(derived G2 ^^ n) ((derived G2 ^^ m) J) \ (derived G2 ^^ n) (h ` I)" - using group.mono_derived[OF G2, of "h ` I" "(derived G2 ^^ m) J" n] - subgroup.subset[OF group_hom.subgroup_img_is_subgroup[OF assms(1) assms(3)]] by blast - also have " ... = h ` { \\<^bsub>G1\<^esub> }" - using group_hom.derived_of_img[OF assms(1) subgroup.subset[OF assms(3)], of n] n by simp - also have " ... = { \\<^bsub>G2\<^esub> }" - using assms(1) by (simp add: group_hom.hom_one) - finally have "(derived G2 ^^ n) ((derived G2 ^^ m) J) \ { \\<^bsub>G2\<^esub> }" . - hence "(derived G2 ^^ (n + m)) J \ { \\<^bsub>G2\<^esub> }" - by (metis comp_eq_dest_lhs funpow_add) - moreover have "{ \\<^bsub>G2\<^esub> } \ (derived G2 ^^ (n + m)) J" - using subgroup.one_closed[OF group.exp_of_derived_is_subgroup[OF G2 assms(4), of "n + m"]] by simp + obtain n m + where n: "(derived G ^^ n) I = { \\<^bsub>G\<^esub> }" and m: "(derived K ^^ m) (g ` J) = { \\<^bsub>K\<^esub> }" + using G.solvable_imp_trivial_derived_seq[OF assms(6)] + K.solvable_imp_trivial_derived_seq[OF assms(7)] + by auto + have "(derived H ^^ m) J \ f ` I" + using m H.exp_of_derived_in_carrier[OF J.subset, of m] assms(4) + by (auto simp add: group_hom.exp_of_derived_img[OF assms(2) J.subset] kernel_def) + hence "(derived H ^^ n) ((derived H ^^ m) J) \ f ` ((derived G ^^ n) I)" + using n H.mono_exp_of_derived unfolding sym[OF group_hom.exp_of_derived_img[OF assms(1) I.subset, of n]] by simp + hence "(derived H ^^ (n + m)) J \ { \\<^bsub>H\<^esub> }" + using group_hom.hom_one[OF assms(1)] unfolding n by (simp add: funpow_add) + moreover have "{ \\<^bsub>H\<^esub> } \ (derived H ^^ (n + m)) J" + using subgroup.one_closed[OF H.exp_of_derived_is_subgroup[OF assms(5), of "n + m"]] by blast ultimately show ?thesis - using group.trivial_derived_seq_imp_solvable[OF G2 assms(4), of "n + m"] by auto + using H.trivial_derived_seq_imp_solvable[OF assms(5)] by simp qed lemma solvable_condition: - assumes "group_hom G1 G2 h" "group_hom G2 G3 f" - and "f ` (carrier G2) = (carrier G3)" - and "kernel G2 G3 f \ h ` (carrier G1)" - shows "\ solvable G1; solvable G3 \ \ solvable G2" -proof - - assume "solvable G1" "solvable G3" - moreover have "\g. \ g \ carrier G2; f g = \\<^bsub>G3\<^esub> \ \ g \ h ` (carrier G1)" - using assms(4) unfolding kernel_def by auto - moreover have "h ` (carrier G1 ) \ (carrier G2)" - using group_hom.hom_closed[OF assms(1)] image_subsetI by blast - ultimately show ?thesis - using solvable_seq_condition[OF assms(1-2), of "carrier G1" "carrier G2"] assms(1-3) - unfolding solvable_def group_hom_def by (simp add: group.subgroup_self) -qed + assumes "group_hom G H f" "group_hom H K g" + and "g ` (carrier H) = carrier K" and "kernel H K g \ f ` (carrier G)" + and "solvable G" "solvable K" shows "solvable H" + using solvable_seq_condition[OF assms(1-2) _ assms(4) group.subgroup_self] assms(3,5-6) + subgroup.subset[OF group_hom.img_is_subgroup[OF assms(1)]] group_hom.axioms(2)[OF assms(1)] + by (simp add: solvable_def) -end +end \ No newline at end of file diff -r 12dce58bcd3f -r 1b5178abaf97 src/HOL/Algebra/Subrings.thy --- a/src/HOL/Algebra/Subrings.thy Thu Oct 04 11:18:39 2018 +0200 +++ b/src/HOL/Algebra/Subrings.thy Thu Oct 04 15:25:47 2018 +0100 @@ -377,6 +377,8 @@ assumes "h \ ring_hom R S" and "subring K R" shows "ring (S \ carrier := h ` K, one := h \, zero := h \ \)" proof - + have [simp]: "h \ = \\<^bsub>S\<^esub>" + using assms ring_hom_one by blast have "ring (R \ carrier := K \)" by (simp add: assms(2) subring_is_ring) moreover have "h \ ring_hom (R \ carrier := K \) S" diff -r 12dce58bcd3f -r 1b5178abaf97 src/HOL/Algebra/Sym_Groups.thy --- a/src/HOL/Algebra/Sym_Groups.thy Thu Oct 04 11:18:39 2018 +0200 +++ b/src/HOL/Algebra/Sym_Groups.thy Thu Oct 04 15:25:47 2018 +0100 @@ -3,32 +3,72 @@ *) theory Sym_Groups - imports Cycles Group Coset Generated_Groups Solvable_Groups + imports Cycles Solvable_Groups + begin +section \Symmetric Groups\ + +subsection \Definitions\ + abbreviation inv' :: "('a \ 'b) \ ('b \ 'a)" where "inv' f \ Hilbert_Choice.inv f" definition sym_group :: "nat \ (nat \ nat) monoid" where "sym_group n = \ carrier = { p. p permutes {1..n} }, mult = (\), one = id \" +definition alt_group :: "nat \ (nat \ nat) monoid" + where "alt_group n = (sym_group n) \ carrier := { p. p permutes {1..n} \ evenperm p } \" + definition sign_img :: "int monoid" where "sign_img = \ carrier = { -1, 1 }, mult = (*), one = 1 \" +subsection \Basic Properties\ + +lemma sym_group_carrier: "p \ carrier (sym_group n) \ p permutes {1..n}" + unfolding sym_group_def by simp + +lemma sym_group_mult: "mult (sym_group n) = (\)" + unfolding sym_group_def by simp + +lemma sym_group_one: "one (sym_group n) = id" + unfolding sym_group_def by simp + +lemma sym_group_carrier': "p \ carrier (sym_group n) \ permutation p" + unfolding sym_group_carrier permutation_permutes by auto + +lemma alt_group_carrier: "p \ carrier (alt_group n) \ p permutes {1..n} \ evenperm p" + unfolding alt_group_def by auto + +lemma alt_group_mult: "mult (alt_group n) = (\)" + unfolding alt_group_def using sym_group_mult by simp + +lemma alt_group_one: "one (alt_group n) = id" + unfolding alt_group_def using sym_group_one by simp + +lemma alt_group_carrier': "p \ carrier (alt_group n) \ permutation p" + unfolding alt_group_carrier permutation_permutes by auto + lemma sym_group_is_group: "group (sym_group n)" - apply (rule groupI) - apply (simp_all add: sym_group_def permutes_compose permutes_id comp_assoc) - using permutes_inv permutes_inv_o(2) by blast + using permutes_inv permutes_inv_o(2) + by (auto intro!: groupI + simp add: sym_group_def permutes_compose + permutes_id comp_assoc, blast) + +lemma sign_img_is_group: "group sign_img" + unfolding sign_img_def by (unfold_locales, auto simp add: Units_def) lemma sym_group_inv_closed: - assumes "p \ carrier (sym_group n)" - shows "inv' p \ carrier (sym_group n)" + assumes "p \ carrier (sym_group n)" shows "inv' p \ carrier (sym_group n)" using assms permutes_inv sym_group_def by auto -lemma sym_group_inv_equality: - assumes "p \ carrier (sym_group n)" - shows "inv\<^bsub>(sym_group n)\<^esub> p = inv' p" +lemma alt_group_inv_closed: + assumes "p \ carrier (alt_group n)" shows "inv' p \ carrier (alt_group n)" + using evenperm_inv[OF alt_group_carrier'] permutes_inv assms alt_group_carrier by auto + +lemma sym_group_inv_equality [simp]: + assumes "p \ carrier (sym_group n)" shows "inv\<^bsub>(sym_group n)\<^esub> p = inv' p" proof - have "inv' p \ p = id" using assms permutes_inv_o(2) sym_group_def by auto @@ -38,54 +78,51 @@ by (simp add: assms sym_group_inv_closed) qed -lemma sign_is_hom: - "group_hom (sym_group n) sign_img sign" - unfolding group_hom_def -proof (auto) - show "group (sym_group n)" - by (simp add: sym_group_is_group) -next - show "group sign_img" - unfolding sign_img_def group_def monoid_def group_axioms_def Units_def by auto -next - show "group_hom_axioms (sym_group n) sign_img sign" - unfolding sign_img_def group_hom_axioms_def sym_group_def hom_def - proof auto - show "\x. sign x \ - 1 \ x permutes {Suc 0..n} \ sign x = 1" - by (meson sign_def) - show "\x y. \ x permutes {Suc 0..n}; y permutes {Suc 0..n} \ \ - sign (x \ y) = sign x * sign y" - using sign_compose permutation_permutes by blast - qed -qed +lemma sign_is_hom: "sign \ hom (sym_group n) sign_img" + unfolding hom_def sign_img_def sym_group_mult using sym_group_carrier'[of _ n] + by (auto simp add: sign_compose, meson sign_def) + +lemma sign_group_hom: "group_hom (sym_group n) sign_img sign" + using group_hom.intro[OF sym_group_is_group sign_img_is_group] sign_is_hom + by (simp add: group_hom_axioms_def) +lemma sign_is_surj: + assumes "n \ 2" shows "sign ` (carrier (sym_group n)) = carrier sign_img" +proof - + have "swapidseq (Suc 0) (Fun.swap (1 :: nat) 2 id)" + using comp_Suc[OF id, of "1 :: nat" "2"] by auto + hence "sign (Fun.swap (1 :: nat) 2 id) = (-1 :: int)" + by (simp add: sign_swap_id) + moreover have "Fun.swap (1 :: nat) 2 id \ carrier (sym_group n)" and "id \ carrier (sym_group n)" + using assms permutes_swap_id[of "1 :: nat" "{1..n}" 2] permutes_id + unfolding sym_group_carrier by auto + ultimately have "carrier sign_img \ sign ` (carrier (sym_group n))" + using sign_id mk_disjoint_insert unfolding sign_img_def by fastforce + moreover have "sign ` (carrier (sym_group n)) \ carrier sign_img" + using sign_is_hom unfolding hom_def by auto + ultimately show ?thesis + by blast +qed -definition alt_group :: "nat \ (nat \ nat) monoid" - where "alt_group n = (sym_group n) \ carrier := { p. p permutes {1..n} \ evenperm p } \" - -lemma alt_group_is_kernel_from_sign: +lemma alt_group_is_sign_kernel: "carrier (alt_group n) = kernel (sym_group n) sign_img sign" unfolding alt_group_def sym_group_def sign_img_def kernel_def sign_def by auto -lemma alt_group_is_group: - "group (alt_group n)" -proof - - have "subgroup (kernel (sym_group n) sign_img sign) (sym_group n)" - using group_hom.subgroup_kernel sign_is_hom by blast - thus ?thesis - using alt_group_def alt_group_is_kernel_from_sign group.subgroup_imp_group - sym_group_is_group by fastforce -qed +lemma alt_group_is_subgroup: "subgroup (carrier (alt_group n)) (sym_group n)" + using group_hom.subgroup_kernel[OF sign_group_hom] + unfolding alt_group_is_sign_kernel by blast -lemma alt_group_inv_closed: - assumes "p \ carrier (alt_group n)" - shows "inv' p \ carrier (alt_group n)" - using assms permutes_inv alt_group_def - using evenperm_inv permutation_permutes by fastforce +lemma alt_group_is_group: "group (alt_group n)" + using group.subgroup_imp_group[OF sym_group_is_group alt_group_is_subgroup] + by (simp add: alt_group_def) + +lemma sign_iso: + assumes "n \ 2" shows "(sym_group n) Mod (carrier (alt_group n)) \ sign_img" + using group_hom.FactGroup_iso[OF sign_group_hom sign_is_surj[OF assms]] + unfolding alt_group_is_sign_kernel . lemma alt_group_inv_equality: - assumes "p \ carrier (alt_group n)" - shows "inv\<^bsub>(alt_group n)\<^esub> p = inv' p" + assumes "p \ carrier (alt_group n)" shows "inv\<^bsub>(alt_group n)\<^esub> p = inv' p" proof - have "inv' p \ p = id" using assms permutes_inv_o(2) alt_group_def by auto @@ -95,52 +132,61 @@ by (simp add: assms alt_group_inv_closed) qed +lemma sym_group_card_carrier: "card (carrier (sym_group n)) = fact n" + using card_permutations[of "{1..n}" n] unfolding sym_group_def by simp + +lemma alt_group_card_carrier: + assumes "n \ 2" shows "2 * card (carrier (alt_group n)) = fact n" +proof - + have "card (rcosets\<^bsub>sym_group n\<^esub> (carrier (alt_group n))) = 2" + using iso_same_card[OF sign_iso[OF assms]] unfolding FactGroup_def sign_img_def by auto + thus ?thesis + using group.lagrange[OF sym_group_is_group alt_group_is_subgroup, of n] + unfolding order_def sym_group_card_carrier by simp +qed + subsection \Transposition Sequences\ text \In order to prove that the Alternating Group can be generated by 3-cycles, we need a stronger decomposition of permutations as transposition sequences than the one - proposed found at Permutations.thy\ + proposed at Permutations.thy. \ -inductive swapidseq_ext :: "'a set \ nat \ ('a \ 'a) \ bool" where -empty: "swapidseq_ext {} 0 id" | -single: "\ swapidseq_ext S n p; a \ S \ \ swapidseq_ext (insert a S) n p" | -comp: "\ swapidseq_ext S n p; a \ b \ \ - swapidseq_ext (insert a (insert b S)) (Suc n) ((Fun.swap a b id) \ p)" +inductive swapidseq_ext :: "'a set \ nat \ ('a \ 'a) \ bool" + where + empty: "swapidseq_ext {} 0 id" + | single: "\ swapidseq_ext S n p; a \ S \ \ swapidseq_ext (insert a S) n p" + | comp: "\ swapidseq_ext S n p; a \ b \ \ + swapidseq_ext (insert a (insert b S)) (Suc n) ((Fun.swap a b id) \ p)" lemma swapidseq_ext_finite: - assumes "swapidseq_ext S n p" - shows "finite S" using assms - apply (induction) by auto + assumes "swapidseq_ext S n p" shows "finite S" + using assms by (induction) (auto) + +lemma swapidseq_ext_zero: + assumes "finite S" shows "swapidseq_ext S 0 id" + using assms empty by (induct set: "finite", fastforce, simp add: single) + +lemma swapidseq_ext_imp_swapidseq: + assumes "swapidseq_ext S n p" shows "swapidseq n p" + using assms by (induction, simp, simp, meson comp_Suc) lemma swapidseq_ext_zero_imp_id: - assumes "swapidseq_ext S 0 p" - shows "p = id" + assumes "swapidseq_ext S 0 p" shows "p = id" proof - - { fix S n and p :: "'a \ 'a" assume "swapidseq_ext S n p" "n = 0" - hence "p = id" - apply (induction) by auto } - thus ?thesis using assms by auto -qed - -lemma swapidseq_ext_zero: - assumes "finite S" - shows "swapidseq_ext S 0 id" using assms -proof (induct set: "finite") - case empty thus ?case using swapidseq_ext.empty . -next - case insert show ?case using swapidseq_ext.single[OF insert(3) insert(2)] . + have "\ swapidseq_ext S n p; n = 0 \ \ p = id" for n + by (induction rule: swapidseq_ext.induct, auto) + thus ?thesis + using assms by simp qed lemma swapidseq_ext_finite_expansion: - assumes "finite B" "swapidseq_ext A n p" - shows "swapidseq_ext (A \ B) n p" using assms -proof (induct set: "finite") - case empty thus ?case by simp -next - case insert show ?case - by (metis Un_insert_right insert.hyps(3) insert.prems insert_absorb single) + assumes "finite B" and "swapidseq_ext A n p" shows "swapidseq_ext (A \ B) n p" + using assms +proof (induct set: "finite", simp) + case (insert b B) show ?case + using insert single[OF insert(3), of b] by (metis Un_insert_right insert_absorb) qed lemma swapidseq_ext_backwards: @@ -148,32 +194,40 @@ shows "\a b A' p'. a \ b \ A = (insert a (insert b A')) \ swapidseq_ext A' n p' \ p = (Fun.swap a b id) \ p'" proof - - { fix A n k and p :: "'a \ 'a" assume "swapidseq_ext A n p" "n = Suc k" + { fix A n k and p :: "'a \ 'a" + assume "swapidseq_ext A n p" "n = Suc k" hence "\a b A' p'. a \ b \ A = (insert a (insert b A')) \ swapidseq_ext A' k p' \ p = (Fun.swap a b id) \ p'" - proof (induction) - case empty thus ?case by simp - next + proof (induction, simp) case single thus ?case by (metis Un_insert_right insert_iff insert_is_Un swapidseq_ext.single) next - case comp thus ?case by blast + case comp thus ?case + by blast qed } - thus ?thesis using assms by simp + thus ?thesis + using assms by simp qed +lemma swapidseq_ext_backwards': + assumes "swapidseq_ext A (Suc n) p" + shows "\a b A' p'. a \ A \ b \ A \ a \ b \ swapidseq_ext A n p' \ p = (Fun.swap a b id) \ p'" + using swapidseq_ext_backwards[OF assms] swapidseq_ext_finite_expansion + by (metis Un_insert_left assms insertI1 sup.idem sup_commute swapidseq_ext_finite) + lemma swapidseq_ext_endswap: assumes "swapidseq_ext S n p" "a \ b" - shows "swapidseq_ext (insert a (insert b S)) (Suc n) (p \ (Fun.swap a b id))" using assms + shows "swapidseq_ext (insert a (insert b S)) (Suc n) (p \ (Fun.swap a b id))" + using assms proof (induction n arbitrary: S p a b) case 0 hence "p = id" using swapidseq_ext_zero_imp_id by blast - thus ?case using 0 by (metis comp_id id_comp swapidseq_ext.comp) + thus ?case + using 0 by (metis comp_id id_comp swapidseq_ext.comp) next case (Suc n) then obtain c d S' and p' :: "'a \ 'a" - where cd: "c \ d" - and S: "S = (insert c (insert d S'))" "swapidseq_ext S' n p'" + where cd: "c \ d" and S: "S = (insert c (insert d S'))" "swapidseq_ext S' n p'" and p: "p = (Fun.swap c d id) \ p'" using swapidseq_ext_backwards[OF Suc(2)] by blast hence "swapidseq_ext (insert a (insert b S')) (Suc n) (p' \ (Fun.swap a b id))" @@ -181,83 +235,43 @@ hence "swapidseq_ext (insert c (insert d (insert a (insert b S')))) (Suc (Suc n)) ((Fun.swap c d id) \ p' \ (Fun.swap a b id))" by (metis cd fun.map_comp swapidseq_ext.comp) - then show ?case by (metis S(1) p insert_commute) -qed - -lemma swapidseq_ext_imp_swapiseq: - assumes "swapidseq_ext S n p" - shows "swapidseq n p" using assms -proof (induction) - case empty thus ?case by simp - case single show ?case using single(3) . -next - case comp thus ?case by (meson comp_Suc) + thus ?case + by (metis S(1) p insert_commute) qed lemma swapidseq_ext_extension: - assumes "swapidseq_ext A n p" "swapidseq_ext B m q" "A \ B = {}" + assumes "swapidseq_ext A n p" and "swapidseq_ext B m q" and "A \ B = {}" shows "swapidseq_ext (A \ B) (n + m) (p \ q)" -proof - - { fix m and q :: "'a \ 'a" and A B :: "'a set" assume "finite A" "swapidseq_ext B m q" - hence "swapidseq_ext (A \ B) m q" - proof (induct set: "finite") - case empty thus ?case by simp - next - case (insert a A') thus ?case - using swapidseq_ext.single[of "A' \ B" m q a] - by (metis Un_insert_left insert_absorb) - qed } note aux_lemma = this - - from assms show ?thesis - proof (induct n arbitrary: p A) - case 0 thus ?case - using swapidseq_ext_zero_imp_id[OF 0(1)] aux_lemma[of A B m q] by (simp add: swapidseq_ext_finite) - next - case (Suc n) - obtain a b A' and p' :: "'a \ 'a" - where A': "a \ b" "A = (insert a (insert b A'))" - and p': "swapidseq_ext A' n p'" - and p: "p = (Fun.swap a b id) \ p'" - using swapidseq_ext_backwards[OF Suc(2)] by blast - hence "swapidseq_ext (A' \ B) (n + m) (p' \ q)" - using Suc.hyps Suc.prems(3) assms(2) by fastforce - thus ?case using swapidseq_ext.comp[of "A' \ B" "n + m" "p' \ q" a b] - by (metis Un_insert_left p A' add_Suc rewriteR_comp_comp) - qed + using assms(1,3) +proof (induction, simp add: assms(2)) + case single show ?case + using swapidseq_ext.single[OF single(3)] single(2,4) by auto +next + case comp show ?case + using swapidseq_ext.comp[OF comp(3,2)] comp(4) + by (metis Un_insert_left add_Suc insert_disjoint(1) o_assoc) qed lemma swapidseq_ext_of_cycles: - assumes "cycle cs" - shows "swapidseq_ext (set cs) (length cs - 1) (cycle_of_list cs)" using assms -proof (induction cs rule: cycle_of_list.induct) + assumes "cycle cs" shows "swapidseq_ext (set cs) (length cs - 1) (cycle_of_list cs)" + using assms +proof (induct cs rule: cycle_of_list.induct) case (1 i j cs) show ?case - proof (cases) - assume "cs = []" thus ?case - using swapidseq_ext.comp[OF swapidseq_ext.empty, of i j] "1.prems" by auto - next - assume "cs \ []" hence "length (j # cs) \ 2" - using not_less_eq_eq by fastforce - hence IH: "swapidseq_ext (set (j # cs)) (length (j # cs) - 1) (cycle_of_list (j # cs))" - using "1.IH" "1.prems" by auto - thus ?case using swapidseq_ext.comp[OF IH, of i j] - by (metis "1.prems" cycle_of_list.simps(1) diff_Suc_1 distinct.simps(2) - distinct_length_2_or_more insert_absorb length_Cons list.simps(15)) - qed + using comp[OF 1(1), of i j] 1(2) by (simp add: o_def) next - case "2_1" thus ?case using swapidseq_ext.empty - by (metis cycle_of_list.simps(2) diff_0_eq_0 empty_set list.size(3)) + case "2_1" show ?case + by (simp, metis eq_id_iff empty) next - case ("2_2" v) thus ?case using swapidseq_ext.single[OF swapidseq_ext.empty, of v] - by (metis cycle_of_list.simps(3) diff_Suc_1 distinct.simps(2) - empty_set length_Cons list.simps(15) list.size(3)) + case ("2_2" v) show ?case + using single[OF empty, of v] by (simp, metis eq_id_iff) qed lemma cycle_decomp_imp_swapidseq_ext: - assumes "cycle_decomp S p" - shows "\n. swapidseq_ext S n p" using assms + assumes "cycle_decomp S p" shows "\n. swapidseq_ext S n p" + using assms proof (induction) - case empty - then show ?case using swapidseq_ext.empty by blast + case empty show ?case + using swapidseq_ext.empty by blast next case (comp I p cs) then obtain m where m: "swapidseq_ext I m p" by blast @@ -267,415 +281,274 @@ using comp.hyps(3) by blast qed -lemma swapidseq_ext_of_permutations: - assumes "p permutes S" "finite S" - shows "\n. swapidseq_ext S n p" - using assms cycle_decomp_imp_swapidseq_ext cycle_decomposition by blast - -lemma split_swapidseq: - assumes "k \ n" "swapidseq n p" - shows "\q r. swapidseq k q \ swapidseq (n - k) r \ p = q \ r" -proof - - { fix n :: "nat" and p :: "'a \ 'a" assume "swapidseq (Suc n) p" - hence "\a b q. a \ b \ swapidseq n q \ p = (Fun.swap a b id) \ q" - proof (cases) - case comp_Suc thus ?thesis by auto - qed } note aux_lemma = this - - from assms show ?thesis - proof (induction k) - case 0 thus ?case by simp - next - case (Suc k) - then obtain r q where 1: "swapidseq k q" "swapidseq (n - k) r" "p = q \ r" - using Suc_leD by blast - then obtain a b r' where 2: "a \ b" "swapidseq (n - (Suc k)) r'" "r = (Fun.swap a b id) \ r'" - using aux_lemma[of "n - (Suc k)" r] by (metis Suc.prems(1) Suc_diff_le diff_Suc_Suc) - have "swapidseq (Suc k) (q \ (Fun.swap a b id))" using 1 2 by (simp add: swapidseq_endswap) - moreover have "p = (q \ (Fun.swap a b id)) \ r'" - using 1 2 fun.map_comp by blast - ultimately show ?case using 2 by blast - qed -qed +lemma swapidseq_ext_of_permutation: + assumes "p permutes S" and "finite S" shows "\n. swapidseq_ext S n p" + using cycle_decomp_imp_swapidseq_ext[OF cycle_decomposition[OF assms]] . lemma split_swapidseq_ext: - assumes "k \ n" "swapidseq_ext S n p" - shows "\q r S1 S2. swapidseq_ext S1 k q \ swapidseq_ext S2 (n - k) r \ p = q \ r \ S1 \ S2 = S" - using assms -proof (induction k) - case 0 have "finite S" - using "0.prems"(2) swapidseq_ext_finite by auto - have "swapidseq_ext {} 0 id \ swapidseq_ext S (n - 0) p \ p = id \ p" - using swapidseq_ext.empty by (simp add: assms(2)) - thus ?case by blast -next - case (Suc k) - then obtain q r S1 S2 where "swapidseq_ext S1 k q" "swapidseq_ext S2 (n - k) r" "p = q \ r" "S1 \ S2 = S" - using Suc_leD by blast - then obtain a b S2' and r' :: "'a \ 'a" - where S2': "a \ b" "S2 = (insert a (insert b S2'))" - and r': "swapidseq_ext S2' (n - (Suc k)) r'" - and r: "r = (Fun.swap a b id) \ r'" - by (metis Suc.prems(1) Suc_diff_le diff_Suc_Suc swapidseq_ext_backwards) - have "swapidseq_ext (insert a (insert b S1)) (Suc k) (q \ (Fun.swap a b id))" - by (simp add: S2'(1) \swapidseq_ext S1 k q\ swapidseq_ext_endswap) - moreover have "p = (q \ (Fun.swap a b id)) \ r'" - by (simp add: \p = q \ r\ fun.map_comp r) - moreover have "(insert a (insert b S1)) \ S2' = S" - using S2'(2) \S1 \ S2 = S\ by auto - ultimately show ?case using r r' by blast + assumes "k \ n" and "swapidseq_ext S n p" + obtains q r U V where "swapidseq_ext U (n - k) q" and "swapidseq_ext V k r" and "p = q \ r" and "U \ V = S" +proof - + from assms + have "\q r U V. swapidseq_ext U (n - k) q \ swapidseq_ext V k r \ p = q \ r \ U \ V = S" + (is "\q r U V. ?split k q r U V") + proof (induct k rule: inc_induct) + case base thus ?case + by (metis diff_self_eq_0 id_o sup_bot.left_neutral empty) + next + case (step m) + then obtain q r U V + where q: "swapidseq_ext U (n - Suc m) q" and r: "swapidseq_ext V (Suc m) r" + and p: "p = q \ r" and S: "U \ V = S" + by blast + obtain a b r' V' + where "a \ b" and r': "V = (insert a (insert b V'))" "swapidseq_ext V' m r'" "r = (Fun.swap a b id) \ r'" + using swapidseq_ext_backwards[OF r] by blast + have "swapidseq_ext (insert a (insert b U)) (n - m) (q \ (Fun.swap a b id))" + using swapidseq_ext_endswap[OF q \a \ b\] step(2) by (metis Suc_diff_Suc) + hence "?split m (q \ (Fun.swap a b id)) r' (insert a (insert b U)) V'" + using r' S unfolding p by fastforce + thus ?case by blast + qed + thus ?thesis + using that by blast qed +subsection \Unsolvability of Symmetric Groups\ -definition three_cycles :: "nat \ (nat \ nat) set" +text \We show that symmetric groups (@{term\sym_group n\}) are unsolvable for @{term\n \ 5\}.\ + +abbreviation three_cycles :: "nat \ (nat \ nat) set" where "three_cycles n \ { cycle_of_list cs | cs. cycle cs \ length cs = 3 \ set cs \ {1..n} }" lemma stupid_lemma: - assumes "length cs = 3" - shows "cs = [(cs ! 0), (cs ! 1), (cs ! 2)]" -proof (intro nth_equalityI) - show "length cs = length [(cs ! 0), (cs ! 1), (cs ! 2)]" - using assms by simp - show "\i. i < length cs \ cs ! i = [(cs ! 0), (cs ! 1), (cs ! 2)] ! i" - by (metis Suc_1 Suc_eq_plus1 add.left_neutral assms less_antisym - less_one nth_Cons' nth_Cons_Suc numeral_3_eq_3) + assumes "length cs = 3" shows "cs = [ (cs ! 0), (cs ! 1), (cs ! 2) ]" + using assms by (auto intro!: nth_equalityI) + (metis Suc_lessI less_2_cases not_less_eq nth_Cons_0 + nth_Cons_Suc numeral_2_eq_2 numeral_3_eq_3) + +lemma three_cycles_incl: "three_cycles n \ carrier (alt_group n)" +proof + fix p assume "p \ three_cycles n" + then obtain cs where cs: "p = cycle_of_list cs" "cycle cs" "length cs = 3" "set cs \ {1..n}" + by auto + obtain a b c where cs_def: "cs = [ a, b, c ]" + using stupid_lemma[OF cs(3)] by auto + have "swapidseq (Suc (Suc 0)) ((Fun.swap a b id) \ (Fun.swap b c id))" + using comp_Suc[OF comp_Suc[OF id], of b c a b] cs(2) unfolding cs_def by simp + hence "evenperm p" + using cs(1) unfolding cs_def by (simp add: evenperm_unique) + thus "p \ carrier (alt_group n)" + using permutes_subset[OF cycle_permutes cs(4)] + unfolding alt_group_carrier cs(1) by simp qed -lemma alt_group_as_three_cycles: +lemma alt_group_carrier_as_three_cycles: "carrier (alt_group n) = generate (alt_group n) (three_cycles n)" -proof - show "generate (alt_group n) (three_cycles n) \ carrier (alt_group n)" - proof - { fix p assume "p \ three_cycles n" - have "p \ carrier (alt_group n)" - proof - - from \p \ three_cycles n\ - obtain cs where p: "p = cycle_of_list cs" - and cs: "cycle cs" "length cs = 3" "set cs \ {1..n}" - using three_cycles_def by blast - hence "p = (Fun.swap (cs ! 0) (cs ! 1) id) \ (Fun.swap (cs ! 1) (cs ! 2) id)" - using stupid_lemma[OF cs(2)] p - by (metis comp_id cycle_of_list.simps(1) cycle_of_list.simps(3)) - - hence "evenperm p" - by (metis cs(1) distinct_length_2_or_more evenperm_comp - evenperm_swap permutation_swap_id stupid_lemma[OF cs(2)]) - - moreover have "permutation p" using p cs(1) cycle_permutes by simp - hence "p permutes {1..n}" - using id_outside_supp[OF cs(1)] p cs permutation_permutes unfolding permutes_def - using permutation_permutes permutes_def subsetCE by metis - - ultimately show ?thesis by (simp add: alt_group_def) - qed } note aux_lemma = this +proof - + interpret A: group "alt_group n" + using alt_group_is_group by simp - fix p assume "p \ generate (alt_group n) (three_cycles n)" - thus "p \ carrier (alt_group n)" - proof (induction) - case one thus ?case by (simp add: alt_group_is_group group.is_monoid) - case incl thus ?case using aux_lemma unfolding alt_group_def by auto - case inv thus ?case using aux_lemma by (simp add: alt_group_is_group) next - case eng thus ?case by (simp add: alt_group_is_group group.is_monoid monoid.m_closed) - qed - qed - -next - show "carrier (alt_group n) \ generate (alt_group n) (three_cycles n)" + show ?thesis proof - fix p assume p: "p \ carrier (alt_group n)" - show "p \ generate (alt_group n) (three_cycles n)" - proof - + show "generate (alt_group n) (three_cycles n) \ carrier (alt_group n)" + using A.generate_subgroup_incl[OF three_cycles_incl A.subgroup_self] . + next + show "carrier (alt_group n) \ generate (alt_group n) (three_cycles n)" + proof { fix q :: "nat \ nat" and a b c - assume A: "a \ b" "b \ c" "{ a, b, c } \ {1..n}" "q = cycle_of_list [a, b, c]" - have "q \ generate (alt_group n) (three_cycles n)" + assume "a \ b" "b \ c" "{ a, b, c } \ {1..n}" + have "cycle_of_list [a, b, c] \ generate (alt_group n) (three_cycles n)" proof (cases) - assume "c = a" hence "q = id" by (simp add: A(4) swap_commute) - thus "q \ generate (alt_group n) (three_cycles n)" - using generate.one[of "alt_group n"] by (simp add: alt_group_def sym_group_def) + assume "c = a" + hence "cycle_of_list [ a, b, c ] = id" + by (simp add: swap_commute) + thus "cycle_of_list [ a, b, c ] \ generate (alt_group n) (three_cycles n)" + using one[of "alt_group n"] unfolding alt_group_one by simp next - assume "c \ a" - have "q \ (three_cycles n)" - unfolding three_cycles_def mem_Collect_eq - proof (intro exI conjI) - show "cycle [a,b,c]" - using A \c \ a\ by auto - qed (use A in auto) - thus "q \ generate (alt_group n) (three_cycles n)" - by (simp add: generate.incl) - qed } note gen3 = this - - { fix S :: "nat set" and q :: "nat \ nat" assume A: "swapidseq_ext S 2 q" "S \ {1..n}" + assume "c \ a" + have "distinct [a, b, c]" + using \a \ b\ and \b \ c\ and \c \ a\ by auto + with \{ a, b, c } \ {1..n}\ + show "cycle_of_list [ a, b, c ] \ generate (alt_group n) (three_cycles n)" + by (intro incl, fastforce) + qed } note aux_lemma1 = this + + { fix S :: "nat set" and q :: "nat \ nat" + assume seq: "swapidseq_ext S (Suc (Suc 0)) q" and S: "S \ {1..n}" have "q \ generate (alt_group n) (three_cycles n)" proof - - obtain a b S' q' where ab: "a \ b" "S = (insert a (insert b S'))" - and q': "swapidseq_ext S' 1 q'" "q = (Fun.swap a b id) \ q'" - using swapidseq_ext_backwards[of S 1 q] A(1) Suc_1 by metis - then obtain c d S'' where cd: "c \ d" "S' = (insert c (insert d S''))" - and q: "q = (Fun.swap a b id) \ (Fun.swap c d id)" - using swapidseq_ext_backwards[of S' 0 q'] - by (metis One_nat_def comp_id swapidseq_ext_zero_imp_id) - hence incl: "{ a, b, c, d } \ {1..n}" using A(2) ab(2) by blast + obtain a b q' where ab: "a \ b" "a \ S" "b \ S" + and q': "swapidseq_ext S (Suc 0) q'" "q = (Fun.swap a b id) \ q'" + using swapidseq_ext_backwards'[OF seq] by auto + obtain c d where cd: "c \ d" "c \ S" "d \ S" + and q: "q = (Fun.swap a b id) \ (Fun.swap c d id)" + using swapidseq_ext_backwards'[OF q'(1)] + swapidseq_ext_zero_imp_id + unfolding q'(2) + by fastforce + + consider (eq) "b = c" | (ineq) "b \ c" by auto thus ?thesis - proof (cases) - assume Eq: "b = c" hence "q = cycle_of_list [a, b, d]" by (simp add: q) - thus ?thesis using gen3 ab cd Eq incl by simp + proof cases + case eq then have "q = cycle_of_list [ a, b, d ]" + unfolding q by simp + moreover have "{ a, b, d } \ {1..n}" + using ab cd S by blast + ultimately show ?thesis + using aux_lemma1[OF ab(1)] cd(1) eq by simp next - assume In: "b \ c" - have "q = (cycle_of_list [a, b, c]) \ (cycle_of_list [b, c, d])" - by (metis (no_types, lifting) comp_id cycle_of_list.simps(1) - cycle_of_list.simps(3) fun.map_comp q swap_id_idempotent) - moreover have "... = cycle_of_list [a, b, c] \\<^bsub>alt_group n\<^esub> cycle_of_list [b, c, d]" - by (simp add: alt_group_def sym_group_def) + case ineq + hence "q = cycle_of_list [ a, b, c ] \ cycle_of_list [ b, c, d ]" + unfolding q by (simp add: comp_swap) + moreover have "{ a, b, c } \ {1..n}" and "{ b, c, d } \ {1..n}" + using ab cd S by blast+ ultimately show ?thesis - by (metis (no_types) In generate.eng[where ?h1.0 = "cycle_of_list [a, b, c]" - and ?h2.0 = "cycle_of_list [b, c, d]"] - gen3[of a b c] gen3[of b c d] \a \ b\ \c \ d\ insert_subset incl) + using eng[OF aux_lemma1[OF ab(1) ineq] aux_lemma1[OF ineq cd(1)]] + unfolding alt_group_mult by simp qed - qed } note gen3swap = this + qed } note aux_lemma2 = this - have "p permutes {1..n}" - using p permutation_permutes unfolding alt_group_def by auto - then obtain l where "swapidseq_ext {1..n} l p" "swapidseq l p" - using swapidseq_ext_of_permutations swapidseq_ext_imp_swapiseq by blast - - have "evenperm p" using p by (simp add: alt_group_def) - hence "even l" using \swapidseq l p\ evenperm_unique by blast - - then obtain k where "swapidseq_ext {1..n} (2 * k) p" - using dvd_def by (metis \swapidseq_ext {1..n} l p\) - thus "p \ generate (alt_group n) (three_cycles n)" + fix p assume "p \ carrier (alt_group n)" then have p: "p permutes {1..n}" "evenperm p" + unfolding alt_group_carrier by auto + obtain m where m: "swapidseq_ext {1..n} m p" + using swapidseq_ext_of_permutation[OF p(1)] by auto + have "even m" + using swapidseq_ext_imp_swapidseq[OF m] p(2) evenperm_unique by blast + then obtain k where k: "m = 2 * k" + by auto + show "p \ generate (alt_group n) (three_cycles n)" + using m unfolding k proof (induct k arbitrary: p) - case 0 - hence "p = id" by (simp add: swapidseq_ext_zero_imp_id) - moreover have "id \ generate (alt_group n) (three_cycles n)" - using generate.one[of "alt_group n"] by (simp add: alt_group_def sym_group_def) - ultimately show ?case by blast + case 0 then have "p = id" + using swapidseq_ext_zero_imp_id by simp + show ?case + using generate.one[of "alt_group n" "three_cycles n"] + unfolding alt_group_one \p = id\ . next - case (Suc k) - then obtain S1 S2 q r - where split: "swapidseq_ext S1 2 q" "swapidseq_ext S2 (2 * k) r" "p = q \ r" "S1 \ S2 = {1..n}" - using split_swapidseq_ext[of 2 "2 * Suc k" "{1..n}" p] by auto - - hence "r \ generate (alt_group n) (three_cycles n)" - using Suc.hyps swapidseq_ext_finite_expansion[of "{1..n}" S2 "2 * k" r] - by (metis (no_types, lifting) Suc.prems Un_commute sup.right_idem swapidseq_ext_finite) - - moreover have "q \ generate (alt_group n) (three_cycles n)" - using gen3swap[OF split(1)] \S1 \ S2 = {1..n}\ by auto - ultimately show ?case - using split generate.eng[of q "alt_group n" "three_cycles n" r] - by (metis (full_types) alt_group_def monoid.simps(1) partial_object.simps(3) sym_group_def) + case (Suc m) + have arith: "2 * (Suc m) - (Suc (Suc 0)) = 2 * m" and "Suc (Suc 0) \ 2 * Suc m" + by auto + then obtain q r U V + where q: "swapidseq_ext U (2 * m) q" and r: "swapidseq_ext V (Suc (Suc 0)) r" + and p: "p = q \ r" and UV: "U \ V = {1..n}" + using split_swapidseq_ext[OF _ Suc(2), of "Suc (Suc 0)"] unfolding arith by metis + have "swapidseq_ext {1..n} (2 * m) q" + using UV q swapidseq_ext_finite_expansion[OF swapidseq_ext_finite[OF r] q] by simp + thus ?case + using eng[OF Suc(1) aux_lemma2[OF r], of q] UV unfolding alt_group_mult p by blast qed qed qed qed -lemma elts_from_card: - assumes "card S \ n" - obtains f where "inj_on f {1..n}" "f ` {1..n} \ S" -proof - - have "\f. inj_on f {1..n} \ f ` {1..n} \ S" using assms - proof (induction n) - case 0 thus ?case by simp - next - case (Suc n) - then obtain f where f: "inj_on f {1..n}" "f ` {1..n} \ S" - using Suc_leD by blast - hence "card (f ` {1..n}) = n" by (simp add: card_image) - then obtain y where y: "y \ S - (f ` {1..n})" - by (metis Diff_eq_empty_iff Suc.prems \f ` {1..n} \ S\ not_less_eq_eq order_refl some_in_eq subset_antisym) - define f' where f': "f' = (\x. (if x \ {1..n} then f x else y))" - hence "\i j. \ i \ {1..Suc n}; j \ {1..Suc n} \ \ f' i = f' j \ i = j" - by (metis (no_types, lifting) DiffD2 f(1) y atLeastAtMostSuc_conv atLeastatMost_empty_iff2 - card_0_eq card_atLeastAtMost diff_Suc_1 finite_atLeastAtMost image_eqI inj_onD insertE nat.simps(3)) - moreover have "f' ` {1..n} \ S \ f' (Suc n) \ S" - using f f' y by (simp add: image_subset_iff) - hence "f' ` {1..Suc n} \ S" using f' by auto - ultimately show ?case unfolding inj_on_def by blast - qed - thus ?thesis using that by auto -qed - -theorem derived_alt_group_is_cons: - assumes "n \ 5" - shows "derived (alt_group n) (carrier (alt_group n)) = carrier (alt_group n)" +theorem derived_alt_group_const: + assumes "n \ 5" shows "derived (alt_group n) (carrier (alt_group n)) = carrier (alt_group n)" proof show "derived (alt_group n) (carrier (alt_group n)) \ carrier (alt_group n)" - by (simp add: alt_group_is_group group.derived_incl group.subgroup_self) + using group.derived_in_carrier[OF alt_group_is_group] by simp next - show "carrier (alt_group n) \ derived (alt_group n) (carrier (alt_group n))" - proof - - have derived_set_in_carrier: - "derived_set (alt_group n) (carrier (alt_group n)) \ carrier (alt_group n)" - proof - fix p assume "p \ derived_set (alt_group n) (carrier (alt_group n))" - then obtain q r - where q: "q \ carrier (alt_group n)" - and r: "r \ carrier (alt_group n)" - and "p = q \\<^bsub>(alt_group n)\<^esub> r \\<^bsub>(alt_group n)\<^esub> (inv' q) \\<^bsub>(alt_group n)\<^esub> (inv' r)" - using alt_group_inv_equality by auto - hence p: "p = q \ r \ (inv' q) \ (inv' r)" - by (simp add: alt_group_def sym_group_def) - - have "q permutes {1..n}" using q by (simp add: alt_group_def) - moreover have r_perm: "r permutes {1..n}" using r by (simp add: alt_group_def) - ultimately have "p permutes {1..n} \ evenperm p" using p - apply (simp add: permutes_compose permutes_inv) - by (metis evenperm_comp evenperm_inv finite_atLeastAtMost - permutation_compose permutation_inverse permutation_permutes) - thus "p \ carrier (alt_group n)" by (simp add: alt_group_def) - qed + { fix p assume "p \ three_cycles n" have "p \ derived (alt_group n) (carrier (alt_group n))" + proof - + obtain cs where cs: "p = cycle_of_list cs" "cycle cs" "length cs = 3" "set cs \ {1..n}" + using \p \ three_cycles n\ by auto + then obtain a b c where cs_def: "cs = [ a, b, c ]" + using stupid_lemma[OF cs(3)] by blast + have "card (set cs) = 3" + using cs(2-3) + by (simp add: distinct_card) - have "three_cycles n \ derived_set (alt_group n) (carrier (alt_group n))" - proof - fix p :: "nat \ nat" assume "p \ three_cycles n" - then obtain cs - where cs: "cycle cs" "length cs = 3" "set cs \ {1..n}" and p: "p = cycle_of_list cs" - unfolding three_cycles_def by blast - then obtain i j k where i: "i = cs ! 0" and j: "j = cs ! 1" and k: "k = cs ! 2" - and ijk: "cs = [i, j, k]" using stupid_lemma[OF cs(2)] by blast + have "set cs \ {1..n}" + using assms cs(3) unfolding sym[OF distinct_card[OF cs(2)]] by auto + then obtain d where d: "d \ set cs" "d \ {1..n}" + using cs(4) by blast - have "p ^^ 2 = cycle_of_list [i, k, j]" - proof - fix l show "(p ^^ 2) l = cycle_of_list [i, k, j] l" - proof (cases) - assume "l \ {i, j, k}" - hence "l \ set cs \ l \ set [i, k, j]" using ijk by auto - thus ?thesis - using id_outside_supp[of cs l] id_outside_supp[of "[i, j, k]" l] p o_apply - by (simp add: ijk numeral_2_eq_2) - next - assume "\ l \ {i, j, k}" hence "l \ {i, j, k}" by simp - have "map ((cycle_of_list cs) ^^ 2) cs = rotate 2 cs" - using cyclic_rotation[of cs 2] cs by simp - also have " ... = rotate1 (rotate1 [i, j , k])" - by (metis One_nat_def Suc_1 funpow_0 ijk rotate_Suc rotate_def) - also have " ... = [k, i, j]" by simp - finally have "map ((cycle_of_list cs) ^^ 2) cs = [k, i, j]" . - hence "map (p ^^ 2) [i, j, k] = [k, i, j]" using p ijk by simp - - moreover have "map (cycle_of_list [i, k, j]) [i, j, k] = [k, i, j]" - using cs(1) ijk by auto - - ultimately show ?thesis using \l \ {i, j, k}\ by auto - qed - qed - hence p2: "p ^^ 2 = (Fun.swap j k id) \ (cycle_of_list [i, j, k]) \ (inv' (Fun.swap j k id))" - using conjugation_of_cycle[of "[i, j, k]" "Fun.swap j k id"] cs(1) ijk by auto + hence "cycle (d # cs)" and "length (d # cs) = 4" and "card {1..n} = n" + using cs(2-3) by auto + hence "set (d # cs) \ {1..n}" + using assms unfolding sym[OF distinct_card[OF \cycle (d # cs)\]] + by (metis Suc_n_not_le_n eval_nat_numeral(3)) + then obtain e where e: "e \ set (d # cs)" "e \ {1..n}" + using d cs(4) by (metis insert_subset list.simps(15) subsetI subset_antisym) - have "card ({1..n} - {i, j, k}) \ card {1..n} - card {i, j, k}" - by (meson diff_card_le_card_Diff finite.emptyI finite.insertI) - hence card_ge_two: "card ({1..n} - {i, j, k}) \ 2" - using assms cs ijk by simp - then obtain f :: "nat \ nat" where f: "inj_on f {1..2}" "f ` {1..2} \ {1..n} - {i, j, k}" - using elts_from_card[OF card_ge_two] by blast - then obtain g h where gh: "g = f 1" "h = f 2" "g \ h" - by (metis Suc_1 atLeastAtMost_iff diff_Suc_1 diff_Suc_Suc inj_onD nat.simps(3) one_le_numeral order_refl) - hence g: "g \ {1..n} - {i, j, k}" and h: "h \ {1..n} - {i, j, k}" using f(2) gh(2) by force+ - hence gh_simps: "g \ h \ g \ {1..n} \ h \ {1..n} \ g \ {i, j, k} \ h \ {i, j, k}" - using g gh(3) by blast - moreover have ijjk: "Fun.swap i j id = Fun.swap j k id \ Fun.swap i j (Fun.swap j k id)" - and jkij: "Fun.swap j k id \ (Fun.swap i j id \ Fun.swap j k id) \ inv' (Fun.swap j k id) = Fun.swap g h (Fun.swap g h (Fun.swap i j (Fun.swap j k id)))" - by (simp_all add: comp_swap inv_swap_id) - moreover have "Fun.swap g h (Fun.swap i j id) = Fun.swap i j (Fun.swap g h id)" - by (metis (no_types) comp_id comp_swap gh_simps insert_iff swap_id_independent) - moreover have "Fun.swap i j (Fun.swap g h (Fun.swap j k id \ id)) = Fun.swap g h (Fun.swap i j (Fun.swap j k id))" - by (metis (no_types) calculation(4) comp_id comp_swap) - moreover have "inj (Fun.swap j k id)" "bij (Fun.swap g h id)" "bij (Fun.swap j k id)" - by auto - moreover have "Fun.swap j k id \ inv' (Fun.swap j k id \ Fun.swap g h id) = Fun.swap g h id" - by (metis (no_types) bij_betw_id bij_swap_iff comp_id comp_swap gh_simps insert_iff inv_swap_id o_inv_distrib swap_id_independent swap_nilpotent) - moreover have "Fun.swap j k id \ (Fun.swap j k id \ (Fun.swap j k id \ Fun.swap i j (Fun.swap j k id) \ Fun.swap j k id)) \ inv' (Fun.swap j k id) = Fun.swap j k id \ Fun.swap i j (Fun.swap j k id)" - by (simp add: comp_swap inv_swap_id) - moreover have "Fun.swap j k id \ Fun.swap i j (Fun.swap j k id) \ Fun.swap j k id = Fun.swap j k id \ (Fun.swap j k id \ (Fun.swap j k id \ Fun.swap i j (Fun.swap j k id) \ Fun.swap j k id))" - by (simp add: comp_swap inv_swap_id) - moreover have "Fun.swap g h id \ (Fun.swap j k id \ Fun.swap i j (Fun.swap j k id) \ Fun.swap j k id) \ inv' (Fun.swap j k id \ Fun.swap g h id) = Fun.swap j k id \ (Fun.swap j k id \ (Fun.swap j k id \ Fun.swap i j (Fun.swap j k id) \ Fun.swap j k id)) \ inv' (Fun.swap j k id)" - by (metis calculation(10) calculation(4) calculation(9) comp_assoc comp_id comp_swap swap_nilpotent) - ultimately have "Fun.swap i j (Fun.swap j k id) = Fun.swap j k id \ Fun.swap g h id \ (Fun.swap j k id \ Fun.swap i j (Fun.swap j k id) \ Fun.swap j k id) \ inv' (Fun.swap j k id \ Fun.swap g h id)" - by (simp add: comp_assoc) - then have final_step: - "p ^^ 2 = ((Fun.swap j k id) \ (Fun.swap g h id)) \ - (cycle_of_list [i, j, k]) \ - (inv' ((Fun.swap j k id) \ (Fun.swap g h id)))" - using ijjk jkij by (auto simp: p2) + define q where "q = (Fun.swap d e id) \ (Fun.swap b c id)" + hence "bij q" + by (simp add: bij_comp) + moreover have "q b = c" and "q c = b" + using d(1) e(1) unfolding q_def cs_def by simp+ + moreover have "q a = a" + using d(1) e(1) cs(2) unfolding q_def cs_def by auto + ultimately have "q \ p \ (inv' q) = cycle_of_list [ a, c, b ]" + using conjugation_of_cycle[OF cs(2), of q] + unfolding sym[OF cs(1)] unfolding cs_def by simp + also have " ... = p \ p" + using cs(2) unfolding cs(1) cs_def + by (auto, metis comp_id comp_swap swap_commute swap_triple) + finally have "q \ p \ (inv' q) = p \ p" . + moreover have "bij p" + unfolding cs(1) cs_def by (simp add: bij_comp) + ultimately have p: "q \ p \ (inv' q) \ (inv' p) = p" + by (simp add: bijection.intro bijection.inv_comp_right comp_assoc) - define q where "q \ (Fun.swap j k id) \ (Fun.swap g h id)" - hence "(p \ p) = q \ p \ (inv' q)" - by (metis final_step One_nat_def Suc_1 comp_id funpow.simps(2) funpow_simps_right(1) ijk p) - hence "(p \ p) \ (inv' p) = q \ p \ (inv' q) \ (inv' p)" by simp - hence 1: "p = q \ p \ (inv' q) \ (inv' p)" - using p cycle_permutes[OF cs(1)] o_assoc[of p p "inv' p"] - by (simp add: permutation_inverse_works(2)) + have "swapidseq (Suc (Suc 0)) q" + using comp_Suc[OF comp_Suc[OF id], of b c d e] e(1) cs(2) unfolding q_def cs_def by auto + hence "evenperm q" + using even_Suc_Suc_iff evenperm_unique by blast + moreover have "q permutes { d, e, b, c }" + unfolding q_def by (simp add: permutes_compose permutes_swap_id) + hence "q permutes {1..n}" + using cs(4) d(2) e(2) permutes_subset unfolding cs_def by fastforce + ultimately have "q \ carrier (alt_group n)" + unfolding alt_group_carrier by simp + moreover have "p \ carrier (alt_group n)" + using \p \ three_cycles n\ three_cycles_incl by blast + ultimately have "p \ derived_set (alt_group n) (carrier (alt_group n))" + using p alt_group_inv_equality unfolding alt_group_mult + by (metis (no_types, lifting) UN_iff singletonI) + thus "p \ derived (alt_group n) (carrier (alt_group n))" + unfolding derived_def by (rule incl) + qed } note aux_lemma = this - have "(Fun.swap j k id) \ (Fun.swap g h id) permutes {1..n}" - by (metis cs(3) gh_simps ijk insert_subset list.simps(15) permutes_compose permutes_swap_id) - moreover have "evenperm ((Fun.swap j k id) \ (Fun.swap g h id))" - by (metis cs(1) distinct_length_2_or_more evenperm_comp evenperm_swap gh(3) ijk permutation_swap_id) - ultimately have 2: "q \ carrier (alt_group n)" - by (simp add: alt_group_def q_def) - - have 3: "p \ carrier (alt_group n)" - using alt_group_as_three_cycles generate.incl[OF \p \ three_cycles n\] by simp + interpret A: group "alt_group n" + using alt_group_is_group . - from 1 2 3 show "p \ derived_set (alt_group n) (carrier (alt_group n))" - using alt_group_is_group[of n] alt_group_inv_equality[OF 2] alt_group_inv_equality[OF 3] - unfolding alt_group_def sym_group_def by fastforce - qed - hence "generate (alt_group n) (three_cycles n) \ derived (alt_group n) (carrier (alt_group n))" - unfolding derived_def - using group.mono_generate[OF alt_group_is_group[of n]] derived_set_in_carrier by simp - thus ?thesis using alt_group_as_three_cycles by blast - qed + have "generate (alt_group n) (three_cycles n) \ derived (alt_group n) (carrier (alt_group n))" + using A.generate_subgroup_incl[OF _ A.derived_is_subgroup] aux_lemma by (meson subsetI) + thus "carrier (alt_group n) \ derived (alt_group n) (carrier (alt_group n))" + using alt_group_carrier_as_three_cycles by simp qed -corollary alt_group_not_solvable: - assumes "n \ 5" - shows "\ solvable (alt_group n)" +corollary alt_group_is_unsolvable: + assumes "n \ 5" shows "\ solvable (alt_group n)" proof (rule ccontr) - assume "\ \ solvable (alt_group n)" hence "solvable (alt_group n)" by simp - then obtain k - where trivial_seq: "(derived (alt_group n) ^^ k) (carrier (alt_group n)) = { \\<^bsub>alt_group n\<^esub> }" - using group.solvable_iff_trivial_derived_seq[OF alt_group_is_group[of n]] by blast - - have "(derived (alt_group n) ^^ k) (carrier (alt_group n)) = (carrier (alt_group n))" - apply (induction k) using derived_alt_group_is_cons[OF assms] by auto - hence "carrier (alt_group n) = { \\<^bsub>alt_group n\<^esub> }" - using trivial_seq by auto - hence singleton: "carrier (alt_group n) = { id }" - by (simp add: alt_group_def sym_group_def) - - have "set [1 :: nat, 2, 3] \ {1..n}" using assms by auto - moreover have "cycle [1 :: nat, 2, 3]" by simp - moreover have "length [1 :: nat, 2, 3] = 3" by simp - ultimately have "cycle_of_list [1 :: nat, 2, 3] \ three_cycles n" - unfolding three_cycles_def by blast - hence "cycle_of_list [1 :: nat, 2, 3] \ carrier (alt_group n)" - using alt_group_as_three_cycles by (simp add: generate.incl) - - moreover have "map (cycle_of_list [1 :: nat, 2, 3]) [1 :: nat, 2, 3] = [2 :: nat, 3, 1]" - using cyclic_rotation[OF \cycle [1 :: nat, 2, 3]\, of 1] by simp - hence "cycle_of_list [1 :: nat, 2, 3] \ id" - by (metis list.map_id list.sel(1) numeral_One numeral_eq_iff semiring_norm(85)) - - ultimately show False using singleton by blast + assume "\ \ solvable (alt_group n)" + then obtain m where "(derived (alt_group n) ^^ m) (carrier (alt_group n)) = { id }" + using group.solvable_iff_trivial_derived_seq[OF alt_group_is_group] unfolding alt_group_one by blast + moreover have "(derived (alt_group n) ^^ m) (carrier (alt_group n)) = carrier (alt_group n)" + using derived_alt_group_const[OF assms] by (induct m) (auto) + ultimately have card_eq_1: "card (carrier (alt_group n)) = 1" + by simp + have ge_2: "n \ 2" + using assms by simp + moreover have "2 = fact n" + using alt_group_card_carrier[OF ge_2] unfolding card_eq_1 + by (metis fact_2 mult.right_neutral of_nat_fact) + ultimately have "n = 2" + by (metis antisym_conv fact_ge_self) + thus False + using assms by simp qed -corollary sym_group_not_solvable: - assumes "n \ 5" - shows "\ solvable (sym_group n)" +corollary sym_group_is_unsolvable: + assumes "n \ 5" shows "\ solvable (sym_group n)" proof - - have "subgroup (kernel (sym_group n) sign_img sign) (sym_group n)" - using group_hom.subgroup_kernel sign_is_hom by blast - hence "subgroup (carrier (alt_group n)) (sym_group n)" - using alt_group_is_kernel_from_sign[of n] by simp - hence "group_hom (alt_group n) (sym_group n) id" - using group.canonical_inj_is_hom[OF sym_group_is_group[of n]] by (simp add: alt_group_def) - thus ?thesis - using group_hom.not_solvable[of "alt_group n" "sym_group n" id] - alt_group_not_solvable[OF assms] inj_on_id by blast + interpret Id: group_hom "alt_group n" "sym_group n" id + using group.canonical_inj_is_hom[OF sym_group_is_group alt_group_is_subgroup] alt_group_def by simp + show ?thesis + using Id.inj_hom_imp_solvable alt_group_is_unsolvable[OF assms] by auto qed -end +end \ No newline at end of file diff -r 12dce58bcd3f -r 1b5178abaf97 src/HOL/Algebra/Weak_Morphisms.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Algebra/Weak_Morphisms.thy Thu Oct 04 15:25:47 2018 +0100 @@ -0,0 +1,498 @@ +(* Title: HOL/Algebra/Weak_Morphisms.thy + Author: Paulo Emílio de Vilhena +*) + +theory Weak_Morphisms + imports QuotRing + +begin + +section \Weak Morphisms\ + +text \The definition of ring isomorphism, as well as the definition of group isomorphism, doesn't + enforce any algebraic constraint to the structure of the schemes involved. This seems + unnatural, but it turns out to be very useful: in order to prove that a scheme B satisfy + certain algebraic constraints, it's sufficient to prove those for a scheme A and show + the existence of an isomorphism between A and B. In this section, we explore this idea + in a different way: given a scheme A and a function f, we build a scheme B with an + algebraic structure of same strength as A where f is an homomorphism from A to B.\ + + +subsection \Definitions\ + +locale weak_group_morphism = normal H G for f and H and G (structure) + + assumes inj_mod_subgroup: "\ a \ carrier G; b \ carrier G \ \ f a = f b \ a \ (inv b) \ H" + +locale weak_ring_morphism = ideal I R for f and I and R (structure) + + assumes inj_mod_ideal: "\ a \ carrier R; b \ carrier R \ \ f a = f b \ a \ b \ I" + + +definition image_group :: "('a \ 'b) \ ('a, 'c) monoid_scheme \ 'b monoid" + where "image_group f G \ + \ carrier = f ` (carrier G), + mult = (\a b. f ((inv_into (carrier G) f a) \\<^bsub>G\<^esub> (inv_into (carrier G) f b))), + one = f \\<^bsub>G\<^esub> \" + +definition image_ring :: "('a \ 'b) \ ('a, 'c) ring_scheme \ 'b ring" + where "image_ring f R \ monoid.extend (image_group f R) + \ zero = f \\<^bsub>R\<^esub>, + add = (\a b. f ((inv_into (carrier R) f a) \\<^bsub>R\<^esub> (inv_into (carrier R) f b))) \" + + +subsection \Weak Group Morphisms\ + +lemma image_group_carrier: "carrier (image_group f G) = f ` (carrier G)" + unfolding image_group_def by simp + +lemma image_group_one: "one (image_group f G) = f \\<^bsub>G\<^esub>" + unfolding image_group_def by simp + +lemma weak_group_morphismsI: + assumes "H \ G" and "\a b. \ a \ carrier G; b \ carrier G \ \ f a = f b \ a \\<^bsub>G\<^esub> (inv\<^bsub>G\<^esub> b) \ H" + shows "weak_group_morphism f H G" + using assms unfolding weak_group_morphism_def weak_group_morphism_axioms_def by auto + +lemma image_group_truncate: + fixes R :: "('a, 'b) monoid_scheme" + shows "monoid.truncate (image_group f R) = image_group f (monoid.truncate R)" + by (simp add: image_group_def monoid.defs) + +lemma image_ring_truncate: "monoid.truncate (image_ring f R) = image_group f R" + by (simp add: image_ring_def monoid.defs) + +lemma (in ring) weak_add_group_morphism: + assumes "weak_ring_morphism f I R" shows "weak_group_morphism f I (add_monoid R)" +proof - + have is_normal: "I \ (add_monoid R)" + using ideal_is_normal[OF weak_ring_morphism.axioms(1)[OF assms]] . + show ?thesis + using weak_group_morphism.intro[OF is_normal] + weak_ring_morphism.inj_mod_ideal[OF assms] + unfolding weak_group_morphism_axioms_def a_minus_def a_inv_def + by auto +qed + +lemma (in group) weak_group_morphism_range: + assumes "weak_group_morphism f H G" and "a \ carrier G" shows "f ` (H #> a) = { f a }" +proof - + interpret H: subgroup H G + using weak_group_morphism.axioms(1)[OF assms(1)] unfolding normal_def by simp + show ?thesis + proof + show "{ f a } \ f ` (H #> a)" + using H.one_closed assms(2) unfolding r_coset_def by force + next + show "f ` (H #> a) \ { f a }" + proof + fix b assume "b \ f ` (H #> a)" then obtain h where "h \ H" "h \ carrier G" "b = f (h \ a)" + unfolding r_coset_def using H.subset by auto + thus "b \ { f a }" + using weak_group_morphism.inj_mod_subgroup[OF assms(1)] assms(2) + by (metis inv_solve_right m_closed singleton_iff) + qed + qed +qed + +lemma (in group) vimage_eq_rcoset: + assumes "weak_group_morphism f H G" and "a \ carrier G" + shows "{ b \ carrier G. f b = f a } = H #> a" and "{ b \ carrier G. f b = f a } = a <# H" +proof - + interpret H: normal H G + using weak_group_morphism.axioms(1)[OF assms(1)] by simp + show "{ b \ carrier G. f b = f a } = H #> a" + proof + show "H #> a \ { b \ carrier G. f b = f a }" + using r_coset_subset_G[OF H.subset assms(2)] weak_group_morphism_range[OF assms] by auto + next + show "{ b \ carrier G. f b = f a } \ H #> a" + proof + fix b assume b: "b \ { b \ carrier G. f b = f a }" then obtain h where "h \ H" "b \ (inv a) = h" + using weak_group_morphism.inj_mod_subgroup[OF assms(1)] assms(2) by fastforce + thus "b \ H #> a" + using H.rcos_module[OF is_group] b assms(2) by blast + qed + qed + thus "{ b \ carrier G. f b = f a } = a <# H" + by (simp add: assms(2) H.coset_eq) +qed + +lemma (in group) weak_group_morphism_ker: + assumes "weak_group_morphism f H G" shows "kernel G (image_group f G) f = H" + using vimage_eq_rcoset(1)[OF assms one_closed] weak_group_morphism.axioms(1)[OF assms(1)] + by (simp add: image_group_def kernel_def normal_def subgroup.subset) + +lemma (in group) weak_group_morphism_inv_into: + assumes "weak_group_morphism f H G" and "a \ carrier G" + obtains h h' where "h \ H" "inv_into (carrier G) f (f a) = h \ a" + and "h' \ H" "inv_into (carrier G) f (f a) = a \ h'" +proof - + have "inv_into (carrier G) f (f a) \ { b \ carrier G. f b = f a }" + using assms(2) by (auto simp add: inv_into_into f_inv_into_f) + thus thesis + using that vimage_eq_rcoset[OF assms] unfolding r_coset_def l_coset_def by blast +qed + +proposition (in group) weak_group_morphism_is_iso: + assumes "weak_group_morphism f H G" shows "(\x. the_elem (f ` x)) \ iso (G Mod H) (image_group f G)" +proof (auto simp add: iso_def hom_def image_group_def) + interpret H: normal H G + using weak_group_morphism.axioms(1)[OF assms] . + + show "\x. x \ carrier (G Mod H) \ the_elem (f ` x) \ f ` carrier G" + unfolding FactGroup_def RCOSETS_def using weak_group_morphism_range[OF assms] by auto + + thus "bij_betw (\x. the_elem (f ` x)) (carrier (G Mod H)) (f ` carrier G)" + unfolding bij_betw_def + proof (auto) + fix a assume "a \ carrier G" + hence "the_elem (f ` (H #> a)) = f a" and "H #> a \ carrier (G Mod H)" + using weak_group_morphism_range[OF assms] unfolding FactGroup_def RCOSETS_def by auto + thus "f a \ (\x. the_elem (f ` x)) ` carrier (G Mod H)" + using image_iff by fastforce + next + show "inj_on (\x. the_elem (f ` x)) (carrier (G Mod H))" + proof (rule inj_onI) + fix x y assume "x \ (carrier (G Mod H))" and "y \ (carrier (G Mod H))" + then obtain a b where a: "a \ carrier G" "x = H #> a" and b: "b \ carrier G" "y = H #> b" + unfolding FactGroup_def RCOSETS_def by auto + assume "the_elem (f ` x) = the_elem (f ` y)" + hence "a \ (inv b) \ H" + using weak_group_morphism.inj_mod_subgroup[OF assms] + weak_group_morphism_range[OF assms] a b by auto + thus "x = y" + using a(1) b(1) unfolding a b + by (meson H.rcos_const H.subset group.coset_mult_inv1 is_group) + qed + qed + + fix x y assume "x \ carrier (G Mod H)" "y \ carrier (G Mod H)" + then obtain a b where a: "a \ carrier G" "x = H #> a" and b: "b \ carrier G" "y = H #> b" + unfolding FactGroup_def RCOSETS_def by auto + + show "the_elem (f ` (x <#> y)) = f (inv_into (carrier G) f (the_elem (f ` x)) \ + inv_into (carrier G) f (the_elem (f ` y)))" + proof (simp add: weak_group_morphism_range[OF assms] a b) + obtain h1 h2 + where h1: "h1 \ H" "inv_into (carrier G) f (f a) = a \ h1" + and h2: "h2 \ H" "inv_into (carrier G) f (f b) = h2 \ b" + using weak_group_morphism_inv_into[OF assms] a(1) b(1) by metis + have "the_elem (f ` ((H #> a) <#> (H #> b))) = the_elem (f ` (H #> (a \ b)))" + by (simp add: a b H.rcos_sum) + hence "the_elem (f ` ((H #> a) <#> (H #> b))) = f (a \ b)" + using weak_group_morphism_range[OF assms] a(1) b(1) by auto + moreover + have "(a \ h1) \ (h2 \ b) = a \ (h1 \ h2 \ b)" + by (simp add: a(1) b(1) h1(1) h2(1) H.subset m_assoc) + hence "(a \ h1) \ (h2 \ b) \ a <# (H #> b)" + using h1(1) h2(1) unfolding l_coset_def r_coset_def by auto + hence "(a \ h1) \ (h2 \ b) \ (a \ b) <# H" + by (simp add: H.subset H.coset_eq a(1) b(1) lcos_m_assoc) + hence "f (inv_into (carrier G) f (f a) \ inv_into (carrier G) f (f b)) = f (a \ b)" + using vimage_eq_rcoset(2)[OF assms] a(1) b(1) unfolding h1 h2 by auto + ultimately + show "the_elem (f ` ((H #> a) <#> (H #> b))) = f (inv_into (carrier G) f (f a) \ + inv_into (carrier G) f (f b))" + by simp + qed +qed + +corollary (in group) image_group_is_group: + assumes "weak_group_morphism f H G" shows "group (image_group f G)" +proof - + interpret H: normal H G + using weak_group_morphism.axioms(1)[OF assms] . + + have "group ((image_group f G) \ one := the_elem (f ` \\<^bsub>G Mod H\<^esub>) \)" + using group.iso_imp_img_group[OF H.factorgroup_is_group weak_group_morphism_is_iso[OF assms]] . + moreover have "\\<^bsub>G Mod H\<^esub> = H #> \" + unfolding FactGroup_def using H.subset by force + hence "the_elem (f ` \\<^bsub>G Mod H\<^esub>) = f \" + using weak_group_morphism_range[OF assms one_closed] by simp + ultimately show ?thesis by (simp add: image_group_def) +qed + +corollary (in group) weak_group_morphism_is_hom: + assumes "weak_group_morphism f H G" shows "f \ hom G (image_group f G)" +proof - + interpret H: normal H G + using weak_group_morphism.axioms(1)[OF assms] . + + have the_elem_hom: "(\x. the_elem (f ` x)) \ hom (G Mod H) (image_group f G)" + using weak_group_morphism_is_iso[OF assms] by (simp add: iso_def) + have hom: "(\x. the_elem (f ` x)) \ (#>) H \ hom G (image_group f G)" + using hom_trans[OF H.r_coset_hom_Mod the_elem_hom] by simp + have restrict: "\a. a \ carrier G \ ((\x. the_elem (f ` x)) \ (#>) H) a = f a" + using weak_group_morphism_range[OF assms] by auto + show ?thesis + using hom_restrict[OF hom restrict] by simp +qed + +corollary (in group) weak_group_morphism_group_hom: + assumes "weak_group_morphism f H G" shows "group_hom G (image_group f G) f" + using image_group_is_group[OF assms] weak_group_morphism_is_hom[OF assms] group_axioms + unfolding group_hom_def group_hom_axioms_def by simp + + +subsection \Weak Ring Morphisms\ + +lemma image_ring_carrier: "carrier (image_ring f R) = f ` (carrier R)" + unfolding image_ring_def image_group_def by (simp add: monoid.defs) + +lemma image_ring_one: "one (image_ring f R) = f \\<^bsub>R\<^esub>" + unfolding image_ring_def image_group_def by (simp add: monoid.defs) + +lemma image_ring_zero: "zero (image_ring f R) = f \\<^bsub>R\<^esub>" + unfolding image_ring_def image_group_def by (simp add: monoid.defs) + +lemma weak_ring_morphismI: + assumes "ideal I R" and "\a b. \ a \ carrier R; b \ carrier R \ \ f a = f b \ a \\<^bsub>R\<^esub> b \ I" + shows "weak_ring_morphism f I R" + using assms unfolding weak_ring_morphism_def weak_ring_morphism_axioms_def by auto + +lemma (in ring) weak_ring_morphism_range: + assumes "weak_ring_morphism f I R" and "a \ carrier R" shows "f ` (I +> a) = { f a }" + using add.weak_group_morphism_range[OF weak_add_group_morphism[OF assms(1)] assms(2)] + unfolding a_r_coset_def . + +lemma (in ring) vimage_eq_a_rcoset: + assumes "weak_ring_morphism f I R" and "a \ carrier R" shows "{ b \ carrier R. f b = f a } = I +> a" + using add.vimage_eq_rcoset[OF weak_add_group_morphism[OF assms(1)] assms(2)] + unfolding a_r_coset_def by simp + +lemma (in ring) weak_ring_morphism_ker: + assumes "weak_ring_morphism f I R" shows "a_kernel R (image_ring f R) f = I" + using add.weak_group_morphism_ker[OF weak_add_group_morphism[OF assms]] + unfolding kernel_def a_kernel_def' image_ring_def image_group_def by (simp add: monoid.defs) + +lemma (in ring) weak_ring_morphism_inv_into: + assumes "weak_ring_morphism f I R" and "a \ carrier R" + obtains i where "i \ I" "inv_into (carrier R) f (f a) = i \ a" + using add.weak_group_morphism_inv_into(1)[OF weak_add_group_morphism[OF assms(1)] assms(2)] by auto + +proposition (in ring) weak_ring_morphism_is_iso: + assumes "weak_ring_morphism f I R" shows "(\x. the_elem (f ` x)) \ ring_iso (R Quot I) (image_ring f R)" +proof (rule ring_iso_memI) + show "bij_betw (\x. the_elem (f ` x)) (carrier (R Quot I)) (carrier (image_ring f R))" + and add_hom: "\x y. \ x \ carrier (R Quot I); y \ carrier (R Quot I) \ \ + the_elem (f ` (x \\<^bsub>R Quot I\<^esub> y)) = the_elem (f ` x) \\<^bsub>image_ring f R\<^esub> the_elem (f ` y)" + using add.weak_group_morphism_is_iso[OF weak_add_group_morphism[OF assms]] + unfolding iso_def hom_def FactGroup_def FactRing_def A_RCOSETS_def set_add_def + by (auto simp add: image_ring_def image_group_def monoid.defs) +next + interpret I: ideal I R + using weak_ring_morphism.axioms(1)[OF assms] . + + show "the_elem (f ` \\<^bsub>R Quot I\<^esub>) = \\<^bsub>image_ring f R\<^esub>" + and "\x. x \ carrier (R Quot I) \ the_elem (f ` x) \ carrier (image_ring f R)" + using weak_ring_morphism_range[OF assms] one_closed I.Icarr + by (auto simp add: image_ring_def image_group_def monoid.defs FactRing_def A_RCOSETS_def') + + fix x y assume "x \ carrier (R Quot I)" "y \ carrier (R Quot I)" + then obtain a b where a: "a \ carrier R" "x = I +> a" and b: "b \ carrier R" "y = I +> b" + unfolding FactRing_def A_RCOSETS_def' by auto + hence prod: "x \\<^bsub>R Quot I\<^esub> y = I +> (a \ b)" + unfolding FactRing_def by (simp add: I.rcoset_mult_add) + + show "the_elem (f ` (x \\<^bsub>R Quot I\<^esub> y)) = the_elem (f ` x) \\<^bsub>image_ring f R\<^esub> the_elem (f ` y)" + unfolding prod + proof (simp add: weak_ring_morphism_range[OF assms] a b image_ring_def image_group_def monoid.defs) + obtain i j + where i: "i \ I" "inv_into (carrier R) f (f a) = i \ a" + and j: "j \ I" "inv_into (carrier R) f (f b) = j \ b" + using weak_ring_morphism_inv_into[OF assms] a(1) b(1) by metis + have "i \ carrier R" and "j \ carrier R" + using I.Icarr i(1) j(1) by auto + hence "(i \ a) \ (j \ b) = (i \ a) \ j \ (i \ b) \ (a \ b)" + using a(1) b(1) by algebra + hence "(i \ a) \ (j \ b) \ I +> (a \ b)" + using i(1) j(1) a(1) b(1) unfolding a_r_coset_def' + by (simp add: I.I_l_closed I.I_r_closed) + thus "f (a \ b) = f (inv_into (carrier R) f (f a) \ inv_into (carrier R) f (f b))" + unfolding i j using weak_ring_morphism_range[OF assms m_closed[OF a(1) b(1)]] + by (metis imageI singletonD) + qed +qed + +corollary (in ring) image_ring_zero': + assumes "weak_ring_morphism f I R" shows "the_elem (f ` \\<^bsub>R Quot I\<^esub>) = \\<^bsub>image_ring f R\<^esub>" +proof - + interpret I: ideal I R + using weak_ring_morphism.axioms(1)[OF assms] . + + have "\\<^bsub>R Quot I\<^esub> = I +> \" + unfolding FactRing_def a_r_coset_def' by force + thus ?thesis + using weak_ring_morphism_range[OF assms zero_closed] unfolding image_ring_zero by simp +qed + +corollary (in ring) image_ring_is_ring: + assumes "weak_ring_morphism f I R" shows "ring (image_ring f R)" +proof - + interpret I: ideal I R + using weak_ring_morphism.axioms(1)[OF assms] . + + have "ring ((image_ring f R) \ zero := the_elem (f ` \\<^bsub>R Quot I\<^esub>) \)" + using ring.ring_iso_imp_img_ring[OF I.quotient_is_ring weak_ring_morphism_is_iso[OF assms]] by simp + thus ?thesis + unfolding image_ring_zero'[OF assms] by simp +qed + +corollary (in ring) image_ring_is_field: + assumes "weak_ring_morphism f I R" and "field (R Quot I)" shows "field (image_ring f R)" + using field.ring_iso_imp_img_field[OF assms(2) weak_ring_morphism_is_iso[OF assms(1)]] + unfolding image_ring_zero'[OF assms(1)] by simp + +corollary (in ring) weak_ring_morphism_is_hom: + assumes "weak_ring_morphism f I R" shows "f \ ring_hom R (image_ring f R)" +proof - + interpret I: ideal I R + using weak_ring_morphism.axioms(1)[OF assms] . + + have the_elem_hom: "(\x. the_elem (f ` x)) \ ring_hom (R Quot I) (image_ring f R)" + using weak_ring_morphism_is_iso[OF assms] by (simp add: ring_iso_def) + have ring_hom: "(\x. the_elem (f ` x)) \ (+>) I \ ring_hom R (image_ring f R)" + using ring_hom_trans[OF I.rcos_ring_hom the_elem_hom] . + have restrict: "\a. a \ carrier R \ ((\x. the_elem (f ` x)) \ (+>) I) a = f a" + using weak_ring_morphism_range[OF assms] by auto + show ?thesis + using ring_hom_restrict[OF ring_hom restrict] by simp +qed + +corollary (in ring) weak_ring_morphism_ring_hom: + assumes "weak_ring_morphism f I R" shows "ring_hom_ring R (image_ring f R) f" + using ring_hom_ringI2[OF ring_axioms image_ring_is_ring[OF assms] weak_ring_morphism_is_hom[OF assms]] . + + +subsection \Injective Functions\ + +text \If the fuction is injective, we don't need to impose any algebraic restriction to the input + scheme in order to state an isomorphism.\ + +lemma inj_imp_image_group_iso: + assumes "inj_on f (carrier G)" shows "f \ iso G (image_group f G)" + using assms by (auto simp add: image_group_def iso_def bij_betw_def hom_def) + +lemma inj_imp_image_group_inv_iso: + assumes "inj f" shows "Hilbert_Choice.inv f \ iso (image_group f G) G" + using assms by (auto simp add: image_group_def iso_def bij_betw_def hom_def inj_on_def) + +lemma inj_imp_image_ring_iso: + assumes "inj_on f (carrier R)" shows "f \ ring_iso R (image_ring f R)" + using assms by (auto simp add: image_ring_def image_group_def ring_iso_def + bij_betw_def ring_hom_def monoid.defs) + +lemma inj_imp_image_ring_inv_iso: + assumes "inj f" shows "Hilbert_Choice.inv f \ ring_iso (image_ring f R) R" + using assms by (auto simp add: image_ring_def image_group_def ring_iso_def + bij_betw_def ring_hom_def inj_on_def monoid.defs) + +lemma (in group) inj_imp_image_group_is_group: + assumes "inj_on f (carrier G)" shows "group (image_group f G)" + using iso_imp_img_group[OF inj_imp_image_group_iso[OF assms]] by (simp add: image_group_def) + +lemma (in ring) inj_imp_image_ring_is_ring: + assumes "inj_on f (carrier R)" shows "ring (image_ring f R)" + using ring_iso_imp_img_ring[OF inj_imp_image_ring_iso[OF assms]] + by (simp add: image_ring_def image_group_def monoid.defs) + +lemma (in domain) inj_imp_image_ring_is_domain: + assumes "inj_on f (carrier R)" shows "domain (image_ring f R)" + using ring_iso_imp_img_domain[OF inj_imp_image_ring_iso[OF assms]] + by (simp add: image_ring_def image_group_def monoid.defs) + +lemma (in field) inj_imp_image_ring_is_field: + assumes "inj_on f (carrier R)" shows "field (image_ring f R)" + using ring_iso_imp_img_field[OF inj_imp_image_ring_iso[OF assms]] + by (simp add: image_ring_def image_group_def monoid.defs) + + +section \Examples\ + +text \In a lot of different contexts, the lack of dependent types make some definitions quite + complicated. The tools developed in this theory give us a way to change the type of a + scheme and preserve all of its algebraic properties. We show, in this section, how to + make use of this feature in order to solve the problem mentioned above. \ + + +subsection \Direct Product\ + +abbreviation nil_monoid :: "('a list) monoid" + where "nil_monoid \ \ carrier = { [] }, mult = (\a b. []), one = [] \" + +definition DirProd_list :: "(('a, 'b) monoid_scheme) list \ ('a list) monoid" + where "DirProd_list Gs = foldr (\G H. image_group (\(x, xs). x # xs) (G \\ H)) Gs nil_monoid" + + +subsubsection \Basic Properties\ + +lemma DirProd_list_carrier: + shows "carrier (DirProd_list (G # Gs)) = (\(x, xs). x # xs) ` (carrier G \ carrier (DirProd_list Gs))" + unfolding DirProd_list_def image_group_def by auto + +lemma DirProd_list_one: + shows "one (DirProd_list Gs) = foldr (\G tl. (one G) # tl) Gs []" + unfolding DirProd_list_def DirProd_def image_group_def by (induct Gs) (auto) + +lemma DirProd_list_carrier_mem: + assumes "gs \ carrier (DirProd_list Gs)" + shows "length gs = length Gs" and "\i. i < length Gs \ (gs ! i) \ carrier (Gs ! i)" +proof - + let ?same_length = "\xs ys. length xs = length ys" + let ?in_carrier = "\i gs Gs. (gs ! i) \ carrier (Gs ! i)" + from assms have "?same_length gs Gs \ (\i < length Gs. ?in_carrier i gs Gs)" + proof (induct Gs arbitrary: gs, simp add: DirProd_list_def) + case (Cons G Gs) + then obtain g' gs' + where g': "g' \ carrier G" and gs': "gs' \ carrier (DirProd_list Gs)" and gs: "gs = g' # gs'" + unfolding DirProd_list_carrier by auto + hence "?same_length gs (G # Gs)" and "\i. i \ {(Suc 0)..< length (G # Gs)} \ ?in_carrier i gs (G # Gs)" + using Cons(1) by auto + moreover have "?in_carrier 0 gs (G # Gs)" + unfolding gs using g' by simp + ultimately show ?case + by (metis atLeastLessThan_iff eq_imp_le less_Suc0 linorder_neqE_nat nat_less_le) + qed + thus "?same_length gs Gs" and "\i. i < length Gs \ ?in_carrier i gs Gs" + by simp+ +qed + +lemma DirProd_list_carrier_memI: + assumes "length gs = length Gs" and "\i. i < length Gs \ (gs ! i) \ carrier (Gs ! i)" + shows "gs \ carrier (DirProd_list Gs)" + using assms +proof (induct Gs arbitrary: gs, simp add: DirProd_list_def) + case (Cons G Gs) + then obtain g' gs' where gs: "gs = g' # gs'" + by (metis length_Suc_conv) + have "g' \ carrier G" + using Cons(3)[of 0] unfolding gs by auto + moreover have "gs' \ carrier (DirProd_list Gs)" + using Cons unfolding gs by force + ultimately show ?case + unfolding DirProd_list_carrier gs by blast +qed + +lemma inj_on_DirProd_carrier: + shows "inj_on (\(g, gs). g # gs) (carrier (G \\ (DirProd_list Gs)))" + unfolding DirProd_def inj_on_def by auto + +lemma DirProd_list_is_group: + assumes "\i. i < length Gs \ group (Gs ! i)" shows "group (DirProd_list Gs)" + using assms +proof (induct Gs) + case Nil thus ?case + unfolding DirProd_list_def by (unfold_locales, auto simp add: Units_def) +next + case (Cons G Gs) + hence is_group: "group (G \\ (DirProd_list Gs))" + using DirProd_group[of G "DirProd_list Gs"] by force + show ?case + using group.inj_imp_image_group_is_group[OF is_group inj_on_DirProd_carrier] + unfolding DirProd_list_def by auto +qed + +lemma DirProd_list_iso: + "(\(g, gs). g # gs) \ iso (G \\ (DirProd_list Gs)) (DirProd_list (G # Gs))" + using inj_imp_image_group_iso[OF inj_on_DirProd_carrier] unfolding DirProd_list_def by auto + +end \ No newline at end of file