--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Algebra/Algebra.thy Mon Jul 02 15:43:22 2018 +0100
@@ -0,0 +1,8 @@
+(* Title: HOL/Algebra/Algebra *)
+
+theory Algebra
+ imports Sylow Chinese_Remainder Zassenhaus Galois_Connection Generated_Fields
+ Divisibility Embedded_Algebras IntRing Sym_Groups
+(*Frobenius Exact_Sequence Polynomials*)
+begin
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Algebra/Chinese_Remainder.thy Mon Jul 02 15:43:22 2018 +0100
@@ -0,0 +1,1207 @@
+(* ************************************************************************** *)
+(* Title: Chinese_Remainder.thy *)
+(* Author: Paulo EmÃlio de Vilhena *)
+(* ************************************************************************** *)
+
+theory Chinese_Remainder
+ imports QuotRing Ideal_Product
+
+begin
+
+section \<open>Chinese Remainder Theorem\<close>
+
+subsection \<open>Direct Product of Rings\<close>
+
+definition
+ RDirProd :: "[ ('a, 'n) ring_scheme, ('b, 'm) ring_scheme ] \<Rightarrow> ('a \<times> 'b) ring"
+ where "RDirProd R S =
+ \<lparr> carrier = carrier R \<times> carrier S,
+ mult = (\<lambda>(r, s). \<lambda>(r', s'). (r \<otimes>\<^bsub>R\<^esub> r', s \<otimes>\<^bsub>S\<^esub> s')),
+ one = (\<one>\<^bsub>R\<^esub>, \<one>\<^bsub>S\<^esub>),
+ zero = (\<zero>\<^bsub>R\<^esub>, \<zero>\<^bsub>S\<^esub>),
+ add = (\<lambda>(r, s). \<lambda>(r', s'). (r \<oplus>\<^bsub>R\<^esub> r', s \<oplus>\<^bsub>S\<^esub> s')) \<rparr>"
+
+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)
+
+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)"
+proof -
+ 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)
+qed
+
+lemma RDirProd_isomorphism1:
+ "(\<lambda>(x, y). (y, x)) \<in> 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
+
+lemma RDirProd_isomorphism2:
+ "(\<lambda>(x, (y, z)). ((x, y), z)) \<in> 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)
+
+lemma RDirProd_isomorphism3:
+ "(\<lambda>((x, y), z). (x, (y, z))) \<in> 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)
+
+lemma RDirProd_isomorphism4:
+ assumes "f \<in> ring_iso R S"
+ shows "(\<lambda>(r, t). (f r, t)) \<in> 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_isomorphism5:
+ assumes "f \<in> ring_iso S T"
+ shows "(\<lambda>(r, s). (r, f s)) \<in> 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]]
+ by (simp add: case_prod_unfold case_prod_unfold comp_def comp_def)
+
+lemma RDirProd_isomorphism6:
+ assumes "f \<in> ring_iso R R'" and "g \<in> ring_iso S S'"
+ shows "(\<lambda>(r, s). (f r, g s)) \<in> ring_iso (RDirProd R S) (RDirProd R' S')"
+ using ring_iso_set_trans[OF RDirProd_isomorphism4[OF assms(1)] RDirProd_isomorphism5[OF assms(2)]]
+ by (simp add: case_prod_beta' comp_def)
+
+
+subsection \<open>Simple Version of The Theorem\<close>
+
+text \<open>We start by proving a simpler version of the theorem. The rest of the theory is
+ dedicated to its generalization\<close>
+
+lemma (in ideal) set_add_zero:
+ assumes "i \<in> I"
+ shows "I +> i = I"
+ by (simp add: a_rcos_const assms)
+
+lemma (in ideal) set_add_zero_imp_mem:
+ assumes "i \<in> carrier R" "I +> i = I"
+ shows "i \<in> I"
+ using a_rcos_self assms(1-2) by auto
+
+lemma (in ring) canonical_proj_is_surj:
+ assumes "ideal I R" "ideal J R" "I <+> J = carrier R"
+ shows "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow>
+ \<exists>a \<in> carrier R. I +> a = I +> x \<and> J +> a = J +> y"
+proof -
+ { fix I J i j assume A: "ideal I R" "ideal J R" "i \<in> I" "j \<in> J" "\<one> = i \<oplus> j"
+ have "I +> \<one> = I +> j"
+ proof -
+ have "I +> \<one> = I +> (i \<oplus> 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 +> \<zero>) <+> (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 +> (\<zero> \<oplus> 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 +> \<one> = I +> j"
+ using A(2) A(4) ideal.Icarr by fastforce
+ qed } note aux_lemma = this
+
+ fix x y assume x: "x \<in> carrier R" and y: "y \<in> carrier R"
+
+ have "\<one> \<in> I <+> J" using assms by simp
+ then obtain i j where i: "i \<in> I" and j: "j \<in> J" and ij: "\<one> = i \<oplus> j"
+ using set_add_def'[of R I J] by auto
+ have mod_I: "I +> j = I +> \<one>" and mod_J: "J +> i = J +> \<one>"
+ 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 \<otimes> x) \<oplus> (i \<otimes> y)) = (I +> (j \<otimes> x)) <+> (I +> (i \<otimes> 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 \<otimes> x)) <+> (I +> \<zero>)"
+ using ideal.set_add_zero[OF assms(1), of "i \<otimes> y"] i assms(1)
+ by (simp add: additive_subgroup.a_subset ideal.I_r_closed ideal.axioms(1) y)
+ also have " ... = I +> (j \<otimes> 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 \<otimes> x) \<oplus> (i \<otimes> y)) = I +> x" using mod_I
+ by (smt additive_subgroup.a_Hcarr assms ideal.axioms(1) ideal.rcoset_mult_add j
+ monoid.l_one monoid_axioms one_closed x)
+
+ have "J +> ((j \<otimes> x) \<oplus> (i \<otimes> y)) = (J +> (j \<otimes> x)) <+> (J +> (i \<otimes> 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 +> \<zero>) <+> (J +> (i \<otimes> y))"
+ using ideal.set_add_zero[OF assms(2), of "j \<otimes> x"] j assms(2)
+ by (simp add: additive_subgroup.a_subset ideal.I_r_closed ideal.axioms(1) x)
+ also have " ... = J +> (i \<otimes> 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 \<otimes> x) \<oplus> (i \<otimes> y)) = J +> y" using mod_J
+ by (smt additive_subgroup.a_Hcarr assms ideal.axioms(1) ideal.rcoset_mult_add i j
+ monoid.l_one monoid_axioms one_closed y x)
+
+ have "(j \<otimes> x) \<oplus> (i \<otimes> y) \<in> carrier R"
+ by (meson x y i j assms add.m_closed additive_subgroup.a_Hcarr ideal.axioms(1) m_closed)
+ thus "\<exists>a \<in> carrier R. I +> a = I +> x \<and> J +> a = J +> y" using Ix Jy by blast
+qed
+
+lemma (in ring) canonical_proj_is_hom:
+ assumes "ideal I R" "ideal J R" "I <+> J = carrier R"
+ shows "(\<lambda>a. (I +> a, J +> a)) \<in> ring_hom R (RDirProd (R Quot I) (R Quot J))"
+proof (rule ring_hom_memI)
+ fix x y assume x: "x \<in> carrier R" and y: "y \<in> carrier R"
+ show "(I +> x, J +> x) \<in> 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 \<otimes> y, J +> x \<otimes> y) =
+ (I +> x, J +> x) \<otimes>\<^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 \<oplus> y, J +> x \<oplus> y) =
+ (I +> x, J +> x) \<oplus>\<^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 +> \<one>, J +> \<one>) = \<one>\<^bsub>RDirProd (R Quot I) (R Quot J)\<^esub>"
+ unfolding RDirProd_def FactRing_def by simp
+qed
+
+theorem (in ring) chinese_remainder_simple:
+ assumes "ideal I R" "ideal J R" "I <+> J = carrier R"
+ shows "(R Quot (I \<inter> J)) \<simeq> (RDirProd (R Quot I) (R Quot J))"
+proof -
+ let ?\<phi> = "\<lambda>a. (I +> a, J +> a)"
+
+ have phi_hom: "?\<phi> \<in> ring_hom R (RDirProd (R Quot I) (R Quot J))"
+ using canonical_proj_is_hom[OF assms] .
+
+ moreover have "?\<phi> ` (carrier R) = carrier (RDirProd (R Quot I) (R Quot J))"
+ proof
+ show "carrier (RDirProd (R Quot I) (R Quot J)) \<subseteq> ?\<phi> ` (carrier R)"
+ proof
+ fix t assume "t \<in> carrier (RDirProd (R Quot I) (R Quot J))"
+ then obtain x y where x: "x \<in> carrier R" and y: "y \<in> 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 \<in> carrier R" "I +> a = I +> x" "J +> a = J +> y"
+ using canonical_proj_is_surj[OF assms x y] by blast
+ hence "?\<phi> a = t" using t by simp
+ thus "t \<in> (?\<phi> ` carrier R)" using \<open>a \<in> carrier R\<close> by blast
+ qed
+ next
+ show "?\<phi> ` carrier R \<subseteq> carrier (RDirProd (R Quot I) (R Quot J))"
+ using phi_hom unfolding ring_hom_def by blast
+ qed
+
+ moreover have "a_kernel R (RDirProd (R Quot I) (R Quot J)) ?\<phi> = I \<inter> J"
+ proof
+ show "I \<inter> J \<subseteq> a_kernel R (RDirProd (R Quot I) (R Quot J)) ?\<phi>"
+ proof
+ fix s assume s: "s \<in> I \<inter> J" hence "I +> s = I \<and> J +> s = J"
+ by (simp add: additive_subgroup.zero_closed assms ideal.axioms(1) ideal.set_add_zero)
+ thus "s \<in> a_kernel R (RDirProd (R Quot I) (R Quot J)) ?\<phi>"
+ 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)) ?\<phi> \<subseteq> I \<inter> 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)" ?\<phi>] is_ring
+ by (simp add: ring_hom_ringI2)
+qed
+
+
+subsection \<open>First Generalization - The Extended Canonical Projection is Surjective\<close>
+
+lemma (in cring) canonical_proj_ext_is_surj:
+ assumes "\<And>i. i \<in> {..n :: nat} \<Longrightarrow> x i \<in> carrier R"
+ and "\<And>i. i \<in> {..n} \<Longrightarrow> ideal (I i) R"
+ and "\<And>i j. \<lbrakk> i \<in> {..n}; j \<in> {..n}; i \<noteq> j \<rbrakk> \<Longrightarrow> I i <+> I j = carrier R"
+ shows "\<exists> a \<in> carrier R. \<forall> i \<in> {..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 \<in> carrier R" "\<And>i. i \<in> {..n :: nat} \<Longrightarrow> (I i) +> a = (I i) +> (x i)"
+ by force
+
+ have inter_is_ideal: "ideal (\<Inter> i \<le> n. I i) R"
+ by (metis (mono_tags, lifting) Suc.prems(2) atMost_Suc i_Intersect image_iff
+ image_is_empty insert_iff not_empty_eq_Iic_eq_empty)
+
+ have "(\<Inter> i \<le> n. I i) <+> I (Suc n) = carrier R"
+ using inter_plus_ideal_eq_carrier Suc by simp
+ then obtain b where b: "b \<in> carrier R"
+ and "(\<Inter> i \<le> n. I i) +> b = (\<Inter> i \<le> n. I i) +> \<zero>"
+ and S: "I (Suc n) +> b = I (Suc n) +> (x (Suc n) \<ominus> a)"
+ using canonical_proj_is_surj[OF inter_is_ideal, of "I (Suc n)" \<zero> "x (Suc n) \<ominus> a"] Suc a by auto
+ hence b_inter: "b \<in> (\<Inter> i \<le> 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: "\<And>i. i \<in> {..n :: nat} \<Longrightarrow> (I i) +> b = (I i) +> \<zero>"
+ proof -
+ fix i assume i: "i \<in> {..n :: nat}"
+ hence "b \<in> 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) +> \<zero>"
+ by (metis b ideal.I_r_closed ideal.set_add_zero r_null zero_closed)
+ qed
+
+ have "\<And>i. i \<in> {..Suc n} \<Longrightarrow> (I i) +> (a \<oplus> b) = (I i) +> (x i)"
+ proof -
+ fix i assume i: "i \<in> {..Suc n}" thus "(I i) +> (a \<oplus> b) = (I i) +> (x i)"
+ proof (cases)
+ assume 1: "i \<in> {..n}"
+ hence "(I i) +> (a \<oplus> b) = ((I i) +> (x i)) <+> ((I i) +> b)"
+ by (metis Suc.prems(2) a b i abelian_subgroup.a_rcos_sum
+ abelian_subgroupI3 ideal.axioms(1) is_abelian_group)
+ also have " ... = ((I i) +> (x i)) <+> ((I i) +> \<zero>)"
+ using eq_zero[OF 1] by simp
+ also have " ... = I i +> ((x i) \<oplus> \<zero>)"
+ by (meson Suc abelian_subgroup.a_rcos_sum abelian_subgroupI3 i
+ ideal.axioms(1) is_abelian_group zero_closed)
+ finally show "(I i) +> (a \<oplus> b) = (I i) +> (x i)"
+ using Suc.prems(1) i by auto
+ next
+ assume "i \<notin> {..n}" hence 2: "i = Suc n" using i by simp
+ hence "I i +> (a \<oplus> b) = I (Suc n) +> (a \<oplus> b)" by simp
+ also have " ... = (I (Suc n) +> a) <+> (I (Suc n) +> (x (Suc n) \<ominus> a))"
+ using S a b Suc.prems(2)[of "Suc n"] 2 abelian_subgroup.a_rcos_sum
+ abelian_subgroupI3 i ideal.axioms(1) is_abelian_group by metis
+ also have " ... = I (Suc n) +> (a \<oplus> (x (Suc n) \<ominus> 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 \<oplus> b) = (I i) +> (x i)" using 2 by simp
+ qed
+ qed
+ thus ?case using a b by auto
+qed
+
+
+subsection \<open>Direct Product of a List of Monoid Structures\<close>
+
+fun DirProd_list :: "(('a, 'b) monoid_scheme) list \<Rightarrow> (('a list), 'b) monoid_scheme"
+ where
+ "DirProd_list [] = \<lparr> carrier = {[]}, mult = (\<lambda>a b. []), one = [], \<dots> = (undefined :: 'b) \<rparr>"
+ | "DirProd_list (Cons R Rs) =
+ \<lparr> carrier = { r # rs | r rs. r \<in> carrier R \<and> rs \<in> carrier (DirProd_list Rs) },
+ mult = (\<lambda>r1 r2. ((hd r1) \<otimes>\<^bsub>R\<^esub> (hd r2)) # ((tl r1) \<otimes>\<^bsub>(DirProd_list Rs)\<^esub> (tl r2))),
+ one = (\<one>\<^bsub>R\<^esub>) # (\<one>\<^bsub>(DirProd_list Rs)\<^esub>), \<dots> = (undefined :: 'b) \<rparr>"
+
+
+lemma DirProd_list_carrier_elts:
+ assumes "rs \<in> 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' \<in> carrier R" "rs' \<in> 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 "\<And>i. i \<in> {..<(length rs)} \<Longrightarrow> rs ! i \<in> carrier (Rs ! i)"
+ and "length rs = length Rs"
+ shows "rs \<in> 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' \<in> carrier R" "rs = r' # rs'"
+ by (metis Suc_length_conv lessThan_iff nth_Cons_0 zero_less_Suc)
+ hence "rs' \<in> 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 \<in> carrier (DirProd_list Rs)"
+ shows "\<And>i. i \<in> {..<(length rs)} \<Longrightarrow> rs ! i \<in> 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' \<in> carrier R"
+ and rs': "rs' \<in> carrier (DirProd_list Rs)"
+ and rs: "rs = r' # rs'" by auto
+ hence "\<And>i. i \<in> {..<(length rs')} \<Longrightarrow> rs' ! i \<in> carrier (Rs ! i)"
+ using "2.hyps"(1) by blast
+ hence "\<And>i. i \<in> {(Suc 0 :: nat)..<(length rs)} \<Longrightarrow> rs ! i \<in> carrier ((R # Rs) ! i)"
+ by (simp add: less_eq_Suc_le rs)
+ moreover have "i = 0 \<Longrightarrow> rs ! i \<in> carrier ((R # Rs) ! i)"
+ using r' rs r' by simp
+ ultimately show ?case
+ using "2.prems"(1) by fastforce
+qed
+
+lemma DirProd_list_m_closed:
+ assumes "r1 \<in> carrier (DirProd_list Rs)" "r2 \<in> carrier (DirProd_list Rs)"
+ and "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow> monoid (Rs ! i)"
+ shows "r1 \<otimes>\<^bsub>(DirProd_list Rs)\<^esub> r2 \<in> 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' \<in> carrier R" "r2' \<in> carrier R"
+ and "rs1' \<in> carrier (DirProd_list Rs)"
+ and "rs2' \<in> carrier (DirProd_list Rs)"
+ and r1: "r1 = r1' # rs1'"
+ and r2: "r2 = r2' # rs2'" by auto
+ moreover have "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow> monoid (Rs ! i)"
+ using "2.prems"(3) by force
+ ultimately have "rs1' \<otimes>\<^bsub>(DirProd_list Rs)\<^esub> rs2' \<in> carrier (DirProd_list Rs)"
+ using "2.hyps"(1) by blast
+ moreover have "monoid R"
+ using "2.prems"(3) by force
+ hence "r1' \<otimes>\<^bsub>R\<^esub> r2' \<in> 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 \<in> carrier (DirProd_list Rs)" "r2 \<in> carrier (DirProd_list Rs)"
+ shows "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow>
+ (r1 \<otimes>\<^bsub>(DirProd_list Rs)\<^esub> r2) ! i = (r1 ! i) \<otimes>\<^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 "\<And>i. i \<in> {(Suc 0)..<(length (R # Rs))} \<Longrightarrow>
+ (r1 \<otimes>\<^bsub>(DirProd_list (R # Rs))\<^esub> r2) ! i = (r1 ! i) \<otimes>\<^bsub>((R # Rs) ! i)\<^esub> (r2 ! i)"
+ using "2"(5) "2"(6) by auto
+ moreover have "(r1 \<otimes>\<^bsub>(DirProd_list (R # Rs))\<^esub> r2) ! 0 = (r1 ! 0) \<otimes>\<^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 lessThan_iff not_less_eq_eq nth_Cons')
+qed
+
+lemma DirProd_list_m_comm:
+ assumes "r1 \<in> carrier (DirProd_list Rs)" "r2 \<in> carrier (DirProd_list Rs)"
+ and "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow> comm_monoid (Rs ! i)"
+ shows "r1 \<otimes>\<^bsub>(DirProd_list Rs)\<^esub> r2 = r2 \<otimes>\<^bsub>(DirProd_list Rs)\<^esub> r1"
+ apply (rule nth_equalityI) apply auto
+proof -
+ show "length (r1 \<otimes>\<^bsub>DirProd_list Rs\<^esub> r2) = length (r2 \<otimes>\<^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 \<otimes>\<^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 \<otimes>\<^bsub>DirProd_list Rs\<^esub> r2) ! i = (r1 ! i) \<otimes>\<^bsub>(Rs ! i)\<^esub> (r2 ! i)"
+ using i DirProd_list_m_output[OF assms(1-2)] by simp
+ also have " ... = (r2 ! i) \<otimes>\<^bsub>(Rs ! i)\<^esub> (r1 ! i)"
+ by (metis DirProd_list_carrier_elts DirProd_list_in_carrierE assms comm_monoid.m_comm i lessThan_iff)
+ also have " ... = (r2 \<otimes>\<^bsub>DirProd_list Rs\<^esub> r1) ! i"
+ using i DirProd_list_m_output[OF assms(2) assms(1)] by simp
+ finally show "(r1 \<otimes>\<^bsub>DirProd_list Rs\<^esub> r2) ! i = (r2 \<otimes>\<^bsub>DirProd_list Rs\<^esub> r1) ! i" .
+qed
+
+lemma DirProd_list_m_assoc:
+ assumes "r1 \<in> carrier (DirProd_list Rs)"
+ and "r2 \<in> carrier (DirProd_list Rs)"
+ and "r3 \<in> carrier (DirProd_list Rs)"
+ and "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow> monoid (Rs ! i)"
+ shows "(r1 \<otimes>\<^bsub>(DirProd_list Rs)\<^esub> r2) \<otimes>\<^bsub>(DirProd_list Rs)\<^esub> r3 =
+ r1 \<otimes>\<^bsub>(DirProd_list Rs)\<^esub> (r2 \<otimes>\<^bsub>(DirProd_list Rs)\<^esub> r3)"
+ apply (rule nth_equalityI) apply auto
+proof -
+ show "length ((r1 \<otimes>\<^bsub>DirProd_list Rs\<^esub> r2) \<otimes>\<^bsub>DirProd_list Rs\<^esub> r3) =
+ length (r1 \<otimes>\<^bsub>DirProd_list Rs\<^esub> (r2 \<otimes>\<^bsub>DirProd_list Rs\<^esub> r3))"
+ by (metis DirProd_list_carrier_elts DirProd_list_m_closed assms)
+
+ fix i assume "i < length (r1 \<otimes>\<^bsub>DirProd_list Rs\<^esub> r2 \<otimes>\<^bsub>DirProd_list Rs\<^esub> r3)"
+ hence i: "i < length Rs"
+ by (metis DirProd_list_carrier_elts DirProd_list_m_closed assms)
+ have "((r1 \<otimes>\<^bsub>DirProd_list Rs\<^esub> r2) \<otimes>\<^bsub>DirProd_list Rs\<^esub> r3) ! i =
+ ((r1 ! i) \<otimes>\<^bsub>(Rs ! i)\<^esub> (r2 ! i)) \<otimes>\<^bsub>(Rs ! i)\<^esub> (r3 ! i)"
+ by (metis DirProd_list_m_closed DirProd_list_m_output i assms lessThan_iff)
+ also have " ... = (r1 ! i) \<otimes>\<^bsub>(Rs ! i)\<^esub> ((r2 ! i) \<otimes>\<^bsub>(Rs ! i)\<^esub> (r3 ! i))"
+ by (metis DirProd_list_carrier_elts DirProd_list_in_carrierE assms i lessThan_iff monoid.m_assoc)
+ also have " ... = (r1 \<otimes>\<^bsub>DirProd_list Rs\<^esub> (r2 \<otimes>\<^bsub>DirProd_list Rs\<^esub> r3)) ! i"
+ by (metis DirProd_list_m_closed DirProd_list_m_output i assms lessThan_iff)
+ finally show "((r1 \<otimes>\<^bsub>DirProd_list Rs\<^esub> r2) \<otimes>\<^bsub>DirProd_list Rs\<^esub> r3) ! i =
+ (r1 \<otimes>\<^bsub>DirProd_list Rs\<^esub> (r2 \<otimes>\<^bsub>DirProd_list Rs\<^esub> r3))! i" .
+qed
+
+lemma DirProd_list_one:
+ "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow> (\<one>\<^bsub>(DirProd_list Rs)\<^esub>) ! i = \<one>\<^bsub>(Rs ! i)\<^esub>"
+ by (induct Rs rule: DirProd_list.induct) (simp_all add: nth_Cons')
+
+lemma DirProd_list_one_closed:
+ assumes "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow> monoid (Rs ! i)"
+ shows "\<one>\<^bsub>(DirProd_list Rs)\<^esub> \<in> carrier (DirProd_list Rs)"
+proof (rule DirProd_list_in_carrierI)
+ show eq_len: "length \<one>\<^bsub>DirProd_list Rs\<^esub> = length Rs"
+ by (induct Rs rule: DirProd_list.induct) (simp_all)
+ show "\<And>i. i \<in> {..<length \<one>\<^bsub>DirProd_list Rs\<^esub>} \<Longrightarrow> \<one>\<^bsub>DirProd_list Rs\<^esub> ! i \<in> carrier (Rs ! i)"
+ using eq_len DirProd_list_one[where ?Rs = Rs] monoid.one_closed by (simp add: assms)
+qed
+
+lemma DirProd_list_l_one:
+ assumes "r1 \<in> carrier (DirProd_list Rs)"
+ and "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow> monoid (Rs ! i)"
+ shows "\<one>\<^bsub>(DirProd_list Rs)\<^esub> \<otimes>\<^bsub>(DirProd_list Rs)\<^esub> r1 = r1"
+ apply (rule nth_equalityI) apply auto
+proof -
+ show eq_len: "length (\<one>\<^bsub>DirProd_list Rs\<^esub> \<otimes>\<^bsub>DirProd_list Rs\<^esub> r1) = length r1"
+ using DirProd_list_carrier_elts[of "\<one>\<^bsub>DirProd_list Rs\<^esub> \<otimes>\<^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 (\<one>\<^bsub>DirProd_list Rs\<^esub> \<otimes>\<^bsub>DirProd_list Rs\<^esub> r1)"
+ hence i: "i < length Rs" using DirProd_list_carrier_elts[OF assms(1)] eq_len by simp
+ hence "(\<one>\<^bsub>DirProd_list Rs\<^esub> \<otimes>\<^bsub>DirProd_list Rs\<^esub> r1) ! i =
+ (\<one>\<^bsub>DirProd_list Rs\<^esub> ! i) \<otimes>\<^bsub>(Rs ! i)\<^esub> (r1 ! i)"
+ using DirProd_list_m_output DirProd_list_one_closed assms by blast
+ also have " ... = \<one>\<^bsub>(Rs ! i)\<^esub> \<otimes>\<^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 "(\<one>\<^bsub>DirProd_list Rs\<^esub> \<otimes>\<^bsub>DirProd_list Rs\<^esub> r1) ! i = (r1 ! i)" .
+qed
+
+lemma DirProd_list_r_one:
+ assumes "r1 \<in> carrier (DirProd_list Rs)"
+ and "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow> monoid (Rs ! i)"
+ shows "r1 \<otimes>\<^bsub>(DirProd_list Rs)\<^esub> \<one>\<^bsub>(DirProd_list Rs)\<^esub> = r1"
+proof -
+ have "r1 \<otimes>\<^bsub>(DirProd_list Rs)\<^esub> \<one>\<^bsub>(DirProd_list Rs)\<^esub> =
+ \<one>\<^bsub>(DirProd_list Rs)\<^esub> \<otimes>\<^bsub>(DirProd_list Rs)\<^esub> r1"
+ apply (rule nth_equalityI) apply auto
+ proof -
+ show " length (r1 \<otimes>\<^bsub>DirProd_list Rs\<^esub> \<one>\<^bsub>DirProd_list Rs\<^esub>) =
+ length (\<one>\<^bsub>DirProd_list Rs\<^esub> \<otimes>\<^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 \<otimes>\<^bsub>DirProd_list Rs\<^esub> \<one>\<^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 \<otimes>\<^bsub>DirProd_list Rs\<^esub> \<one>\<^bsub>DirProd_list Rs\<^esub>) ! i = (r1 ! i) \<otimes>\<^bsub>(Rs ! i)\<^esub> \<one>\<^bsub>(Rs ! i)\<^esub>"
+ by (metis DirProd_list_m_output DirProd_list_one DirProd_list_one_closed assms lessThan_iff)
+ also have " ... = \<one>\<^bsub>(Rs ! i)\<^esub> \<otimes>\<^bsub>(Rs ! i)\<^esub> (r1 ! i)"
+ using DirProd_list_carrier_elts DirProd_list_in_carrierE assms i by fastforce
+ also have " ... = (\<one>\<^bsub>DirProd_list Rs\<^esub> \<otimes>\<^bsub>DirProd_list Rs\<^esub> r1) ! i"
+ by (metis DirProd_list_m_output DirProd_list_one DirProd_list_one_closed assms i lessThan_iff)
+ finally show "(r1 \<otimes>\<^bsub>DirProd_list Rs\<^esub> \<one>\<^bsub>DirProd_list Rs\<^esub>) ! i =
+ (\<one>\<^bsub>DirProd_list Rs\<^esub> \<otimes>\<^bsub>DirProd_list Rs\<^esub> r1) ! i" .
+ qed
+ thus ?thesis using DirProd_list_l_one assms by auto
+qed
+
+lemma DirProd_list_monoid:
+ assumes "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow> monoid (Rs ! i)"
+ shows "monoid (DirProd_list Rs)"
+ unfolding monoid_def apply auto
+proof -
+ show "\<one>\<^bsub>DirProd_list Rs\<^esub> \<in> carrier (DirProd_list Rs)"
+ using DirProd_list_one_closed[of Rs] assms by simp
+
+ fix x y z
+ assume x: "x \<in> carrier (DirProd_list Rs)"
+ and y: "y \<in> carrier (DirProd_list Rs)"
+ and z: "z \<in> carrier (DirProd_list Rs)"
+ show "x \<otimes>\<^bsub>DirProd_list Rs\<^esub> y \<in> carrier (DirProd_list Rs)"
+ using DirProd_list_m_closed[OF x y] assms by simp
+ show "x \<otimes>\<^bsub>DirProd_list Rs\<^esub> y \<otimes>\<^bsub>DirProd_list Rs\<^esub> z =
+ x \<otimes>\<^bsub>DirProd_list Rs\<^esub> (y \<otimes>\<^bsub>DirProd_list Rs\<^esub> z)"
+ using DirProd_list_m_assoc[OF x y z] assms by simp
+ show "\<one>\<^bsub>DirProd_list Rs\<^esub> \<otimes>\<^bsub>DirProd_list Rs\<^esub> x = x"
+ using DirProd_list_l_one[OF x] assms by simp
+ show "x \<otimes>\<^bsub>DirProd_list Rs\<^esub> \<one>\<^bsub>DirProd_list Rs\<^esub> = x"
+ using DirProd_list_r_one[OF x] assms by simp
+qed
+
+lemma DirProd_list_comm_monoid:
+ assumes "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow> 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:
+ "(\<lambda>(hd, tl). hd # tl) \<in> iso (R \<times>\<times> (DirProd_list Rs)) (DirProd_list (R # Rs))"
+ unfolding iso_def hom_def bij_betw_def inj_on_def by auto
+
+lemma DirProd_list_isomorphism2:
+ "(\<lambda>r. (hd r, tl r)) \<in> iso (DirProd_list (R # Rs)) (R \<times>\<times> (DirProd_list Rs))" (is "?\<phi> \<in> ?A")
+ unfolding iso_def hom_def bij_betw_def inj_on_def apply auto
+proof -
+ fix a b assume "a \<in> carrier R" "b \<in> carrier (DirProd_list Rs)"
+ hence "a # b \<in> {r # rs |r rs. r \<in> carrier R \<and> rs \<in> carrier (DirProd_list Rs)} \<and> ?\<phi> (a # b) = (a, b)"
+ by simp
+ thus "(a, b) \<in> ?\<phi> ` {r # rs |r rs. r \<in> carrier R \<and> rs \<in> carrier (DirProd_list Rs)}"
+ by (metis (no_types, lifting) image_iff)
+qed
+
+lemma DirProd_list_group:
+ assumes "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow> 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 \<times>\<times> DirProd_list Rs \<cong> DirProd_list (R # Rs)"
+ unfolding is_iso_def using DirProd_list_isomorphism1 by auto
+ ultimately show ?case
+ using group.iso_imp_group[of "R \<times>\<times> (DirProd_list Rs)" "DirProd_list (R # Rs)"]
+ DirProd_group[of R "DirProd_list Rs"] by simp
+qed
+
+lemma DirProd_list_comm_group:
+ assumes "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow> 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 "\<And>i. i \<in> {..<(length (R # Rs))} \<Longrightarrow> group ((R # Rs) ! i)"
+ shows "group_hom (R \<times>\<times> DirProd_list Rs) (DirProd_list (R # Rs)) (\<lambda>(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 \<times>\<times> 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 "(\<lambda>(x, y). x # y) \<in> hom (R \<times>\<times> 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 \<in> carrier (DirProd_list Rs)"
+ and "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow> group (Rs ! i)"
+ shows "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow> (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' \<in> carrier R" and rs': "rs' \<in> carrier (DirProd_list Rs)"
+ and r: "r = r' # rs'" by auto
+ hence "(r', rs') \<in> carrier (R \<times>\<times> DirProd_list Rs)" by simp
+ moreover have "group_hom (R \<times>\<times> DirProd_list Rs) (DirProd_list (R # Rs)) (\<lambda>(hd, tl). hd # tl)"
+ using DirProd_list_group_hom[of R Rs] 2 by auto
+ moreover have "inv\<^bsub>(R \<times>\<times> 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 \<times>\<times> DirProd_list Rs" "DirProd_list (R # Rs)"
+ "\<lambda>(hd, tl). hd # tl" "(r', rs')"] r by simp
+ thus ?case
+ by (smt "2.hyps"(1) "2.prems"(1) "2.prems"(3) One_nat_def Suc_less_eq Suc_pred length_Cons
+ lessThan_iff list.sel(3) not_gr0 nth_Cons' nth_tl r rs')
+qed
+
+
+subsection \<open>Direct Product for of a List of Rings\<close>
+
+text \<open>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.\<close>
+
+fun RDirProd_list :: "('a ring) list \<Rightarrow> ('a list) ring"
+ where "RDirProd_list Rs =
+ monoid.extend (monoid.truncate (DirProd_list Rs))
+ \<lparr> zero = one (DirProd_list (map add_monoid Rs)),
+ add = mult (DirProd_list (map add_monoid Rs)) \<rparr>"
+
+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 RDirProd_list_monoid:
+ assumes "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow> 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
+
+lemma RDirProd_list_comm_monoid:
+ assumes "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow> 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 "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow> abelian_monoid (Rs ! i)"
+ shows "abelian_monoid (RDirProd_list Rs)"
+proof -
+ have "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow> 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
+
+lemma RDirProd_list_abelian_group:
+ assumes "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow> abelian_group (Rs ! i)"
+ shows "abelian_group (RDirProd_list Rs)"
+proof -
+ have "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow> 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 \<in> carrier (RDirProd_list Rs)"
+ shows "length rs = length Rs"
+ using assms by (simp add: DirProd_list_carrier_elts monoid.defs)
+
+lemma RDirProd_list_in_carrierE:
+ assumes "rs \<in> carrier (RDirProd_list Rs)"
+ shows "\<And>i. i \<in> {..<(length rs)} \<Longrightarrow> rs ! i \<in> carrier (Rs ! i)"
+ using assms by (simp add: DirProd_list_in_carrierE monoid.defs)
+
+lemma RDirProd_list_in_carrierI:
+ assumes "\<And>i. i \<in> {..<(length rs)} \<Longrightarrow> rs ! i \<in> carrier (Rs ! i)"
+ and "length rs = length Rs"
+ shows "rs \<in> carrier (RDirProd_list Rs)"
+ using DirProd_list_in_carrierI assms by (simp add: monoid.defs, blast)
+
+lemma RDirProd_list_one:
+ "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow> (\<one>\<^bsub>(RDirProd_list Rs)\<^esub>) ! i = \<one>\<^bsub>(Rs ! i)\<^esub>"
+ by (simp add: DirProd_list_one monoid.defs)
+
+lemma RDirProd_list_zero:
+ "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow> (\<zero>\<^bsub>(RDirProd_list Rs)\<^esub>) ! i = \<zero>\<^bsub>(Rs ! i)\<^esub>"
+ by (induct Rs rule: DirProd_list.induct) (simp_all add: monoid.defs nth_Cons')
+
+lemma RDirProd_list_m_output:
+ assumes "r1 \<in> carrier (RDirProd_list Rs)" "r2 \<in> carrier (RDirProd_list Rs)"
+ shows "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow>
+ (r1 \<otimes>\<^bsub>(RDirProd_list Rs)\<^esub> r2) ! i = (r1 ! i) \<otimes>\<^bsub>(Rs ! i)\<^esub> (r2 ! i)"
+ using assms by (simp add: DirProd_list_m_output monoid.defs)
+
+lemma RDirProd_list_a_output:
+ assumes "r1 \<in> carrier (RDirProd_list Rs)" "r2 \<in> carrier (RDirProd_list Rs)"
+ shows "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow>
+ (r1 \<oplus>\<^bsub>(RDirProd_list Rs)\<^esub> r2) ! i = (r1 ! i) \<oplus>\<^bsub>(Rs ! i)\<^esub> (r2 ! i)"
+ using RDirProd_list_add_monoid[of Rs] monoid.defs(3)
+ by (smt DirProd_list_m_output assms length_map lessThan_iff
+ monoid.select_convs(1) nth_map partial_object.select_convs(1))
+
+lemma RDirProd_list_a_inv:
+ assumes "r \<in> carrier (RDirProd_list Rs)"
+ and "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow> abelian_group (Rs ! i)"
+ shows "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow> (\<ominus>\<^bsub>(RDirProd_list Rs)\<^esub> r) ! i = \<ominus>\<^bsub>(Rs ! i)\<^esub> (r ! i)"
+ using RDirProd_list_add_monoid[of Rs] monoid.defs(3) DirProd_list_m_inv
+ by (smt a_inv_def abelian_group.a_group assms length_map lessThan_iff
+ monoid.surjective nth_map partial_object.ext_inject)
+
+lemma RDirProd_list_l_distr:
+ assumes "x \<in> carrier (RDirProd_list Rs)"
+ and "y \<in> carrier (RDirProd_list Rs)"
+ and "z \<in> carrier (RDirProd_list Rs)"
+ and "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow> ring (Rs ! i)"
+ shows "(x \<oplus>\<^bsub>(RDirProd_list Rs)\<^esub> y) \<otimes>\<^bsub>(RDirProd_list Rs)\<^esub> z =
+ (x \<otimes>\<^bsub>(RDirProd_list Rs)\<^esub> z) \<oplus>\<^bsub>(RDirProd_list Rs)\<^esub> (y \<otimes>\<^bsub>(RDirProd_list Rs)\<^esub> z)"
+proof -
+ have "length ((x \<oplus>\<^bsub>(RDirProd_list Rs)\<^esub> y) \<otimes>\<^bsub>(RDirProd_list Rs)\<^esub> z) =
+ length ((x \<otimes>\<^bsub>(RDirProd_list Rs)\<^esub> z) \<oplus>\<^bsub>(RDirProd_list Rs)\<^esub> (y \<otimes>\<^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 "\<And>i. i < length Rs \<Longrightarrow>
+ ((x \<oplus>\<^bsub>(RDirProd_list Rs)\<^esub> y) \<otimes>\<^bsub>(RDirProd_list Rs)\<^esub> z) ! i =
+ ((x \<otimes>\<^bsub>(RDirProd_list Rs)\<^esub> z) \<oplus>\<^bsub>(RDirProd_list Rs)\<^esub> (y \<otimes>\<^bsub>(RDirProd_list Rs)\<^esub> z)) ! i"
+ proof -
+ fix i assume i: "i < length Rs"
+ hence "((x \<oplus>\<^bsub>(RDirProd_list Rs)\<^esub> y) \<otimes>\<^bsub>(RDirProd_list Rs)\<^esub> z) ! i =
+ ((x ! i) \<oplus>\<^bsub>(Rs ! i)\<^esub> (y ! i)) \<otimes>\<^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) \<otimes>\<^bsub>(Rs ! i)\<^esub> (z ! i)) \<oplus>\<^bsub>(Rs ! i)\<^esub> ((y ! i) \<otimes>\<^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 \<otimes>\<^bsub>(RDirProd_list Rs)\<^esub> z) \<oplus>\<^bsub>(RDirProd_list Rs)\<^esub> (y \<otimes>\<^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 \<oplus>\<^bsub>(RDirProd_list Rs)\<^esub> y) \<otimes>\<^bsub>(RDirProd_list Rs)\<^esub> z) ! i =
+ ((x \<otimes>\<^bsub>(RDirProd_list Rs)\<^esub> z) \<oplus>\<^bsub>(RDirProd_list Rs)\<^esub> (y \<otimes>\<^bsub>(RDirProd_list Rs)\<^esub> z)) ! i" .
+ qed
+
+ moreover have "length ((x \<oplus>\<^bsub>(RDirProd_list Rs)\<^esub> y) \<otimes>\<^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 \<in> carrier (RDirProd_list Rs)"
+ and "y \<in> carrier (RDirProd_list Rs)"
+ and "z \<in> carrier (RDirProd_list Rs)"
+ and "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow> ring (Rs ! i)"
+ shows "z \<otimes>\<^bsub>(RDirProd_list Rs)\<^esub> (x \<oplus>\<^bsub>(RDirProd_list Rs)\<^esub> y) =
+ (z \<otimes>\<^bsub>(RDirProd_list Rs)\<^esub> x) \<oplus>\<^bsub>(RDirProd_list Rs)\<^esub> (z \<otimes>\<^bsub>(RDirProd_list Rs)\<^esub> y)"
+proof -
+ have "length (z \<otimes>\<^bsub>(RDirProd_list Rs)\<^esub> (x \<oplus>\<^bsub>(RDirProd_list Rs)\<^esub> y)) =
+ length ((z \<otimes>\<^bsub>(RDirProd_list Rs)\<^esub> x) \<oplus>\<^bsub>(RDirProd_list Rs)\<^esub> (z \<otimes>\<^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 "\<And>i. i < length Rs \<Longrightarrow>
+ (z \<otimes>\<^bsub>(RDirProd_list Rs)\<^esub> (x \<oplus>\<^bsub>(RDirProd_list Rs)\<^esub> y)) ! i =
+ ((z \<otimes>\<^bsub>(RDirProd_list Rs)\<^esub> x) \<oplus>\<^bsub>(RDirProd_list Rs)\<^esub> (z \<otimes>\<^bsub>(RDirProd_list Rs)\<^esub> y)) ! i"
+ proof -
+ fix i assume i: "i < length Rs"
+ hence "(z \<otimes>\<^bsub>(RDirProd_list Rs)\<^esub> (x \<oplus>\<^bsub>(RDirProd_list Rs)\<^esub> y)) ! i =
+ (z ! i) \<otimes>\<^bsub>(Rs ! i)\<^esub> ((x ! i) \<oplus>\<^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) \<otimes>\<^bsub>(Rs ! i)\<^esub> (x ! i)) \<oplus>\<^bsub>(Rs ! i)\<^esub> ((z ! i) \<otimes>\<^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 \<otimes>\<^bsub>(RDirProd_list Rs)\<^esub> x) \<oplus>\<^bsub>(RDirProd_list Rs)\<^esub> (z \<otimes>\<^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 \<otimes>\<^bsub>(RDirProd_list Rs)\<^esub> (x \<oplus>\<^bsub>(RDirProd_list Rs)\<^esub> y)) ! i =
+ ((z \<otimes>\<^bsub>(RDirProd_list Rs)\<^esub> x) \<oplus>\<^bsub>(RDirProd_list Rs)\<^esub> (z \<otimes>\<^bsub>(RDirProd_list Rs)\<^esub> y)) ! i" .
+ qed
+
+ moreover have "length (z \<otimes>\<^bsub>(RDirProd_list Rs)\<^esub> (x \<oplus>\<^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 "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow> 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 "\<And>i. i \<in> {..<(length Rs)} \<Longrightarrow> 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 "\<And>i. i \<in> {..< n} \<Longrightarrow> ideal (I i) R"
+ shows "cring (RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< n]))"
+ (is "cring (RDirProd_list ?Rs)")
+proof -
+ have "\<And>i. i \<in> {..<(length ?Rs)} \<Longrightarrow> cring (?Rs ! i)"
+ by (simp add: assms ideal.quotient_is_cring is_cring)
+ thus ?thesis
+ using RDirProd_list_cring by blast
+qed
+
+lemma RDirProd_list_isomorphism1:
+ "(\<lambda>(hd, tl). hd # tl) \<in> 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)) \<simeq> (RDirProd_list (R # Rs))"
+ unfolding is_ring_iso_def using RDirProd_list_isomorphism1 by blast
+
+lemma RDirProd_list_isomorphism2:
+ "(\<lambda>r. (hd r, tl r)) \<in> 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 ?\<phi> = "\<lambda>r. (hd r, tl r)"
+ fix a b assume "a \<in> carrier R" "b \<in> carrier (DirProd_list Rs)"
+ hence "a # b \<in> {r # rs |r rs. r \<in> carrier R \<and> rs \<in> carrier (DirProd_list Rs)} \<and> ?\<phi> (a # b) = (a, b)"
+ by simp
+ thus "(a, b) \<in> ?\<phi> ` {r # rs |r rs. r \<in> carrier R \<and> rs \<in> carrier (DirProd_list Rs)}"
+ by (metis (no_types, lifting) image_iff)
+qed
+
+lemma RDirProd_list_isomorphism3:
+ "(\<lambda>(r, l). r @ [l]) \<in> 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)
+next
+ case (2 R Rs)
+
+ { fix r1 r2 assume A0: "r1 \<in> carrier (RDirProd_list (R # Rs))"
+ and A1: "r2 \<in> carrier (RDirProd_list (R # Rs))"
+ have "length r1 \<ge> 1"
+ and "length r2 \<ge> 1"
+ and "length (r1 \<otimes>\<^bsub>(RDirProd_list (R # Rs))\<^esub> r2) \<ge> 1"
+ and "length (r1 \<oplus>\<^bsub>(RDirProd_list (R # Rs))\<^esub> r2) \<ge> 1"
+ and "length (\<one>\<^bsub>(RDirProd_list (R # Rs))\<^esub>) \<ge> 1"
+ proof -
+ show len_r1: "length r1 \<ge> 1"
+ and len_r2: "length r2 \<ge> 1"
+ by (metis RDirProd_list_carrier_elts A0 A1 length_Cons less_one nat.simps(3) not_less)+
+ show "length (r1 \<otimes>\<^bsub>(RDirProd_list (R # Rs))\<^esub> r2) \<ge> 1"
+ and "length (r1 \<oplus>\<^bsub>(RDirProd_list (R # Rs))\<^esub> r2) \<ge> 1"
+ and "length (\<one>\<^bsub>(RDirProd_list (R # Rs))\<^esub>) \<ge> 1"
+ using len_r1 len_r2 by (simp_all add: monoid.defs)
+ qed } note aux_lemma = this
+
+ moreover
+ have "(\<lambda>(r, s). (hd r, (tl r, s))) \<in>
+ 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 "(\<lambda>(hd, (tl, s)). hd # (tl @ [s])) \<in>
+ 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 "(\<lambda>(r, s). (hd r) # ((tl r) @ [s])) \<in>
+ 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: "(\<lambda>(r, s). (hd r # tl r) @ [s]) \<in>
+ 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 "\<lambda>r. length (fst r) \<ge> 1" "\<lambda>(r, s). r @ [s]"])
+ show "\<And>r. 1 \<le> length (fst r) \<Longrightarrow>
+ (case r of (r, s) \<Rightarrow> (hd r # tl r) @ [s]) = (case r of (r, s) \<Rightarrow> r @ [s])"
+ by (simp add: Suc_le_eq case_prod_beta')
+ show "morphic_prop (RDirProd (RDirProd_list (R # Rs)) S) (\<lambda>r. 1 \<le> length (fst r))"
+ unfolding RDirProd_def by (rule morphic_propI) (auto simp add: monoid.defs aux_lemma)
+ qed
+qed
+
+
+subsection \<open>Second Generalization - The Extended Canonical Projection is a Homomorphism and
+ Description of its Kernel\<close>
+
+theorem (in cring) canonical_proj_ext_is_hom:
+ assumes "\<And>i. i \<in> {..< (n :: nat)} \<Longrightarrow> ideal (I i) R"
+ and "\<And>i j. \<lbrakk> i \<in> {..< n}; j \<in> {..< n}; i \<noteq> j \<rbrakk> \<Longrightarrow> I i <+> I j = carrier R"
+ shows "(\<lambda>a. (map (\<lambda>i. (I i) +> a) [0..< n])) \<in>
+ ring_hom R (RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< n]))" (is "?\<phi> \<in> ?ring_hom")
+proof (rule ring_hom_memI)
+ { fix x assume x: "x \<in> carrier R"
+ have "?\<phi> x \<in> carrier (RDirProd_list (map (\<lambda>i. R Quot I i) [0..<n]))"
+ apply (rule RDirProd_list_in_carrierI)
+ by (simp_all add: FactRing_def a_rcosetsI additive_subgroup.a_subset assms(1) ideal.axioms(1) x) }
+ note aux_lemma = this
+
+ fix x y assume x: "x \<in> carrier R" and y: "y \<in> carrier R"
+
+ show x': "?\<phi> x \<in> carrier (RDirProd_list (map (\<lambda>i. R Quot I i) [0..<n]))"
+ using aux_lemma[OF x] .
+ hence x'': "?\<phi> x \<in> carrier (DirProd_list (map (\<lambda>i. R Quot I i) [0..<n]))"
+ by (simp add: monoid.defs)
+
+ have y': "?\<phi> y \<in> carrier (RDirProd_list (map (\<lambda>i. R Quot I i) [0..<n]))"
+ using aux_lemma[OF y] .
+ hence y'': "map (\<lambda>i. I i +> y) [0..<n] \<in> carrier (DirProd_list (map (\<lambda>i. R Quot I i) [0..<n]))"
+ by (simp add: monoid.defs)
+
+ show "?\<phi> (x \<otimes> y) = ?\<phi> x \<otimes>\<^bsub>RDirProd_list (map (\<lambda>i. R Quot I i) [0..<n])\<^esub> ?\<phi> 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 "?\<phi> x" "(map (\<lambda>i. R Quot I i) [0..<n])" "?\<phi> y"] x'' y''
+ by (simp add: x'' y'' FactRing_def add.left_neutral assms(1) diff_zero ideal.rcoset_mult_add
+ length_map length_upt lessThan_iff monoid.simps(1) nth_map_upt x y)
+
+ show "?\<phi> (x \<oplus> y) = ?\<phi> x \<oplus>\<^bsub>RDirProd_list (map (\<lambda>i. R Quot I i) [0..<n])\<^esub> ?\<phi> y"
+ proof -
+ have "length (?\<phi> (x \<oplus> y)) =
+ length ((?\<phi> x) \<oplus>\<^bsub>RDirProd_list (map (\<lambda>i. R Quot I i) [0..<n])\<^esub> (?\<phi> y))"
+ by (metis RDirProd_list_carrier_elts RDirProd_list_of_quot_is_cring assms(1)
+ cring.cring_simprules(1) length_map x' y')
+
+ moreover
+ have "\<And>j. j < n \<Longrightarrow>
+ (?\<phi> (x \<oplus> y)) ! j = ((?\<phi> x) \<oplus>\<^bsub>RDirProd_list (map (\<lambda>i. R Quot I i) [0..<n])\<^esub> (?\<phi> y)) ! j"
+ proof -
+ fix j assume j: "j < n"
+ have "(?\<phi> (x \<oplus> y)) ! j = I j +> x \<oplus> y" using j by simp
+ also have " ... = (I j +> x) \<oplus>\<^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 " ... = ((?\<phi> x) \<oplus>\<^bsub>RDirProd_list (map (\<lambda>i. R Quot I i) [0..<n])\<^esub> (?\<phi> y)) ! j"
+ by (smt RDirProd_list_a_output add.left_neutral diff_zero j
+ length_map length_upt lessThan_iff nth_map nth_upt x' y')
+ finally show "(?\<phi> (x \<oplus> y)) ! j =
+ ((?\<phi> x) \<oplus>\<^bsub>RDirProd_list (map (\<lambda>i. R Quot I i) [0..<n])\<^esub> (?\<phi> y)) ! j" .
+ qed
+ ultimately show "?\<phi> (x \<oplus> y) = ?\<phi> x \<oplus>\<^bsub>RDirProd_list (map (\<lambda>i. R Quot I i) [0..<n])\<^esub> ?\<phi> y"
+ by (simp add: nth_equalityI)
+ qed
+next
+ show "(?\<phi> \<one>) = \<one>\<^bsub>RDirProd_list (map (\<lambda>i. R Quot I i) [0..<n])\<^esub>"
+ 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 (\<lambda>i. R Quot I i) [0..<n]"]
+ by (simp add: FactRing_def monoid.defs)
+qed
+
+
+theorem (in cring) canonical_proj_ext_kernel:
+ assumes "\<And>i. i \<in> {..(n :: nat)} \<Longrightarrow> ideal (I i) R"
+ and "\<And>i j. \<lbrakk> i \<in> {..n}; j \<in> {..n}; i \<noteq> j \<rbrakk> \<Longrightarrow> I i <+> I j = carrier R"
+ shows "(\<Inter>i \<le> n. I i) = a_kernel R (RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< Suc n]))
+ (\<lambda>a. (map (\<lambda>i. (I i) +> a) [0..< Suc n]))"
+proof -
+ let ?\<phi> = "\<lambda>a. (map (\<lambda>i. (I i) +> a) [0..< Suc n])"
+ show ?thesis
+ proof
+ show "(\<Inter>i \<le> n. I i) \<subseteq> a_kernel R (RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< Suc n])) ?\<phi>"
+ proof
+ fix s assume s: "s \<in> (\<Inter>i \<le> n. I i)"
+ hence "\<And>i. i \<le> n \<Longrightarrow> (I i) +> s = I i"
+ by (simp add: additive_subgroup.zero_closed assms ideal.axioms(1) ideal.set_add_zero)
+ hence "\<And>i. i \<le> n \<Longrightarrow> (?\<phi> s) ! i = I i"
+ by (metis add.left_neutral diff_zero le_imp_less_Suc nth_map_upt)
+ moreover have
+ "\<And>i. i \<le> n \<Longrightarrow> (\<zero>\<^bsub>(RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< Suc n]))\<^esub>) ! i =
+ \<zero>\<^bsub>(R Quot (I i))\<^esub>"
+ using RDirProd_list_zero[where ?Rs = "map (\<lambda>i. R Quot I i) [0..<Suc n]"]
+ by (metis (no_types, lifting) add.left_neutral atMost_iff le_imp_less_Suc
+ length_map length_upt lessThan_Suc_atMost nth_map_upt diff_zero)
+ hence
+ "\<And>i. i \<le> n \<Longrightarrow> (\<zero>\<^bsub>(RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< Suc n]))\<^esub>) ! i = I i"
+ unfolding FactRing_def by simp
+ moreover
+ have "length (\<zero>\<^bsub>(RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< Suc n]))\<^esub>) = Suc n"
+ using RDirProd_list_carrier_elts RDirProd_list_cring
+ by (smt add.left_neutral assms(1) cring.cring_simprules(2) diff_zero nth_map_upt
+ ideal.quotient_is_cring is_cring length_map length_upt lessThan_Suc_atMost lessThan_iff)
+ moreover have "length (?\<phi> s) = Suc n" by simp
+ ultimately have "?\<phi> s = \<zero>\<^bsub>(RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< Suc n]))\<^esub>"
+ by (simp add: nth_equalityI)
+
+ moreover have "s \<in> carrier R"
+ using additive_subgroup.a_Hcarr assms(1) ideal.axioms(1) s by fastforce
+ ultimately show "s \<in> a_kernel R (RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< Suc n])) ?\<phi>"
+ using a_kernel_def'[of R "RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< Suc n])"] by simp
+ qed
+ next
+ show "a_kernel R (RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< Suc n])) ?\<phi> \<subseteq> (\<Inter>i \<le> n. I i)"
+ proof
+ fix s assume s: "s \<in> a_kernel R (RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< Suc n])) ?\<phi>"
+ hence "?\<phi> s = \<zero>\<^bsub>(RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< Suc n]))\<^esub>"
+ unfolding a_kernel_def kernel_def by (simp add: monoid.defs)
+ hence "\<And>i. i \<le> n \<Longrightarrow> (I i) +> s = \<zero>\<^bsub>(R Quot (I i))\<^esub>"
+ using RDirProd_list_zero[where ?Rs = "map (\<lambda>i. R Quot I i) [0..<Suc n]"]
+ by (smt add.left_neutral atMost_iff diff_zero le_imp_less_Suc
+ length_map length_upt lessThan_Suc_atMost nth_map_upt)
+ hence "\<And>i. i \<le> n \<Longrightarrow> (I i) +> s = I i"
+ unfolding FactRing_def by simp
+ moreover have "s \<in> carrier R"
+ using s unfolding a_kernel_def kernel_def by simp
+ ultimately show "s \<in> (\<Inter>i \<le> n. I i)"
+ using ideal.set_add_zero_imp_mem[where ?i = s and ?R = R] by (simp add: assms(1))
+ qed
+ qed
+qed
+
+
+subsection \<open>Final Generalization\<close>
+
+theorem (in cring) chinese_remainder:
+ assumes "\<And>i. i \<in> {..(n :: nat)} \<Longrightarrow> ideal (I i) R"
+ and "\<And>i j. \<lbrakk> i \<in> {..n}; j \<in> {..n}; i \<noteq> j \<rbrakk> \<Longrightarrow> I i <+> I j = carrier R"
+ shows "R Quot (\<Inter>i \<le> n. I i) \<simeq> RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< Suc n])"
+ using assms
+proof (induct n)
+ case 0
+ have "(\<lambda>r. (r, [])) \<in> 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)) \<simeq> (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)) \<simeq> (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 (\<Inter> i \<le> n. I i) R"
+ using Suc.prems(1) i_Intersect[of "I ` {..n}"] atMost_Suc atLeast1_atMost_eq_remove0 by auto
+ hence "R Quot (\<Inter> i \<le> Suc n. I i) \<simeq> RDirProd (R Quot (\<Inter> i \<le> n. I i)) (R Quot (I (Suc n)))"
+ using chinese_remainder_simple[of "\<Inter> i \<le> 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 (\<Inter> i \<le> n. I i) \<simeq> RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< Suc n])"
+ using Suc.hyps Suc.prems(1-2) by auto
+ hence "RDirProd (R Quot (\<Inter> i \<le> n. I i)) (R Quot (I (Suc n))) \<simeq>
+ RDirProd (RDirProd_list (map (\<lambda>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 (\<Inter> i \<le> Suc n. I i) \<simeq>
+ RDirProd (RDirProd_list (map (\<lambda>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 (\<lambda>i. R Quot (I i)) [0..< Suc n])) (R Quot (I (Suc n))) \<simeq>
+ RDirProd_list ((map (\<lambda>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 (\<lambda>i. R Quot (I i)) [0..< Suc n])) (R Quot (I (Suc n))) \<simeq>
+ RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< Suc (Suc n)])" by simp
+
+ ultimately show ?case using ring_iso_trans by simp
+qed
+
+theorem (in cring) (* chinese_remainder: another proof *)
+ assumes "\<And>i. i \<in> {..(n :: nat)} \<Longrightarrow> ideal (I i) R"
+ and "\<And>i j. \<lbrakk> i \<in> {..n}; j \<in> {..n}; i \<noteq> j \<rbrakk> \<Longrightarrow> I i <+> I j = carrier R"
+ shows "R Quot (\<Inter>i \<le> n. I i) \<simeq> RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< Suc n])"
+proof -
+ let ?\<phi> = "\<lambda>a. (map (\<lambda>i. (I i) +> a) [0..< Suc n])"
+
+ have phi_hom: "?\<phi> \<in> ring_hom R (RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< Suc n]))"
+ using canonical_proj_ext_is_hom[of "Suc n"] assms by simp
+
+ moreover have "?\<phi> ` (carrier R) = carrier (RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< Suc n]))"
+ proof
+ show "carrier (RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< Suc n])) \<subseteq> ?\<phi> ` (carrier R)"
+ proof
+ fix x assume x: "x \<in> carrier (RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< Suc n]))"
+ hence x_nth_eq: "\<And>i. i \<in> {..< Suc n} \<Longrightarrow> x ! i \<in> carrier (R Quot (I i))"
+ using RDirProd_list_in_carrierE
+ by (smt RDirProd_list_carrier_elts add.left_neutral diff_zero length_map
+ length_upt lessThan_iff nth_map nth_upt)
+ then obtain y where y: "\<And>i. i \<in> {..< Suc n} \<Longrightarrow> x ! i = (I i) +> (y i)"
+ "\<And>i. i \<in> {..< Suc n} \<Longrightarrow> y i \<in> carrier R"
+ proof -
+ from x_nth_eq have "\<exists>y. (\<forall>i \<in> {..< Suc n}. x ! i = (I i) +> (y i)) \<and>
+ (\<forall>i \<in> {..< Suc n}. y i \<in> carrier R)"
+ proof (induct n)
+ case 0
+ have "x ! 0 \<in> carrier (R Quot (I 0))" using x_nth_eq by simp
+ then obtain x_0 where x_0: "x ! 0 = (I 0) +> x_0" "x_0 \<in> carrier R"
+ unfolding FactRing_def using A_RCOSETS_def'[of R "I 0"] by auto
+ define y :: "nat \<Rightarrow> 'a" where "y = (\<lambda>i. x_0)"
+ hence "x ! 0 = (I 0) +> (y 0) \<and> (y 0) \<in> carrier R"
+ using x_0 by simp
+ thus ?case using x_0 by blast
+ next
+ case (Suc n)
+ then obtain y' where y': "\<And>i. i \<in> {..<Suc n} \<Longrightarrow> x ! i = I i +> y' i"
+ "\<And>i. i \<in> {..<Suc n} \<Longrightarrow> y' i \<in> carrier R" by auto
+
+ have "x ! (Suc n) \<in> carrier (R Quot (I (Suc n)))" using Suc by simp
+ then obtain x_Sn where x_Sn: "x ! (Suc n) = (I (Suc n)) +> x_Sn"
+ "x_Sn \<in> carrier R"
+ unfolding FactRing_def using A_RCOSETS_def'[of R "I (Suc n)"] by auto
+
+ define y where "y = (\<lambda>i. if i = (Suc n) then x_Sn else (y' i))"
+ thus ?case using y' x_Sn
+ by (metis (full_types) insert_iff lessThan_Suc)
+ qed
+ thus ?thesis using that by blast
+ qed
+
+ then obtain a where a: "a \<in> carrier R"
+ and "\<And>i. i \<in> {..< Suc n} \<Longrightarrow> (I i) +> a = (I i) +> (y i)"
+ using canonical_proj_ext_is_surj[of n y I] assms(1-2) by auto
+ hence a_x: "\<And>i. i \<in> {..< Suc n} \<Longrightarrow> (I i) +> a = x ! i"
+ using y by simp
+ have "?\<phi> a = x"
+ apply (rule nth_equalityI)
+ using RDirProd_list_carrier_elts x apply fastforce
+ by (metis a_x add.left_neutral diff_zero length_map length_upt lessThan_iff nth_map_upt)
+ thus "x \<in> ?\<phi> ` (carrier R)"
+ using a by blast
+ qed
+ next
+ show "?\<phi> ` carrier R \<subseteq> carrier (RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< Suc n]))"
+ using phi_hom unfolding ring_hom_def by blast
+ qed
+
+ moreover have "a_kernel R (RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< Suc n])) ?\<phi> = (\<Inter>i \<le> n. I i)"
+ using canonical_proj_ext_kernel assms by blast
+
+ moreover have "ring (RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< Suc n]))"
+ using RDirProd_list_of_quot_is_cring assms(1) cring_def lessThan_Suc_atMost by blast
+
+ ultimately show ?thesis
+ using ring_hom_ring.FactRing_iso[of R "RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< Suc n])" ?\<phi>]
+ is_ring by (simp add: ring_hom_ringI2)
+qed
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Algebra/Cycles.thy Mon Jul 02 15:43:22 2018 +0100
@@ -0,0 +1,756 @@
+(* ************************************************************************** *)
+(* Title: Cycles.thy *)
+(* Author: Paulo EmÃlio de Vilhena *)
+(* ************************************************************************** *)
+
+theory Cycles
+ imports "HOL-Library.Permutations" "HOL-Library.FuncSet"
+
+begin
+
+section \<open>Cycles\<close>
+
+abbreviation cycle :: "'a list \<Rightarrow> bool" where
+ "cycle cs \<equiv> distinct cs"
+
+fun cycle_of_list :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a"
+ where
+ "cycle_of_list (i # j # cs) = (Fun.swap i j id) \<circ> (cycle_of_list (j # cs))"
+ | "cycle_of_list cs = id"
+
+
+subsection\<open>Cycles as Rotations\<close>
+
+text\<open>We start proving that the function derived from a cycle rotates its support list.\<close>
+
+lemma id_outside_supp:
+ assumes "cycle cs" and "x \<notin> 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
+
+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
+
+theorem cyclic_rotation:
+ 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)
+ case (1 i j cs) thus ?case
+ proof (cases)
+ assume "cs = Nil" thus ?thesis by simp
+ next
+ assume "cs \<noteq> Nil" hence ge_two: "length (j # cs) \<ge> 2"
+ using not_less by auto
+ have "map (cycle_of_list (i # j # cs)) (i # j # cs) =
+ map (Fun.swap i j id) (map (cycle_of_list (j # cs)) (i # j # cs))" by simp
+ also have " ... = map (Fun.swap i j id) (i # (rotate1 (j # cs)))"
+ by (metis "1.IH" "1.prems" distinct.simps(2) id_outside_supp list.simps(9))
+ also have " ... = map (Fun.swap i j id) (i # (cs @ [j]))" by simp
+ also have " ... = j # (map (Fun.swap i j id) cs) @ [i]" by simp
+ also have " ... = j # cs @ [i]"
+ by (metis "1.prems" distinct.simps(2) list.set_intros(2) map_idI swap_id_eq)
+ 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 }
+ 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
+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)
+
+corollary cycle_is_id_root:
+ 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 "\<And>x. x \<in> (set cs) \<Longrightarrow> ((cycle_of_list cs) ^^ (length cs)) x = x"
+ using map_eq_conv by fastforce
+ moreover have "\<And>n x. x \<notin> (set cs) \<Longrightarrow> ((cycle_of_list cs) ^^ n) x = x"
+ proof -
+ fix n show "\<And>x. x \<notin> (set cs) \<Longrightarrow> ((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 "\<And>x. x \<notin> (set cs) \<Longrightarrow> ((cycle_of_list cs) ^^ (length cs)) x = x" by simp
+ ultimately show ?thesis
+ by (meson eq_id_iff)
+qed
+
+corollary
+ assumes "cycle cs"
+ shows "(cycle_of_list cs) = (cycle_of_list (rotate n cs))"
+proof -
+ { fix cs :: "'a list" assume A: "cycle cs"
+ have "(cycle_of_list cs) = (cycle_of_list (rotate1 cs))"
+ proof
+ have "\<And>x. x \<in> set cs \<Longrightarrow> (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 \<in> set cs"
+ ultimately show "(cycle_of_list cs) x = (cycle_of_list (rotate1 cs)) x" by auto
+ qed
+ moreover have "\<And>x. x \<notin> set cs \<Longrightarrow> (cycle_of_list cs) x = (cycle_of_list (rotate1 cs)) x"
+ using A by (simp add: id_outside_supp)
+ ultimately show "\<And>x. (cycle_of_list cs) x = (cycle_of_list (rotate1 cs)) x" 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
+qed
+
+
+subsection\<open>Conjugation of cycles\<close>
+
+lemma conjugation_of_cycle:
+ assumes "cycle cs" and "bij p"
+ shows "p \<circ> (cycle_of_list cs) \<circ> (inv p) = cycle_of_list (map p cs)"
+ using assms
+proof (induction cs rule: cycle_of_list.induct)
+ case (1 i j cs)
+ have "p \<circ> cycle_of_list (i # j # cs) \<circ> inv p =
+ (p \<circ> (Fun.swap i j id) \<circ> inv p) \<circ> (p \<circ> cycle_of_list (j # cs) \<circ> inv p)"
+ by (simp add: assms(2) bij_is_inj fun.map_comp)
+ also have " ... = (Fun.swap (p i) (p j) id) \<circ> (p \<circ> cycle_of_list (j # cs) \<circ> inv p)"
+ by (simp add: "1.prems"(2) bij_is_inj bij_swap_comp comp_swap o_assoc)
+ finally have "p \<circ> cycle_of_list (i # j # cs) \<circ> inv p =
+ (Fun.swap (p i) (p j) id) \<circ> (cycle_of_list (map p (j # cs)))"
+ using "1.IH" "1.prems"(1) assms(2) by fastforce
+ thus ?case by (metis cycle_of_list.simps(1) list.simps(9))
+next
+ case "2_1" thus ?case
+ by (metis bij_is_surj comp_id cycle_of_list.simps(2) list.simps(8) surj_iff)
+next
+ case "2_2" thus ?case
+ by (metis bij_is_surj comp_id cycle_of_list.simps(3) list.simps(8) list.simps(9) surj_iff)
+qed
+
+
+subsection\<open>When Cycles Commute\<close>
+
+lemma cycles_commute:
+ assumes "cycle \<sigma>1" "cycle \<sigma>2" and "set \<sigma>1 \<inter> set \<sigma>2 = {}"
+ shows "(cycle_of_list \<sigma>1) \<circ> (cycle_of_list \<sigma>2) = (cycle_of_list \<sigma>2) \<circ> (cycle_of_list \<sigma>1)"
+proof -
+ { fix \<pi>1 :: "'a list" and \<pi>2 :: "'a list" and x :: "'a"
+ assume A: "cycle \<pi>1" "cycle \<pi>2" "set \<pi>1 \<inter> set \<pi>2 = {}" "x \<in> set \<pi>1" "x \<notin> set \<pi>2"
+ have "((cycle_of_list \<pi>1) \<circ> (cycle_of_list \<pi>2)) x =
+ ((cycle_of_list \<pi>2) \<circ> (cycle_of_list \<pi>1)) x"
+ proof -
+ have "((cycle_of_list \<pi>1) \<circ> (cycle_of_list \<pi>2)) x = (cycle_of_list \<pi>1) x"
+ using id_outside_supp[OF A(2) A(5)] by simp
+ also have " ... = ((cycle_of_list \<pi>2) \<circ> (cycle_of_list \<pi>1)) x"
+ using id_outside_supp[OF A(2), of "(cycle_of_list \<pi>1) x"]
+ cycle_is_surj[OF A(1)] A(3) A(4) by fastforce
+ finally show ?thesis .
+ qed }
+ note aux_lemma = this
+
+ let ?\<sigma>12 = "\<lambda>x. ((cycle_of_list \<sigma>1) \<circ> (cycle_of_list \<sigma>2)) x"
+ let ?\<sigma>21 = "\<lambda>x. ((cycle_of_list \<sigma>2) \<circ> (cycle_of_list \<sigma>1)) x"
+
+ show ?thesis
+ proof
+ fix x have "x \<in> set \<sigma>1 \<union> set \<sigma>2 \<or> x \<notin> set \<sigma>1 \<union> set \<sigma>2" by blast
+ from this show "?\<sigma>12 x = ?\<sigma>21 x"
+ proof
+ assume "x \<in> set \<sigma>1 \<union> set \<sigma>2"
+ hence "(x \<in> set \<sigma>1 \<and> x \<notin> set \<sigma>2) \<or> (x \<notin> set \<sigma>1 \<and> x \<in> set \<sigma>2)" using assms(3) by blast
+ from this show "?\<sigma>12 x = ?\<sigma>21 x"
+ proof
+ assume "x \<in> set \<sigma>1 \<and> x \<notin> set \<sigma>2" thus ?thesis
+ using aux_lemma[OF assms(1-3)] by simp
+ next
+ assume "x \<notin> set \<sigma>1 \<and> x \<in> set \<sigma>2" thus ?thesis
+ using assms aux_lemma inf_commute by metis
+ qed
+ next
+ assume "x \<notin> set \<sigma>1 \<union> set \<sigma>2" thus ?thesis using id_outside_supp assms(1-2)
+ by (metis UnCI comp_apply)
+ qed
+ qed
+qed
+
+
+subsection\<open>Cycles from Permutations\<close>
+
+subsubsection\<open>Exponentiation of permutations\<close>
+
+text\<open>Some important properties of permutations before defining how to extract its cycles\<close>
+
+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 exp_of_permutation2:
+ assumes "p permutes S"
+ and "i < j" "(p ^^ j) = (p ^^ i)"
+ shows "(p ^^ (j - i)) = id" using assms
+proof -
+ have "(p ^^ i) \<circ> (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) \<circ> (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)
+qed
+
+lemma exp_of_permutation3:
+ assumes "p permutes S" "finite S"
+ shows "\<exists>n. (p ^^ n) = id \<and> n > 0"
+proof (rule ccontr)
+ assume "\<nexists>n. (p ^^ n) = id \<and> 0 < n"
+ hence S: "\<And>n. n > 0 \<Longrightarrow> (p ^^ n) \<noteq> id" by auto
+ hence "\<And>i j. \<lbrakk> i \<ge> 0; j \<ge> 0 \<rbrakk> \<Longrightarrow> i \<noteq> j \<Longrightarrow> (p ^^ i) \<noteq> (p ^^ j)"
+ proof -
+ fix i :: "nat" and j :: "nat" assume "i \<ge> 0" "j \<ge> 0" and Ineq: "i \<noteq> j"
+ show "(p ^^ i) \<noteq> (p ^^ j)"
+ proof (rule ccontr)
+ assume "\<not> (p ^^ i) \<noteq> (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
+ qed
+ hence "bij_betw (\<lambda>i. (p ^^ i)) {i :: nat . i \<ge> 0} {(p ^^ i) | i :: nat . i \<ge> 0}"
+ unfolding bij_betw_def inj_on_def by blast
+ hence "infinite {(p ^^ i) | i :: nat . i \<ge> 0}"
+ using bij_betw_finite by auto
+ moreover have "{(p ^^ i) | i :: nat . i \<ge> 0} \<subseteq> {\<pi>. \<pi> permutes S}"
+ using exp_of_permutation1[OF assms(1)] by blast
+ hence "finite {(p ^^ i) | i :: nat . i \<ge> 0}"
+ by (simp add: assms(2) finite_permutations finite_subset)
+ ultimately show False ..
+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)) \<circ> (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 "\<exists>n. (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
+ using funpow_mult[of n k p] id_funpow[of n] mult.commute[of k n] by smt
+qed
+
+
+subsubsection\<open>Extraction of cycles from permutations\<close>
+
+definition
+ least_power :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> nat"
+ where "least_power f x = (LEAST n. (f ^^ n) x = x \<and> n > 0)"
+
+abbreviation
+ support :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a list"
+ where "support p x \<equiv> map (\<lambda>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 "\<exists>n. (p ^^ n) x = x \<and> 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_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: "\<exists>n. (p ^^ n) x = x \<and> n > 0"
+ using eq_id_iff exp_of_permutation3 by metis
+ assume "\<not> 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_gt_one:
+ assumes "permutation p" "p x \<noteq> 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: "\<exists>n. (p ^^ n) x = x \<and> n > 0"
+ using eq_id_iff exp_of_permutation3 by metis
+ assume "\<not> 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 \<open>p x = x\<close> and \<open>p x \<noteq> x\<close> show False by simp
+qed
+
+lemma least_power_bound:
+ assumes "permutation p" shows "\<exists>m > 0. (least_power p x) \<le> m"
+proof -
+ obtain S where "p permutes S" "finite S"
+ using assms permutation_permutes by blast
+ hence "\<exists>n. (p ^^ n) x = x \<and> 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
+qed
+
+lemma lt_least_power:
+ assumes "Suc n = least_power p x"
+ and "0 < i" "i \<le> n"
+ shows "(p ^^ i) x \<noteq> x"
+proof (rule ccontr)
+ assume "\<not> (p ^^ i) x \<noteq> x" hence "(p ^^ i) x = x" by simp
+ hence "i \<in> {n. (p ^^ n) x = x \<and> 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
+qed
+
+lemma least_power_welldefined:
+ assumes "permutation p" and "y \<in> {(p ^^ k) x | k. k \<ge> 0}"
+ shows "least_power p x = least_power p y"
+proof -
+ have aux_lemma: "\<And>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 \<ge> least_power p (p z)"
+ by (metis assms(1) inc_induct le_SucE least_power_gt_zero lt_least_power nat_le_linear)
+
+ 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 \<le> 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 \<ge> 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
+ qed
+qed
+
+theorem cycle_of_permutation:
+ assumes "permutation p"
+ shows "cycle (support p x)"
+proof (rule ccontr)
+ assume "\<not> cycle (support p x)"
+ hence "\<exists> i j. i \<in> {0..<least_power p x} \<and> j \<in> {0..<least_power p x} \<and> i \<noteq> j \<and> (p ^^ i) x = (p ^^ j) x"
+ using atLeast_upt by (simp add: distinct_conv_nth)
+ then obtain i j where ij: "0 \<le> 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 \<le> j - i \<and> j - i < least_power p x"
+ by (simp add: ij(3) less_imp_diff_less)
+ hence "(p ^^ (j - i)) x \<noteq> 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
+qed
+
+
+subsection\<open>Decomposition on Cycles\<close>
+
+text\<open>We show that a permutation can be decomposed on cycles\<close>
+
+subsubsection\<open>Preliminaries\<close>
+
+lemma support_set:
+ assumes "permutation p"
+ shows "set (support p x) = {(p ^^ k) x | k. k \<ge> 0}"
+proof -
+ have "{(p ^^ k) x | k. k \<ge> 0} = {(p ^^ k) x | k. 0 \<le> k \<and> k < (least_power p x)}" (is "?A = ?B")
+ proof
+ show "?B \<subseteq> ?A" by blast
+ next
+ show "?A \<subseteq> ?B"
+ proof
+ fix y assume "y \<in> ?A"
+ then obtain k :: "nat" where k: "k \<ge> 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 \<in> ?B"
+ by blast
+ qed
+ qed
+
+ moreover have "{(p ^^ k) x | k. 0 \<le> k \<and> k < (least_power p x)} = set (support p x)" (is "?B = ?C")
+ proof
+ show "?B \<subseteq> ?C"
+ proof
+ fix y assume "y \<in> {(p ^^ k) x | k. 0 \<le> k \<and> k < (least_power p x)}"
+ then obtain k where "k \<ge> 0" "k < (least_power p x)" "y = (p ^^ k) x" by blast
+ thus "y \<in> ?C" by auto
+ qed
+ next
+ show "?C \<subseteq> ?B"
+ proof
+ fix y assume "y \<in> ?C"
+ then obtain k where "k \<ge> 0" "k < (least_power p x)" "(support p x) ! k = y" by auto
+ thus "y \<in> ?B" by auto
+ qed
+ qed
+
+ ultimately show ?thesis by simp
+qed
+
+lemma disjoint_support:
+ assumes "p permutes S" "finite S"
+ shows "disjoint {{(p ^^ k) x | k. k \<ge> 0} | x. x \<in> S}" (is "disjoint ?A")
+proof (rule disjointI)
+ { fix a b assume "a \<in> ?A" "b \<in> ?A" "a \<inter> b \<noteq> {}"
+ then obtain x y where x: "x \<in> S" "a = {(p ^^ k) x | k. k \<ge> 0}"
+ and y: "y \<in> S" "b = {(p ^^ k) y | k. k \<ge> 0}" by blast
+ assume "a \<inter> b \<noteq> {}"
+ then obtain z kx ky where z: "kx \<ge> 0" "ky \<ge> 0" "z = (p ^^ kx) x" "z = (p ^^ ky) y"
+ using x(2) y(2) by blast
+ have "a \<subseteq> b"
+ proof
+ fix w assume "w \<in> a"
+ then obtain k where k: "k \<ge> 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 \<in> b" using y by blast
+ qed } note aux_lemma = this
+
+ fix a b assume ab: "a \<in> ?A" "b \<in> ?A" "a \<noteq> b"
+ show "a \<inter> b = {}"
+ proof (rule ccontr)
+ assume "a \<inter> b \<noteq> {}" 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 "\<Union>{{(p ^^ k) x | k. k \<ge> 0} | x. x \<in> S} = S"
+proof
+ show "\<Union>{{(p ^^ k) x |k. 0 \<le> k} |x. x \<in> S} \<subseteq> S"
+ proof
+ fix y assume "y \<in> \<Union>{{(p ^^ k) x |k. 0 \<le> k} |x. x \<in> S}"
+ then obtain x k where x: "x \<in> S" and k: "k \<ge> 0" and y: "y = (p ^^ k) x" by blast
+ have "(p ^^ k) x \<in> 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 \<in> S" using y by simp
+ qed
+next
+ show "S \<subseteq> \<Union>{{(p ^^ k) x |k. 0 \<le> k} |x. x \<in> S}"
+ proof
+ fix x assume x: "x \<in> S"
+ hence "x \<in> {(p ^^ k) x |k. 0 \<le> k}"
+ by (metis (mono_tags, lifting) CollectI funpow_0 le_numeral_extra(3))
+ thus "x \<in> \<Union>{{(p ^^ k) x |k. 0 \<le> k} |x. x \<in> S}" using x by blast
+ qed
+qed
+
+theorem cycle_restrict:
+ assumes "permutation p" "y \<in> set (support p x)"
+ shows "p y = (cycle_of_list (support p x)) y"
+proof -
+ have "\<And> i. \<lbrakk> 0 \<le> i; i < length (support p x) - 1 \<rbrakk> \<Longrightarrow>
+ p ((support p x) ! i) = (support p x) ! (i + 1)"
+ proof -
+ fix i assume i: "0 \<le> 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)" .
+ qed
+ hence 1: "map p (butlast (support p x)) = tl (support p x)"
+ using nth_equalityI[of "map p (butlast (support p x))" "tl (support p x)"]
+ by (smt Suc_eq_plus1 le0 length_butlast length_map length_tl nth_butlast nth_map nth_tl)
+
+ have "p ((support p x) ! (length (support p x) - 1)) = p ((p ^^ (length (support p x) - 1)) x)"
+ 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\<open>Decomposition\<close>
+
+inductive cycle_decomp :: "'a set \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool" where
+empty: "cycle_decomp {} id" |
+comp: "\<lbrakk> cycle_decomp I p; cycle cs; set cs \<inter> I = {} \<rbrakk> \<Longrightarrow>
+ cycle_decomp (set cs \<union> I) ((cycle_of_list cs) \<circ> p)"
+
+
+lemma semidecomposition:
+ assumes "p permutes S" "finite S"
+ shows "(\<lambda>y. if y \<in> (S - set (support p x)) then p y else y) permutes (S - set (support p x))"
+ (is "?q permutes ?S'")
+proof -
+ have "\<And>y. y \<notin> S \<Longrightarrow> 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 "\<And>y. y \<in> set (support p x) \<Longrightarrow> p y \<in> set (support p x)"
+ using cycle_restrict assms permutation_permutes by (metis imageI)
+
+ ultimately
+ have 1: "\<And>y. y \<notin> ?S' \<Longrightarrow> p y \<notin> ?S'" by auto
+ have 2: "\<And>y. y \<in> ?S' \<Longrightarrow> p y \<in> ?S'"
+ proof -
+ fix y assume y: "y \<in> ?S'" show "p y \<in> ?S'"
+ proof (rule ccontr)
+ assume "p y \<notin> ?S'" hence "p y \<in> set (support p x)"
+ using assms(1) permutes_in_image y by fastforce
+ then obtain y' where y': "y' \<in> 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 "\<And> y. y \<in> ?S' \<Longrightarrow> \<exists>!x. p x = y"
+ by (metis assms(1) permutes_def)
+ hence "\<And> y. y \<in> ?S' \<Longrightarrow> \<exists>x \<in> ?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 (smt bij_betw_cong)
+ moreover have "\<And>y. y \<notin> ?S' \<Longrightarrow> ?q y = y" by auto
+ ultimately show ?thesis
+ using bij_imp_permutes by blast
+qed
+
+
+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
+ (* using less empty by auto *)
+ by (metis empty less.prems(1) permutes_empty)
+ next
+ case False
+ then obtain x where x: "x \<in> S" by blast
+ define S' :: "'a set" where S': "S' = S - set (support p x)"
+ define q :: "'a \<Rightarrow> 'a" where q: "q = (\<lambda>x. if x \<in> 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 \<in> set (support p x)"
+ by (smt 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' \<open>x \<in> S\<close> 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)) \<circ> q"
+ proof
+ fix y show "p y = ((cycle_of_list (support p x)) \<circ> q) y"
+ proof (cases)
+ assume y: "y \<in> set (support p x)" hence "y \<notin> S'" using S' by simp
+ hence "q y = y" using q by simp
+ thus ?thesis
+ (* using cycle_restrict less.prems(1) less.prems(2) permutation_permutes y by fastforce *)
+ using comp_apply cycle_restrict less.prems permutation_permutes y by fastforce
+ next
+ assume y: "y \<notin> set (support p x)" thus ?thesis
+ proof (cases)
+ assume "y \<in> S'"
+ hence "q y \<in> S'" using q_permutes
+ by (simp add: permutes_in_image)
+ hence "q y \<notin> 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: \<open>y \<in> S'\<close> q)
+ next
+ assume "y \<notin> S'" hence "y \<notin> S" using y S' by blast
+ hence "(cycle_of_list (support p x) \<circ> q) y = (cycle_of_list (support p x)) y"
+ by (simp add: \<open>y \<notin> S'\<close> 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 \<open>y \<notin> S\<close> less.prems(1) permutes_def)
+ finally show ?thesis by simp
+ qed
+ qed
+ qed
+ moreover have "cycle (support p x)"
+ using cycle_of_permutation less.prems permutation_permutes by fastforce
+ moreover have "set (support p x) \<inter> S' = {}" using S' by simp
+ moreover have "set (support p x) \<subseteq> 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) \<union> S'" using S' by blast
+ ultimately show ?thesis using comp[of S' q "support p x"] 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
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Algebra/Embedded_Algebras.thy Mon Jul 02 15:43:22 2018 +0100
@@ -0,0 +1,1318 @@
+(* ************************************************************************** *)
+(* Title: Embedded_Algebras.thy *)
+(* Author: Paulo EmÃlio de Vilhena *)
+(* ************************************************************************** *)
+
+theory Embedded_Algebras
+ imports Subrings Generated_Groups
+
+begin
+
+section \<open>Definitions\<close>
+
+locale embedded_algebra =
+ K?: subfield K R + R?: ring R for K :: "'a set" and R :: "('a, 'b) ring_scheme" (structure)
+
+definition (in ring) line_extension :: "'a set \<Rightarrow> 'a \<Rightarrow> 'a set \<Rightarrow> 'a set"
+ where "line_extension K a E = (K #> a) <+>\<^bsub>R\<^esub> E"
+
+fun (in ring) Span :: "'a set \<Rightarrow> 'a list \<Rightarrow> 'a set"
+ where "Span K Us = foldr (line_extension K) Us { \<zero> }"
+
+fun (in ring) combine :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a"
+ where
+ "combine (k # Ks) (u # Us) = (k \<otimes> u) \<oplus> (combine Ks Us)"
+ | "combine Ks Us = \<zero>"
+
+inductive (in ring) independent :: "'a set \<Rightarrow> 'a list \<Rightarrow> bool"
+ where
+ li_Nil [simp, intro]: "independent K []"
+ | li_Cons: "\<lbrakk> u \<in> carrier R; u \<notin> Span K Us; independent K Us \<rbrakk> \<Longrightarrow> independent K (u # Us)"
+
+inductive (in ring) dimension :: "nat \<Rightarrow> 'a set \<Rightarrow> 'a set \<Rightarrow> bool"
+ where
+ zero_dim [simp, intro]: "dimension 0 K { \<zero> }"
+ | Suc_dim: "\<lbrakk> v \<in> carrier R; v \<notin> E; dimension n K E \<rbrakk> \<Longrightarrow> dimension (Suc n) K (line_extension K v E)"
+
+
+subsubsection \<open>Syntactic Definitions\<close>
+
+abbreviation (in ring) dependent :: "'a set \<Rightarrow> 'a list \<Rightarrow> bool"
+ where "dependent K U \<equiv> \<not> independent K U"
+
+definition over :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" (infixl "over" 65)
+ where "f over a = f a"
+
+
+
+context ring
+begin
+
+
+subsection \<open>Basic Properties - First Part\<close>
+
+lemma combine_in_carrier [simp, intro]:
+ "\<lbrakk> set Ks \<subseteq> carrier R; set Us \<subseteq> carrier R \<rbrakk> \<Longrightarrow> combine Ks Us \<in> carrier R"
+ by (induct Ks Us rule: combine.induct) (auto)
+
+lemma combine_r_distr:
+ "\<lbrakk> set Ks \<subseteq> carrier R; set Us \<subseteq> carrier R \<rbrakk> \<Longrightarrow>
+ k \<in> carrier R \<Longrightarrow> k \<otimes> (combine Ks Us) = combine (map ((\<otimes>) k) Ks) Us"
+ by (induct Ks Us rule: combine.induct) (auto simp add: m_assoc r_distr)
+
+lemma combine_l_distr:
+ "\<lbrakk> set Ks \<subseteq> carrier R; set Us \<subseteq> carrier R \<rbrakk> \<Longrightarrow>
+ u \<in> carrier R \<Longrightarrow> (combine Ks Us) \<otimes> u = combine Ks (map (\<lambda>u'. u' \<otimes> u) Us)"
+ by (induct Ks Us rule: combine.induct) (auto simp add: m_assoc l_distr)
+
+lemma combine_eq_foldr:
+ "combine Ks Us = foldr (\<lambda>(k, u). \<lambda>l. (k \<otimes> u) \<oplus> l) (zip Ks Us) \<zero>"
+ by (induct Ks Us rule: combine.induct) (auto)
+
+lemma combine_replicate:
+ "set Us \<subseteq> carrier R \<Longrightarrow> combine (replicate (length Us) \<zero>) Us = \<zero>"
+ by (induct Us) (auto)
+
+lemma combine_append:
+ assumes "length Ks = length Us"
+ and "set Ks \<subseteq> carrier R" "set Us \<subseteq> carrier R"
+ and "set Ks' \<subseteq> carrier R" "set Vs \<subseteq> carrier R"
+ shows "(combine Ks Us) \<oplus> (combine Ks' Vs) = combine (Ks @ Ks') (Us @ Vs)"
+ using assms
+proof (induct Ks arbitrary: Us)
+ case Nil thus ?case by auto
+next
+ case (Cons k Ks)
+ then obtain u Us' where Us: "Us = u # Us'"
+ by (metis length_Suc_conv)
+ hence u: "u \<in> carrier R" and Us': "set Us' \<subseteq> carrier R"
+ using Cons(4) by auto
+ then show ?case
+ using combine_in_carrier[OF _ Us', of Ks] Cons
+ combine_in_carrier[OF Cons(5-6)] unfolding Us
+ by (auto, simp add: add.m_assoc)
+qed
+
+lemma combine_add:
+ assumes "length Ks = length Us" and "length Ks' = length Us"
+ and "set Ks \<subseteq> carrier R" "set Ks' \<subseteq> carrier R" "set Us \<subseteq> carrier R"
+ shows "(combine Ks Us) \<oplus> (combine Ks' Us) = combine (map2 (\<oplus>) Ks Ks') Us"
+ using assms
+proof (induct Us arbitrary: Ks Ks')
+ case Nil thus ?case by simp
+next
+ case (Cons u Us)
+ then obtain c c' Cs Cs' where Ks: "Ks = c # Cs" and Ks': "Ks' = c' # Cs'"
+ by (metis length_Suc_conv)
+ hence in_carrier:
+ "c \<in> carrier R" "set Cs \<subseteq> carrier R"
+ "c' \<in> carrier R" "set Cs' \<subseteq> carrier R"
+ "u \<in> carrier R" "set Us \<subseteq> carrier R"
+ using Cons(4-6) by auto
+ hence lc_in_carrier: "combine Cs Us \<in> carrier R" "combine Cs' Us \<in> carrier R"
+ using combine_in_carrier by auto
+ have "combine Ks (u # Us) \<oplus> combine Ks' (u # Us) =
+ ((c \<otimes> u) \<oplus> combine Cs Us) \<oplus> ((c' \<otimes> u) \<oplus> combine Cs' Us)"
+ unfolding Ks Ks' by auto
+ also have " ... = ((c \<oplus> c') \<otimes> u \<oplus> (combine Cs Us \<oplus> combine Cs' Us))"
+ using lc_in_carrier in_carrier(1,3,5) by (simp add: l_distr ring_simprules(7,22))
+ also have " ... = combine (map2 (\<oplus>) Ks Ks') (u # Us)"
+ using Cons unfolding Ks Ks' by auto
+ finally show ?case .
+qed
+
+
+subsection \<open>Some Basic Properties of Linear_Ind\<close>
+
+lemma independent_in_carrier: "independent K Us \<Longrightarrow> set Us \<subseteq> carrier R"
+ by (induct Us rule: independent.induct) (simp_all)
+
+lemma independent_backwards:
+ "independent K (u # Us) \<Longrightarrow> u \<notin> Span K Us"
+ "independent K (u # Us) \<Longrightarrow> independent K Us"
+ "independent K (u # Us) \<Longrightarrow> u \<in> carrier R"
+ by (cases rule: independent.cases, auto)+
+
+
+text \<open>Now, we fix K, a subfield of the ring. Many lemmas would also be true for weaker
+ structures, but our interest is to work with subfields, so generalization could
+ be the subjuct of a future work.\<close>
+
+context
+ fixes K :: "'a set" assumes K: "subfield K R"
+begin
+
+
+subsection \<open>Basic Properties - Second Part\<close>
+
+lemmas subring_props [simp] =
+ subringE[OF subfieldE(1)[OF K]]
+
+lemma line_extension_mem_iff: "u \<in> line_extension K a E \<longleftrightarrow> (\<exists>k \<in> K. \<exists>v \<in> E. u = k \<otimes> a \<oplus> v)"
+ unfolding line_extension_def set_add_def'[of R "K #> a" E] unfolding r_coset_def by blast
+
+lemma line_extension_is_subgroup:
+ assumes "subgroup E (add_monoid R)" "a \<in> carrier R"
+ shows "subgroup (line_extension K a E) (add_monoid R)"
+proof (rule add.subgroupI)
+ show "line_extension K a E \<subseteq> carrier R"
+ by (simp add: assms add.subgroupE(1) line_extension_def r_coset_subset_G set_add_closed)
+next
+ have "\<zero> = \<zero> \<otimes> a \<oplus> \<zero>"
+ using assms(2) by simp
+ hence "\<zero> \<in> line_extension K a E"
+ using line_extension_mem_iff subgroup.one_closed[OF assms(1)] by auto
+ thus "line_extension K a E \<noteq> {}" by auto
+next
+ fix u1 u2
+ assume "u1 \<in> line_extension K a E" and "u2 \<in> line_extension K a E"
+ then obtain k1 k2 v1 v2
+ where u1: "k1 \<in> K" "v1 \<in> E" "u1 = (k1 \<otimes> a) \<oplus> v1"
+ and u2: "k2 \<in> K" "v2 \<in> E" "u2 = (k2 \<otimes> a) \<oplus> v2"
+ and in_carr: "k1 \<in> carrier R" "v1 \<in> carrier R" "k2 \<in> carrier R" "v2 \<in> carrier R"
+ using line_extension_mem_iff by (meson add.subgroupE(1)[OF assms(1)] subring_props(1) subsetCE)
+
+ hence "u1 \<oplus> u2 = ((k1 \<oplus> k2) \<otimes> a) \<oplus> (v1 \<oplus> v2)"
+ using assms(2) by algebra
+ moreover have "k1 \<oplus> k2 \<in> K" and "v1 \<oplus> v2 \<in> E"
+ using add.subgroupE(4)[OF assms(1)] u1 u2 by auto
+ ultimately show "u1 \<oplus> u2 \<in> line_extension K a E"
+ using line_extension_mem_iff by auto
+
+ have "\<ominus> u1 = ((\<ominus> k1) \<otimes> a) \<oplus> (\<ominus> v1)"
+ using in_carr(1-2) u1(3) assms(2) by algebra
+ moreover have "\<ominus> k1 \<in> K" and "\<ominus> v1 \<in> E"
+ using add.subgroupE(3)[OF assms(1)] u1 by auto
+ ultimately show "(\<ominus> u1) \<in> line_extension K a E"
+ using line_extension_mem_iff by auto
+qed
+
+corollary Span_is_add_subgroup:
+ "set Us \<subseteq> carrier R \<Longrightarrow> subgroup (Span K Us) (add_monoid R)"
+ using line_extension_is_subgroup add.normal_invE(1)[OF add.one_is_normal] by (induct Us) (auto)
+
+lemma line_extension_smult_closed:
+ assumes "\<And>k v. \<lbrakk> k \<in> K; v \<in> E \<rbrakk> \<Longrightarrow> k \<otimes> v \<in> E" and "E \<subseteq> carrier R" "a \<in> carrier R"
+ shows "\<And>k u. \<lbrakk> k \<in> K; u \<in> line_extension K a E \<rbrakk> \<Longrightarrow> k \<otimes> u \<in> line_extension K a E"
+proof -
+ fix k u assume A: "k \<in> K" "u \<in> line_extension K a E"
+ then obtain k' v'
+ where u: "k' \<in> K" "v' \<in> E" "u = k' \<otimes> a \<oplus> v'"
+ and in_carr: "k \<in> carrier R" "k' \<in> carrier R" "v' \<in> carrier R"
+ using line_extension_mem_iff assms(2) by (meson subring_props(1) subsetCE)
+ hence "k \<otimes> u = (k \<otimes> k') \<otimes> a \<oplus> (k \<otimes> v')"
+ using assms(3) by algebra
+ thus "k \<otimes> u \<in> line_extension K a E"
+ using assms(1)[OF A(1) u(2)] line_extension_mem_iff u(1) A(1) by auto
+qed
+
+lemma Span_subgroup_props [simp]:
+ assumes "set Us \<subseteq> carrier R"
+ shows "Span K Us \<subseteq> carrier R"
+ and "\<zero> \<in> Span K Us"
+ and "\<And>v1 v2. \<lbrakk> v1 \<in> Span K Us; v2 \<in> Span K Us \<rbrakk> \<Longrightarrow> (v1 \<oplus> v2) \<in> Span K Us"
+ and "\<And>v. v \<in> Span K Us \<Longrightarrow> (\<ominus> v) \<in> Span K Us"
+ using add.subgroupE subgroup.one_closed[of _ "add_monoid R"]
+ Span_is_add_subgroup[OF assms(1)] by auto
+
+lemma Span_smult_closed [simp]:
+ assumes "set Us \<subseteq> carrier R"
+ shows "\<And>k v. \<lbrakk> k \<in> K; v \<in> Span K Us \<rbrakk> \<Longrightarrow> k \<otimes> v \<in> Span K Us"
+ using assms
+proof (induct Us)
+ case Nil thus ?case
+ using r_null subring_props(1) by (auto, blast)
+next
+ case Cons thus ?case
+ using Span_subgroup_props(1) line_extension_smult_closed by auto
+qed
+
+lemma Span_m_inv_simprule [simp]:
+ assumes "set Us \<subseteq> carrier R"
+ shows "\<lbrakk> k \<in> K - { \<zero> }; a \<in> carrier R \<rbrakk> \<Longrightarrow> k \<otimes> a \<in> Span K Us \<Longrightarrow> a \<in> Span K Us"
+proof -
+ assume k: "k \<in> K - { \<zero> }" and a: "a \<in> carrier R" and ka: "k \<otimes> a \<in> Span K Us"
+ have inv_k: "inv k \<in> K" "inv k \<otimes> k = \<one>"
+ using subfield_m_inv[OF K k] by simp+
+ hence "inv k \<otimes> (k \<otimes> a) \<in> Span K Us"
+ using Span_smult_closed[OF assms _ ka] by simp
+ thus ?thesis
+ using inv_k subring_props(1)a k by (smt DiffD1 l_one m_assoc set_rev_mp)
+qed
+
+
+subsection \<open>Span as Linear Combinations\<close>
+
+text \<open>We show that Span is the set of linear combinations\<close>
+
+lemma line_extension_of_combine_set:
+ assumes "u \<in> carrier R"
+ shows "line_extension K u { combine Ks Us | Ks. set Ks \<subseteq> K } =
+ { combine Ks (u # Us) | Ks. set Ks \<subseteq> K }"
+ (is "?line_extension = ?combinations")
+proof
+ show "?line_extension \<subseteq> ?combinations"
+ proof
+ fix v assume "v \<in> ?line_extension"
+ then obtain k Ks
+ where "k \<in> K" "set Ks \<subseteq> K" and "v = combine (k # Ks) (u # Us)"
+ using line_extension_mem_iff by auto
+ thus "v \<in> ?combinations"
+ by (metis (mono_tags, lifting) insert_subset list.simps(15) mem_Collect_eq)
+ qed
+next
+ show "?combinations \<subseteq> ?line_extension"
+ proof
+ fix v assume "v \<in> ?combinations"
+ then obtain Ks where v: "set Ks \<subseteq> K" "v = combine Ks (u # Us)"
+ by auto
+ thus "v \<in> ?line_extension"
+ proof (cases Ks)
+ case Cons thus ?thesis
+ using v line_extension_mem_iff by auto
+ next
+ case Nil
+ hence "v = \<zero>"
+ using v by simp
+ moreover have "combine [] Us = \<zero>" by simp
+ hence "\<zero> \<in> { combine Ks Us | Ks. set Ks \<subseteq> K }"
+ by (metis (mono_tags, lifting) local.Nil mem_Collect_eq v(1))
+ hence "(\<zero> \<otimes> u) \<oplus> \<zero> \<in> ?line_extension"
+ using line_extension_mem_iff subring_props(2) by blast
+ hence "\<zero> \<in> ?line_extension"
+ using assms by auto
+ ultimately show ?thesis by auto
+ qed
+ qed
+qed
+
+lemma Span_eq_combine_set:
+ assumes "set Us \<subseteq> carrier R" shows "Span K Us = { combine Ks Us | Ks. set Ks \<subseteq> K }"
+ using assms line_extension_of_combine_set
+ by (induct Us) (auto, metis empty_set empty_subsetI)
+
+lemma line_extension_of_combine_set_length_version:
+ assumes "u \<in> carrier R"
+ shows "line_extension K u { combine Ks Us | Ks. length Ks = length Us \<and> set Ks \<subseteq> K } =
+ { combine Ks (u # Us) | Ks. length Ks = length (u # Us) \<and> set Ks \<subseteq> K }"
+ (is "?line_extension = ?combinations")
+proof
+ show "?line_extension \<subseteq> ?combinations"
+ proof
+ fix v assume "v \<in> ?line_extension"
+ then obtain k Ks
+ where "v = combine (k # Ks) (u # Us)" "length (k # Ks) = length (u # Us)" "set (k # Ks) \<subseteq> K"
+ using line_extension_mem_iff by auto
+ thus "v \<in> ?combinations" by blast
+ qed
+next
+ show "?combinations \<subseteq> ?line_extension"
+ proof
+ fix c assume "c \<in> ?combinations"
+ then obtain Ks where c: "c = combine Ks (u # Us)" "length Ks = length (u # Us)" "set Ks \<subseteq> K"
+ by blast
+ then obtain k Ks' where k: "Ks = k # Ks'"
+ by (metis length_Suc_conv)
+ thus "c \<in> ?line_extension"
+ using c line_extension_mem_iff unfolding k by auto
+ qed
+qed
+
+lemma Span_eq_combine_set_length_version:
+ assumes "set Us \<subseteq> carrier R"
+ shows "Span K Us = { combine Ks Us | Ks. length Ks = length Us \<and> set Ks \<subseteq> K }"
+ using assms line_extension_of_combine_set_length_version by (induct Us) (auto)
+
+
+subsubsection \<open>Corollaries\<close>
+
+corollary Span_mem_iff:
+ assumes "set Us \<subseteq> carrier R" and "a \<in> carrier R"
+ shows "a \<in> Span K Us \<longleftrightarrow> (\<exists>k \<in> K - { \<zero> }. \<exists>Ks. set Ks \<subseteq> K \<and> combine (k # Ks) (a # Us) = \<zero>)"
+ (is "?in_Span \<longleftrightarrow> ?exists_combine")
+proof
+ assume "?in_Span"
+ then obtain Ks where Ks: "set Ks \<subseteq> K" "a = combine Ks Us"
+ using Span_eq_combine_set[OF assms(1)] by auto
+ hence "((\<ominus> \<one>) \<otimes> a) \<oplus> a = combine ((\<ominus> \<one>) # Ks) (a # Us)"
+ by auto
+ moreover have "((\<ominus> \<one>) \<otimes> a) \<oplus> a = \<zero>"
+ using assms(2) l_minus l_neg by auto
+ moreover have "\<ominus> \<one> \<noteq> \<zero>"
+ using subfieldE(6)[OF K] l_neg by force
+ ultimately show "?exists_combine"
+ using subring_props(3,5) Ks(1) by (force simp del: combine.simps)
+next
+ assume "?exists_combine"
+ then obtain k Ks
+ where k: "k \<in> K" "k \<noteq> \<zero>" and Ks: "set Ks \<subseteq> K" and a: "(k \<otimes> a) \<oplus> combine Ks Us = \<zero>"
+ by auto
+ hence "combine Ks Us \<in> Span K Us"
+ using Span_eq_combine_set[OF assms(1)] by auto
+ hence "k \<otimes> a \<in> Span K Us"
+ using Span_subgroup_props[OF assms(1)] k Ks a
+ by (metis (no_types, lifting) assms(2) contra_subsetD m_closed minus_equality subring_props(1))
+ thus "?in_Span"
+ using Span_m_inv_simprule[OF assms(1) _ assms(2), of k] k by auto
+qed
+
+corollary Span_mem_iff_length_version:
+ assumes "set Us \<subseteq> carrier R"
+ shows "a \<in> Span K Us \<longleftrightarrow> (\<exists>Ks. set Ks \<subseteq> K \<and> length Ks = length Us \<and> a = combine Ks Us)"
+ using Span_eq_combine_set_length_version[OF assms] by blast
+
+
+subsection \<open>Span as the minimal subgroup that contains K <#> (set Us)\<close>
+
+text \<open>Now we show the link between Span and Group.generate\<close>
+
+lemma mono_Span:
+ assumes "set Us \<subseteq> carrier R" and "u \<in> carrier R"
+ shows "Span K Us \<subseteq> Span K (u # Us)"
+proof
+ fix v assume v: "v \<in> Span K Us"
+ hence "(\<zero> \<otimes> u) \<oplus> v \<in> Span K (u # Us)"
+ using line_extension_mem_iff by auto
+ thus "v \<in> Span K (u # Us)"
+ using Span_subgroup_props(1)[OF assms(1)] assms(2) v
+ by (auto simp del: Span.simps)
+qed
+
+lemma Span_min:
+ assumes "set Us \<subseteq> carrier R" and "subgroup E (add_monoid R)"
+ shows "K <#> (set Us) \<subseteq> E \<Longrightarrow> Span K Us \<subseteq> E"
+proof -
+ assume "K <#> (set Us) \<subseteq> E" show "Span K Us \<subseteq> E"
+ proof
+ fix v assume "v \<in> Span K Us"
+ then obtain Ks where v: "set Ks \<subseteq> K" "v = combine Ks Us"
+ using Span_eq_combine_set[OF assms(1)] by auto
+ from \<open>set Ks \<subseteq> K\<close> \<open>set Us \<subseteq> carrier R\<close> and \<open>K <#> (set Us) \<subseteq> E\<close>
+ show "v \<in> E" unfolding v(2)
+ proof (induct Ks Us rule: combine.induct)
+ case (1 k Ks u Us)
+ hence "k \<in> K" and "u \<in> set (u # Us)" by auto
+ hence "k \<otimes> u \<in> E"
+ using 1(4) unfolding set_mult_def by auto
+ moreover have "K <#> set Us \<subseteq> E"
+ using 1(4) unfolding set_mult_def by auto
+ hence "combine Ks Us \<in> E"
+ using 1 by auto
+ ultimately show ?case
+ using add.subgroupE(4)[OF assms(2)] by auto
+ next
+ case "2_1" thus ?case
+ using subgroup.one_closed[OF assms(2)] by auto
+ next
+ case "2_2" thus ?case
+ using subgroup.one_closed[OF assms(2)] by auto
+ qed
+ qed
+qed
+
+lemma Span_eq_generate:
+ assumes "set Us \<subseteq> carrier R" shows "Span K Us = generate (add_monoid R) (K <#> (set Us))"
+proof (rule add.generateI)
+ show "K <#> set Us \<subseteq> 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
+ show "\<And>E. \<lbrakk> subgroup E (add_monoid R); K <#> set Us \<subseteq> E \<rbrakk> \<Longrightarrow> Span K Us \<subseteq> E"
+ using Span_min assms by blast
+next
+ show "K <#> set Us \<subseteq> Span K Us"
+ using assms
+ proof (induct Us)
+ case Nil thus ?case
+ unfolding set_mult_def by auto
+ next
+ case (Cons u Us)
+ have "K <#> set (u # Us) = (K <#> { u }) \<union> (K <#> set Us)"
+ unfolding set_mult_def by auto
+ moreover have "\<And>k. k \<in> K \<Longrightarrow> k \<otimes> u \<in> Span K (u # Us)"
+ proof -
+ fix k assume k: "k \<in> K"
+ hence "combine [ k ] (u # Us) \<in> Span K (u # Us)"
+ using Span_eq_combine_set[OF Cons(2)] by (auto simp del: combine.simps)
+ moreover have "k \<in> carrier R" and "u \<in> carrier R"
+ using Cons(2) k subring_props(1) by (blast, auto)
+ ultimately show "k \<otimes> u \<in> Span K (u # Us)"
+ by (auto simp del: Span.simps)
+ qed
+ hence "K <#> { u } \<subseteq> Span K (u # Us)"
+ unfolding set_mult_def by auto
+ moreover have "K <#> set Us \<subseteq> Span K (u # Us)"
+ using mono_Span[of Us u] Cons by (auto simp del: Span.simps)
+ ultimately show ?case
+ using Cons by (auto simp del: Span.simps)
+ qed
+qed
+
+
+subsubsection \<open>Corollaries\<close>
+
+corollary Span_same_set:
+ assumes "set Us \<subseteq> carrier R"
+ shows "set Us = set Vs \<Longrightarrow> Span K Us = Span K Vs"
+ using Span_eq_generate assms by auto
+
+corollary Span_incl: "set Us \<subseteq> carrier R \<Longrightarrow> K <#> (set Us) \<subseteq> Span K Us"
+ using Span_eq_generate generate.incl[of _ _ "add_monoid R"] by auto
+
+corollary Span_base_incl: "set Us \<subseteq> carrier R \<Longrightarrow> set Us \<subseteq> Span K Us"
+proof -
+ assume A: "set Us \<subseteq> carrier R"
+ hence "{ \<one> } <#> set Us = set Us"
+ unfolding set_mult_def by force
+ moreover have "{ \<one> } <#> set Us \<subseteq> K <#> set Us"
+ using subring_props(3) unfolding set_mult_def by blast
+ ultimately show ?thesis
+ using Span_incl[OF A] by auto
+qed
+
+corollary mono_Span_sublist:
+ assumes "set Us \<subseteq> set Vs" "set Vs \<subseteq> carrier R"
+ shows "Span K Us \<subseteq> 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)]]
+ Span_eq_generate[OF assms(2)] Span_eq_generate[of Us] assms by auto
+
+corollary mono_Span_append:
+ assumes "set Us \<subseteq> carrier R" "set Vs \<subseteq> carrier R"
+ shows "Span K Us \<subseteq> Span K (Us @ Vs)"
+ and "Span K Us \<subseteq> Span K (Vs @ Us)"
+ using mono_Span_sublist[of Us "Us @ Vs"] assms
+ Span_same_set[of "Us @ Vs" "Vs @ Us"] by auto
+
+corollary mono_Span_subset:
+ assumes "set Us \<subseteq> Span K Vs" "set Vs \<subseteq> carrier R"
+ shows "Span K Us \<subseteq> Span K Vs"
+proof (rule Span_min[OF _ Span_is_add_subgroup[OF assms(2)]])
+ show "set Us \<subseteq> carrier R"
+ using Span_subgroup_props(1)[OF assms(2)] assms by auto
+ show "K <#> set Us \<subseteq> Span K Vs"
+ using Span_smult_closed[OF assms(2)] assms(1) unfolding set_mult_def by blast
+qed
+
+lemma Span_strict_incl:
+ assumes "set Us \<subseteq> carrier R" "set Vs \<subseteq> carrier R"
+ shows "Span K Us \<subset> Span K Vs \<Longrightarrow> (\<exists>v \<in> set Vs. v \<notin> Span K Us)"
+proof -
+ assume "Span K Us \<subset> Span K Vs" show "\<exists>v \<in> set Vs. v \<notin> Span K Us"
+ proof (rule ccontr)
+ assume "\<not> (\<exists>v \<in> set Vs. v \<notin> Span K Us)"
+ hence "Span K Vs \<subseteq> Span K Us"
+ using mono_Span_subset[OF _ assms(1), of Vs] by auto
+ from \<open>Span K Us \<subset> Span K Vs\<close> and \<open>Span K Vs \<subseteq> Span K Us\<close>
+ show False by simp
+ qed
+qed
+
+lemma Span_append_eq_set_add:
+ assumes "set Us \<subseteq> carrier R" and "set Vs \<subseteq> carrier R"
+ shows "Span K (Us @ Vs) = (Span K Us <+>\<^bsub>R\<^esub> Span K Vs)"
+ using assms
+proof (induct Us)
+ case Nil thus ?case
+ using Span_subgroup_props(1)[OF Nil(2)] unfolding set_add_def' by force
+next
+ case (Cons u Us)
+ hence in_carrier:
+ "u \<in> carrier R" "set Us \<subseteq> carrier R" "set Vs \<subseteq> carrier R"
+ by auto
+
+ have "line_extension K u (Span K Us <+>\<^bsub>R\<^esub> Span K Vs) = (Span K (u # Us) <+>\<^bsub>R\<^esub> Span K Vs)"
+ proof
+ show "line_extension K u (Span K Us <+>\<^bsub>R\<^esub> Span K Vs) \<subseteq> (Span K (u # Us) <+>\<^bsub>R\<^esub> Span K Vs)"
+ proof
+ fix v assume "v \<in> line_extension K u (Span K Us <+>\<^bsub>R\<^esub> Span K Vs)"
+ then obtain k u' v'
+ where v: "k \<in> K" "u' \<in> Span K Us" "v' \<in> Span K Vs" "v = k \<otimes> u \<oplus> (u' \<oplus> v')"
+ using line_extension_mem_iff[of v u "Span K Us <+>\<^bsub>R\<^esub> Span K Vs"]
+ unfolding set_add_def' by blast
+ hence "v = (k \<otimes> u \<oplus> u') \<oplus> v'"
+ using in_carrier(2-3)[THEN Span_subgroup_props(1)] in_carrier(1) subring_props(1)
+ by (metis (no_types, lifting) rev_subsetD ring_simprules(7) semiring_simprules(3))
+ moreover have "k \<otimes> u \<oplus> u' \<in> Span K (u # Us)"
+ using line_extension_mem_iff v(1-2) by auto
+ ultimately show "v \<in> Span K (u # Us) <+>\<^bsub>R\<^esub> Span K Vs"
+ unfolding set_add_def' using v(3) by auto
+ qed
+ next
+ show "Span K (u # Us) <+>\<^bsub>R\<^esub> Span K Vs \<subseteq> line_extension K u (Span K Us <+>\<^bsub>R\<^esub> Span K Vs)"
+ proof
+ fix v assume "v \<in> Span K (u # Us) <+>\<^bsub>R\<^esub> Span K Vs"
+ then obtain k u' v'
+ where v: "k \<in> K" "u' \<in> Span K Us" "v' \<in> Span K Vs" "v = (k \<otimes> u \<oplus> u') \<oplus> v'"
+ using line_extension_mem_iff[of _ u "Span K Us"] unfolding set_add_def' by auto
+ hence "v = (k \<otimes> u) \<oplus> (u' \<oplus> v')"
+ using in_carrier(2-3)[THEN Span_subgroup_props(1)] in_carrier(1) subring_props(1)
+ by (metis (no_types, lifting) rev_subsetD ring_simprules(5,7))
+ thus "v \<in> line_extension K u (Span K Us <+>\<^bsub>R\<^esub> Span K Vs)"
+ using line_extension_mem_iff[of "(k \<otimes> u) \<oplus> (u' \<oplus> v')" u "Span K Us <+>\<^bsub>R\<^esub> Span K Vs"]
+ unfolding set_add_def' using v by auto
+ qed
+ qed
+ thus ?case
+ using Cons by auto
+qed
+
+
+subsection \<open>Characterisation of Linearly Independent "Sets"\<close>
+
+declare independent_backwards [intro]
+declare independent_in_carrier [intro]
+
+lemma independent_distinct: "independent K Us \<Longrightarrow> distinct Us"
+proof (induct Us rule: list.induct)
+ case Nil thus ?case by simp
+next
+ case Cons thus ?case
+ using independent_backwards[OF Cons(2)]
+ independent_in_carrier[OF Cons(2)]
+ Span_base_incl
+ by auto
+qed
+
+lemma independent_strinct_incl:
+ assumes "independent K (u # Us)" shows "Span K Us \<subset> Span K (u # Us)"
+proof -
+ have "u \<in> Span K (u # Us)"
+ using Span_base_incl[OF independent_in_carrier[OF assms]] by auto
+ moreover have "Span K Us \<subseteq> Span K (u # Us)"
+ using mono_Span independent_in_carrier[OF assms] by auto
+ ultimately show ?thesis
+ using independent_backwards(1)[OF assms] by auto
+qed
+
+corollary independent_replacement:
+ assumes "independent K (u # Us)" and "independent K Vs"
+ shows "Span K (u # Us) \<subseteq> Span K Vs \<Longrightarrow> (\<exists>v \<in> set Vs. independent K (v # Us))"
+proof -
+ assume "Span K (u # Us) \<subseteq> Span K Vs"
+ hence "Span K Us \<subset> Span K Vs"
+ using independent_strinct_incl[OF assms(1)] by auto
+ then obtain v where v: "v \<in> set Vs" "v \<notin> Span K Us"
+ using Span_strict_incl[of Us Vs] assms[THEN independent_in_carrier] by auto
+ thus ?thesis
+ using li_Cons[of v K Us] assms independent_in_carrier[OF assms(2)] by auto
+qed
+
+lemma independent_split:
+ assumes "independent K (Us @ Vs)"
+ shows "independent K Vs"
+ and "independent K Us"
+ and "Span K Us \<inter> Span K Vs = { \<zero> }"
+proof -
+ from assms show "independent K Vs"
+ by (induct Us) (auto)
+next
+ from assms show "independent K Us"
+ proof (induct Us)
+ case Nil thus ?case by simp
+ next
+ case (Cons u Us')
+ hence u: "u \<in> carrier R" and "set Us' \<subseteq> carrier R" "set Vs \<subseteq> carrier R"
+ using independent_in_carrier[of K "(u # Us') @ Vs"] by auto
+ hence "Span K Us' \<subseteq> Span K (Us' @ Vs)"
+ using mono_Span_append(1) by simp
+ thus ?case
+ using independent_backwards[of K u "Us' @ Vs"] Cons li_Cons[OF u] by auto
+ qed
+next
+ from assms show "Span K Us \<inter> Span K Vs = { \<zero> }"
+ proof (induct Us rule: list.induct)
+ case Nil thus ?case
+ using Span_subgroup_props(2)[OF independent_in_carrier[of K Vs]] by simp
+ next
+ case (Cons u Us)
+ hence IH: "Span K Us \<inter> Span K Vs = {\<zero>}" by auto
+ have in_carrier:
+ "u \<in> carrier R" "set Us \<subseteq> carrier R" "set Vs \<subseteq> carrier R" "set (u # Us) \<subseteq> carrier R"
+ using Cons(2)[THEN independent_in_carrier] by auto
+ hence "{ \<zero> } \<subseteq> Span K (u # Us) \<inter> Span K Vs"
+ using in_carrier(3-4)[THEN Span_subgroup_props(2)] by auto
+
+ moreover have "Span K (u # Us) \<inter> Span K Vs \<subseteq> { \<zero> }"
+ proof (rule ccontr)
+ assume "\<not> Span K (u # Us) \<inter> Span K Vs \<subseteq> {\<zero>}"
+ hence "\<exists>a. a \<noteq> \<zero> \<and> a \<in> Span K (u # Us) \<and> a \<in> Span K Vs" by auto
+ then obtain k u' v'
+ where u': "u' \<in> Span K Us" "u' \<in> carrier R"
+ and v': "v' \<in> Span K Vs" "v' \<in> carrier R" "v' \<noteq> \<zero>"
+ and k: "k \<in> K" "(k \<otimes> u \<oplus> u') = v'"
+ using line_extension_mem_iff[of _ u "Span K Us"] in_carrier(2-3)[THEN Span_subgroup_props(1)]
+ subring_props(1) by force
+ hence "v' = \<zero>" if "k = \<zero>"
+ using in_carrier(1) that IH by auto
+ hence diff_zero: "k \<noteq> \<zero>" using v'(3) by auto
+
+ have "k \<in> carrier R"
+ using subring_props(1) k(1) by blast
+ hence "k \<otimes> u = (\<ominus> u') \<oplus> v'"
+ using in_carrier(1) k(2) u'(2) v'(2) add.m_comm r_neg1 by auto
+ hence "k \<otimes> u \<in> Span K (Us @ Vs)"
+ using Span_subgroup_props(4)[OF in_carrier(2) u'(1)] v'(1)
+ Span_append_eq_set_add[OF in_carrier(2-3)] unfolding set_add_def' by blast
+ hence "u \<in> Span K (Us @ Vs)"
+ using Cons(2) Span_m_inv_simprule[OF _ _ in_carrier(1), of "Us @ Vs" k]
+ diff_zero k(1) in_carrier(2-3) by auto
+ moreover have "u \<notin> Span K (Us @ Vs)"
+ using independent_backwards(1)[of K u "Us @ Vs"] Cons(2) by auto
+ ultimately show False by simp
+ qed
+
+ ultimately show ?case by auto
+ qed
+qed
+
+lemma independent_append:
+ assumes "independent K Us" and "independent K Vs" and "Span K Us \<inter> Span K Vs = { \<zero> }"
+ shows "independent K (Us @ Vs)"
+ using assms
+proof (induct Us rule: list.induct)
+ case Nil thus ?case by simp
+next
+ case (Cons u Us)
+ hence in_carrier:
+ "u \<in> carrier R" "set Us \<subseteq> carrier R" "set Vs \<subseteq> carrier R" "set (u # Us) \<subseteq> carrier R"
+ using Cons(2-3)[THEN independent_in_carrier] by auto
+ hence "Span K Us \<subseteq> Span K (u # Us)"
+ using mono_Span by auto
+ hence "Span K Us \<inter> Span K Vs = { \<zero> }"
+ using Cons(4) Span_subgroup_props(2)[OF in_carrier(2)] by auto
+ hence "independent K (Us @ Vs)"
+ using Cons by auto
+ moreover have "u \<notin> Span K (Us @ Vs)"
+ proof (rule ccontr)
+ assume "\<not> u \<notin> Span K (Us @ Vs)"
+ then obtain u' v'
+ where u': "u' \<in> Span K Us" "u' \<in> carrier R"
+ and v': "v' \<in> Span K Vs" "v' \<in> carrier R" and u:"u = u' \<oplus> v'"
+ using Span_append_eq_set_add[OF in_carrier(2-3)] in_carrier(2-3)[THEN Span_subgroup_props(1)]
+ unfolding set_add_def' by blast
+ hence "u \<oplus> (\<ominus> u') = v'"
+ using in_carrier(1) by algebra
+ moreover have "u \<in> Span K (u # Us)" and "u' \<in> Span K (u # Us)"
+ using Span_base_incl[OF in_carrier(4)] mono_Span[OF in_carrier(2,1)] u'(1)
+ by (auto simp del: Span.simps)
+ hence "u \<oplus> (\<ominus> u') \<in> Span K (u # Us)"
+ using Span_subgroup_props(3-4)[OF in_carrier(4)] by (auto simp del: Span.simps)
+ ultimately have "u \<oplus> (\<ominus> u') = \<zero>"
+ using Cons(4) v'(1) by auto
+ hence "u = u'"
+ using Cons(4) v'(1) in_carrier(1) u'(2) \<open>u \<oplus> \<ominus> u' = v'\<close> u by auto
+ thus False
+ using u'(1) independent_backwards(1)[OF Cons(2)] by simp
+ qed
+ ultimately show ?case
+ using in_carrier(1) li_Cons by simp
+qed
+
+lemma independent_imp_trivial_combine:
+ assumes "independent K Us"
+ shows "\<And>Ks. \<lbrakk> set Ks \<subseteq> K; combine Ks Us = \<zero> \<rbrakk> \<Longrightarrow> set (take (length Us) Ks) \<subseteq> { \<zero> }"
+ using assms
+proof (induct Us rule: list.induct)
+ case Nil thus ?case by simp
+next
+ case (Cons u Us) thus ?case
+ proof (cases "Ks = []")
+ assume "Ks = []" thus ?thesis by auto
+ next
+ assume "Ks \<noteq> []"
+ then obtain k Ks' where k: "k \<in> K" and Ks': "set Ks' \<subseteq> K" and Ks: "Ks = k # Ks'"
+ using Cons(2) by (metis insert_subset list.exhaust_sel list.simps(15))
+ hence Us: "set Us \<subseteq> carrier R" and u: "u \<in> carrier R"
+ using independent_in_carrier[OF Cons(4)] by auto
+ have "u \<in> Span K Us" if "k \<noteq> \<zero>"
+ using that Span_mem_iff[OF Us u] Cons(3-4) Ks' k unfolding Ks by blast
+ hence k_zero: "k = \<zero>"
+ using independent_backwards[OF Cons(4)] by blast
+ hence "combine Ks' Us = \<zero>"
+ using combine_in_carrier[OF _ Us, of Ks'] Ks' u Cons(3) subring_props(1) unfolding Ks by auto
+ hence "set (take (length Us) Ks') \<subseteq> { \<zero> }"
+ using Cons(1)[OF Ks' _ independent_backwards(2)[OF Cons(4)]] by simp
+ thus ?thesis
+ using k_zero unfolding Ks by auto
+ qed
+qed
+
+lemma trivial_combine_imp_independent:
+ assumes "set Us \<subseteq> carrier R"
+ and "\<And>Ks. \<lbrakk> set Ks \<subseteq> K; combine Ks Us = \<zero> \<rbrakk> \<Longrightarrow> set (take (length Us) Ks) \<subseteq> { \<zero> }"
+ shows "independent K Us"
+ using assms
+proof (induct Us)
+ case Nil thus ?case by simp
+next
+ case (Cons u Us)
+ hence Us: "set Us \<subseteq> carrier R" and u: "u \<in> carrier R" by auto
+
+ have "\<And>Ks. \<lbrakk> set Ks \<subseteq> K; combine Ks Us = \<zero> \<rbrakk> \<Longrightarrow> set (take (length Us) Ks) \<subseteq> { \<zero> }"
+ proof -
+ fix Ks assume Ks: "set Ks \<subseteq> K" and lin_c: "combine Ks Us = \<zero>"
+ hence "combine (\<zero> # Ks) (u # Us) = \<zero>"
+ using u subring_props(1) combine_in_carrier[OF _ Us] by auto
+ hence "set (take (length (u # Us)) (\<zero> # Ks)) \<subseteq> { \<zero> }"
+ using Cons(3)[of "\<zero> # Ks"] subring_props(2) Ks by auto
+ thus "set (take (length Us) Ks) \<subseteq> { \<zero> }" by auto
+ qed
+ hence "independent K Us"
+ using Cons(1)[OF Us] by simp
+
+ moreover have "u \<notin> Span K Us"
+ proof (rule ccontr)
+ assume "\<not> u \<notin> Span K Us"
+ then obtain k Ks where k: "k \<in> K" "k \<noteq> \<zero>" and Ks: "set Ks \<subseteq> K" and u: "combine (k # Ks) (u # Us) = \<zero>"
+ using Span_mem_iff[OF Us u] by auto
+ have "set (take (length (u # Us)) (k # Ks)) \<subseteq> { \<zero> }"
+ using Cons(3)[OF _ u] k(1) Ks by auto
+ hence "k = \<zero>" by auto
+ from \<open>k = \<zero>\<close> and \<open>k \<noteq> \<zero>\<close> show False by simp
+ qed
+
+ ultimately show ?case
+ using li_Cons[OF u] by simp
+qed
+
+corollary unique_decomposition:
+ assumes "independent K Us"
+ shows "a \<in> Span K Us \<Longrightarrow> \<exists>!Ks. set Ks \<subseteq> K \<and> length Ks = length Us \<and> a = combine Ks Us"
+proof -
+ note in_carrier = independent_in_carrier[OF assms]
+
+ assume "a \<in> Span K Us"
+ then obtain Ks where Ks: "set Ks \<subseteq> K" "length Ks = length Us" "a = combine Ks Us"
+ using Span_mem_iff_length_version[OF in_carrier] by blast
+
+ moreover
+ have "\<And>Ks'. \<lbrakk> set Ks' \<subseteq> K; length Ks' = length Us; a = combine Ks' Us \<rbrakk> \<Longrightarrow> Ks = Ks'"
+ proof -
+ fix Ks' assume Ks': "set Ks' \<subseteq> K" "length Ks' = length Us" "a = combine Ks' Us"
+ hence set_Ks: "set Ks \<subseteq> carrier R" and set_Ks': "set Ks' \<subseteq> carrier R"
+ using subring_props(1) Ks(1) by blast+
+ have same_length: "length Ks = length Ks'"
+ using Ks Ks' by simp
+
+ have "(combine Ks Us) \<oplus> ((\<ominus> \<one>) \<otimes> (combine Ks' Us)) = \<zero>"
+ using combine_in_carrier[OF set_Ks in_carrier]
+ combine_in_carrier[OF set_Ks' in_carrier] Ks(3) Ks'(3) by algebra
+ hence "(combine Ks Us) \<oplus> (combine (map ((\<otimes>) (\<ominus> \<one>)) Ks') Us) = \<zero>"
+ using combine_r_distr[OF set_Ks' in_carrier, of "\<ominus> \<one>"] subring_props by auto
+ moreover have set_map: "set (map ((\<otimes>) (\<ominus> \<one>)) Ks') \<subseteq> K"
+ using Ks'(1) subring_props by (induct Ks') (auto)
+ hence "set (map ((\<otimes>) (\<ominus> \<one>)) Ks') \<subseteq> carrier R"
+ using subring_props(1) by blast
+ ultimately have "combine (map2 (\<oplus>) Ks (map ((\<otimes>) (\<ominus> \<one>)) Ks')) Us = \<zero>"
+ using combine_add[OF Ks(2) _ set_Ks _ in_carrier, of "map ((\<otimes>) (\<ominus> \<one>)) Ks'"] Ks'(2) by auto
+ moreover have "set (map2 (\<oplus>) Ks (map ((\<otimes>) (\<ominus> \<one>)) Ks')) \<subseteq> K"
+ using Ks(1) set_map subring_props(7)
+ by (induct Ks) (auto, metis contra_subsetD in_set_zipE local.set_map set_ConsD subring_props(7))
+ ultimately have "set (take (length Us) (map2 (\<oplus>) Ks (map ((\<otimes>) (\<ominus> \<one>)) Ks'))) \<subseteq> { \<zero> }"
+ using independent_imp_trivial_combine[OF assms] by auto
+ hence "set (map2 (\<oplus>) Ks (map ((\<otimes>) (\<ominus> \<one>)) Ks')) \<subseteq> { \<zero> }"
+ using Ks(2) Ks'(2) by auto
+ thus "Ks = Ks'"
+ using set_Ks set_Ks' same_length
+ proof (induct Ks arbitrary: Ks')
+ case Nil thus?case by simp
+ next
+ case (Cons k Ks)
+ then obtain k' Ks'' where k': "Ks' = k' # Ks''"
+ by (metis Suc_length_conv)
+ have "Ks = Ks''"
+ using Cons unfolding k' by auto
+ moreover have "k = k'"
+ using Cons(2-4) l_minus minus_equality unfolding k' by (auto, fastforce)
+ ultimately show ?case
+ unfolding k' by simp
+ qed
+ qed
+
+ ultimately show ?thesis by blast
+qed
+
+
+subsection \<open>Replacement Theorem\<close>
+
+lemma independent_rotate1_aux:
+ "independent K (u # Us @ Vs) \<Longrightarrow> independent K ((Us @ [u]) @ Vs)"
+proof -
+ assume "independent K (u # Us @ Vs)"
+ hence li: "independent K [u]" "independent K Us" "independent K Vs"
+ and inter: "Span K [u] \<inter> Span K Us = { \<zero> }"
+ "Span K (u # Us) \<inter> Span K Vs = { \<zero> }"
+ using independent_split[of "u # Us" Vs] independent_split[of "[u]" Us] by auto
+ hence "independent K (Us @ [u])"
+ using independent_append[OF li(2,1)] by auto
+ moreover have "Span K (Us @ [u]) \<inter> Span K Vs = { \<zero> }"
+ using Span_same_set[of "u # Us" "Us @ [u]"] li(1-2)[THEN independent_in_carrier] inter(2) by auto
+ ultimately show "independent K ((Us @ [u]) @ Vs)"
+ using independent_append[OF _ li(3), of "Us @ [u]"] by simp
+qed
+
+corollary independent_rotate1:
+ "independent K (Us @ Vs) \<Longrightarrow> independent K ((rotate1 Us) @ Vs)"
+ using independent_rotate1_aux by (cases Us) (auto)
+
+(*
+corollary independent_rotate:
+ "independent K (Us @ Vs) \<Longrightarrow> independent K ((rotate n Us) @ Vs)"
+ using independent_rotate1 by (induct n) auto
+
+lemma rotate_append: "rotate (length l) (l @ q) = q @ l"
+ by (induct l arbitrary: q) (auto simp add: rotate1_rotate_swap)
+*)
+
+corollary independent_same_set:
+ assumes "set Us = set Vs" and "length Us = length Vs"
+ shows "independent K Us \<Longrightarrow> independent K Vs"
+proof -
+ assume "independent K Us" thus ?thesis
+ using assms
+ proof (induct Us arbitrary: Vs rule: list.induct)
+ case Nil thus ?case by simp
+ next
+ case (Cons u Us)
+ then obtain Vs' Vs'' where Vs: "Vs = Vs' @ (u # Vs'')"
+ by (metis list.set_intros(1) split_list)
+
+ have in_carrier: "u \<in> carrier R" "set Us \<subseteq> carrier R"
+ using independent_in_carrier[OF Cons(2)] by auto
+
+ have "distinct Vs"
+ using Cons(3-4) independent_distinct[OF Cons(2)]
+ by (metis card_distinct distinct_card)
+ hence "u \<notin> set (Vs' @ Vs'')" and "u \<notin> set Us"
+ using independent_distinct[OF Cons(2)] unfolding Vs by auto
+ hence set_eq: "set Us = set (Vs' @ Vs'')" and "length (Vs' @ Vs'') = length Us"
+ using Cons(3-4) unfolding Vs by auto
+ hence "independent K (Vs' @ Vs'')"
+ using Cons(1)[OF independent_backwards(2)[OF Cons(2)]] unfolding Vs by simp
+ hence "independent K (u # (Vs' @ Vs''))"
+ using li_Cons Span_same_set[OF _ set_eq] independent_backwards(1)[OF Cons(2)] in_carrier by auto
+ hence "independent K (Vs' @ (u # Vs''))"
+ using independent_rotate1[of "u # Vs'" Vs''] by auto
+ thus ?case unfolding Vs .
+ qed
+qed
+
+lemma replacement_theorem:
+ assumes "independent K (Us' @ Us)" and "independent K Vs"
+ and "Span K (Us' @ Us) \<subseteq> Span K Vs"
+ shows "\<exists>Vs'. set Vs' \<subseteq> set Vs \<and> length Vs' = length Us' \<and> independent K (Vs' @ Us)"
+ using assms
+proof (induct "length Us'" arbitrary: Us' Us)
+ case 0 thus ?case by auto
+next
+ case (Suc n)
+ then obtain u Us'' where Us'': "Us' = Us'' @ [u]"
+ by (metis list.size(3) nat.simps(3) rev_exhaust)
+ then obtain Vs' where Vs': "set Vs' \<subseteq> set Vs" "length Vs' = n" "independent K (Vs' @ (u # Us))"
+ using Suc(1)[of Us'' "u # Us"] Suc(2-5) by auto
+ hence li: "independent K ((u # Vs') @ Us)"
+ using independent_same_set[OF _ _ Vs'(3), of "(u # Vs') @ Us"] by auto
+ moreover have in_carrier:
+ "u \<in> carrier R" "set Us \<subseteq> carrier R" "set Us' \<subseteq> carrier R" "set Vs \<subseteq> carrier R"
+ using Suc(3-4)[THEN independent_in_carrier] Us'' by auto
+ moreover have "Span K ((u # Vs') @ Us) \<subseteq> Span K Vs"
+ proof -
+ have "set Us \<subseteq> Span K Vs" "u \<in> Span K Vs"
+ using Suc(5) Span_base_incl[of "Us' @ Us"] Us'' in_carrier(2-3) by auto
+ moreover have "set Vs' \<subseteq> Span K Vs"
+ using Span_base_incl[OF in_carrier(4)] Vs'(1) by auto
+ ultimately have "set ((u # Vs') @ Us) \<subseteq> Span K Vs" by auto
+ thus ?thesis
+ using mono_Span_subset[OF _ in_carrier(4)] by (simp del: Span.simps)
+ qed
+ ultimately obtain v where "v \<in> set Vs" "independent K ((v # Vs') @ Us)"
+ using independent_replacement[OF _ Suc(4), of u "Vs' @ Us"] by auto
+ thus ?case
+ using Vs'(1-2) Suc(2)
+ by (metis (mono_tags, lifting) insert_subset length_Cons list.simps(15))
+qed
+
+corollary independent_length_le:
+ assumes "independent K Us" and "independent K Vs"
+ shows "set Us \<subseteq> Span K Vs \<Longrightarrow> length Us \<le> length Vs"
+proof -
+ assume "set Us \<subseteq> Span K Vs"
+ hence "Span K Us \<subseteq> Span K Vs"
+ using mono_Span_subset[OF _ independent_in_carrier[OF assms(2)]] by simp
+ then obtain Vs' where Vs': "set Vs' \<subseteq> set Vs" "length Vs' = length Us" "independent K Vs'"
+ using replacement_theorem[OF _ assms(2), of Us "[]"] assms(1) by auto
+ hence "card (set Vs') \<le> card (set Vs)"
+ by (simp add: card_mono)
+ thus "length Us \<le> length Vs"
+ using independent_distinct assms(2) Vs'(2-3) by (simp add: distinct_card)
+qed
+
+
+subsection \<open>Dimension\<close>
+
+lemma exists_base:
+ assumes "dimension n K E"
+ shows "\<exists>Vs. set Vs \<subseteq> carrier R \<and> independent K Vs \<and> length Vs = n \<and> Span K Vs = E"
+ (is "\<exists>Vs. ?base K Vs E n")
+ using assms
+proof (induct E rule: dimension.induct)
+ case zero_dim thus ?case by auto
+next
+ case (Suc_dim v E n K)
+ then obtain Vs where Vs: "set Vs \<subseteq> carrier R" "independent K Vs" "length Vs = n" "Span K Vs = E"
+ by auto
+ hence "?base K (v # Vs) (line_extension K v E) (Suc n)"
+ using Suc_dim li_Cons by auto
+ thus ?case by blast
+qed
+
+lemma dimension_zero [intro]: "dimension 0 K E \<Longrightarrow> E = { \<zero> }"
+proof -
+ assume "dimension 0 K E"
+ then obtain Vs where "length Vs = 0" "Span K Vs = E"
+ using exists_base by blast
+ thus ?thesis
+ by auto
+qed
+
+lemma dimension_independent [intro]: "independent K Us \<Longrightarrow> dimension (length Us) K (Span K Us)"
+proof (induct Us)
+ case Nil thus ?case by simp
+next
+ case Cons thus ?case
+ using Suc_dim[OF independent_backwards(3,1)[OF Cons(2)]] by auto
+qed
+
+lemma dimensionI:
+ assumes "independent K Us" "Span K Us = E"
+ shows "dimension (length Us) K E"
+ using dimension_independent[OF assms(1)] assms(2) by simp
+
+lemma space_subgroup_props:
+ assumes "dimension n K E"
+ shows "E \<subseteq> carrier R"
+ and "\<zero> \<in> E"
+ and "\<And>v1 v2. \<lbrakk> v1 \<in> E; v2 \<in> E \<rbrakk> \<Longrightarrow> (v1 \<oplus> v2) \<in> E"
+ and "\<And>v. v \<in> E \<Longrightarrow> (\<ominus> v) \<in> E"
+ and "\<And>k v. \<lbrakk> k \<in> K; v \<in> E \<rbrakk> \<Longrightarrow> k \<otimes> v \<in> E"
+ and "\<lbrakk> k \<in> K - { \<zero> }; a \<in> carrier R \<rbrakk> \<Longrightarrow> k \<otimes> a \<in> E \<Longrightarrow> a \<in> E"
+ using exists_base[OF assms] Span_subgroup_props Span_smult_closed Span_m_inv_simprule by auto
+
+lemma independent_length_le_dimension:
+ assumes "dimension n K E" and "independent K Us" "set Us \<subseteq> E"
+ shows "length Us \<le> n"
+proof -
+ obtain Vs where Vs: "set Vs \<subseteq> carrier R" "independent K Vs" "length Vs = n" "Span K Vs = E"
+ using exists_base[OF assms(1)] by auto
+ thus ?thesis
+ using independent_length_le assms(2-3) by auto
+qed
+
+lemma dimension_is_inj:
+ assumes "dimension n K E" and "dimension m K E"
+ shows "n = m"
+proof -
+ { fix n m assume n: "dimension n K E" and m: "dimension m K E"
+ then obtain Vs
+ where Vs: "set Vs \<subseteq> carrier R" "independent K Vs" "length Vs = n" "Span K Vs = E"
+ using exists_base by meson
+ hence "n \<le> m"
+ using independent_length_le_dimension[OF m Vs(2)] Span_base_incl[OF Vs(1)] by auto
+ } note aux_lemma = this
+
+ show ?thesis
+ using aux_lemma[OF assms] aux_lemma[OF assms(2,1)] by simp
+qed
+
+corollary independent_length_eq_dimension:
+ assumes "dimension n K E" and "independent K Us" "set Us \<subseteq> E"
+ shows "length Us = n \<longleftrightarrow> Span K Us = E"
+proof
+ assume len: "length Us = n" show "Span K Us = E"
+ proof (rule ccontr)
+ assume "Span K Us \<noteq> E"
+ hence "Span K Us \<subset> E"
+ using mono_Span_subset[of Us] exists_base[OF assms(1)] assms(3) by blast
+ then obtain v where v: "v \<in> E" "v \<notin> Span K Us"
+ using Span_strict_incl exists_base[OF assms(1)] space_subgroup_props(1)[OF assms(1)] assms by blast
+ hence "independent K (v # Us)"
+ using li_Cons[OF _ _ assms(2)] space_subgroup_props(1)[OF assms(1)] by auto
+ hence "length (v # Us) \<le> n"
+ using independent_length_le_dimension[OF assms(1)] v(1) assms(2-3) by fastforce
+ moreover have "length (v # Us) = Suc n"
+ using len by simp
+ ultimately show False by simp
+ qed
+next
+ assume "Span K Us = E"
+ hence "dimension (length Us) K E"
+ using dimensionI assms by auto
+ thus "length Us = n"
+ using dimension_is_inj[OF assms(1)] by auto
+qed
+
+lemma complete_base:
+ assumes "dimension n K E" and "independent K Us" "set Us \<subseteq> E"
+ shows "\<exists>Vs. length (Vs @ Us) = n \<and> independent K (Vs @ Us) \<and> Span K (Vs @ Us) = E"
+proof -
+ { fix Us k assume "k \<le> n" "independent K Us" "set Us \<subseteq> E" "length Us = k"
+ hence "\<exists>Vs. length (Vs @ Us) = n \<and> independent K (Vs @ Us) \<and> Span K (Vs @ Us) = E"
+ proof (induct arbitrary: Us rule: inc_induct)
+ case base thus ?case
+ using independent_length_eq_dimension[OF assms(1) base(1-2)] by auto
+ next
+ case (step m)
+ have "Span K Us \<subseteq> E"
+ using mono_Span_subset step(4-6) exists_base[OF assms(1)] by blast
+ hence "Span K Us \<subset> E"
+ using independent_length_eq_dimension[OF assms(1) step(4-5)] step(2,6) assms(1) by blast
+ then obtain v where v: "v \<in> E" "v \<notin> Span K Us"
+ using Span_strict_incl exists_base[OF assms(1)] by blast
+ hence "independent K (v # Us)"
+ using space_subgroup_props(1)[OF assms(1)] li_Cons[OF _ v(2) step(4)] by auto
+ then obtain Vs
+ where "length (Vs @ (v # Us)) = n" "independent K (Vs @ (v # Us))" "Span K (Vs @ (v # Us)) = E"
+ using step(3)[of "v # Us"] step(1-2,4-6) v by auto
+ thus ?case
+ by (metis append.assoc append_Cons append_Nil)
+ qed } note aux_lemma = this
+
+ have "length Us \<le> n"
+ using independent_length_le_dimension[OF assms] .
+ thus ?thesis
+ using aux_lemma[OF _ assms(2-3)] by auto
+qed
+
+lemma dimension_backwards:
+ "dimension (Suc n) K E \<Longrightarrow> \<exists>v \<in> carrier R. \<exists>E'. dimension n K E' \<and> v \<notin> E' \<and> E = line_extension K v E'"
+ by (cases rule: dimension.cases) (auto)
+
+lemma dimension_direct_sum_space:
+ assumes "dimension n K E" and "dimension m K F" and "E \<inter> F = { \<zero> }"
+ shows "dimension (n + m) K (E <+>\<^bsub>R\<^esub> F)"
+proof -
+ obtain Us Vs
+ where Vs: "set Vs \<subseteq> carrier R" "independent K Vs" "length Vs = n" "Span K Vs = E"
+ and Us: "set Us \<subseteq> carrier R" "independent K Us" "length Us = m" "Span K Us = F"
+ using assms(1-2)[THEN exists_base] by auto
+ hence "Span K (Vs @ Us) = E <+>\<^bsub>R\<^esub> F"
+ using Span_append_eq_set_add by auto
+ moreover have "independent K (Vs @ Us)"
+ using assms(3) independent_append[OF Vs(2) Us(2)] unfolding Vs(4) Us(4) by simp
+ ultimately show "dimension (n + m) K (E <+>\<^bsub>R\<^esub> F)"
+ using dimensionI[of "Vs @ Us"] Vs(3) Us(3) by auto
+qed
+
+lemma dimension_sum_space:
+ assumes "dimension n K E" and "dimension m K F" and "dimension k K (E \<inter> F)"
+ shows "dimension (n + m - k) K (E <+>\<^bsub>R\<^esub> F)"
+proof -
+ obtain Bs
+ where Bs: "set Bs \<subseteq> carrier R" "length Bs = k" "independent K Bs" "Span K Bs = E \<inter> F"
+ using exists_base[OF assms(3)] by blast
+ then obtain Us Vs
+ where Us: "length (Us @ Bs) = n" "independent K (Us @ Bs)" "Span K (Us @ Bs) = E"
+ and Vs: "length (Vs @ Bs) = m" "independent K (Vs @ Bs)" "Span K (Vs @ Bs) = F"
+ using Span_base_incl[OF Bs(1)] assms(1-2)[THEN complete_base] by (metis le_infE)
+ hence in_carrier: "set Us \<subseteq> carrier R" "set (Vs @ Bs) \<subseteq> carrier R"
+ using independent_in_carrier[OF Us(2)] independent_in_carrier[OF Vs(2)] by auto
+ hence "Span K Us \<inter> (Span K (Vs @ Bs)) \<subseteq> Span K Bs"
+ using Bs(4) Us(3) Vs(3) mono_Span_append(1)[OF _ Bs(1), of Us] by auto
+ hence "Span K Us \<inter> (Span K (Vs @ Bs)) \<subseteq> { \<zero> }"
+ using independent_split(3)[OF Us(2)] by blast
+ hence "Span K Us \<inter> (Span K (Vs @ Bs)) = { \<zero> }"
+ using in_carrier[THEN Span_subgroup_props(2)] by auto
+
+ hence dim: "dimension (n + m - k) K (Span K (Us @ (Vs @ Bs)))"
+ using independent_append[OF independent_split(2)[OF Us(2)] Vs(2)] Us(1) Vs(1) Bs(2)
+ dimension_independent[of "Us @ (Vs @ Bs)"] by auto
+
+ have "(Span K Us) <+>\<^bsub>R\<^esub> F \<subseteq> E <+>\<^bsub>R\<^esub> F"
+ using mono_Span_append(1)[OF in_carrier(1) Bs(1)] Us(3) unfolding set_add_def' by auto
+ moreover have "E <+>\<^bsub>R\<^esub> F \<subseteq> (Span K Us) <+>\<^bsub>R\<^esub> F"
+ proof
+ fix v assume "v \<in> E <+>\<^bsub>R\<^esub> F"
+ then obtain u' v' where v: "u' \<in> E" "v' \<in> F" "v = u' \<oplus> v'"
+ unfolding set_add_def' by auto
+ then obtain u1' u2' where u1': "u1' \<in> Span K Us" and u2': "u2' \<in> Span K Bs" and u': "u' = u1' \<oplus> u2'"
+ using Span_append_eq_set_add[OF in_carrier(1) Bs(1)] Us(3) unfolding set_add_def' by blast
+
+ have "v = u1' \<oplus> (u2' \<oplus> v')"
+ using Span_subgroup_props(1)[OF Bs(1)] Span_subgroup_props(1)[OF in_carrier(1)]
+ space_subgroup_props(1)[OF assms(2)] u' v u1' u2' a_assoc[of u1' u2' v'] by auto
+ moreover have "u2' \<oplus> v' \<in> F"
+ using space_subgroup_props(3)[OF assms(2) _ v(2)] u2' Bs(4) by auto
+ ultimately show "v \<in> (Span K Us) <+>\<^bsub>R\<^esub> F"
+ using u1' unfolding set_add_def' by auto
+ qed
+ ultimately have "Span K (Us @ (Vs @ Bs)) = E <+>\<^bsub>R\<^esub> F"
+ using Span_append_eq_set_add[OF in_carrier] Vs(3) by auto
+
+ thus ?thesis using dim by simp
+qed
+
+end
+
+end
+
+lemma (in ring) telescopic_base_aux:
+ assumes "subfield K R" "subfield F R"
+ and "dimension n K F" and "dimension 1 F E"
+ shows "dimension n K E"
+proof -
+ obtain Us u
+ where Us: "set Us \<subseteq> carrier R" "length Us = n" "independent K Us" "Span K Us = F"
+ and u: "u \<in> carrier R" "independent F [u]" "Span F [u] = E"
+ using exists_base[OF assms(2,4)] exists_base[OF assms(1,3)] independent_backwards(3) assms(2)
+ by (metis One_nat_def length_0_conv length_Suc_conv)
+ have in_carrier: "set (map (\<lambda>u'. u' \<otimes> u) Us) \<subseteq> carrier R"
+ using Us(1) u(1) by (induct Us) (auto)
+
+ have li: "independent K (map (\<lambda>u'. u' \<otimes> u) Us)"
+ proof (rule trivial_combine_imp_independent[OF assms(1) in_carrier])
+ fix Ks assume Ks: "set Ks \<subseteq> K" and "combine Ks (map (\<lambda>u'. u' \<otimes> u) Us) = \<zero>"
+ hence "(combine Ks Us) \<otimes> u = \<zero>"
+ using combine_l_distr[OF _ Us(1) u(1)] subring_props(1)[OF assms(1)] by auto
+ hence "combine [ combine Ks Us ] [ u ] = \<zero>"
+ by simp
+ moreover have "combine Ks Us \<in> F"
+ using Us(4) Ks(1) Span_eq_combine_set[OF assms(1) Us(1)] by auto
+ ultimately have "combine Ks Us = \<zero>"
+ using independent_imp_trivial_combine[OF assms(2) u(2), of "[ combine Ks Us ]"] by auto
+ hence "set (take (length Us) Ks) \<subseteq> { \<zero> }"
+ using independent_imp_trivial_combine[OF assms(1) Us(3) Ks(1)] by simp
+ thus "set (take (length (map (\<lambda>u'. u' \<otimes> u) Us)) Ks) \<subseteq> { \<zero> }" by simp
+ qed
+
+ have "E \<subseteq> Span K (map (\<lambda>u'. u' \<otimes> u) Us)"
+ proof
+ fix v assume "v \<in> E"
+ then obtain f where f: "f \<in> F" "v = f \<otimes> u \<oplus> \<zero>"
+ using u(1,3) line_extension_mem_iff[OF assms(2)] by auto
+ then obtain Ks where Ks: "set Ks \<subseteq> K" "f = combine Ks Us"
+ using Span_eq_combine_set[OF assms(1) Us(1)] Us(4) by auto
+ have "v = f \<otimes> u"
+ using subring_props(1)[OF assms(2)] f u(1) by auto
+ hence "v = combine Ks (map (\<lambda>u'. u' \<otimes> u) Us)"
+ using combine_l_distr[OF _ Us(1) u(1), of Ks] Ks(1-2)
+ subring_props(1)[OF assms(1)] by blast
+ thus "v \<in> Span K (map (\<lambda>u'. u' \<otimes> u) Us)"
+ unfolding Span_eq_combine_set[OF assms(1) in_carrier] using Ks(1) by blast
+ qed
+ moreover have "Span K (map (\<lambda>u'. u' \<otimes> u) Us) \<subseteq> E"
+ proof
+ fix v assume "v \<in> Span K (map (\<lambda>u'. u' \<otimes> u) Us)"
+ then obtain Ks where Ks: "set Ks \<subseteq> K" "v = combine Ks (map (\<lambda>u'. u' \<otimes> u) Us)"
+ unfolding Span_eq_combine_set[OF assms(1) in_carrier] by blast
+ hence "v = (combine Ks Us) \<otimes> u"
+ using combine_l_distr[OF _ Us(1) u(1), of Ks] subring_props(1)[OF assms(1)] by auto
+ moreover have "combine Ks Us \<in> F"
+ using Us(4) Span_eq_combine_set[OF assms(1) Us(1)] Ks(1) by blast
+ ultimately have "v = (combine Ks Us) \<otimes> u \<oplus> \<zero>" and "combine Ks Us \<in> F"
+ using subring_props(1)[OF assms(2)] u(1) by auto
+ thus "v \<in> E"
+ using u(3) line_extension_mem_iff[OF assms(2)] by auto
+ qed
+ ultimately have "Span K (map (\<lambda>u'. u' \<otimes> u) Us) = E" by auto
+ thus ?thesis
+ using dimensionI[OF assms(1) li] Us(2) by simp
+qed
+
+lemma (in ring) telescopic_base:
+ assumes "subfield K R" "subfield F R"
+ and "dimension n K F" and "dimension m F E"
+ shows "dimension (n * m) K E"
+ using assms(4)
+proof (induct m arbitrary: E)
+ case 0 thus ?case
+ using dimension_zero[OF assms(2)] zero_dim by auto
+next
+ case (Suc m)
+ obtain Vs
+ where Vs: "set Vs \<subseteq> carrier R" "length Vs = Suc m" "independent F Vs" "Span F Vs = E"
+ using exists_base[OF assms(2) Suc(2)] by blast
+ then obtain v Vs' where v: "Vs = v # Vs'"
+ by (meson length_Suc_conv)
+ hence li: "independent F [ v ]" "independent F Vs'" and inter: "Span F [ v ] \<inter> Span F Vs' = { \<zero> }"
+ using Vs(3) independent_split[OF assms(2), of "[ v ]" Vs'] by auto
+ have "dimension n K (Span F [ v ])"
+ using dimension_independent[OF assms(2) li(1)] telescopic_base_aux[OF assms(1-3)] by simp
+ moreover have "dimension (n * m) K (Span F Vs')"
+ using Suc(1) dimension_independent[OF assms(2) li(2)] Vs(2) unfolding v by auto
+ ultimately have "dimension (n * Suc m) K (Span F [ v ] <+>\<^bsub>R\<^esub> Span F Vs')"
+ using dimension_direct_sum_space[OF assms(1) _ _ inter] by auto
+ thus "dimension (n * Suc m) K E"
+ using Span_append_eq_set_add[OF assms(2) li[THEN independent_in_carrier]] Vs(4) v by auto
+qed
+
+
+(*
+lemma combine_take:
+ assumes "set Ks \<subseteq> carrier R" "set Us \<subseteq> carrier R"
+ shows "length Ks \<le> length Us \<Longrightarrow> combine Ks Us = combine Ks (take (length Ks) Us)"
+ and "length Us \<le> length Ks \<Longrightarrow> combine Ks Us = combine (take (length Us) Ks) Us"
+proof -
+ assume len: "length Ks \<le> length Us"
+ hence Us: "Us = (take (length Ks) Us) @ (drop (length Ks) Us)" by auto
+ hence set_t: "set (take (length Ks) Us) \<subseteq> carrier R" and set_d: "set (drop (length Ks) Us) \<subseteq> carrier R"
+ using assms(2) len by (metis le_sup_iff set_append)+
+ hence "combine Ks Us = (combine Ks (take (length Ks) Us)) \<oplus> \<zero>"
+ using combine_append[OF _ assms(1), of "take (length Ks) Us" "[]" "drop (length Ks) Us"] len by auto
+ also have " ... = combine Ks (take (length Ks) Us)"
+ using combine_in_carrier[OF assms(1) set_t] by auto
+ finally show "combine Ks Us = combine Ks (take (length Ks) Us)" .
+next
+ assume len: "length Us \<le> length Ks"
+ hence Us: "Ks = (take (length Us) Ks) @ (drop (length Us) Ks)" by auto
+ hence set_t: "set (take (length Us) Ks) \<subseteq> carrier R" and set_d: "set (drop (length Us) Ks) \<subseteq> carrier R"
+ using assms(1) len by (metis le_sup_iff set_append)+
+ hence "combine Ks Us = (combine (take (length Us) Ks) Us) \<oplus> \<zero>"
+ using combine_append[OF _ _ assms(2), of "take (length Us) Ks" "drop (length Us) Ks" "[]"] len by auto
+ also have " ... = combine (take (length Us) Ks) Us"
+ using combine_in_carrier[OF set_t assms(2)] by auto
+ finally show "combine Ks Us = combine (take (length Us) Ks) Us" .
+qed
+*)
+
+(*
+lemma combine_normalize:
+ assumes "set Ks \<subseteq> K" "set Us \<subseteq> carrier R" "a = combine Ks Us"
+ shows "\<exists>Ks'. set Ks' \<subseteq> K \<and> length Ks' = length Us \<and> a = combine Ks' Us"
+proof (cases "length Ks \<le> length Us")
+ assume "\<not> length Ks \<le> length Us"
+ hence len: "length Us < length Ks" by simp
+ hence "length (take (length Us) Ks) = length Us" and "set (take (length Us) Ks) \<subseteq> K"
+ using assms(1) by (auto, metis contra_subsetD in_set_takeD)
+ thus ?thesis
+ using combine_take(2)[OF _ assms(2), of Ks] assms(1,3) subring_props(1) len
+ by (metis dual_order.trans nat_less_le)
+next
+ assume len: "length Ks \<le> length Us"
+ have Ks: "set Ks \<subseteq> carrier R" and set_r: "set (replicate (length Us - length Ks) \<zero>) \<subseteq> carrier R"
+ using assms subring_props(1) zero_closed by (metis dual_order.trans, auto)
+ moreover
+ have set_t: "set (take (length Ks) Us) \<subseteq> carrier R"
+ and set_d: "set (drop (length Ks) Us) \<subseteq> carrier R"
+ using assms(2) len dual_order.trans by (metis set_take_subset, metis set_drop_subset)
+ ultimately
+ have "combine (Ks @ (replicate (length Us - length Ks) \<zero>)) Us =
+ (combine Ks (take (length Ks) Us)) \<oplus>
+ (combine (replicate (length Us - length Ks) \<zero>) (drop (length Ks) Us))"
+ using combine_append[OF _ Ks set_t set_r set_d] len by auto
+ also have " ... = combine Ks (take (length Ks) Us)"
+ using combine_replicate[OF set_d] combine_in_carrier[OF Ks set_t] by auto
+ also have " ... = a"
+ using combine_take(1)[OF Ks assms(2) len] assms(3) by simp
+ finally have "combine (Ks @ (replicate (length Us - length Ks) \<zero>)) Us = a" .
+ moreover have "set (Ks @ (replicate (length Us - length Ks) \<zero>)) \<subseteq> K"
+ using assms(1) subring_props(2) by auto
+ moreover have "length (Ks @ (replicate (length Us - length Ks) \<zero>)) = length Us"
+ using len by simp
+ ultimately show ?thesis by blast
+qed
+*)
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Algebra/Generated_Fields.thy Mon Jul 02 15:43:22 2018 +0100
@@ -0,0 +1,190 @@
+(* ************************************************************************** *)
+(* Title: Generated_Fields.thy *)
+(* Author: Martin Baillon *)
+(* ************************************************************************** *)
+
+theory Generated_Fields
+
+imports Generated_Rings Subrings Multiplicative_Group
+
+begin
+
+inductive_set
+ generate_field :: "('a, 'b) ring_scheme \<Rightarrow> 'a set \<Rightarrow> 'a set"
+ for R and H where
+ one : "\<one>\<^bsub>R\<^esub> \<in> generate_field R H"
+ | incl : "h \<in> H \<Longrightarrow> h \<in> generate_field R H"
+ | a_inv: "h \<in> generate_field R H \<Longrightarrow> \<ominus>\<^bsub>R\<^esub> h \<in> generate_field R H"
+ | m_inv: "\<lbrakk> h \<in> generate_field R H; h \<noteq> \<zero>\<^bsub>R\<^esub> \<rbrakk> \<Longrightarrow> inv\<^bsub>R\<^esub> h \<in> generate_field R H"
+ | eng_add : "\<lbrakk> h1 \<in> generate_field R H; h2 \<in> generate_field R H \<rbrakk> \<Longrightarrow> h1 \<oplus>\<^bsub>R\<^esub> h2 \<in> generate_field R H"
+ | eng_mult: "\<lbrakk> h1 \<in> generate_field R H; h2 \<in> generate_field R H \<rbrakk> \<Longrightarrow> h1 \<otimes>\<^bsub>R\<^esub> h2 \<in> generate_field R H"
+
+
+subsection\<open>Basic Properties of Generated Rings - First Part\<close>
+
+lemma (in field) generate_field_in_carrier:
+ assumes "H \<subseteq> carrier R"
+ shows "h \<in> generate_field R H \<Longrightarrow> h \<in> carrier R"
+ apply (induction rule: generate_field.induct)
+ using assms field_Units
+ by blast+
+
+lemma (in field) generate_field_incl:
+ assumes "H \<subseteq> carrier R"
+ shows "generate_field R H \<subseteq> carrier R"
+ using generate_field_in_carrier[OF assms] by auto
+
+lemma (in field) zero_in_generate: "\<zero>\<^bsub>R\<^esub> \<in> generate_field R H"
+ using one a_inv generate_field.eng_add one_closed r_neg
+ by metis
+
+lemma (in field) generate_field_is_subfield:
+ assumes "H \<subseteq> carrier R"
+ shows "subfield (generate_field R H) R"
+proof (intro subfieldI', simp_all add: m_inv)
+ show "subring (generate_field R H) R"
+ by (auto intro: subringI[of "generate_field R H"]
+ simp add: eng_add a_inv eng_mult one generate_field_in_carrier[OF assms])
+qed
+
+lemma (in field) generate_field_is_add_subgroup:
+ assumes "H \<subseteq> carrier R"
+ shows "subgroup (generate_field R H) (add_monoid R)"
+ using subring.axioms(1)[OF subfieldE(1)[OF generate_field_is_subfield[OF assms]]] .
+
+lemma (in field) generate_field_is_field :
+ assumes "H \<subseteq> carrier R"
+ shows "field (R \<lparr> carrier := generate_field R H \<rparr>)"
+ using subfield_iff generate_field_is_subfield assms by simp
+
+lemma (in field) generate_field_min_subfield1:
+ assumes "H \<subseteq> carrier R"
+ and "subfield E R" "H \<subseteq> E"
+ shows "generate_field R H \<subseteq> E"
+proof
+ fix h
+ assume h: "h \<in> generate_field R H"
+ show "h \<in> E"
+ using h and assms(3) and subfield_m_inv[OF assms(2)]
+ by (induct rule: generate_field.induct)
+ (auto simp add: subringE(3,5-7)[OF subfieldE(1)[OF assms(2)]])
+qed
+
+lemma (in field) generate_fieldI:
+ assumes "H \<subseteq> carrier R"
+ and "subfield E R" "H \<subseteq> E"
+ and "\<And>K. \<lbrakk> subfield K R; H \<subseteq> K \<rbrakk> \<Longrightarrow> E \<subseteq> K"
+ shows "E = generate_field R H"
+proof
+ show "E \<subseteq> generate_field R H"
+ using assms generate_field_is_subfield generate_field.incl by (metis subset_iff)
+ show "generate_field R H \<subseteq> E"
+ using generate_field_min_subfield1[OF assms(1-3)] by simp
+qed
+
+lemma (in field) generate_fieldE:
+ assumes "H \<subseteq> carrier R" and "E = generate_field R H"
+ shows "subfield E R" and "H \<subseteq> E" and "\<And>K. \<lbrakk> subfield K R; H \<subseteq> K \<rbrakk> \<Longrightarrow> E \<subseteq> K"
+proof -
+ show "subfield E R" using assms generate_field_is_subfield by simp
+ show "H \<subseteq> E" using assms(2) by (simp add: generate_field.incl subsetI)
+ show "\<And>K. subfield K R \<Longrightarrow> H \<subseteq> K \<Longrightarrow> E \<subseteq> K"
+ using assms generate_field_min_subfield1 by auto
+qed
+
+lemma (in field) generate_field_min_subfield2:
+ assumes "H \<subseteq> carrier R"
+ shows "generate_field R H = \<Inter>{K. subfield K R \<and> H \<subseteq> K}"
+proof
+ have "subfield (generate_field R H) R \<and> H \<subseteq> generate_field R H"
+ by (simp add: assms generate_fieldE(2) generate_field_is_subfield)
+ thus "\<Inter>{K. subfield K R \<and> H \<subseteq> K} \<subseteq> generate_field R H" by blast
+next
+ have "\<And>K. subfield K R \<and> H \<subseteq> K \<Longrightarrow> generate_field R H \<subseteq> K"
+ by (simp add: assms generate_field_min_subfield1)
+ thus "generate_field R H \<subseteq> \<Inter>{K. subfield K R \<and> H \<subseteq> K}" by blast
+qed
+
+lemma (in field) mono_generate_field:
+ assumes "I \<subseteq> J" and "J \<subseteq> carrier R"
+ shows "generate_field R I \<subseteq> generate_field R J"
+proof-
+ have "I \<subseteq> generate_field R J "
+ using assms generate_fieldE(2) by blast
+ thus "generate_field R I \<subseteq> generate_field R J"
+ using generate_field_min_subfield1[of I "generate_field R J"] assms generate_field_is_subfield[OF assms(2)]
+ by blast
+qed
+
+
+lemma (in field) subfield_gen_incl :
+ assumes "subfield H R"
+ and "subfield K R"
+ and "I \<subseteq> H"
+ and "I \<subseteq> K"
+ shows "generate_field (R\<lparr>carrier := K\<rparr>) I \<subseteq> generate_field (R\<lparr>carrier := H\<rparr>) I"
+proof
+ {fix J assume J_def : "subfield J R" "I \<subseteq> J"
+ have "generate_field (R \<lparr> carrier := J \<rparr>) I \<subseteq> J"
+ using field.mono_generate_field[of "(R\<lparr>carrier := J\<rparr>)" I J] subfield_iff(2)[OF J_def(1)]
+ field.generate_field_in_carrier[of "R\<lparr>carrier := J\<rparr>"] field_axioms J_def
+ by auto}
+ note incl_HK = this
+ {fix x have "x \<in> generate_field (R\<lparr>carrier := K\<rparr>) I \<Longrightarrow> x \<in> generate_field (R\<lparr>carrier := H\<rparr>) I"
+ proof (induction rule : generate_field.induct)
+ case one
+ have "\<one>\<^bsub>R\<lparr>carrier := H\<rparr>\<^esub> \<otimes> \<one>\<^bsub>R\<lparr>carrier := K\<rparr>\<^esub> = \<one>\<^bsub>R\<lparr>carrier := H\<rparr>\<^esub>" by simp
+ moreover have "\<one>\<^bsub>R\<lparr>carrier := H\<rparr>\<^esub> \<otimes> \<one>\<^bsub>R\<lparr>carrier := K\<rparr>\<^esub> = \<one>\<^bsub>R\<lparr>carrier := K\<rparr>\<^esub>" by simp
+ ultimately show ?case using assms generate_field.one by metis
+ next
+ case (incl h) thus ?case using generate_field.incl by force
+ next
+ case (a_inv h)
+ note hyp = this
+ have "a_inv (R\<lparr>carrier := K\<rparr>) h = a_inv R h"
+ using assms group.m_inv_consistent[of "add_monoid R" K] a_comm_group incl_HK[of K] hyp
+ subring.axioms(1)[OF subfieldE(1)[OF assms(2)]]
+ unfolding comm_group_def a_inv_def by auto
+ moreover have "a_inv (R\<lparr>carrier := H\<rparr>) h = a_inv R h"
+ using assms group.m_inv_consistent[of "add_monoid R" H] a_comm_group incl_HK[of H] hyp
+ subring.axioms(1)[OF subfieldE(1)[OF assms(1)]]
+ unfolding comm_group_def a_inv_def by auto
+ ultimately show ?case using generate_field.a_inv a_inv.IH by fastforce
+ next
+ case (m_inv h)
+ note hyp = this
+ have h_K : "h \<in> (K - {\<zero>})" using incl_HK[OF assms(2) assms(4)] hyp by auto
+ hence "m_inv (R\<lparr>carrier := K\<rparr>) h = m_inv R h"
+ using field.m_inv_mult_of[OF subfield_iff(2)[OF assms(2)]]
+ group.m_inv_consistent[of "mult_of R" "K - {\<zero>}"] field_mult_group units_of_inv
+ subgroup_mult_of subfieldE[OF assms(2)] unfolding mult_of_def apply simp
+ by (metis h_K mult_of_def mult_of_is_Units subgroup.mem_carrier units_of_carrier assms(2))
+ moreover have h_H : "h \<in> (H - {\<zero>})" using incl_HK[OF assms(1) assms(3)] hyp by auto
+ hence "m_inv (R\<lparr>carrier := H\<rparr>) h = m_inv R h"
+ using field.m_inv_mult_of[OF subfield_iff(2)[OF assms(1)]]
+ group.m_inv_consistent[of "mult_of R" "H - {\<zero>}"] field_mult_group
+ subgroup_mult_of[OF assms(1)] unfolding mult_of_def apply simp
+ by (metis h_H field_Units m_inv_mult_of mult_of_is_Units subgroup.mem_carrier units_of_def)
+ ultimately show ?case using generate_field.m_inv m_inv.IH h_H by fastforce
+ next
+ case (eng_add h1 h2)
+ thus ?case using incl_HK assms generate_field.eng_add by force
+ next
+ case (eng_mult h1 h2)
+ thus ?case using generate_field.eng_mult by force
+ qed}
+ thus "\<And>x. x \<in> generate_field (R\<lparr>carrier := K\<rparr>) I \<Longrightarrow> x \<in> generate_field (R\<lparr>carrier := H\<rparr>) I"
+ by auto
+qed
+
+lemma (in field) subfield_gen_equality:
+ assumes "subfield H R" "K \<subseteq> H"
+ shows "generate_field R K = generate_field (R \<lparr> carrier := H \<rparr>) K"
+ using subfield_gen_incl[OF assms(1) carrier_is_subfield assms(2)] assms subringE(1)
+ subfield_gen_incl[OF carrier_is_subfield assms(1) _ assms(2)] subfieldE(1)[OF assms(1)]
+ by force
+
+
+
+end
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Algebra/Generated_Groups.thy Mon Jul 02 15:43:22 2018 +0100
@@ -0,0 +1,631 @@
+(* ************************************************************************** *)
+(* Title: Generated_Groups.thy *)
+(* Author: Paulo EmÃlio de Vilhena *)
+(* ************************************************************************** *)
+
+theory Generated_Groups
+ imports Group Coset
+
+begin
+
+section\<open>Generated Groups\<close>
+
+inductive_set
+ generate :: "('a, 'b) monoid_scheme \<Rightarrow> 'a set \<Rightarrow> 'a set"
+ for G and H where
+ one: "\<one>\<^bsub>G\<^esub> \<in> generate G H"
+ | incl: "h \<in> H \<Longrightarrow> h \<in> generate G H"
+ | inv: "h \<in> H \<Longrightarrow> inv\<^bsub>G\<^esub> h \<in> generate G H"
+ | eng: "h1 \<in> generate G H \<Longrightarrow> h2 \<in> generate G H \<Longrightarrow> h1 \<otimes>\<^bsub>G\<^esub> h2 \<in> generate G H"
+
+
+subsection\<open>Basic Properties of Generated Groups - First Part\<close>
+
+lemma (in group) generate_in_carrier:
+ assumes "H \<subseteq> carrier G"
+ shows "h \<in> generate G H \<Longrightarrow> h \<in> carrier G"
+ apply (induction rule: generate.induct) using assms by blast+
+
+lemma (in group) generate_m_inv_closed:
+ assumes "H \<subseteq> carrier G"
+ shows "h \<in> generate G H \<Longrightarrow> (inv h) \<in> 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 \<otimes> h2) = (inv h2) \<otimes> (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 \<subseteq> carrier G"
+ shows "subgroup (generate G H) G"
+proof (intro subgroupI)
+ show "generate G H \<subseteq> carrier G" using generate_in_carrier[OF assms] by blast
+ show "generate G H \<noteq> {}" using generate.one by auto
+ show "\<And>h. h \<in> generate G H \<Longrightarrow> inv h \<in> generate G H"
+ using generate_m_inv_closed[OF assms] by blast
+ show "\<And>h1 h2. \<lbrakk> h1 \<in> generate G H; h2 \<in> generate G H \<rbrakk> \<Longrightarrow> h1 \<otimes> h2 \<in> generate G H"
+ by (simp add: generate.eng)
+qed
+
+
+subsection\<open>Characterisations of Generated Groups\<close>
+
+lemma (in group) generate_min_subgroup1:
+ assumes "H \<subseteq> carrier G"
+ and "subgroup E G" "H \<subseteq> E"
+ shows "generate G H \<subseteq> E"
+proof
+ fix h show "h \<in> generate G H \<Longrightarrow> h \<in> 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 \<subseteq> carrier G"
+ and "subgroup E G" "H \<subseteq> E"
+ and "\<And>K. \<lbrakk> subgroup K G; H \<subseteq> K \<rbrakk> \<Longrightarrow> E \<subseteq> K"
+ shows "E = generate G H"
+proof
+ show "E \<subseteq> generate G H"
+ using assms generate_is_subgroup generate.incl by (metis subset_iff)
+ show "generate G H \<subseteq> E"
+ using generate_min_subgroup1[OF assms(1-3)] by simp
+qed
+
+lemma (in group) generateE:
+ assumes "H \<subseteq> carrier G" and "E = generate G H"
+ shows "subgroup E G" and "H \<subseteq> E" and "\<And>K. \<lbrakk> subgroup K G; H \<subseteq> K \<rbrakk> \<Longrightarrow> E \<subseteq> K"
+proof -
+ show "subgroup E G" using assms generate_is_subgroup by simp
+ show "H \<subseteq> E" using assms(2) by (simp add: generate.incl subsetI)
+ show "\<And>K. subgroup K G \<Longrightarrow> H \<subseteq> K \<Longrightarrow> E \<subseteq> K"
+ using assms generate_min_subgroup1 by auto
+qed
+
+lemma (in group) generate_min_subgroup2:
+ assumes "H \<subseteq> carrier G"
+ shows "generate G H = \<Inter>{K. subgroup K G \<and> H \<subseteq> K}"
+proof
+ have "subgroup (generate G H) G \<and> H \<subseteq> generate G H"
+ by (simp add: assms generateE(2) generate_is_subgroup)
+ thus "\<Inter>{K. subgroup K G \<and> H \<subseteq> K} \<subseteq> generate G H" by blast
+next
+ have "\<And>K. subgroup K G \<and> H \<subseteq> K \<Longrightarrow> generate G H \<subseteq> K"
+ by (simp add: assms generate_min_subgroup1)
+ thus "generate G H \<subseteq> \<Inter>{K. subgroup K G \<and> H \<subseteq> K}" by blast
+qed
+
+
+subsection\<open>Representation of Elements from a Generated Group\<close>
+
+text\<open>We define a sort of syntax tree to allow induction arguments with elements of a generated group\<close>
+
+datatype 'a repr =
+ One | Inv "'a" | Leaf "'a" | Mult "'a repr" "'a repr"
+
+fun norm :: "('a, 'b) monoid_scheme \<Rightarrow> 'a repr \<Rightarrow> 'a"
+ where
+ "norm G (One) = \<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) \<otimes>\<^bsub>G\<^esub> (norm G h2)"
+
+fun elts :: "'a repr \<Rightarrow> 'a set"
+ where
+ "elts (One) = {}"
+ | "elts (Inv h) = { h }"
+ | "elts (Leaf h) = { h }"
+ | "elts (Mult h1 h2) = (elts h1) \<union> (elts h2)"
+
+lemma (in group) generate_repr_iff:
+ assumes "H \<subseteq> carrier G"
+ shows "(h \<in> generate G H) \<longleftrightarrow> (\<exists>r. (elts r) \<subseteq> H \<and> norm G r = h)"
+proof
+ show "h \<in> generate G H \<Longrightarrow> \<exists>r. (elts r) \<subseteq> H \<and> 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) \<subseteq> H \<and> norm G (Leaf h) = h" by simp
+ thus ?case by blast
+ next
+ case (inv h)
+ hence "elts (Inv h) \<subseteq> H \<and> 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 \<subseteq> H" "norm G r1 = h1"
+ and r2: "elts r2 \<subseteq> H" "norm G r2 = h2" by blast
+ hence "elts (Mult r1 r2) \<subseteq> H \<and> norm G (Mult r1 r2) = h1 \<otimes> h2" by simp
+ thus ?case by blast
+ qed
+
+ show "\<exists>r. elts r \<subseteq> H \<and> norm G r = h \<Longrightarrow> h \<in> generate G H"
+ proof -
+ assume "\<exists>r. elts r \<subseteq> H \<and> norm G r = h"
+ then obtain r where "elts r \<subseteq> H" "norm G r = h" by blast
+ thus "h \<in> 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
+ qed
+ qed
+qed
+
+corollary (in group) generate_repr_set:
+ assumes "H \<subseteq> carrier G"
+ shows "generate G H = {norm G r | r. (elts r) \<subseteq> H}" (is "?A = ?B")
+proof
+ show "?A \<subseteq> ?B"
+ proof
+ fix h assume "h \<in> generate G H" thus "h \<in> {norm G r |r. elts r \<subseteq> H}"
+ using generate_repr_iff[OF assms] by auto
+ qed
+next
+ show "?B \<subseteq> ?A"
+ proof
+ fix h assume "h \<in> {norm G r |r. elts r \<subseteq> H}" thus "h \<in> generate G H"
+ using generate_repr_iff[OF assms] by auto
+ qed
+qed
+
+corollary (in group) mono_generate:
+ assumes "I \<subseteq> J" and "J \<subseteq> carrier G"
+ shows "generate G I \<subseteq> generate G J"
+ using assms generate_repr_iff by fastforce
+
+lemma (in group) subgroup_gen_equality:
+ assumes "subgroup H G" "K \<subseteq> H"
+ shows "generate G K = generate (G \<lparr> carrier := H \<rparr>) K"
+proof -
+ have "generate G K \<subseteq> H"
+ by (meson assms generate_min_subgroup1 order.trans subgroup.subset)
+ have mult_eq: "\<And>k1 k2. \<lbrakk> k1 \<in> generate G K; k2 \<in> generate G K \<rbrakk> \<Longrightarrow>
+ k1 \<otimes>\<^bsub>G \<lparr> carrier := H \<rparr>\<^esub> k2 = k1 \<otimes> k2"
+ using \<open>generate G K \<subseteq> H\<close> subgroup_mult_equality by simp
+
+ { fix r assume A: "elts r \<subseteq> K"
+ hence "norm G r = norm (G \<lparr> carrier := H \<rparr>) r"
+ proof (induction r rule: repr.induct)
+ case One thus ?case by simp
+ next
+ case (Inv k) hence "k \<in> K" using A by simp
+ thus ?case using m_inv_consistent[OF assms(1)] assms(2) by auto
+ next
+ case (Leaf k) hence "k \<in> 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
+
+ show ?thesis
+ proof
+ show "generate G K \<subseteq> generate (G\<lparr>carrier := H\<rparr>) K"
+ proof
+ fix h assume "h \<in> generate G K" then obtain r where r: "elts r \<subseteq> K" "h = norm G r"
+ using generate_repr_iff assms by (metis order.trans subgroup.subset)
+ hence "h = norm (G \<lparr> carrier := H \<rparr>) r" using aux_lemma by simp
+ thus "h \<in> generate (G\<lparr>carrier := H\<rparr>) K"
+ using r assms group.generate_repr_iff [of "G \<lparr> carrier := H \<rparr>" K]
+ subgroup.subgroup_is_group[OF assms(1) is_group] by auto
+ qed
+ show "generate (G\<lparr>carrier := H\<rparr>) K \<subseteq> generate G K"
+ proof
+ fix h assume "h \<in> generate (G\<lparr>carrier := H\<rparr>) K"
+ then obtain r where r: "elts r \<subseteq> K" "h = norm (G\<lparr>carrier := H\<rparr>) r"
+ using group.generate_repr_iff [of "G \<lparr> carrier := H \<rparr>" 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 \<in> generate G K"
+ by (meson assms generate_repr_iff order.trans r(1) subgroup.subset)
+ qed
+ qed
+qed
+
+corollary (in group) gen_equality_betw_subgroups:
+ assumes "subgroup I G" "subgroup J G" "K \<subseteq> (I \<inter> J)"
+ shows "generate (G \<lparr> carrier := I \<rparr>) K = generate (G \<lparr> carrier := J \<rparr>) K"
+ by (metis Int_subset_iff assms subgroup_gen_equality)
+
+lemma (in group) normal_generateI:
+ assumes "H \<subseteq> carrier G"
+ and "\<And>h g. \<lbrakk> h \<in> H; g \<in> carrier G \<rbrakk> \<Longrightarrow> g \<otimes> h \<otimes> (inv g) \<in> H"
+ shows "generate G H \<lhd> G"
+proof (rule normal_invI)
+ show "subgroup (generate G H) G"
+ by (simp add: assms(1) generate_is_subgroup)
+next
+ have "\<And>r g. \<lbrakk> elts r \<subseteq> H; g \<in> carrier G \<rbrakk> \<Longrightarrow> (g \<otimes> (norm G r) \<otimes> (inv g)) \<in> (generate G H)"
+ proof -
+ fix r g assume "elts r \<subseteq> H" "g \<in> carrier G"
+ thus "(g \<otimes> (norm G r) \<otimes> (inv g)) \<in> (generate G H)"
+ proof (induction r rule: repr.induct)
+ case One thus ?case
+ by (simp add: generate.one)
+ next
+ case (Inv h)
+ hence "g \<otimes> h \<otimes> (inv g) \<in> H" using assms(2) by simp
+ moreover have "norm G (Inv (g \<otimes> h \<otimes> (inv g))) = g \<otimes> (inv h) \<otimes> (inv g)"
+ using Inv.prems(1) Inv.prems(2) assms(1) inv_mult_group m_assoc by auto
+ ultimately have "\<exists>r. elts r \<subseteq> H \<and> norm G r = g \<otimes> (inv h) \<otimes> (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 \<otimes> h \<otimes> inv g" H] by simp
+ next
+ case (Mult h1 h2)
+ hence "elts h1 \<subseteq> H \<and> elts h2 \<subseteq> H" using Mult(3) by simp
+ hence in_gen: "norm G h1 \<in> generate G H \<and> norm G h2 \<in> generate G H"
+ using assms(1) generate_repr_iff by auto
+
+ have "g \<otimes> norm G (Mult h1 h2) \<otimes> inv g = g \<otimes> (norm G h1 \<otimes> norm G h2) \<otimes> inv g" by simp
+ also have " ... = g \<otimes> (norm G h1 \<otimes> (inv g \<otimes> g) \<otimes> norm G h2) \<otimes> inv g"
+ using Mult(4) in_gen assms(1) generate_in_carrier by auto
+ also have " ... = (g \<otimes> norm G h1 \<otimes> inv g) \<otimes> (g \<otimes> norm G h2 \<otimes> inv g)"
+ using Mult.prems(2) assms(1) generate_in_carrier in_gen inv_closed m_assoc m_closed by presburger
+ finally have "g \<otimes> norm G (Mult h1 h2) \<otimes> inv g =
+ (g \<otimes> norm G h1 \<otimes> inv g) \<otimes> (g \<otimes> norm G h2 \<otimes> inv g)" .
+
+ thus ?case
+ using generate.eng[of "g \<otimes> norm G h1 \<otimes> inv g" G H "g \<otimes> norm G h2 \<otimes> inv g"]
+ by (simp add: Mult.IH Mult.prems(2) \<open>elts h1 \<subseteq> H \<and> elts h2 \<subseteq> H\<close>)
+ qed
+ qed
+ thus "\<And>x h. x \<in> carrier G \<Longrightarrow> h \<in> generate G H \<Longrightarrow> x \<otimes> h \<otimes> inv x \<in> generate G H"
+ using assms(1) generate_repr_iff by auto
+qed
+
+
+subsection\<open>Basic Properties of Generated Groups - Second Part\<close>
+
+lemma (in group) generate_pow:
+ assumes "a \<in> carrier G"
+ shows "generate G { a } = { a [^] k | k. k \<in> (UNIV :: int set) }"
+proof
+ show "generate G { a } \<subseteq> { a [^] k | k. k \<in> (UNIV :: int set) }"
+ proof
+ fix h show "h \<in> generate G { a } \<Longrightarrow> h \<in> { a [^] k | k. k \<in> (UNIV :: int set) }"
+ proof (induction rule: generate.induct)
+ case one thus ?case by (metis (mono_tags, lifting) UNIV_I int_pow_0 mem_Collect_eq)
+ next
+ case (incl h) hence "h = a" by auto thus ?case
+ by (metis (mono_tags, lifting) CollectI UNIV_I assms group.int_pow_1 is_group)
+ next
+ case (inv h) hence "h = a" by auto thus ?case
+ by (metis (mono_tags, lifting) mem_Collect_eq UNIV_I assms group.int_pow_1 int_pow_neg is_group)
+ next
+ case (eng h1 h2) thus ?case
+ by (smt assms group.int_pow_mult is_group iso_tuple_UNIV_I mem_Collect_eq)
+ qed
+ qed
+
+ show "{ a [^] k | k. k \<in> (UNIV :: int set) } \<subseteq> generate G { a }"
+ proof
+ { fix k :: "nat" have "a [^] k \<in> 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 \<in> { a [^] k | k. k \<in> (UNIV :: int set) }"
+ then obtain k :: "nat" where "h = a [^] k \<or> h = inv (a [^] k)"
+ by (smt group.int_pow_def2 is_group mem_Collect_eq)
+ thus "h \<in> generate G { a }" using aux_lemma
+ using assms generate_m_inv_closed by auto
+ qed
+qed
+
+corollary (in group) generate_one: "generate G { \<one> } = { \<one> }"
+ using generate_pow[of "\<one>", OF one_closed] by simp
+
+corollary (in group) generate_empty: "generate G {} = { \<one> }"
+ using mono_generate[of "{}" "{ \<one> }"] generate_one generate.one one_closed by blast
+
+corollary (in group)
+ assumes "H \<subseteq> carrier G" "h \<in> H"
+ shows "h [^] (k :: int) \<in> generate G H"
+ using mono_generate[of "{ h }" H] generate_pow[of h] assms by auto
+
+
+subsection\<open>Derived Subgroup\<close>
+
+abbreviation derived_set :: "('a, 'b) monoid_scheme \<Rightarrow> 'a set \<Rightarrow> 'a set" where
+ "derived_set G H \<equiv>
+ \<Union>h1 \<in> H. (\<Union>h2 \<in> H. { h1 \<otimes>\<^bsub>G\<^esub> h2 \<otimes>\<^bsub>G\<^esub> (inv\<^bsub>G\<^esub> h1) \<otimes>\<^bsub>G\<^esub> (inv\<^bsub>G\<^esub> h2) })"
+
+definition derived :: "('a, 'b) monoid_scheme \<Rightarrow> 'a set \<Rightarrow> 'a set" where
+ "derived G H = generate G (derived_set G H)"
+
+lemma (in group) derived_set_incl:
+ assumes "subgroup H G"
+ shows "derived_set G H \<subseteq> H"
+ by (auto simp add: m_inv_consistent[OF assms] subgroupE[OF assms])
+
+lemma (in group) derived_incl:
+ assumes "subgroup H G"
+ shows "derived G H \<subseteq> 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 \<subseteq> H"
+ shows "derived (G \<lparr> carrier := H \<rparr>) K = derived G K"
+proof -
+ have "derived_set G K \<subseteq> H"
+ proof
+ fix x assume "x \<in> derived_set G K"
+ then obtain k1 k2
+ where k12: "k1 \<in> K" "k2 \<in> K"
+ and "x = k1 \<otimes> k2 \<otimes> inv k1 \<otimes> inv k2" by blast
+ thus "x \<in> H" using k12 assms by (meson subgroup_def subsetCE)
+ qed
+
+ moreover have "derived_set (G \<lparr> carrier := H \<rparr>) K = derived_set G K"
+ proof
+ show "derived_set G K \<subseteq> derived_set (G\<lparr>carrier := H\<rparr>) K"
+ proof
+ fix x assume "x \<in> derived_set G K"
+ then obtain k1 k2 where k12: "k1 \<in> K" "k2 \<in> K"
+ and "x = k1 \<otimes> k2 \<otimes> inv k1 \<otimes> inv k2" by blast
+ hence "x = k1 \<otimes>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> k2 \<otimes>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> inv\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> k1 \<otimes>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> inv\<^bsub>G\<lparr>carrier := H\<rparr>\<^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 \<in> derived_set (G\<lparr>carrier := H\<rparr>) K" using k12 by blast
+ qed
+ next
+ show "derived_set (G \<lparr> carrier := H \<rparr>) K \<subseteq> derived_set G K"
+ proof
+ fix x assume "x \<in> derived_set (G \<lparr> carrier := H \<rparr>) K"
+ then obtain k1 k2
+ where k12: "k1 \<in> K" "k2 \<in> K"
+ and "x = k1 \<otimes>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> k2 \<otimes>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> inv\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> k1 \<otimes>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> inv\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> k2"
+ by blast
+ hence "x = k1 \<otimes> k2 \<otimes> inv k1 \<otimes> 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 \<in> 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\<lparr>carrier := H\<rparr>) K"] by simp
+qed
+
+lemma (in comm_group) derived_set:
+ assumes "H \<subseteq> carrier G"
+ shows "derived G H = { \<one> }"
+proof -
+ have "derived_set G H = {} \<or> derived_set G H = { \<one> }"
+ proof (cases)
+ assume "H = {}" thus ?thesis by simp
+ next
+ assume "H \<noteq> {}" then obtain h' where h': "h' \<in> H" by blast
+ have "derived_set G H = { \<one> }"
+ proof -
+ { fix h assume A: "h \<in> derived_set G H"
+ have "h = \<one>"
+ proof -
+ obtain h1 h2 where h1: "h1 \<in> H" and h2: "h2 \<in> H" and h: "h = h1 \<otimes> h2 \<otimes> inv h1 \<otimes> inv h2"
+ using A by blast
+ hence "h = (h1 \<otimes> inv h1) \<otimes> (h2 \<otimes> inv h2)" using assms
+ by (smt inv_closed inv_mult m_assoc m_closed r_inv set_rev_mp)
+ thus ?thesis using h1 h2 assms by (metis contra_subsetD one_closed r_inv r_one)
+ qed } note aux_lemma = this
+ show ?thesis
+ proof
+ show "derived_set G H \<subseteq> { \<one> }" using aux_lemma by blast
+ next
+ show "{ \<one> } \<subseteq> derived_set G H"
+ proof -
+ have "h' \<otimes> h' \<otimes> inv h' \<otimes> inv h' \<in> 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
+
+lemma (in group) derived_set_in_carrier:
+ assumes "H \<subseteq> carrier G"
+ shows "derived_set G H \<subseteq> carrier G"
+proof
+ fix h assume "h \<in> derived_set G H"
+ then obtain h1 h2 where "h1 \<in> H" "h2 \<in> H" "h = h1 \<otimes> h2 \<otimes> inv h1 \<otimes> inv h2"
+ by blast
+ thus "h \<in> carrier G" using assms by blast
+qed
+
+lemma (in group) derived_is_normal:
+ assumes "H \<lhd> G"
+ shows "derived G H \<lhd> G" unfolding derived_def
+proof (rule normal_generateI)
+ show "(\<Union>h1 \<in> H. \<Union>h2 \<in> H. { h1 \<otimes> h2 \<otimes> inv h1 \<otimes> inv h2 }) \<subseteq> carrier G"
+ using subgroup.subset assms normal_invE(1) by blast
+next
+ show "\<And>h g. \<lbrakk> h \<in> derived_set G H; g \<in> carrier G \<rbrakk> \<Longrightarrow> g \<otimes> h \<otimes> inv g \<in> derived_set G H"
+ proof -
+ fix h g assume "h \<in> derived_set G H" and g: "g \<in> carrier G"
+ then obtain h1 h2 where h1: "h1 \<in> H" "h1 \<in> carrier G"
+ and h2: "h2 \<in> H" "h2 \<in> carrier G"
+ and h: "h = h1 \<otimes> h2 \<otimes> inv h1 \<otimes> inv h2"
+ using subgroup.subset assms normal_invE(1) by blast
+ hence "g \<otimes> h \<otimes> inv g =
+ g \<otimes> h1 \<otimes> (inv g \<otimes> g) \<otimes> h2 \<otimes> (inv g \<otimes> g) \<otimes> inv h1 \<otimes> (inv g \<otimes> g) \<otimes> inv h2 \<otimes> inv g"
+ by (simp add: g m_assoc)
+ also
+ have " ... =
+ (g \<otimes> h1 \<otimes> inv g) \<otimes> (g \<otimes> h2 \<otimes> inv g) \<otimes> (g \<otimes> inv h1 \<otimes> inv g) \<otimes> (g \<otimes> inv h2 \<otimes> inv g)"
+ using g h1 h2 m_assoc inv_closed m_closed by presburger
+ finally
+ have "g \<otimes> h \<otimes> inv g =
+ (g \<otimes> h1 \<otimes> inv g) \<otimes> (g \<otimes> h2 \<otimes> inv g) \<otimes> inv (g \<otimes> h1 \<otimes> inv g) \<otimes> inv (g \<otimes> h2 \<otimes> inv g)"
+ by (simp add: g h1 h2 inv_mult_group m_assoc)
+ moreover have "g \<otimes> h1 \<otimes> inv g \<in> H" by (simp add: assms normal_invE(2) g h1)
+ moreover have "g \<otimes> h2 \<otimes> inv g \<in> H" by (simp add: assms normal_invE(2) g h2)
+ ultimately show "g \<otimes> h \<otimes> inv g \<in> derived_set G H" by blast
+ qed
+qed
+
+corollary (in group) derived_self_is_normal: "derived G (carrier G) \<lhd> G"
+ by (simp add: group.derived_is_normal group.normal_invI is_group subgroup_self)
+
+corollary (in group) derived_subgroup_is_normal:
+ assumes "subgroup H G"
+ shows "derived G H \<lhd> G \<lparr> carrier := H \<rparr>"
+proof -
+ have "H \<lhd> G \<lparr> carrier := H \<rparr>"
+ 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 \<lparr> carrier := H \<rparr>) H \<lhd> G \<lparr> carrier := H \<rparr>"
+ using group.derived_is_normal[of "G \<lparr> carrier := H \<rparr>" H] normal_def by blast
+ thus ?thesis using subgroup_derived_equality[OF assms] by simp
+qed
+
+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
+
+lemma (in group) derived_quot_is_comm:
+ assumes "H \<in> carrier (G Mod (derived G (carrier G)))"
+ and "K \<in> carrier (G Mod (derived G (carrier G)))"
+ shows "H <#> K = K <#> H"
+proof -
+ { fix H K assume A1: "H \<in> carrier (G Mod (derived G (carrier G)))"
+ and A2: "K \<in> carrier (G Mod (derived G (carrier G)))"
+ have "H <#> K \<subseteq> K <#> H"
+ proof -
+ obtain h k where hk: "h \<in> carrier G" "k \<in> 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 \<otimes> 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 \<otimes> 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 \<otimes> k) <# (derived G (carrier G)) \<subseteq> (k \<otimes> h) <# (derived G (carrier G))"
+ proof
+ fix x assume "x \<in> h \<otimes> k <# derived G (carrier G)"
+ then obtain d where d: "d \<in> derived G (carrier G)" "d \<in> carrier G" "x = h \<otimes> k \<otimes> 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 \<otimes> (h \<otimes> inv h) \<otimes> inv k) \<otimes> h \<otimes> k \<otimes> d"
+ using hk by simp
+ also have " ... = (k \<otimes> h) \<otimes> (inv h \<otimes> inv k) \<otimes> h \<otimes> k \<otimes> d"
+ using d(2) hk m_assoc by (metis subgroup_def subgroup_self)
+ also have " ... = (k \<otimes> h) \<otimes> ((inv h \<otimes> inv k \<otimes> h \<otimes> k) \<otimes> d)"
+ using d(2) hk m_assoc by simp
+ finally have "x = (k \<otimes> h) \<otimes> ((inv h \<otimes> inv k \<otimes> h \<otimes> k) \<otimes> d)" .
+
+ moreover have "inv h \<otimes> inv k \<otimes> inv (inv h) \<otimes> inv (inv k) \<in> derived_set G (carrier G)"
+ using inv_closed[OF hk(1)] inv_closed[OF hk(2)] by blast
+ hence "inv h \<otimes> inv k \<otimes> h \<otimes> k \<in> derived_set G (carrier G)"
+ by (simp add: hk(1) hk(2))
+ hence "(inv h \<otimes> inv k \<otimes> h \<otimes> k) \<otimes> d \<in> derived G (carrier G)"
+ using d(1) unfolding derived_def by (simp add: generate.eng generate.incl)
+
+ ultimately show "x \<in> (k \<otimes> h) <# (derived G (carrier G))"
+ unfolding l_coset_def by blast
+ qed
+ ultimately show ?thesis by simp
+ qed }
+ thus ?thesis using assms by auto
+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 \<lparr> carrier := H \<rparr>) Mod (derived G H))"
+proof -
+ have "group (G \<lparr> carrier := H \<rparr>)"
+ using assms subgroup_imp_group by auto
+ thus ?thesis
+ using group.derived_quot_is_comm_group subgroup_derived_equality[OF assms] by fastforce
+qed
+
+lemma (in group) mono_derived:
+ assumes "J \<subseteq> carrier G" "I \<subseteq> J"
+ shows "(derived G ^^ n) I \<subseteq> (derived G ^^ n) J"
+proof -
+ { fix I J assume A: "J \<subseteq> carrier G" "I \<subseteq> J" have "derived G I \<subseteq> derived G J"
+ proof -
+ have "derived_set G I \<subseteq> 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 \<subseteq> carrier G" have "(derived G ^^ n) I \<subseteq> carrier G"
+ proof (induction n)
+ case 0 thus ?case using A by simp
+ next
+ case (Suc n) thus ?case
+ using aux_lemma1 derived_self_is_normal funpow_simps_right(2) funpow_swap1
+ normal_def o_apply order.trans order_refl subgroup.subset by smt
+ qed } note aux_lemma2 = this
+
+ 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
+ qed
+qed
+
+lemma (in group) exp_of_derived_in_carrier:
+ assumes "H \<subseteq> carrier G"
+ shows "(derived G ^^ n) H \<subseteq> 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 " ... \<subseteq> (derived G) (carrier G)"
+ using mono_derived[of "carrier G" "(derived G ^^ n) H" 1] Suc by simp
+ also have " ... \<subseteq> carrier G" unfolding derived_def
+ by (simp add: derived_set_incl generate_min_subgroup1 subgroup_self)
+ finally show ?case .
+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 \<subseteq> 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
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Algebra/Generated_Rings.thy Mon Jul 02 15:43:22 2018 +0100
@@ -0,0 +1,154 @@
+(* ************************************************************************** *)
+(* Title: Generated_Rings.thy *)
+(* Author: Martin Baillon *)
+(* ************************************************************************** *)
+
+theory Generated_Rings
+ imports Subrings
+begin
+
+section\<open>Generated Rings\<close>
+
+inductive_set
+ generate_ring :: "('a, 'b) ring_scheme \<Rightarrow> 'a set \<Rightarrow> 'a set"
+ for R and H where
+ one: "\<one>\<^bsub>R\<^esub> \<in> generate_ring R H"
+ | incl: "h \<in> H \<Longrightarrow> h \<in> generate_ring R H"
+ | a_inv: "h \<in> generate_ring R H \<Longrightarrow> \<ominus>\<^bsub>R\<^esub> h \<in> generate_ring R H"
+ | eng_add : "\<lbrakk> h1 \<in> generate_ring R H; h2 \<in> generate_ring R H \<rbrakk> \<Longrightarrow> h1 \<oplus>\<^bsub>R\<^esub> h2 \<in> generate_ring R H"
+ | eng_mult: "\<lbrakk> h1 \<in> generate_ring R H; h2 \<in> generate_ring R H \<rbrakk> \<Longrightarrow> h1 \<otimes>\<^bsub>R\<^esub> h2 \<in> generate_ring R H"
+
+subsection\<open>Basic Properties of Generated Rings - First Part\<close>
+
+lemma (in ring) generate_ring_in_carrier:
+ assumes "H \<subseteq> carrier R"
+ shows "h \<in> generate_ring R H \<Longrightarrow> h \<in> carrier R"
+ apply (induction rule: generate_ring.induct) using assms
+ by blast+
+
+lemma (in ring) generate_ring_incl:
+ assumes "H \<subseteq> carrier R"
+ shows "generate_ring R H \<subseteq> carrier R"
+ using generate_ring_in_carrier[OF assms] by auto
+
+lemma (in ring) zero_in_generate: "\<zero>\<^bsub>R\<^esub> \<in> generate_ring R H"
+ using one a_inv by (metis generate_ring.eng_add one_closed r_neg)
+
+lemma (in ring) generate_ring_is_subring:
+ assumes "H \<subseteq> carrier R"
+ shows "subring (generate_ring R H) R"
+ by (auto intro!: subringI[of "generate_ring R H"]
+ simp add: generate_ring_in_carrier[OF assms] one a_inv eng_mult eng_add)
+
+lemma (in ring) generate_ring_is_ring:
+ assumes "H \<subseteq> carrier R"
+ shows "ring (R \<lparr> carrier := generate_ring R H \<rparr>)"
+ using subring_iff[OF generate_ring_incl[OF assms]] generate_ring_is_subring[OF assms] by simp
+
+lemma (in ring) generate_ring_min_subring1:
+ assumes "H \<subseteq> carrier R" and "subring E R" "H \<subseteq> E"
+ shows "generate_ring R H \<subseteq> E"
+proof
+ fix h assume h: "h \<in> generate_ring R H"
+ show "h \<in> E"
+ using h and assms(3)
+ by (induct rule: generate_ring.induct)
+ (auto simp add: subringE(3,5-7)[OF assms(2)])
+qed
+
+lemma (in ring) generate_ringI:
+ assumes "H \<subseteq> carrier R"
+ and "subring E R" "H \<subseteq> E"
+ and "\<And>K. \<lbrakk> subring K R; H \<subseteq> K \<rbrakk> \<Longrightarrow> E \<subseteq> K"
+ shows "E = generate_ring R H"
+proof
+ show "E \<subseteq> generate_ring R H"
+ using assms generate_ring_is_subring generate_ring.incl by (metis subset_iff)
+ show "generate_ring R H \<subseteq> E"
+ using generate_ring_min_subring1[OF assms(1-3)] by simp
+qed
+
+lemma (in ring) generate_ringE:
+ assumes "H \<subseteq> carrier R" and "E = generate_ring R H"
+ shows "subring E R" and "H \<subseteq> E" and "\<And>K. \<lbrakk> subring K R; H \<subseteq> K \<rbrakk> \<Longrightarrow> E \<subseteq> K"
+proof -
+ show "subring E R" using assms generate_ring_is_subring by simp
+ show "H \<subseteq> E" using assms(2) by (simp add: generate_ring.incl subsetI)
+ show "\<And>K. subring K R \<Longrightarrow> H \<subseteq> K \<Longrightarrow> E \<subseteq> K"
+ using assms generate_ring_min_subring1 by auto
+qed
+
+lemma (in ring) generate_ring_min_subring2:
+ assumes "H \<subseteq> carrier R"
+ shows "generate_ring R H = \<Inter>{K. subring K R \<and> H \<subseteq> K}"
+proof
+ have "subring (generate_ring R H) R \<and> H \<subseteq> generate_ring R H"
+ by (simp add: assms generate_ringE(2) generate_ring_is_subring)
+ thus "\<Inter>{K. subring K R \<and> H \<subseteq> K} \<subseteq> generate_ring R H" by blast
+next
+ have "\<And>K. subring K R \<and> H \<subseteq> K \<Longrightarrow> generate_ring R H \<subseteq> K"
+ by (simp add: assms generate_ring_min_subring1)
+ thus "generate_ring R H \<subseteq> \<Inter>{K. subring K R \<and> H \<subseteq> K}" by blast
+qed
+
+lemma (in ring) mono_generate_ring:
+ assumes "I \<subseteq> J" and "J \<subseteq> carrier R"
+ shows "generate_ring R I \<subseteq> generate_ring R J"
+proof-
+ have "I \<subseteq> generate_ring R J "
+ using assms generate_ringE(2) by blast
+ thus "generate_ring R I \<subseteq> generate_ring R J"
+ using generate_ring_min_subring1[of I "generate_ring R J"] assms generate_ring_is_subring[OF assms(2)]
+ by blast
+qed
+
+lemma (in ring) subring_gen_incl :
+ assumes "subring H R"
+ and "subring K R"
+ and "I \<subseteq> H"
+ and "I \<subseteq> K"
+ shows "generate_ring (R\<lparr>carrier := K\<rparr>) I \<subseteq> generate_ring (R\<lparr>carrier := H\<rparr>) I"
+proof
+ {fix J assume J_def : "subring J R" "I \<subseteq> J"
+ have "generate_ring (R \<lparr> carrier := J \<rparr>) I \<subseteq> J"
+ using ring.mono_generate_ring[of "(R\<lparr>carrier := J\<rparr>)" I J ] subring.subring_is_ring[OF J_def(1)]
+ ring.generate_ring_in_carrier[of "R\<lparr>carrier := J\<rparr>"] ring_axioms J_def(2)
+ by auto}
+ note incl_HK = this
+ {fix x have "x \<in> generate_ring (R\<lparr>carrier := K\<rparr>) I \<Longrightarrow> x \<in> generate_ring (R\<lparr>carrier := H\<rparr>) I"
+ proof (induction rule : generate_ring.induct)
+ case one
+ have "\<one>\<^bsub>R\<lparr>carrier := H\<rparr>\<^esub> \<otimes> \<one>\<^bsub>R\<lparr>carrier := K\<rparr>\<^esub> = \<one>\<^bsub>R\<lparr>carrier := H\<rparr>\<^esub>" by simp
+ moreover have "\<one>\<^bsub>R\<lparr>carrier := H\<rparr>\<^esub> \<otimes> \<one>\<^bsub>R\<lparr>carrier := K\<rparr>\<^esub> = \<one>\<^bsub>R\<lparr>carrier := K\<rparr>\<^esub>" by simp
+ ultimately show ?case using assms generate_ring.one by metis
+ next
+ case (incl h) thus ?case using generate_ring.incl by force
+ next
+ case (a_inv h)
+ note hyp = this
+ have "a_inv (R\<lparr>carrier := K\<rparr>) h = a_inv R h"
+ using assms group.m_inv_consistent[of "add_monoid R" K] a_comm_group incl_HK[of K] hyp
+ unfolding subring_def comm_group_def a_inv_def by auto
+ moreover have "a_inv (R\<lparr>carrier := H\<rparr>) h = a_inv R h"
+ using assms group.m_inv_consistent[of "add_monoid R" H] a_comm_group incl_HK[of H] hyp
+ unfolding subring_def comm_group_def a_inv_def by auto
+ ultimately show ?case using generate_ring.a_inv a_inv.IH by fastforce
+ next
+ case (eng_add h1 h2)
+ thus ?case using incl_HK assms generate_ring.eng_add by force
+ next
+ case (eng_mult h1 h2)
+ thus ?case using generate_ring.eng_mult by force
+ qed}
+ thus "\<And>x. x \<in> generate_ring (R\<lparr>carrier := K\<rparr>) I \<Longrightarrow> x \<in> generate_ring (R\<lparr>carrier := H\<rparr>) I"
+ by auto
+qed
+
+lemma (in ring) subring_gen_equality:
+ assumes "subring H R" "K \<subseteq> H"
+ shows "generate_ring R K = generate_ring (R \<lparr> carrier := H \<rparr>) K"
+ using subring_gen_incl[OF assms(1)carrier_is_subring assms(2)] assms subringE(1)
+ subring_gen_incl[OF carrier_is_subring assms(1) _ assms(2)]
+ by force
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Algebra/Ideal_Product.thy Mon Jul 02 15:43:22 2018 +0100
@@ -0,0 +1,590 @@
+(* ************************************************************************** *)
+(* Title: Ideal_Product.thy *)
+(* Author: Paulo EmÃlio de Vilhena *)
+(* ************************************************************************** *)
+
+theory Ideal_Product
+ imports Ideal
+
+begin
+
+section \<open>Product of Ideals\<close>
+
+text \<open>In this section, we study the structure of the set of ideals of a given ring.\<close>
+
+inductive_set
+ ideal_prod :: "[ ('a, 'b) ring_scheme, 'a set, 'a set ] \<Rightarrow> 'a set" (infixl "\<cdot>\<index>" 80)
+ for R and I and J (* both I and J are supposed ideals *) where
+ prod: "\<lbrakk> i \<in> I; j \<in> J \<rbrakk> \<Longrightarrow> i \<otimes>\<^bsub>R\<^esub> j \<in> ideal_prod R I J"
+ | sum: "\<lbrakk> s1 \<in> ideal_prod R I J; s2 \<in> ideal_prod R I J \<rbrakk> \<Longrightarrow> s1 \<oplus>\<^bsub>R\<^esub> s2 \<in> ideal_prod R I J"
+
+definition ideals_set :: "('a, 'b) ring_scheme \<Rightarrow> ('a set) ring"
+ where "ideals_set R = \<lparr> carrier = { I. ideal I R },
+ mult = ideal_prod R,
+ one = carrier R,
+ zero = { \<zero>\<^bsub>R\<^esub> },
+ add = set_add R \<rparr>"
+
+
+subsection \<open>Basic Properties\<close>
+
+lemma (in ring) ideal_prod_in_carrier:
+ assumes "ideal I R" "ideal J R"
+ shows "I \<cdot> J \<subseteq> carrier R"
+proof
+ fix s assume "s \<in> I \<cdot> J" thus "s \<in> carrier R"
+ by (induct s rule: ideal_prod.induct) (auto, meson assms ideal.I_l_closed ideal.Icarr)
+qed
+
+lemma (in ring) ideal_prod_inter:
+ assumes "ideal I R" "ideal J R"
+ shows "I \<cdot> J \<subseteq> I \<inter> J"
+proof
+ fix s assume "s \<in> I \<cdot> J" thus "s \<in> I \<inter> J"
+ apply (induct s rule: ideal_prod.induct)
+ apply (auto, (meson assms ideal.I_r_closed ideal.I_l_closed ideal.Icarr)+)
+ apply (simp_all add: additive_subgroup.a_closed assms ideal.axioms(1))
+ done
+qed
+
+lemma (in ring) ideal_prod_is_ideal:
+ assumes "ideal I R" "ideal J R"
+ shows "ideal (I \<cdot> J) R"
+proof (rule idealI)
+ show "ring R" using is_ring .
+next
+ show "subgroup (I \<cdot> J) (add_monoid R)"
+ unfolding subgroup_def
+ proof (auto)
+ show "\<zero> \<in> I \<cdot> J" using ideal_prod.prod[of \<zero> I \<zero> J R]
+ by (simp add: additive_subgroup.zero_closed assms ideal.axioms(1))
+ next
+ fix s1 s2 assume s1: "s1 \<in> I \<cdot> J" and s2: "s2 \<in> I \<cdot> J"
+ show "s1 \<in> carrier R" using ideal_prod_in_carrier[OF assms] s1 by blast
+ show "s1 \<oplus> s2 \<in> I \<cdot> J" by (simp add: ideal_prod.sum[OF s1 s2])
+ show "inv\<^bsub>add_monoid R\<^esub> s1 \<in> I \<cdot> J" using s1
+ proof (induct s1 rule: ideal_prod.induct)
+ case (prod i j)
+ hence "inv\<^bsub>add_monoid R\<^esub> (i \<otimes> j) = (inv\<^bsub>add_monoid R\<^esub> i) \<otimes> j"
+ by (metis a_inv_def assms(1) assms(2) ideal.Icarr l_minus)
+ thus ?case using ideal_prod.prod[of "inv\<^bsub>add_monoid R\<^esub> i" I j J R] assms
+ by (simp add: additive_subgroup.a_subgroup ideal.axioms(1) prod.hyps subgroup.m_inv_closed)
+ next
+ case (sum s1 s2) thus ?case
+ by (smt a_inv_def add.inv_mult_group contra_subsetD
+ assms ideal_prod.sum ideal_prod_in_carrier)
+ qed
+ qed
+next
+ fix s x assume s: "s \<in> I \<cdot> J" and x: "x \<in> carrier R"
+ show "x \<otimes> s \<in> I \<cdot> J" using s
+ proof (induct s rule: ideal_prod.induct)
+ case (prod i j) thus ?case using ideal_prod.prod[of "x \<otimes> i" I j J R] assms
+ by (simp add: x ideal.I_l_closed ideal.Icarr m_assoc)
+ next
+ case (sum s1 s2) thus ?case
+ by (smt assms contra_subsetD ideal_prod.sum ideal_prod_in_carrier r_distr x)
+ qed
+
+ show "s \<otimes> x \<in> I \<cdot> J" using s
+ proof (induct s rule: ideal_prod.induct)
+ case (prod i j) thus ?case using ideal_prod.prod[of i I "j \<otimes> x" J R] assms x
+ by (simp add: x ideal.I_r_closed ideal.Icarr m_assoc)
+ next
+ case (sum s1 s2) thus ?case
+ by (smt assms contra_subsetD ideal_prod.sum ideal_prod_in_carrier l_distr x)
+ qed
+qed
+
+lemma (in ring) ideal_prod_eq_genideal:
+ assumes "ideal I R" "ideal J R"
+ shows "I \<cdot> J = Idl (I <#> J)"
+proof
+ have "I <#> J \<subseteq> I \<cdot> J"
+ proof
+ fix s assume "s \<in> I <#> J"
+ then obtain i j where "i \<in> I" "j \<in> J" "s = i \<otimes> j"
+ unfolding set_mult_def by blast
+ thus "s \<in> I \<cdot> J" using ideal_prod.prod by simp
+ qed
+ thus "Idl (I <#> J) \<subseteq> I \<cdot> J"
+ unfolding genideal_def using ideal_prod_is_ideal[OF assms] by blast
+next
+ show "I \<cdot> J \<subseteq> Idl (I <#> J)"
+ proof
+ fix s assume "s \<in> I \<cdot> J" thus "s \<in> Idl (I <#> J)"
+ proof (induct s rule: ideal_prod.induct)
+ case (prod i j) hence "i \<otimes> j \<in> I <#> J" unfolding set_mult_def by blast
+ thus ?case unfolding genideal_def by blast
+ next
+ case (sum s1 s2) thus ?case
+ by (simp add: additive_subgroup.a_closed additive_subgroup.a_subset
+ assms genideal_ideal ideal.axioms(1) set_mult_closed)
+ qed
+ qed
+qed
+
+
+lemma (in ring) ideal_prod_simp:
+ assumes "ideal I R" "ideal J R" (* the second assumption could be suppressed *)
+ shows "I = I <+> (I \<cdot> J)"
+proof
+ show "I \<subseteq> I <+> I \<cdot> J"
+ proof
+ fix i assume "i \<in> I" hence "i \<oplus> \<zero> \<in> I <+> I \<cdot> J"
+ using set_add_def'[of R I "I \<cdot> J"] ideal_prod_is_ideal[OF assms]
+ additive_subgroup.zero_closed[OF ideal.axioms(1), of "I \<cdot> J" R] by auto
+ thus "i \<in> I <+> I \<cdot> J"
+ using \<open>i \<in> I\<close> assms(1) ideal.Icarr by fastforce
+ qed
+next
+ show "I <+> I \<cdot> J \<subseteq> I"
+ proof
+ fix s assume "s \<in> I <+> I \<cdot> J"
+ then obtain i ij where "i \<in> I" "ij \<in> I \<cdot> J" "s = i \<oplus> ij"
+ using set_add_def'[of R I "I \<cdot> J"] by auto
+ thus "s \<in> I"
+ using ideal_prod_inter[OF assms]
+ by (meson additive_subgroup.a_closed assms(1) ideal.axioms(1) inf_sup_ord(1) subsetCE)
+ qed
+qed
+
+lemma (in ring) ideal_prod_one:
+ assumes "ideal I R"
+ shows "I \<cdot> (carrier R) = I"
+proof
+ show "I \<cdot> (carrier R) \<subseteq> I"
+ proof
+ fix s assume "s \<in> I \<cdot> (carrier R)" thus "s \<in> I"
+ by (induct s rule: ideal_prod.induct)
+ (simp_all add: assms ideal.I_r_closed additive_subgroup.a_closed ideal.axioms(1))
+ qed
+next
+ show "I \<subseteq> I \<cdot> (carrier R)"
+ proof
+ fix i assume "i \<in> I" thus "i \<in> I \<cdot> (carrier R)"
+ by (metis assms ideal.Icarr ideal_prod.simps one_closed r_one)
+ qed
+qed
+
+lemma (in ring) ideal_prod_zero:
+ assumes "ideal I R"
+ shows "I \<cdot> { \<zero> } = { \<zero> }"
+proof
+ show "I \<cdot> { \<zero> } \<subseteq> { \<zero> }"
+ proof
+ fix s assume "s \<in> I \<cdot> {\<zero>}" thus "s \<in> { \<zero> }"
+ using assms ideal.Icarr by (induct s rule: ideal_prod.induct) (fastforce, simp)
+ qed
+next
+ show "{ \<zero> } \<subseteq> I \<cdot> { \<zero> }"
+ by (simp add: additive_subgroup.zero_closed assms
+ ideal.axioms(1) ideal_prod_is_ideal zeroideal)
+qed
+
+lemma (in ring) ideal_prod_assoc:
+ assumes "ideal I R" "ideal J R" "ideal K R"
+ shows "(I \<cdot> J) \<cdot> K = I \<cdot> (J \<cdot> K)"
+proof
+ show "(I \<cdot> J) \<cdot> K \<subseteq> I \<cdot> (J \<cdot> K)"
+ proof
+ fix s assume "s \<in> (I \<cdot> J) \<cdot> K" thus "s \<in> I \<cdot> (J \<cdot> K)"
+ proof (induct s rule: ideal_prod.induct)
+ case (sum s1 s2) thus ?case
+ by (simp add: ideal_prod.sum)
+ next
+ case (prod i k) thus ?case
+ proof (induct i rule: ideal_prod.induct)
+ case (prod i j) thus ?case
+ using ideal_prod.prod[OF prod(1) ideal_prod.prod[OF prod(2-3),of R], of R]
+ by (metis assms ideal.Icarr m_assoc)
+ next
+ case (sum s1 s2) thus ?case
+ by (smt additive_subgroup.a_Hcarr contra_subsetD ideal.axioms(1)
+ assms ideal_prod.sum ideal_prod_in_carrier l_distr)
+ qed
+ qed
+ qed
+next
+ show "I \<cdot> (J \<cdot> K) \<subseteq> (I \<cdot> J) \<cdot> K"
+ proof
+ fix s assume "s \<in> I \<cdot> (J \<cdot> K)" thus "s \<in> (I \<cdot> J) \<cdot> K"
+ proof (induct s rule: ideal_prod.induct)
+ case (sum s1 s2) thus ?case by (simp add: ideal_prod.sum)
+ next
+ case (prod i j) show ?case using prod(2) prod(1)
+ proof (induct j rule: ideal_prod.induct)
+ case (prod j k) thus ?case
+ using ideal_prod.prod[OF ideal_prod.prod[OF prod(3) prod(1), of R] prod (2), of R]
+ by (metis assms ideal.Icarr m_assoc)
+ next
+ case (sum s1 s2) thus ?case
+ by (smt additive_subgroup.a_Hcarr contra_subsetD ideal.axioms(1)
+ assms ideal_prod.sum ideal_prod_in_carrier r_distr)
+ qed
+ qed
+ qed
+qed
+
+lemma (in ring) ideal_prod_r_distr:
+ assumes "ideal I R" "ideal J R" "ideal K R"
+ shows "I \<cdot> (J <+> K) = (I \<cdot> J) <+> (I \<cdot> K)"
+proof
+ show "I \<cdot> (J <+> K) \<subseteq> I \<cdot> J <+> I \<cdot> K"
+ proof
+ fix s assume "s \<in> I \<cdot> (J <+> K)" thus "s \<in> I \<cdot> J <+> I \<cdot> K"
+ proof(induct s rule: ideal_prod.induct)
+ case (prod i jk)
+ then obtain j k where j: "j \<in> J" and k: "k \<in> K" and jk: "jk = j \<oplus> k"
+ using set_add_def'[of R J K] by auto
+ hence "i \<otimes> j \<oplus> i \<otimes> k \<in> I \<cdot> J <+> I \<cdot> K"
+ using ideal_prod.prod[OF prod(1) j,of R]
+ ideal_prod.prod[OF prod(1) k,of R]
+ set_add_def'[of R "I \<cdot> J" "I \<cdot> K"] by auto
+ thus ?case
+ using assms ideal.Icarr r_distr jk j k prod(1) by metis
+ next
+ case (sum s1 s2) thus ?case
+ by (simp add: add_ideals additive_subgroup.a_closed assms ideal.axioms(1)
+ local.ring_axioms ring.ideal_prod_is_ideal)
+ qed
+ qed
+next
+ { fix s J K assume A: "ideal J R" "ideal K R" "s \<in> I \<cdot> J"
+ have "s \<in> I \<cdot> (J <+> K) \<and> s \<in> I \<cdot> (K <+> J)"
+ proof -
+ from \<open>s \<in> I \<cdot> J\<close> have "s \<in> I \<cdot> (J <+> K)"
+ proof (induct s rule: ideal_prod.induct)
+ case (prod i j)
+ hence "(j \<oplus> \<zero>) \<in> J <+> K"
+ using set_add_def'[of R J K]
+ additive_subgroup.zero_closed[OF ideal.axioms(1), of K R] A(2) by auto
+ thus ?case
+ by (metis A(1) additive_subgroup.a_Hcarr ideal.axioms(1) ideal_prod.prod prod r_zero)
+ next
+ case (sum s1 s2) thus ?case
+ by (simp add: ideal_prod.sum)
+ qed
+ thus ?thesis
+ by (metis A(1) A(2) ideal_def ring.union_genideal sup_commute)
+ qed } note aux_lemma = this
+
+ show "I \<cdot> J <+> I \<cdot> K \<subseteq> I \<cdot> (J <+> K)"
+ proof
+ fix s assume "s \<in> I \<cdot> J <+> I \<cdot> K"
+ then obtain s1 s2 where s1: "s1 \<in> I \<cdot> J" and s2: "s2 \<in> I \<cdot> K" and s: "s = s1 \<oplus> s2"
+ using set_add_def'[of R "I \<cdot> J" "I \<cdot> K"] by auto
+ thus "s \<in> I \<cdot> (J <+> K)"
+ using aux_lemma[OF assms(2) assms(3) s1]
+ aux_lemma[OF assms(3) assms(2) s2] by (simp add: ideal_prod.sum)
+ qed
+qed
+
+lemma (in cring) ideal_prod_commute:
+ assumes "ideal I R" "ideal J R"
+ shows "I \<cdot> J = J \<cdot> I"
+proof -
+ { fix I J assume A: "ideal I R" "ideal J R"
+ have "I \<cdot> J \<subseteq> J \<cdot> I"
+ proof
+ fix s assume "s \<in> I \<cdot> J" thus "s \<in> J \<cdot> I"
+ proof (induct s rule: ideal_prod.induct)
+ case (prod i j) thus ?case
+ using m_comm[OF ideal.Icarr[OF A(1) prod(1)] ideal.Icarr[OF A(2) prod(2)]]
+ by (simp add: ideal_prod.prod)
+ next
+ case (sum s1 s2) thus ?case by (simp add: ideal_prod.sum)
+ qed
+ qed }
+ thus ?thesis using assms by blast
+qed
+
+text \<open>The following result would also be true for locale ring\<close>
+lemma (in cring) ideal_prod_distr:
+ assumes "ideal I R" "ideal J R" "ideal K R"
+ shows "I \<cdot> (J <+> K) = (I \<cdot> J) <+> (I \<cdot> K)"
+ and "(J <+> K) \<cdot> I = (J \<cdot> I) <+> (K \<cdot> I)"
+ by (simp_all add: assms ideal_prod_commute local.ring_axioms
+ ring.add_ideals ring.ideal_prod_r_distr)
+
+lemma (in cring) ideal_prod_eq_inter:
+ assumes "ideal I R" "ideal J R"
+ and "I <+> J = carrier R"
+ shows "I \<cdot> J = I \<inter> J"
+proof
+ show "I \<cdot> J \<subseteq> I \<inter> J"
+ using assms ideal_prod_inter by auto
+next
+ show "I \<inter> J \<subseteq> I \<cdot> J"
+ proof
+ have "\<one> \<in> I <+> J" using assms(3) one_closed by simp
+ then obtain i j where ij: "i \<in> I" "j \<in> J" "\<one> = i \<oplus> j"
+ using set_add_def'[of R I J] by auto
+
+ fix s assume s: "s \<in> I \<inter> J"
+ hence "(i \<otimes> s \<in> I \<cdot> J) \<and> (s \<otimes> j \<in> I \<cdot> J)"
+ using ij(1-2) by (simp add: ideal_prod.prod)
+ moreover have "s = (i \<otimes> s) \<oplus> (s \<otimes> j)"
+ using ideal.Icarr[OF assms(1) ij(1)]
+ ideal.Icarr[OF assms(2) ij(2)]
+ ideal.Icarr[OF assms(1), of s]
+ by (metis ij(3) s m_comm[of s i] Int_iff r_distr r_one)
+ ultimately show "s \<in> I \<cdot> J"
+ using ideal_prod.sum by fastforce
+ qed
+qed
+
+
+subsection \<open>Structure of the Set of Ideals\<close>
+
+text \<open>We focus on commutative rings for convenience.\<close>
+
+lemma (in cring) ideals_set_is_semiring: "semiring (ideals_set R)"
+proof -
+ have "abelian_monoid (ideals_set R)"
+ apply (rule abelian_monoidI) unfolding ideals_set_def
+ apply (simp_all add: add_ideals zeroideal)
+ apply (simp add: add.set_mult_assoc additive_subgroup.a_subset ideal.axioms(1) set_add_defs(1))
+ apply (metis Un_absorb1 additive_subgroup.a_subset additive_subgroup.zero_closed
+ cgenideal_minimal cgenideal_self empty_iff genideal_minimal ideal.axioms(1)
+ local.ring_axioms order_refl ring.genideal_self subset_antisym subset_singletonD
+ union_genideal zero_closed zeroideal)
+ by (metis sup_commute union_genideal)
+
+ moreover have "monoid (ideals_set R)"
+ apply (rule monoidI) unfolding ideals_set_def
+ apply (simp_all add: ideal_prod_is_ideal oneideal
+ ideal_prod_commute ideal_prod_one)
+ by (metis ideal_prod_assoc ideal_prod_commute)
+
+ ultimately show ?thesis
+ unfolding semiring_def semiring_axioms_def ideals_set_def
+ by (simp_all add: ideal_prod_distr ideal_prod_commute ideal_prod_zero zeroideal)
+qed
+
+lemma (in cring) ideals_set_is_comm_monoid: "comm_monoid (ideals_set R)"
+proof -
+ have "monoid (ideals_set R)"
+ apply (rule monoidI) unfolding ideals_set_def
+ apply (simp_all add: ideal_prod_is_ideal oneideal
+ ideal_prod_commute ideal_prod_one)
+ by (metis ideal_prod_assoc ideal_prod_commute)
+ thus ?thesis
+ unfolding comm_monoid_def comm_monoid_axioms_def
+ by (simp add: ideal_prod_commute ideals_set_def)
+qed
+
+lemma (in cring) ideal_prod_eq_Inter_aux:
+ assumes "I: {..(Suc n)} \<rightarrow> { J. ideal J R }"
+ and "\<And>i j. \<lbrakk> i \<in> {..(Suc n)}; j \<in> {..(Suc n)} \<rbrakk> \<Longrightarrow>
+ i \<noteq> j \<Longrightarrow> (I i) <+> (I j) = carrier R"
+ shows "(\<Otimes>\<^bsub>(ideals_set R)\<^esub> k \<in> {..n}. I k) <+> (I (Suc n)) = carrier R" using assms
+proof (induct n arbitrary: I)
+ case 0
+ hence "(\<Otimes>\<^bsub>(ideals_set R)\<^esub> k \<in> {..0}. I k) <+> I (Suc 0) = (I 0) <+> (I (Suc 0))"
+ using comm_monoid.finprod_0[OF ideals_set_is_comm_monoid, of I]
+ by (simp add: atMost_Suc ideals_set_def)
+ also have " ... = carrier R"
+ using 0(2)[of 0 "Suc 0"] by simp
+ finally show ?case .
+next
+ interpret ISet: comm_monoid "ideals_set R"
+ by (simp add: ideals_set_is_comm_monoid)
+
+ case (Suc n)
+ let ?I' = "\<lambda>i. I (Suc i)"
+ have "?I': {..(Suc n)} \<rightarrow> { J. ideal J R }"
+ using Suc.prems(1) by auto
+ moreover have "\<And>i j. \<lbrakk> i \<in> {..(Suc n)}; j \<in> {..(Suc n)} \<rbrakk> \<Longrightarrow>
+ i \<noteq> j \<Longrightarrow> (?I' i) <+> (?I' j) = carrier R"
+ by (simp add: Suc.prems(2))
+ ultimately have "(\<Otimes>\<^bsub>(ideals_set R)\<^esub> k \<in> {..n}. ?I' k) <+> (?I' (Suc n)) = carrier R"
+ using Suc.hyps by metis
+
+ moreover have I_carr: "I: {..Suc (Suc n)} \<rightarrow> carrier (ideals_set R)"
+ unfolding ideals_set_def using Suc by simp
+ hence I'_carr: "I \<in> Suc ` {..n} \<rightarrow> carrier (ideals_set R)" by auto
+ ultimately have "(\<Otimes>\<^bsub>(ideals_set R)\<^esub> k \<in> {(Suc 0)..Suc n}. I k) <+> (I (Suc (Suc n))) = carrier R"
+ using ISet.finprod_reindex[of I "\<lambda>i. Suc i" "{..n}"] by (simp add: atMost_atLeast0)
+
+ hence "(carrier R) \<cdot> (I 0) = ((\<Otimes>\<^bsub>(ideals_set R)\<^esub> k \<in> {Suc 0..Suc n}. I k) <+> I (Suc (Suc n))) \<cdot> (I 0)"
+ by auto
+ moreover have fprod_cl1: "ideal (\<Otimes>\<^bsub>(ideals_set R)\<^esub> k \<in> {Suc 0..Suc n}. I k) R"
+ by (metis I'_carr ISet.finprod_closed One_nat_def ideals_set_def image_Suc_atMost
+ mem_Collect_eq partial_object.select_convs(1))
+ ultimately
+ have "I 0 = (\<Otimes>\<^bsub>(ideals_set R)\<^esub> k \<in> {Suc 0..Suc n}. I k) \<cdot> (I 0) <+> I (Suc (Suc n)) \<cdot> (I 0)"
+ by (metis PiE Suc.prems(1) atLeast0_atMost_Suc atLeast0_atMost_Suc_eq_insert_0
+ atMost_atLeast0 ideal_prod_commute ideal_prod_distr(2) ideal_prod_one insertI1
+ mem_Collect_eq oneideal)
+ also have " ... = (I 0) \<cdot> (\<Otimes>\<^bsub>(ideals_set R)\<^esub> k \<in> {Suc 0..Suc n}. I k) <+> I (Suc (Suc n)) \<cdot> (I 0)"
+ using fprod_cl1 ideal_prod_commute Suc.prems(1)
+ by (simp add: atLeast0_atMost_Suc_eq_insert_0 atMost_atLeast0)
+ also have " ... = (I 0) \<otimes>\<^bsub>(ideals_set R)\<^esub> (\<Otimes>\<^bsub>(ideals_set R)\<^esub> k \<in> {Suc 0..Suc n}. I k) <+>
+ I (Suc (Suc n)) \<cdot> (I 0)"
+ by (simp add: ideals_set_def)
+ finally have I0: "I 0 = (\<Otimes>\<^bsub>(ideals_set R)\<^esub> k \<in> {..Suc n}. I k) <+> I (Suc (Suc n)) \<cdot> (I 0)"
+ using ISet.finprod_insert[of "{Suc 0..Suc n}" 0 I]
+ I_carr I'_carr atMost_atLeast0 ISet.finprod_0' atMost_Suc by auto
+
+ have I_SucSuc_I0: "ideal (I (Suc (Suc n))) R \<and> ideal (I 0) R"
+ using Suc.prems(1) by auto
+ have fprod_cl2: "ideal (\<Otimes>\<^bsub>(ideals_set R)\<^esub> k \<in> {..Suc n}. I k) R"
+ by (smt ISet.finprod_closed I_carr Pi_split_insert_domain atMost_Suc ideals_set_def mem_Collect_eq partial_object.select_convs(1))
+
+ have "carrier R = I (Suc (Suc n)) <+> I 0"
+ by (simp add: Suc.prems(2))
+ also have " ... = I (Suc (Suc n)) <+>
+ ((\<Otimes>\<^bsub>(ideals_set R)\<^esub> k \<in> {..Suc n}. I k) <+> I (Suc (Suc n)) \<cdot> (I 0))"
+ using I0 by auto
+ also have " ... = I (Suc (Suc n)) <+>
+ (I (Suc (Suc n)) \<cdot> (I 0) <+> (\<Otimes>\<^bsub>(ideals_set R)\<^esub> k \<in> {..Suc n}. I k))"
+ using fprod_cl2 I_SucSuc_I0 by (metis Un_commute ideal_prod_is_ideal union_genideal)
+ also have " ... = (I (Suc (Suc n)) <+> I (Suc (Suc n)) \<cdot> (I 0)) <+>
+ (\<Otimes>\<^bsub>(ideals_set R)\<^esub> k \<in> {..Suc n}. I k)"
+ using fprod_cl2 I_SucSuc_I0 by (metis add.set_mult_assoc ideal_def ideal_prod_in_carrier
+ oneideal ring.ideal_prod_one set_add_defs(1))
+ also have " ... = I (Suc (Suc n)) <+> (\<Otimes>\<^bsub>(ideals_set R)\<^esub> k \<in> {..Suc n}. I k)"
+ using ideal_prod_simp[of "I (Suc (Suc n))" "I 0"] I_SucSuc_I0 by simp
+ also have " ... = (\<Otimes>\<^bsub>(ideals_set R)\<^esub> k \<in> {..Suc n}. I k) <+> I (Suc (Suc n))"
+ using fprod_cl2 I_SucSuc_I0 by (metis Un_commute union_genideal)
+ finally show ?case by simp
+qed
+
+theorem (in cring) ideal_prod_eq_Inter:
+ assumes "I: {..n :: nat} \<rightarrow> { J. ideal J R }"
+ and "\<And>i j. \<lbrakk> i \<in> {..n}; j \<in> {..n} \<rbrakk> \<Longrightarrow> i \<noteq> j \<Longrightarrow> (I i) <+> (I j) = carrier R"
+ shows "(\<Otimes>\<^bsub>(ideals_set R)\<^esub> k \<in> {..n}. I k) = (\<Inter> k \<in> {..n}. I k)" using assms
+proof (induct n)
+ case 0 thus ?case
+ using comm_monoid.finprod_0[OF ideals_set_is_comm_monoid] by (simp add: ideals_set_def)
+next
+ interpret ISet: comm_monoid "ideals_set R"
+ by (simp add: ideals_set_is_comm_monoid)
+
+ case (Suc n)
+ hence IH: "(\<Otimes>\<^bsub>(ideals_set R)\<^esub> k \<in> {..n}. I k) = (\<Inter> k \<in> {..n}. I k)"
+ by (simp add: atMost_Suc)
+ hence "(\<Otimes>\<^bsub>(ideals_set R)\<^esub> k \<in> {..Suc n}. I k) = I (Suc n) \<otimes>\<^bsub>(ideals_set R)\<^esub> (\<Inter> k \<in> {..n}. I k)"
+ using ISet.finprod_insert[of "{Suc 0..Suc n}" 0 I] atMost_Suc_eq_insert_0[of n]
+ by (metis ISet.finprod_Suc Suc.prems(1) ideals_set_def partial_object.select_convs(1))
+ hence "(\<Otimes>\<^bsub>(ideals_set R)\<^esub> k \<in> {..Suc n}. I k) = I (Suc n) \<cdot> (\<Inter> k \<in> {..n}. I k)"
+ by (simp add: ideals_set_def)
+ moreover have "(\<Inter> k \<in> {..n}. I k) <+> I (Suc n) = carrier R"
+ using ideal_prod_eq_Inter_aux[of I n] by (simp add: Suc.prems IH)
+ moreover have "ideal (\<Inter> k \<in> {..n}. I k) R"
+ using ring.i_Intersect[of R "I ` {..n}"]
+ by (metis IH ISet.finprod_closed Pi_split_insert_domain Suc.prems(1) atMost_Suc
+ ideals_set_def mem_Collect_eq partial_object.select_convs(1))
+ ultimately
+ have "(\<Otimes>\<^bsub>(ideals_set R)\<^esub> k \<in> {..Suc n}. I k) = (\<Inter> k \<in> {..n}. I k) \<inter> I (Suc n)"
+ using ideal_prod_eq_inter[of "\<Inter> k \<in> {..n}. I k" "I (Suc n)"]
+ ideal_prod_commute[of "\<Inter> k \<in> {..n}. I k" "I (Suc n)"]
+ by (metis PiE Suc.prems(1) atMost_iff mem_Collect_eq order_refl)
+ thus ?case by (simp add: Int_commute atMost_Suc)
+qed
+
+corollary (in cring) inter_plus_ideal_eq_carrier:
+ assumes "\<And>i. i \<in> {..(Suc n)} \<Longrightarrow> ideal (I i) R"
+ and "\<And>i j. \<lbrakk> i \<in> {..(Suc n)}; j \<in> {..(Suc n)}; i \<noteq> j \<rbrakk> \<Longrightarrow> I i <+> I j = carrier R"
+ shows "(\<Inter> i \<le> n. I i) <+> (I (Suc n)) = carrier R"
+ using ideal_prod_eq_Inter[of I n] ideal_prod_eq_Inter_aux[of I n] by (auto simp add: assms)
+
+corollary (in cring) inter_plus_ideal_eq_carrier_arbitrary:
+ assumes "\<And>i. i \<in> {..(Suc n)} \<Longrightarrow> ideal (I i) R"
+ and "\<And>i j. \<lbrakk> i \<in> {..(Suc n)}; j \<in> {..(Suc n)}; i \<noteq> j \<rbrakk> \<Longrightarrow> I i <+> I j = carrier R"
+ and "j \<in> {..(Suc n)}"
+ shows "(\<Inter> i \<in> ({..(Suc n)} - { j }). I i) <+> (I j) = carrier R"
+proof -
+ define I' where "I' = (\<lambda>i. if i = Suc n then (I j) else
+ if i = j then (I (Suc n))
+ else (I i))"
+ have "\<And>i. i \<in> {..(Suc n)} \<Longrightarrow> ideal (I' i) R"
+ using I'_def assms(1) assms(3) by auto
+ moreover have "\<And>i j. \<lbrakk> i \<in> {..(Suc n)}; j \<in> {..(Suc n)}; i \<noteq> j \<rbrakk> \<Longrightarrow> I' i <+> I' j = carrier R"
+ using I'_def assms(2-3) by force
+ ultimately have "(\<Inter> i \<le> n. I' i) <+> (I' (Suc n)) = carrier R"
+ using inter_plus_ideal_eq_carrier by simp
+
+ moreover have "I' ` {..n} = I ` ({..(Suc n)} - { j })"
+ proof
+ show "I' ` {..n} \<subseteq> I ` ({..Suc n} - {j})"
+ proof
+ fix x assume "x \<in> I' ` {..n}"
+ then obtain i where i: "i \<in> {..n}" "I' i = x" by blast
+ thus "x \<in> I ` ({..Suc n} - {j})"
+ proof (cases)
+ assume "i = j" thus ?thesis using i I'_def by auto
+ next
+ assume "i \<noteq> j" thus ?thesis using I'_def i insert_iff by auto
+ qed
+ qed
+ next
+ show "I ` ({..Suc n} - {j}) \<subseteq> I' ` {..n}"
+ proof
+ fix x assume "x \<in> I ` ({..Suc n} - {j})"
+ then obtain i where i: "i \<in> {..Suc n}" "i \<noteq> j" "I i = x" by blast
+ thus "x \<in> I' ` {..n}"
+ proof (cases)
+ assume "i = Suc n" thus ?thesis using I'_def assms(3) i(2-3) by auto
+ next
+ assume "i \<noteq> Suc n" thus ?thesis using I'_def i by auto
+ qed
+ qed
+ qed
+ ultimately show ?thesis using I'_def by metis
+qed
+
+
+subsection \<open>Another Characterization of Prime Ideals\<close>
+
+text \<open>With product of ideals being defined, we can give another definition of a prime ideal\<close>
+
+lemma (in ring) primeideal_divides_ideal_prod:
+ assumes "primeideal P R" "ideal I R" "ideal J R"
+ and "I \<cdot> J \<subseteq> P"
+ shows "I \<subseteq> P \<or> J \<subseteq> P"
+proof (cases)
+ assume "\<exists> i \<in> I. i \<notin> P"
+ then obtain i where i: "i \<in> I" "i \<notin> P" by blast
+ have "J \<subseteq> P"
+ proof
+ fix j assume j: "j \<in> J"
+ hence "i \<otimes> j \<in> P"
+ using ideal_prod.prod[OF i(1) j, of R] assms(4) by auto
+ thus "j \<in> P"
+ using primeideal.I_prime[OF assms(1), of i j] i j
+ by (meson assms(2-3) ideal.Icarr)
+ qed
+ thus ?thesis by blast
+next
+ assume "\<not> (\<exists> i \<in> I. i \<notin> P)" thus ?thesis by blast
+qed
+
+lemma (in cring) divides_ideal_prod_imp_primeideal:
+ assumes "ideal P R"
+ and "P \<noteq> carrier R"
+ and "\<And>I J. \<lbrakk> ideal I R; ideal J R; I \<cdot> J \<subseteq> P \<rbrakk> \<Longrightarrow> I \<subseteq> P \<or> J \<subseteq> P"
+ shows "primeideal P R"
+proof -
+ have "\<And>a b. \<lbrakk> a \<in> carrier R; b \<in> carrier R; a \<otimes> b \<in> P \<rbrakk> \<Longrightarrow> a \<in> P \<or> b \<in> P"
+ proof -
+ fix a b assume A: "a \<in> carrier R" "b \<in> carrier R" "a \<otimes> b \<in> P"
+ have "(PIdl a) \<cdot> (PIdl b) = Idl (PIdl (a \<otimes> b))"
+ using ideal_prod_eq_genideal[of "Idl { a }" "Idl { b }"]
+ A(1-2) cgenideal_eq_genideal cgenideal_ideal cgenideal_prod by auto
+ hence "(PIdl a) \<cdot> (PIdl b) = PIdl (a \<otimes> b)"
+ by (simp add: A Idl_subset_ideal cgenideal_ideal cgenideal_minimal
+ genideal_self oneideal subset_antisym)
+ hence "(PIdl a) \<cdot> (PIdl b) \<subseteq> P"
+ by (simp add: A(3) assms(1) cgenideal_minimal)
+ hence "(PIdl a) \<subseteq> P \<or> (PIdl b) \<subseteq> P"
+ by (simp add: A assms(3) cgenideal_ideal)
+ thus "a \<in> P \<or> b \<in> P"
+ using A cgenideal_self by blast
+ qed
+ thus ?thesis
+ using assms is_cring by (simp add: primeidealI)
+qed
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Algebra/Solvable_Groups.thy Mon Jul 02 15:43:22 2018 +0100
@@ -0,0 +1,431 @@
+(* ************************************************************************** *)
+(* Title: Solvable_Groups.thy *)
+(* Author: Paulo EmÃlio de Vilhena *)
+(* ************************************************************************** *)
+
+theory Solvable_Groups
+ imports Group Coset Generated_Groups
+
+begin
+
+inductive solvable_seq :: "('a, 'b) monoid_scheme \<Rightarrow> 'a set \<Rightarrow> bool" for G where
+unity: "solvable_seq G { \<one>\<^bsub>G\<^esub> }" |
+extension: "\<lbrakk> solvable_seq G K; K \<subset> H; subgroup H G; K \<lhd> (G \<lparr> carrier := H \<rparr>);
+ comm_group ((G \<lparr> carrier := H \<rparr>) Mod K) \<rbrakk> \<Longrightarrow> solvable_seq G H"
+
+definition
+ solvable :: "('a, 'b) monoid_scheme \<Rightarrow> bool"
+ where "solvable G \<longleftrightarrow> solvable_seq G (carrier G)"
+
+
+subsection \<open>Solvable Groups and Derived Subgroups\<close>
+
+text \<open>We show that a group G is solvable iff the subgroup (derived G ^^ n) (carrier G)
+ is trivial for a sufficiently large n\<close>
+
+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
+
+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 \<noteq> 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
+
+theorem (in group) trivial_derived_seq_imp_solvable:
+ assumes "subgroup H G" and "((derived G) ^^ n) H = { \<one> }"
+ shows "solvable_seq G H" using assms
+proof (induction n arbitrary: H)
+ case 0 hence "H = { \<one> }" by simp thus ?case by (simp add: unity)
+next
+ case (Suc n)
+ hence "(derived G ^^ n) (derived G H) = { \<one> }"
+ 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)
+qed
+
+theorem (in group) solvable_imp_trivial_derived_seq:
+ assumes "solvable_seq G H" and "subgroup H G"
+ shows "\<exists>n. (derived G ^^ n) H = { \<one> }"
+proof -
+ { fix K H assume A: "K \<subseteq> H" "K \<lhd> (G \<lparr> carrier := H \<rparr>)" "subgroup K G" "subgroup H G"
+ "comm_group ((G \<lparr> carrier := H \<rparr>) Mod K)"
+ have "derived G H \<subseteq> K"
+ proof -
+ have "derived_set G H \<subseteq> K"
+ proof
+ fix h assume "h \<in> derived_set G H"
+ then obtain h1 h2 where h12: "h1 \<in> H" "h2 \<in> H" "h = h1 \<otimes> h2 \<otimes> inv h1 \<otimes> inv h2" by blast
+
+ hence K_h12: "(K #> (h1 \<otimes> h2)) \<in> carrier ((G \<lparr> carrier := H \<rparr>) Mod K)"
+ unfolding FactGroup_def RCOSETS_def r_coset_def apply simp by (metis A(4) subgroup_def)
+ have K_h1: "K #>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> h1 \<in> carrier ((G \<lparr> carrier := H \<rparr>) Mod K)"
+ unfolding FactGroup_def RCOSETS_def r_coset_def apply simp using h12(1) by blast
+ have K_h2: "K #>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> h2 \<in> carrier ((G \<lparr> carrier := H \<rparr>) Mod K)"
+ unfolding FactGroup_def RCOSETS_def r_coset_def apply simp using h12(2) by blast
+
+ hence "K #>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> (h1 \<otimes> h2) =
+ (K #>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> h1) <#>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> (K #>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> h2)"
+ using normal.rcos_sum[OF A(2),of h1 h2] h12(1-2) by simp
+ also have " ... =
+ (K #>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> h2) <#>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> (K #>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> h1)"
+ using comm_groupE(4)[OF A(5) K_h1 K_h2] by simp
+ finally have "K #>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> (h1 \<otimes> h2) = K #>\<^bsub>G\<lparr>carrier := H\<rparr>\<^esub> (h2 \<otimes> h1)"
+ using normal.rcos_sum[OF A(2),of h2 h1] h12(1-2) by simp
+
+ moreover have "h1 \<otimes> h2 \<in> H" and "h2 \<otimes> h1 \<in> H"
+ using h12 subgroupE(4)[OF A(4)] by auto
+ ultimately have "K #> (h1 \<otimes> h2) = K #> (h2 \<otimes> h1)" by auto
+
+ then obtain k where k: "k \<in> K" "\<one> \<otimes> (h1 \<otimes> h2) = k \<otimes> (h2 \<otimes> h1)"
+ using subgroup.one_closed[OF A(3)] unfolding r_coset_def by blast
+ hence "(h1 \<otimes> h2) \<otimes> (inv h1 \<otimes> inv h2) = k"
+ by (smt A(1) A(4) h12(1-2) inv_mult_group inv_solve_right
+ l_one m_closed subgroup.mem_carrier subsetCE)
+ hence "h = k" using h12
+ by (meson A(4) \<open>h1 \<otimes> h2 \<in> H\<close> inv_closed m_assoc subgroup.mem_carrier)
+ thus "h \<in> 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 "\<exists>n. (derived G ^^ n) H = { \<one> }" using assms
+ proof (induct H rule: solvable_seq.induct)
+ case unity have "(derived G ^^ 0) { \<one> } = { \<one> }" by simp thus ?case by blast
+ next
+ case (extension K H)
+ then obtain n where n: "(derived G ^^ n) K = { \<one> }"
+ using solvable_imp_subgroup extension by blast
+ have "derived G H \<subseteq> K" using solvable_imp_subgroup extension aux_lemma by blast
+ hence "(derived G ^^ n) (derived G H) \<subseteq> (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 \<subseteq> { \<one> }"
+ by (metis funpow_simps_right(2) n o_apply)
+ moreover have "\<one> \<in> derived G ((derived G ^^ n) H)"
+ unfolding derived_def using generate.one by auto
+ hence "{ \<one> } \<subseteq> (derived G ^^ (Suc n)) H" by simp
+ ultimately show ?case by blast
+ qed
+qed
+
+theorem (in group) solvable_iff_trivial_derived_seq:
+ "solvable G \<longleftrightarrow> (\<exists>n. (derived G ^^ n) (carrier G) = { \<one> })"
+ unfolding solvable_def
+ using solvable_imp_trivial_derived_seq subgroup_self
+ trivial_derived_seq_imp_solvable by blast
+
+corollary (in group) solvable_subgroup:
+ assumes "subgroup H G"
+ shows "solvable G \<Longrightarrow> solvable_seq G H"
+proof -
+ { fix I J assume A: "subgroup I G" "subgroup J G" "I \<subseteq> J" "solvable_seq G J"
+ have "solvable_seq G I"
+ proof -
+ obtain n where "(derived G ^^ n) J = { \<one> }"
+ using solvable_imp_trivial_derived_seq[OF A(4) A(2)] by auto
+ hence "(derived G ^^ n) I \<subseteq> { \<one> }"
+ using mono_derived[OF subgroup.subset[OF A(2)] A(3)] by auto
+ hence "(derived G ^^ n) I = { \<one> }"
+ 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
+qed
+
+
+subsection \<open>Short Exact Sequences\<close>
+
+text \<open>Even if we don't talk about short exact sequences explicitly, we show that given an
+ injective homomorphism from a group H to a group G, if H isn't solvable the group G
+ isn't neither. \<close>
+
+lemma (in group_hom) generate_of_img:
+ assumes "K \<subseteq> carrier G"
+ shows "generate H (h ` K) = h ` (generate G K)"
+proof
+ have img_in_carrier: "h ` K \<subseteq> carrier H"
+ by (meson assms group_hom.hom_closed group_hom_axioms image_subsetI subsetCE)
+
+ show "generate H (h ` K) \<subseteq> h ` generate G K"
+ proof
+ fix x assume "x \<in> generate H (h ` K)"
+ then obtain r where r: "elts r \<subseteq> (h ` K)" "norm H r = x"
+ using H.generate_repr_iff img_in_carrier by auto
+ from \<open>elts r \<subseteq> (h ` K)\<close> have "norm H r \<in> h ` generate G K"
+ proof (induct r rule: repr.induct)
+ case One
+ have "\<one>\<^bsub>G\<^esub> \<in> generate G K" using generate.one[of G] by simp
+ hence "h \<one>\<^bsub>G\<^esub> \<in> h ` generate G K" by blast
+ thus ?case by simp
+ next
+ case (Inv x) hence "x \<in> h ` K" by simp
+ then obtain k where k: "k \<in> 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 \<in> h ` K" by simp
+ then obtain k where "k \<in> K" "x = h k" by blast
+ thus ?case by (simp add: generate.incl)
+ next
+ case (Mult x1 x2) hence A: "elts x1 \<union> elts x2 \<subseteq> h ` K" by simp
+ have "norm H x1 \<in> h ` (generate G K)" using A Mult by simp
+ moreover have "norm H x2 \<in> h ` (generate G K)" using A Mult by simp
+ ultimately obtain k1 k2 where k1: "k1 \<in> generate G K" "norm H x1 = h k1"
+ and k2: "k2 \<in> generate G K" "norm H x2 = h k2" by blast
+ hence "norm H (Mult x1 x2) = h (k1 \<otimes> k2)"
+ using G.generate_in_carrier assms by auto
+ thus ?case using k1 k2 by (simp add: generate.eng)
+ qed
+ thus "x \<in> h ` generate G K" using r by simp
+ qed
+
+next
+ show "h ` generate G K \<subseteq> generate H (h ` K)"
+ proof
+ fix x assume "x \<in> h ` generate G K"
+ then obtain r where r: "elts r \<subseteq> K" "x = h (norm G r)" using G.generate_repr_iff assms by auto
+ from \<open>elts r \<subseteq> K\<close> have "h (norm G r) \<in> 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 \<in> K" by simp
+ hence "h (norm G (Inv x)) = inv\<^bsub>H\<^esub> (h x)" using assms by auto
+ moreover have "h x \<in> 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 \<union> elts x2 \<subseteq> K" by simp
+ have "norm G x1 \<in> carrier G"
+ by (meson A G.generateE(1) G.generate_repr_iff Un_subset_iff assms subgroup.mem_carrier)
+ moreover have "norm G x2 \<in> carrier G"
+ by (meson A G.generateE(1) G.generate_repr_iff Un_subset_iff assms subgroup.mem_carrier)
+ ultimately have "h (norm G (Mult x1 x2)) = h (norm G x1) \<otimes>\<^bsub>H\<^esub> h (norm G x2)" by simp
+ thus ?case using Mult A by (simp add: generate.eng)
+ qed
+ thus "x \<in> generate H (h ` K)" using r by simp
+ qed
+qed
+
+lemma (in group_hom) derived_of_img:
+ assumes "K \<subseteq> carrier G"
+ shows "(derived H ^^ n) (h ` K) = h ` ((derived G ^^ n) K)"
+proof -
+ { fix K assume A: "K \<subseteq> carrier G"
+ have "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) \<subseteq> h ` derived_set G K"
+ proof
+ fix x assume "x \<in> derived_set H (h ` K)"
+ then obtain k1 k2
+ where k12: "k1 \<in> K" "k2 \<in> K"
+ and "x = (h k1) \<otimes>\<^bsub>H\<^esub> (h k2) \<otimes>\<^bsub>H\<^esub> inv\<^bsub>H\<^esub> (h k1) \<otimes>\<^bsub>H\<^esub> inv\<^bsub>H\<^esub>(h k2)" by blast
+ hence "x = h (k1 \<otimes> k2 \<otimes> inv k1 \<otimes> inv k2)"
+ by (smt G.inv_closed G.m_closed A hom_inv hom_mult subsetCE)
+ thus "x \<in> h ` (derived_set G K)" using k12 by blast
+ qed
+ next
+ show "h ` derived_set G K \<subseteq> derived_set H (h ` K)"
+ proof
+ fix x assume " x \<in> h ` derived_set G K"
+ then obtain k1 k2 where k12: "k1 \<in> K" "k2 \<in> K"
+ and "x = h (k1 \<otimes> k2 \<otimes> inv k1 \<otimes> inv k2)" by blast
+ hence "x = (h k1) \<otimes>\<^bsub>H\<^esub> (h k2) \<otimes>\<^bsub>H\<^esub> inv\<^bsub>H\<^esub> (h k1) \<otimes>\<^bsub>H\<^esub> inv\<^bsub>H\<^esub>(h k2)"
+ by (smt G.inv_closed G.m_closed A hom_inv hom_mult subsetCE)
+ thus "x \<in> 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 \<open>K \<subseteq> carrier G\<close> 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 \<subseteq> I"
+ proof (induction n)
+ case 0 thus ?case by simp
+ next
+ case (Suc n)
+ hence "(derived G) ((derived G ^^ n) I) \<subseteq> (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) = { \<one>\<^bsub>H\<^esub> }"
+ using H.solvable_imp_subgroup H.solvable_imp_trivial_derived_seq assms(3) by blast
+ hence "h ` ((derived G ^^ n) I) = { \<one>\<^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 "\<And>x. \<lbrakk> x \<in> ((derived G ^^ n) I); h x = \<one>\<^bsub>H\<^esub> \<rbrakk> \<Longrightarrow> x = \<one>"
+ by (metis G.exp_of_derived_is_subgroup assms(1) hom_one inj_on_eq_iff subgroup_def)
+ ultimately have "(derived G ^^ n) I = { \<one> }" 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 "\<not> solvable G"
+ shows "\<not> solvable H"
+proof -
+ { fix I J assume A: "subgroup I H" "subgroup J H" "I \<subseteq> J" "solvable_seq H J"
+ have "solvable_seq H I"
+ proof -
+ obtain n where n: "(derived H ^^ n) J = { \<one>\<^bsub>H\<^esub> }"
+ using A(4) H.solvable_imp_subgroup H.solvable_imp_trivial_derived_seq by blast
+ have "(derived H ^^ n) I \<subseteq> (derived H ^^ n) J"
+ using A by (simp add: H.mono_derived subgroupE(1))
+ hence "(derived H ^^ n) I \<subseteq> { \<one>\<^bsub>H\<^esub> }" using n by simp
+ hence "(derived H ^^ n) I = { \<one>\<^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 "\<not> \<not> 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
+qed
+
+corollary (in group_hom) inj_hom_imp_solvable:
+ assumes "inj_on h (carrier G)"
+ shows "solvable H \<Longrightarrow> solvable G"
+ using not_solvable[OF assms] by auto
+
+theorem (in group_hom) solvable_imp_solvable_img:
+ assumes "subgroup I G"
+ and "solvable_seq G I"
+ shows "solvable_seq H (h ` I)"
+proof -
+ obtain n where "(derived G ^^ n) I = { \<one>\<^bsub>G\<^esub> }"
+ using G.solvable_imp_trivial_derived_seq[OF assms(2) assms(1)] ..
+ hence "(derived H ^^ n) (h ` I) = { \<one>\<^bsub>H\<^esub> }"
+ using derived_of_img[OF G.subgroupE(1)[OF assms(1)], of n] by simp
+ thus ?thesis
+ using H.trivial_derived_seq_imp_solvable[OF subgroup_img_is_subgroup[OF assms(1)]] by simp
+qed
+
+corollary (in group_hom) surj_hom_imp_solvable:
+ assumes "h ` (carrier G) = (carrier H)"
+ shows "solvable G \<Longrightarrow> solvable H"
+ using solvable_imp_solvable_img[OF G.subgroup_self] assms unfolding solvable_def by auto
+
+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 \<subseteq> J"
+ and "\<And>g. \<lbrakk> g \<in> carrier G2; f g = \<one>\<^bsub>G3\<^esub> \<rbrakk> \<Longrightarrow> g \<in> h ` I"
+ shows "\<lbrakk> solvable_seq G1 I; solvable_seq G3 (f ` J) \<rbrakk> \<Longrightarrow> solvable_seq G2 J"
+proof -
+ have G1: "group G1" and G2: "group G2" and G3: "group G3"
+ using assms(1-2) 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 = { \<one>\<^bsub>G1\<^esub> }"
+ and m: "(derived G3 ^^ m) (f ` J) = { \<one>\<^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) = { \<one>\<^bsub>G3\<^esub> }"
+ using m by simp
+ hence "((derived G2 ^^ m) J) \<subseteq> 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) \<subseteq> (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 ` { \<one>\<^bsub>G1\<^esub> }"
+ using group_hom.derived_of_img[OF assms(1) subgroup.subset[OF assms(3)], of n] n by simp
+ also have " ... = { \<one>\<^bsub>G2\<^esub> }"
+ using assms(1) by (simp add: group_hom.hom_one)
+ finally have "(derived G2 ^^ n) ((derived G2 ^^ m) J) \<subseteq> { \<one>\<^bsub>G2\<^esub> }" .
+ hence "(derived G2 ^^ (n + m)) J \<subseteq> { \<one>\<^bsub>G2\<^esub> }"
+ by (metis comp_eq_dest_lhs funpow_add)
+ moreover have "{ \<one>\<^bsub>G2\<^esub> } \<subseteq> (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
+ ultimately show ?thesis
+ using group.trivial_derived_seq_imp_solvable[OF G2 assms(4), of "n + m"] by auto
+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 \<subseteq> h ` (carrier G1)"
+ shows "\<lbrakk> solvable G1; solvable G3 \<rbrakk> \<Longrightarrow> solvable G2"
+proof -
+ assume "solvable G1" "solvable G3"
+ moreover have "\<And>g. \<lbrakk> g \<in> carrier G2; f g = \<one>\<^bsub>G3\<^esub> \<rbrakk> \<Longrightarrow> g \<in> h ` (carrier G1)"
+ using assms(4) unfolding kernel_def by auto
+ moreover have "h ` (carrier G1 ) \<subseteq> (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
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Algebra/Subrings.thy Mon Jul 02 15:43:22 2018 +0100
@@ -0,0 +1,428 @@
+(* ************************************************************************** *)
+(* Title: Subrings.thy *)
+(* Authors: Martin Baillon and Paulo EmÃlio de Vilhena *)
+(* ************************************************************************** *)
+
+theory Subrings
+ imports Ring RingHom QuotRing Multiplicative_Group
+
+begin
+
+section \<open>Subrings\<close>
+
+subsection \<open>Definitions\<close>
+
+locale subring =
+ subgroup H "add_monoid R" + submonoid H R for H and R (structure)
+
+locale subcring = subring +
+ assumes sub_m_comm: "\<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<otimes> h2 = h2 \<otimes> h1"
+
+locale subdomain = subcring +
+ assumes sub_one_not_zero [simp]: "\<one> \<noteq> \<zero>"
+ assumes subintegral: "\<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<otimes> h2 = \<zero> \<Longrightarrow> h1 = \<zero> \<or> h2 = \<zero>"
+
+locale subfield = subdomain K R for K and R (structure) +
+ assumes subfield_Units: "Units (R \<lparr> carrier := K \<rparr>) = K - { \<zero> }"
+
+
+subsection \<open>Basic Properties\<close>
+
+subsubsection \<open>Subrings\<close>
+
+lemma (in ring) subringI:
+ assumes "H \<subseteq> carrier R"
+ and "\<one> \<in> H"
+ and "\<And>h. h \<in> H \<Longrightarrow> \<ominus> h \<in> H"
+ and "\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<otimes> h2 \<in> H"
+ and "\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<oplus> h2 \<in> H"
+ shows "subring H R"
+ using add.subgroupI[OF assms(1) _ assms(3, 5)] assms(2)
+ submonoid.intro[OF assms(1, 4, 2)]
+ unfolding subring_def by auto
+
+lemma subringE:
+ assumes "subring H R"
+ shows "H \<subseteq> carrier R"
+ and "\<zero>\<^bsub>R\<^esub> \<in> H"
+ and "\<one>\<^bsub>R\<^esub> \<in> H"
+ and "H \<noteq> {}"
+ and "\<And>h. h \<in> H \<Longrightarrow> \<ominus>\<^bsub>R\<^esub> h \<in> H"
+ and "\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<otimes>\<^bsub>R\<^esub> h2 \<in> H"
+ and "\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<oplus>\<^bsub>R\<^esub> h2 \<in> H"
+ using subring.axioms[OF assms]
+ unfolding submonoid_def subgroup_def a_inv_def by auto
+
+lemma (in ring) carrier_is_subring: "subring (carrier R) R"
+ by (simp add: subringI)
+
+lemma (in ring) subring_inter:
+ assumes "subring I R" and "subring J R"
+ shows "subring (I \<inter> J) R"
+ using subringE[OF assms(1)] subringE[OF assms(2)] subringI[of "I \<inter> J"] by auto
+
+lemma (in ring) subring_Inter:
+ assumes "\<And>I. I \<in> S \<Longrightarrow> subring I R" and "S \<noteq> {}"
+ shows "subring (\<Inter>S) R"
+proof (rule subringI, auto simp add: assms subringE[of _ R])
+ fix x assume "\<forall>I \<in> S. x \<in> I" thus "x \<in> carrier R"
+ using assms subringE(1)[of _ R] by blast
+qed
+
+lemma (in subring) subring_is_ring:
+ assumes "ring R"
+ shows "ring (R \<lparr> carrier := H \<rparr>)"
+ apply unfold_locales
+ using assms subring_axioms submonoid.one_closed[OF subgroup_is_submonoid] subgroup_is_group
+ by (simp_all add: subringE ring.ring_simprules abelian_group.a_group group.Units_eq ring_def)
+
+lemma (in ring) ring_incl_imp_subring:
+ assumes "H \<subseteq> carrier R"
+ and "ring (R \<lparr> carrier := H \<rparr>)"
+ shows "subring H R"
+ using group.group_incl_imp_subgroup[OF add.group_axioms, of H] assms(1)
+ monoid.monoid_incl_imp_submonoid[OF monoid_axioms assms(1)]
+ ring.axioms(1, 2)[OF assms(2)] abelian_group.a_group[of "R \<lparr> carrier := H \<rparr>"]
+ unfolding subring_def by auto
+
+lemma (in ring) subring_iff:
+ assumes "H \<subseteq> carrier R"
+ shows "subring H R \<longleftrightarrow> ring (R \<lparr> carrier := H \<rparr>)"
+ using subring.subring_is_ring[OF _ ring_axioms] ring_incl_imp_subring[OF assms] by auto
+
+
+subsubsection \<open>Subcrings\<close>
+
+lemma (in ring) subcringI:
+ assumes "subring H R"
+ and "\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<otimes> h2 = h2 \<otimes> h1"
+ shows "subcring H R"
+ unfolding subcring_def subcring_axioms_def using assms by simp+
+
+lemma (in cring) subcringI':
+ assumes "subring H R"
+ shows "subcring H R"
+ using subcringI[OF assms] subringE(1)[OF assms] m_comm by auto
+
+lemma subcringE:
+ assumes "subcring H R"
+ shows "H \<subseteq> carrier R"
+ and "\<zero>\<^bsub>R\<^esub> \<in> H"
+ and "\<one>\<^bsub>R\<^esub> \<in> H"
+ and "H \<noteq> {}"
+ and "\<And>h. h \<in> H \<Longrightarrow> \<ominus>\<^bsub>R\<^esub> h \<in> H"
+ and "\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<otimes>\<^bsub>R\<^esub> h2 \<in> H"
+ and "\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<oplus>\<^bsub>R\<^esub> h2 \<in> H"
+ and "\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<otimes>\<^bsub>R\<^esub> h2 = h2 \<otimes>\<^bsub>R\<^esub> h1"
+ using subringE[OF subcring.axioms(1)[OF assms]] subcring.sub_m_comm[OF assms] by simp+
+
+lemma (in cring) carrier_is_subcring: "subcring (carrier R) R"
+ by (simp add: subcringI' carrier_is_subring)
+
+lemma (in ring) subcring_inter:
+ assumes "subcring I R" and "subcring J R"
+ shows "subcring (I \<inter> J) R"
+ using subcringE[OF assms(1)] subcringE[OF assms(2)]
+ subcringI[of "I \<inter> J"] subringI[of "I \<inter> J"] by auto
+
+lemma (in ring) subcring_Inter:
+ assumes "\<And>I. I \<in> S \<Longrightarrow> subcring I R" and "S \<noteq> {}"
+ shows "subcring (\<Inter>S) R"
+proof (rule subcringI)
+ show "subring (\<Inter>S) R"
+ using subcring.axioms(1)[of _ R] subring_Inter[of S] assms by auto
+next
+ fix h1 h2 assume h1: "h1 \<in> \<Inter>S" and h2: "h2 \<in> \<Inter>S"
+ obtain S' where S': "S' \<in> S"
+ using assms(2) by blast
+ hence "h1 \<in> S'" "h2 \<in> S'"
+ using h1 h2 by blast+
+ thus "h1 \<otimes> h2 = h2 \<otimes> h1"
+ using subcring.sub_m_comm[OF assms(1)[OF S']] by simp
+qed
+
+lemma (in ring) subcring_iff:
+ assumes "H \<subseteq> carrier R"
+ shows "subcring H R \<longleftrightarrow> cring (R \<lparr> carrier := H \<rparr>)"
+proof
+ assume A: "subcring H R"
+ hence ring: "ring (R \<lparr> carrier := H \<rparr>)"
+ using subring_iff[OF assms] subcring.axioms(1)[OF A] by simp
+ moreover have "comm_monoid (R \<lparr> carrier := H \<rparr>)"
+ using monoid.monoid_comm_monoidI[OF ring.is_monoid[OF ring]]
+ subcring.sub_m_comm[OF A] by auto
+ ultimately show "cring (R \<lparr> carrier := H \<rparr>)"
+ using cring_def by blast
+next
+ assume A: "cring (R \<lparr> carrier := H \<rparr>)"
+ hence "subring H R"
+ using cring.axioms(1) subring_iff[OF assms] by simp
+ moreover have "comm_monoid (R \<lparr> carrier := H \<rparr>)"
+ using A unfolding cring_def by simp
+ hence"\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<otimes> h2 = h2 \<otimes> h1"
+ using comm_monoid.m_comm[of "R \<lparr> carrier := H \<rparr>"] by auto
+ ultimately show "subcring H R"
+ unfolding subcring_def subcring_axioms_def by auto
+qed
+
+
+subsubsection \<open>Subdomains\<close>
+
+lemma (in ring) subdomainI:
+ assumes "subcring H R"
+ and "\<one> \<noteq> \<zero>"
+ and "\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<otimes> h2 = \<zero> \<Longrightarrow> h1 = \<zero> \<or> h2 = \<zero>"
+ shows "subdomain H R"
+ unfolding subdomain_def subdomain_axioms_def using assms by simp+
+
+lemma (in domain) subdomainI':
+ assumes "subring H R"
+ shows "subdomain H R"
+proof (rule subdomainI[OF subcringI[OF assms]], simp_all)
+ show "\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<otimes> h2 = h2 \<otimes> h1"
+ using m_comm subringE(1)[OF assms] by auto
+ show "\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H; h1 \<otimes> h2 = \<zero> \<rbrakk> \<Longrightarrow> (h1 = \<zero>) \<or> (h2 = \<zero>)"
+ using integral subringE(1)[OF assms] by auto
+qed
+
+lemma subdomainE:
+ assumes "subdomain H R"
+ shows "H \<subseteq> carrier R"
+ and "\<zero>\<^bsub>R\<^esub> \<in> H"
+ and "\<one>\<^bsub>R\<^esub> \<in> H"
+ and "H \<noteq> {}"
+ and "\<And>h. h \<in> H \<Longrightarrow> \<ominus>\<^bsub>R\<^esub> h \<in> H"
+ and "\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<otimes>\<^bsub>R\<^esub> h2 \<in> H"
+ and "\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<oplus>\<^bsub>R\<^esub> h2 \<in> H"
+ and "\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<otimes>\<^bsub>R\<^esub> h2 = h2 \<otimes>\<^bsub>R\<^esub> h1"
+ and "\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<otimes>\<^bsub>R\<^esub> h2 = \<zero>\<^bsub>R\<^esub> \<Longrightarrow> h1 = \<zero>\<^bsub>R\<^esub> \<or> h2 = \<zero>\<^bsub>R\<^esub>"
+ and "\<one>\<^bsub>R\<^esub> \<noteq> \<zero>\<^bsub>R\<^esub>"
+ using subcringE[OF subdomain.axioms(1)[OF assms]] assms
+ unfolding subdomain_def subdomain_axioms_def by auto
+
+lemma (in ring) subdomain_iff:
+ assumes "H \<subseteq> carrier R"
+ shows "subdomain H R \<longleftrightarrow> domain (R \<lparr> carrier := H \<rparr>)"
+proof
+ assume A: "subdomain H R"
+ hence cring: "cring (R \<lparr> carrier := H \<rparr>)"
+ using subcring_iff[OF assms] subdomain.axioms(1)[OF A] by simp
+ thus "domain (R \<lparr> carrier := H \<rparr>)"
+ using domain.intro[OF cring] subdomain.subintegral[OF A] subdomain.sub_one_not_zero[OF A]
+ unfolding domain_axioms_def by auto
+next
+ assume A: "domain (R \<lparr> carrier := H \<rparr>)"
+ hence subcring: "subcring H R"
+ using subcring_iff[OF assms] unfolding domain_def by simp
+ thus "subdomain H R"
+ using subdomain.intro[OF subcring] domain.integral[OF A] domain.one_not_zero[OF A]
+ unfolding subdomain_axioms_def by auto
+qed
+
+
+subsubsection \<open>Subfields\<close>
+
+lemma (in ring) subfieldI:
+ assumes "subcring K R" and "Units (R \<lparr> carrier := K \<rparr>) = K - { \<zero> }"
+ shows "subfield K R"
+proof (rule subfield.intro)
+ show "subfield_axioms K R"
+ using assms(2) unfolding subfield_axioms_def .
+ show "subdomain K R"
+ proof (rule subdomainI[OF assms(1)], auto)
+ have subM: "submonoid K R"
+ using subring.axioms(2)[OF subcring.axioms(1)[OF assms(1)]] .
+
+ show contr: "\<one> = \<zero> \<Longrightarrow> False"
+ proof -
+ assume one_eq_zero: "\<one> = \<zero>"
+ have "\<one> \<in> K" and "\<one> \<otimes> \<one> = \<one>"
+ using submonoid.one_closed[OF subM] by simp+
+ hence "\<one> \<in> Units (R \<lparr> carrier := K \<rparr>)"
+ unfolding Units_def by (simp, blast)
+ hence "\<one> \<noteq> \<zero>"
+ using assms(2) by simp
+ thus False
+ using one_eq_zero by simp
+ qed
+
+ fix k1 k2 assume k1: "k1 \<in> K" and k2: "k2 \<in> K" "k2 \<noteq> \<zero>" and k12: "k1 \<otimes> k2 = \<zero>"
+ obtain k2' where k2': "k2' \<in> K" "k2' \<otimes> k2 = \<one>" "k2 \<otimes> k2' = \<one>"
+ using assms(2) k2 unfolding Units_def by auto
+ have "\<zero> = (k1 \<otimes> k2) \<otimes> k2'"
+ using k12 k2'(1) submonoid.mem_carrier[OF subM] by fastforce
+ also have "... = k1"
+ using k1 k2(1) k2'(1,3) submonoid.mem_carrier[OF subM] by (simp add: m_assoc)
+ finally have "\<zero> = k1" .
+ thus "k1 = \<zero>" by simp
+ qed
+qed
+
+lemma (in field) subfieldI':
+ assumes "subring K R" and "\<And>k. k \<in> K - { \<zero> } \<Longrightarrow> inv k \<in> K"
+ shows "subfield K R"
+proof (rule subfieldI)
+ show "subcring K R"
+ using subcringI[OF assms(1)] m_comm subringE(1)[OF assms(1)] by auto
+ show "Units (R \<lparr> carrier := K \<rparr>) = K - { \<zero> }"
+ proof
+ show "K - { \<zero> } \<subseteq> Units (R \<lparr> carrier := K \<rparr>)"
+ proof
+ fix k assume k: "k \<in> K - { \<zero> }"
+ hence inv_k: "inv k \<in> K"
+ using assms(2) by simp
+ moreover have "k \<in> carrier R - { \<zero> }"
+ using subringE(1)[OF assms(1)] k by auto
+ ultimately have "k \<otimes> inv k = \<one>" "inv k \<otimes> k = \<one>"
+ by (simp add: field_Units)+
+ thus "k \<in> Units (R \<lparr> carrier := K \<rparr>)"
+ unfolding Units_def using k inv_k by auto
+ qed
+ next
+ show "Units (R \<lparr> carrier := K \<rparr>) \<subseteq> K - { \<zero> }"
+ proof
+ fix k assume k: "k \<in> Units (R \<lparr> carrier := K \<rparr>)"
+ then obtain k' where k': "k' \<in> K" "k \<otimes> k' = \<one>"
+ unfolding Units_def by auto
+ hence "k \<in> carrier R" and "k' \<in> carrier R"
+ using k subringE(1)[OF assms(1)] unfolding Units_def by auto
+ hence "\<zero> = \<one>" if "k = \<zero>"
+ using that k'(2) by auto
+ thus "k \<in> K - { \<zero> }"
+ using k unfolding Units_def by auto
+ qed
+ qed
+qed
+
+lemma (in field) carrier_is_subfield: "subfield (carrier R) R"
+ by (auto intro: subfieldI[OF carrier_is_subcring] simp add: field_Units)
+
+lemma subfieldE:
+ assumes "subfield K R"
+ shows "subring K R" and "subcring K R"
+ and "K \<subseteq> carrier R"
+ and "\<And>k1 k2. \<lbrakk> k1 \<in> K; k2 \<in> K \<rbrakk> \<Longrightarrow> k1 \<otimes>\<^bsub>R\<^esub> k2 = k2 \<otimes>\<^bsub>R\<^esub> k1"
+ and "\<And>k1 k2. \<lbrakk> k1 \<in> K; k2 \<in> K \<rbrakk> \<Longrightarrow> k1 \<otimes>\<^bsub>R\<^esub> k2 = \<zero>\<^bsub>R\<^esub> \<Longrightarrow> k1 = \<zero>\<^bsub>R\<^esub> \<or> k2 = \<zero>\<^bsub>R\<^esub>"
+ and "\<one>\<^bsub>R\<^esub> \<noteq> \<zero>\<^bsub>R\<^esub>"
+ using subdomain.axioms(1)[OF subfield.axioms(1)[OF assms]] subcring_def
+ subdomainE(1, 8, 9, 10)[OF subfield.axioms(1)[OF assms]] by auto
+
+lemma (in ring) subfield_m_inv:
+ assumes "subfield K R" and "k \<in> K - { \<zero> }"
+ shows "inv k \<in> K - { \<zero> }" and "k \<otimes> inv k = \<one>" and "inv k \<otimes> k = \<one>"
+proof -
+ have K: "subring K R" "submonoid K R"
+ using subfieldE(1)[OF assms(1)] subring.axioms(2) by auto
+ have monoid: "monoid (R \<lparr> carrier := K \<rparr>)"
+ using submonoid.submonoid_is_monoid[OF subring.axioms(2)[OF K(1)] is_monoid] .
+
+ have "monoid R"
+ by (simp add: monoid_axioms)
+
+ hence k: "k \<in> Units (R \<lparr> carrier := K \<rparr>)"
+ using subfield.subfield_Units[OF assms(1)] assms(2) by blast
+ hence unit_of_R: "k \<in> Units R"
+ using assms(2) subringE(1)[OF subfieldE(1)[OF assms(1)]] unfolding Units_def by auto
+ have "inv\<^bsub>(R \<lparr> carrier := K \<rparr>)\<^esub> k \<in> Units (R \<lparr> carrier := K \<rparr>)"
+ by (simp add: k monoid monoid.Units_inv_Units)
+ hence "inv\<^bsub>(R \<lparr> carrier := K \<rparr>)\<^esub> k \<in> K - { \<zero> }"
+ using subfield.subfield_Units[OF assms(1)] by blast
+ thus "inv k \<in> K - { \<zero> }" and "k \<otimes> inv k = \<one>" and "inv k \<otimes> k = \<one>"
+ using Units_l_inv[OF unit_of_R] Units_r_inv[OF unit_of_R]
+ using monoid.m_inv_monoid_consistent[OF monoid_axioms k K(2)] by auto
+qed
+
+lemma (in ring) subfield_iff:
+ shows "\<lbrakk> field (R \<lparr> carrier := K \<rparr>); K \<subseteq> carrier R \<rbrakk> \<Longrightarrow> subfield K R"
+ and "subfield K R \<Longrightarrow> field (R \<lparr> carrier := K \<rparr>)"
+proof-
+ assume A: "field (R \<lparr> carrier := K \<rparr>)" "K \<subseteq> carrier R"
+ have "\<And>k1 k2. \<lbrakk> k1 \<in> K; k2 \<in> K \<rbrakk> \<Longrightarrow> k1 \<otimes> k2 = k2 \<otimes> k1"
+ using comm_monoid.m_comm[OF cring.axioms(2)[OF fieldE(1)[OF A(1)]]] by simp
+ moreover have "subring K R"
+ using ring_incl_imp_subring[OF A(2) cring.axioms(1)[OF fieldE(1)[OF A(1)]]] .
+ ultimately have "subcring K R"
+ using subcringI by simp
+ thus "subfield K R"
+ using field.field_Units[OF A(1)] subfieldI by auto
+next
+ assume A: "subfield K R"
+ have cring: "cring (R \<lparr> carrier := K \<rparr>)"
+ using subcring_iff[OF subringE(1)[OF subfieldE(1)[OF A]]] subfieldE(2)[OF A] by simp
+ thus "field (R \<lparr> carrier := K \<rparr>)"
+ using cring.cring_fieldI[OF cring] subfield.subfield_Units[OF A] by simp
+qed
+
+lemma (in field) subgroup_mult_of :
+ assumes "subfield K R"
+ shows "subgroup (K - {\<zero>}) (mult_of R)"
+proof (intro group.group_incl_imp_subgroup[OF field_mult_group])
+ show "K - {\<zero>} \<subseteq> carrier (mult_of R)"
+ by (simp add: Diff_mono assms carrier_mult_of subfieldE(3))
+ show "group ((mult_of R) \<lparr> carrier := K - {\<zero>} \<rparr>)"
+ using field.field_mult_group[OF subfield_iff(2)[OF assms]]
+ unfolding mult_of_def by simp
+qed
+
+
+subsection \<open>Subring Homomorphisms\<close>
+
+lemma (in ring) hom_imp_img_subring:
+ assumes "h \<in> ring_hom R S" and "subring K R"
+ shows "ring (S \<lparr> carrier := h ` K, one := h \<one>, zero := h \<zero> \<rparr>)"
+proof -
+ have "ring (R \<lparr> carrier := K \<rparr>)"
+ using subring.subring_is_ring[OF assms(2) ring_axioms] by simp
+ moreover have "h \<in> ring_hom (R \<lparr> carrier := K \<rparr>) S"
+ using assms subringE(1)[OF assms (2)] unfolding ring_hom_def
+ apply simp
+ apply blast
+ done
+ ultimately show ?thesis
+ using ring.ring_hom_imp_img_ring[of "R \<lparr> carrier := K \<rparr>" h S] by simp
+qed
+
+lemma (in ring_hom_ring) img_is_subring:
+ assumes "subring K R" shows "subring (h ` K) S"
+proof -
+ have "ring (S \<lparr> carrier := h ` K \<rparr>)"
+ using R.hom_imp_img_subring[OF homh assms] hom_zero hom_one by simp
+ moreover have "h ` K \<subseteq> carrier S"
+ using ring_hom_memE(1)[OF homh] subringE(1)[OF assms] by auto
+ ultimately show ?thesis
+ using ring_incl_imp_subring by simp
+qed
+
+lemma (in ring_hom_ring) img_is_subfield:
+ assumes "subfield K R" and "\<one>\<^bsub>S\<^esub> \<noteq> \<zero>\<^bsub>S\<^esub>"
+ shows "inj_on h K" and "subfield (h ` K) S"
+proof -
+ have K: "K \<subseteq> carrier R" "subring K R" "subring (h ` K) S"
+ using subfieldE(1)[OF assms(1)] subringE(1) img_is_subring by auto
+ have field: "field (R \<lparr> carrier := K \<rparr>)"
+ and ring: "ring (R \<lparr> carrier := K \<rparr>)" "ring (S \<lparr> carrier := h ` K \<rparr>)"
+ using R.subfield_iff assms(1)
+ subring.subring_is_ring[OF K(2) R.ring_axioms]
+ subring.subring_is_ring[OF K(3) S.ring_axioms] by auto
+
+ hence h: "h \<in> ring_hom (R \<lparr> carrier := K \<rparr>) (S \<lparr> carrier := h ` K \<rparr>)"
+ unfolding ring_hom_def apply auto
+ using ring_hom_memE[OF homh] K
+ by (meson contra_subsetD)+
+ hence ring_hom: "ring_hom_ring (R \<lparr> carrier := K \<rparr>) (S \<lparr> carrier := h ` K \<rparr>) h"
+ using ring_axioms ring ring_hom_ringI2 by blast
+ have "h ` K \<noteq> { \<zero>\<^bsub>S\<^esub> }"
+ using subfieldE(1, 5)[OF assms(1)] subringE(3) assms(2)
+ by (metis hom_one image_eqI singletonD)
+ thus "inj_on h K"
+ using ring_hom_ring.non_trivial_field_hom_imp_inj[OF ring_hom field] by auto
+
+ hence "h \<in> ring_iso (R \<lparr> carrier := K \<rparr>) (S \<lparr> carrier := h ` K \<rparr>)"
+ using h unfolding ring_iso_def bij_betw_def by auto
+ hence "field (S \<lparr> carrier := h ` K \<rparr>)"
+ using field.ring_iso_imp_img_field[OF field, of h "S \<lparr> carrier := h ` K \<rparr>"] by auto
+ thus "subfield (h ` K) S"
+ using S.subfield_iff[of "h ` K"] K(1) ring_hom_memE(1)[OF homh] by blast
+qed
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Algebra/Sym_Groups.thy Mon Jul 02 15:43:22 2018 +0100
@@ -0,0 +1,664 @@
+(* ************************************************************************** *)
+(* Title: Sym_Groups.th *)
+(* Author: Paulo EmÃlio de Vilhena *)
+(* ************************************************************************** *)
+
+theory Sym_Groups
+ imports Cycles Group Coset Generated_Groups Solvable_Groups
+
+begin
+
+abbreviation inv' :: "('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'a)"
+ where "inv' f \<equiv> Hilbert_Choice.inv f"
+
+definition sym_group :: "nat \<Rightarrow> (nat \<Rightarrow> nat) monoid"
+ where "sym_group n = \<lparr> carrier = { p. p permutes {1..n} }, mult = (\<circ>), one = id \<rparr>"
+
+definition sign_img :: "int monoid"
+ where "sign_img = \<lparr> carrier = { -1, 1 }, mult = ( * ), one = 1 \<rparr>"
+
+
+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
+
+lemma sym_group_inv_closed:
+ assumes "p \<in> carrier (sym_group n)"
+ shows "inv' p \<in> carrier (sym_group n)"
+ using assms permutes_inv sym_group_def by auto
+
+lemma sym_group_inv_equality:
+ assumes "p \<in> carrier (sym_group n)"
+ shows "inv\<^bsub>(sym_group n)\<^esub> p = inv' p"
+proof -
+ have "inv' p \<circ> p = id"
+ using assms permutes_inv_o(2) sym_group_def by auto
+ hence "(inv' p) \<otimes>\<^bsub>(sym_group n)\<^esub> p = one (sym_group n)"
+ by (simp add: sym_group_def)
+ thus ?thesis using group.inv_equality[OF sym_group_is_group]
+ 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 "\<And>x. sign x \<noteq> - 1 \<Longrightarrow> x permutes {Suc 0..n} \<Longrightarrow> sign x = 1"
+ by (meson sign_def)
+ show "\<And>x y. \<lbrakk> x permutes {Suc 0..n}; y permutes {Suc 0..n} \<rbrakk> \<Longrightarrow>
+ sign (x \<circ> y) = sign x * sign y"
+ using sign_compose permutation_permutes by blast
+ qed
+qed
+
+
+definition alt_group :: "nat \<Rightarrow> (nat \<Rightarrow> nat) monoid"
+ where "alt_group n = (sym_group n) \<lparr> carrier := { p. p permutes {1..n} \<and> evenperm p } \<rparr>"
+
+lemma alt_group_is_kernel_from_sign:
+ "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_inv_closed:
+ assumes "p \<in> carrier (alt_group n)"
+ shows "inv' p \<in> carrier (alt_group n)"
+ using assms permutes_inv alt_group_def
+ using evenperm_inv permutation_permutes by fastforce
+
+lemma alt_group_inv_equality:
+ assumes "p \<in> carrier (alt_group n)"
+ shows "inv\<^bsub>(alt_group n)\<^esub> p = inv' p"
+proof -
+ have "inv' p \<circ> p = id"
+ using assms permutes_inv_o(2) alt_group_def by auto
+ hence "(inv' p) \<otimes>\<^bsub>(alt_group n)\<^esub> p = one (alt_group n)"
+ by (simp add: alt_group_def sym_group_def)
+ thus ?thesis using group.inv_equality[OF alt_group_is_group]
+ by (simp add: assms alt_group_inv_closed)
+qed
+
+
+subsection \<open>Transposition Sequences\<close>
+
+text \<open>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\<close>
+
+inductive swapidseq_ext :: "'a set \<Rightarrow> nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool" where
+empty: "swapidseq_ext {} 0 id" |
+single: "\<lbrakk> swapidseq_ext S n p; a \<notin> S \<rbrakk> \<Longrightarrow> swapidseq_ext (insert a S) n p" |
+comp: "\<lbrakk> swapidseq_ext S n p; a \<noteq> b \<rbrakk> \<Longrightarrow>
+ swapidseq_ext (insert a (insert b S)) (Suc n) ((Fun.swap a b id) \<circ> p)"
+
+
+lemma swapidseq_ext_finite:
+ assumes "swapidseq_ext S n p"
+ shows "finite S" using assms
+ apply (induction) by auto
+
+lemma swapidseq_ext_zero_imp_id:
+ assumes "swapidseq_ext S 0 p"
+ shows "p = id"
+proof -
+ { fix S n and p :: "'a \<Rightarrow> '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)] .
+qed
+
+lemma swapidseq_ext_finite_expansion:
+ assumes "finite B" "swapidseq_ext A n p"
+ shows "swapidseq_ext (A \<union> 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)
+qed
+
+lemma swapidseq_ext_backwards:
+ assumes "swapidseq_ext A (Suc n) p"
+ shows "\<exists>a b A' p'. a \<noteq> b \<and> A = (insert a (insert b A')) \<and>
+ swapidseq_ext A' n p' \<and> p = (Fun.swap a b id) \<circ> p'"
+proof -
+ { fix A n k and p :: "'a \<Rightarrow> 'a" assume "swapidseq_ext A n p" "n = Suc k"
+ hence "\<exists>a b A' p'. a \<noteq> b \<and> A = (insert a (insert b A')) \<and>
+ swapidseq_ext A' k p' \<and> p = (Fun.swap a b id) \<circ> p'"
+ proof (induction)
+ case empty thus ?case by simp
+ next
+ case single thus ?case
+ by (metis Un_insert_right insert_iff insert_is_Un swapidseq_ext.single)
+ next
+ case comp thus ?case by blast
+ qed }
+ thus ?thesis using assms by simp
+qed
+
+lemma swapidseq_ext_endswap:
+ assumes "swapidseq_ext S n p" "a \<noteq> b"
+ shows "swapidseq_ext (insert a (insert b S)) (Suc n) (p \<circ> (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)
+next
+ case (Suc n)
+ then obtain c d S' and p' :: "'a \<Rightarrow> 'a"
+ where cd: "c \<noteq> d"
+ and S: "S = (insert c (insert d S'))" "swapidseq_ext S' n p'"
+ and p: "p = (Fun.swap c d id) \<circ> p'"
+ using swapidseq_ext_backwards[OF Suc(2)] by blast
+ hence "swapidseq_ext (insert a (insert b S')) (Suc n) (p' \<circ> (Fun.swap a b id))"
+ by (simp add: Suc.IH Suc.prems(2))
+ hence "swapidseq_ext (insert c (insert d (insert a (insert b S')))) (Suc (Suc n))
+ ((Fun.swap c d id) \<circ> p' \<circ> (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)
+qed
+
+lemma swapidseq_ext_extension:
+ assumes "swapidseq_ext A n p" "swapidseq_ext B m q" "A \<inter> B = {}"
+ shows "swapidseq_ext (A \<union> B) (n + m) (p \<circ> q)"
+proof -
+ { fix m and q :: "'a \<Rightarrow> 'a" and A B :: "'a set" assume "finite A" "swapidseq_ext B m q"
+ hence "swapidseq_ext (A \<union> 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' \<union> 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 \<Rightarrow> 'a"
+ where A': "a \<noteq> b" "A = (insert a (insert b A'))"
+ and p': "swapidseq_ext A' n p'"
+ and p: "p = (Fun.swap a b id) \<circ> p'"
+ using swapidseq_ext_backwards[OF Suc(2)] by blast
+ hence "swapidseq_ext (A' \<union> B) (n + m) (p' \<circ> q)"
+ using Suc.hyps Suc.prems(3) assms(2) by fastforce
+ thus ?case using swapidseq_ext.comp[of "A' \<union> B" "n + m" "p' \<circ> q" a b]
+ by (metis Un_insert_left p A' add_Suc rewriteR_comp_comp)
+ qed
+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)
+ 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 \<noteq> []" hence "length (j # cs) \<ge> 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
+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))
+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))
+qed
+
+lemma cycle_decomp_imp_swapidseq_ext:
+ assumes "cycle_decomp S p"
+ shows "\<exists>n. swapidseq_ext S n p" using assms
+proof (induction)
+ case empty
+ then 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
+ hence "swapidseq_ext (set cs) (length cs - 1) (cycle_of_list cs)"
+ using comp.hyps(2) swapidseq_ext_of_cycles by blast
+ thus ?case using swapidseq_ext_extension m
+ using comp.hyps(3) by blast
+qed
+
+lemma swapidseq_ext_of_permutations:
+ assumes "p permutes S" "finite S"
+ shows "\<exists>n. swapidseq_ext S n p"
+ using assms cycle_decomp_imp_swapidseq_ext cycle_decomposition by blast
+
+lemma split_swapidseq:
+ assumes "k \<le> n" "swapidseq n p"
+ shows "\<exists>q r. swapidseq k q \<and> swapidseq (n - k) r \<and> p = q \<circ> r"
+proof -
+ { fix n :: "nat" and p :: "'a \<Rightarrow> 'a" assume "swapidseq (Suc n) p"
+ hence "\<exists>a b q. a \<noteq> b \<and> swapidseq n q \<and> p = (Fun.swap a b id) \<circ> 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 \<circ> r"
+ using Suc_leD by blast
+ then obtain a b r' where 2: "a \<noteq> b" "swapidseq (n - (Suc k)) r'" "r = (Fun.swap a b id) \<circ> 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 \<circ> (Fun.swap a b id))" using 1 2 by (simp add: swapidseq_endswap)
+ moreover have "p = (q \<circ> (Fun.swap a b id)) \<circ> r'"
+ using 1 2 fun.map_comp by blast
+ ultimately show ?case using 2 by blast
+ qed
+qed
+
+lemma split_swapidseq_ext:
+ assumes "k \<le> n" "swapidseq_ext S n p"
+ shows "\<exists>q r S1 S2. swapidseq_ext S1 k q \<and> swapidseq_ext S2 (n - k) r \<and> p = q \<circ> r \<and> S1 \<union> 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 \<and> swapidseq_ext S (n - 0) p \<and> p = id \<circ> 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 \<circ> r" "S1 \<union> S2 = S"
+ using Suc_leD by blast
+ then obtain a b S2' and r' :: "'a \<Rightarrow> 'a"
+ where S2': "a \<noteq> b" "S2 = (insert a (insert b S2'))"
+ and r': "swapidseq_ext S2' (n - (Suc k)) r'"
+ and r: "r = (Fun.swap a b id) \<circ> 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 \<circ> (Fun.swap a b id))"
+ by (simp add: S2'(1) \<open>swapidseq_ext S1 k q\<close> swapidseq_ext_endswap)
+ moreover have "p = (q \<circ> (Fun.swap a b id)) \<circ> r'"
+ by (simp add: \<open>p = q \<circ> r\<close> fun.map_comp r)
+ moreover have "(insert a (insert b S1)) \<union> S2' = S"
+ using S2'(2) \<open>S1 \<union> S2 = S\<close> by auto
+ ultimately show ?case using r r' by blast
+qed
+
+
+
+abbreviation three_cycles :: "nat \<Rightarrow> (nat \<Rightarrow> nat) set"
+ where "three_cycles n \<equiv>
+ { cycle_of_list cs | cs. cycle cs \<and> length cs = 3 \<and> set cs \<subseteq> {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 "\<forall> ia < length cs. cs ! ia = [(cs ! 0), (cs ! 1), (cs ! 2)] ! ia"
+ by (metis Suc_1 Suc_eq_plus1 add.left_neutral assms less_antisym
+ less_one nth_Cons' nth_Cons_Suc numeral_3_eq_3)
+qed
+
+lemma alt_group_as_three_cycles:
+ "carrier (alt_group n) = generate (alt_group n) (three_cycles n)"
+proof
+ show "generate (alt_group n) (three_cycles n) \<subseteq> carrier (alt_group n)"
+ proof
+ { fix p assume "p \<in> three_cycles n"
+ have "p \<in> carrier (alt_group n)"
+ proof -
+ from \<open>p \<in> three_cycles n\<close>
+ obtain cs where p: "p = cycle_of_list cs"
+ and cs: "cycle cs" "length cs = 3" "set cs \<subseteq> {1..n}" by blast
+ hence "p = (Fun.swap (cs ! 0) (cs ! 1) id) \<circ> (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
+
+ fix p assume "p \<in> generate (alt_group n) (three_cycles n)"
+ thus "p \<in> 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) \<subseteq> generate (alt_group n) (three_cycles n)"
+ proof
+ fix p assume p: "p \<in> carrier (alt_group n)"
+ show "p \<in> generate (alt_group n) (three_cycles n)"
+ proof -
+ { fix q :: "nat \<Rightarrow> nat" and a b c
+ assume A: "a \<noteq> b" "b \<noteq> c" "{ a, b, c } \<subseteq> {1..n}" "q = cycle_of_list [a, b, c]"
+ have "q \<in> generate (alt_group n) (three_cycles n)"
+ proof (cases)
+ assume "c = a" hence "q = id" by (simp add: A(4) swap_commute)
+ thus "q \<in> generate (alt_group n) (three_cycles n)"
+ using generate.one[of "alt_group n"] by (simp add: alt_group_def sym_group_def)
+ next
+ assume "c \<noteq> a" hence "q \<in> (three_cycles n)"
+ by (smt A distinct.simps(2) distinct_singleton empty_set length_Cons
+ list.simps(15) list.size(3) mem_Collect_eq numeral_3_eq_3 set_ConsD)
+ thus "q \<in> generate (alt_group n) (three_cycles n)"
+ by (simp add: generate.incl)
+ qed } note aux_lemma1 = this
+
+ { fix S :: "nat set" and q :: "nat \<Rightarrow> nat" assume A: "swapidseq_ext S 2 q" "S \<subseteq> {1..n}"
+ have "q \<in> generate (alt_group n) (three_cycles n)"
+ proof -
+ obtain a b S' q' where ab: "a \<noteq> b" "S = (insert a (insert b S'))"
+ and q': "swapidseq_ext S' 1 q'" "q = (Fun.swap a b id) \<circ> q'"
+ using swapidseq_ext_backwards[of S 1 q] A(1) Suc_1 by metis
+ then obtain c d S'' where cd: "c \<noteq> d" "S' = (insert c (insert d S''))"
+ and q: "q = (Fun.swap a b id) \<circ> (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 } \<subseteq> {1..n}" using A(2) ab(2) by blast
+ thus ?thesis
+ proof (cases)
+ assume Eq: "b = c" hence "q = cycle_of_list [a, b, d]" by (simp add: q)
+ thus ?thesis using aux_lemma1 ab cd Eq incl by simp
+ next
+ assume In: "b \<noteq> c"
+ have "q = (cycle_of_list [a, b, c]) \<circ> (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)
+ thus ?thesis
+ using aux_lemma1[of a b c] aux_lemma1[of b c d]
+ generate.eng[where ?h1.0 = "cycle_of_list [a, b, c]"
+ and ?h2.0 = "cycle_of_list [b, c, d]"]
+ using In ab alt_group_def cd sym_group_def incl
+ by (smt insert_subset monoid.select_convs(1) partial_object.simps(3))
+ qed
+ 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 \<open>swapidseq l p\<close> evenperm_unique by blast
+
+ then obtain k where "swapidseq_ext {1..n} (2 * k) p"
+ using dvd_def by (metis \<open>swapidseq_ext {1..n} l p\<close>)
+ thus "p \<in> generate (alt_group n) (three_cycles n)"
+ proof (induct k arbitrary: p)
+ case 0
+ hence "p = id" by (simp add: swapidseq_ext_zero_imp_id)
+ moreover have "id \<in> 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
+ 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 \<circ> r" "S1 \<union> S2 = {1..n}"
+ using split_swapidseq_ext[of 2 "2 * Suc k" "{1..n}" p]
+ by (smt One_nat_def Suc_1 Suc_leI Suc_le_mono add_diff_cancel_left' le_Suc_eq
+ mult_Suc_right nat_mult_eq_1_iff one_le_mult_iff zero_less_Suc)
+
+ hence "r \<in> 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 \<in> generate (alt_group n) (three_cycles n)"
+ using aux_lemma2[OF split(1)] \<open>S1 \<union> S2 = {1..n}\<close> by auto
+ ultimately show ?case
+ using split generate.eng[of q "alt_group n" "three_cycles n" r]
+ by (smt alt_group_def monoid.select_convs(1) partial_object.simps(3) sym_group_def)
+ qed
+ qed
+ qed
+qed
+
+lemma elts_from_card:
+ assumes "card S \<ge> n"
+ obtains f where "inj_on f {1..n}" "f ` {1..n} \<subseteq> S"
+proof -
+ have "\<exists>f. inj_on f {1..n} \<and> f ` {1..n} \<subseteq> 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} \<subseteq> S"
+ using Suc_leD by blast
+ hence "card (f ` {1..n}) = n" by (simp add: card_image)
+ then obtain y where y: "y \<in> S - (f ` {1..n})"
+ by (metis Diff_eq_empty_iff Suc.prems \<open>f ` {1..n} \<subseteq> S\<close> not_less_eq_eq order_refl some_in_eq subset_antisym)
+ define f' where f': "f' = (\<lambda>x. (if x \<in> {1..n} then f x else y))"
+ hence "\<And>i j. \<lbrakk> i \<in> {1..Suc n}; j \<in> {1..Suc n} \<rbrakk> \<Longrightarrow> f' i = f' j \<Longrightarrow> 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} \<subseteq> S \<and> f' (Suc n) \<in> S"
+ using f f' y by (simp add: image_subset_iff)
+ hence "f' ` {1..Suc n} \<subseteq> 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 \<ge> 5"
+ shows "derived (alt_group n) (carrier (alt_group n)) = carrier (alt_group n)"
+proof
+ show "derived (alt_group n) (carrier (alt_group n)) \<subseteq> carrier (alt_group n)"
+ by (simp add: alt_group_is_group group.derived_incl group.subgroup_self)
+next
+ show "carrier (alt_group n) \<subseteq> derived (alt_group n) (carrier (alt_group n))"
+ proof -
+ have derived_set_in_carrier:
+ "derived_set (alt_group n) (carrier (alt_group n)) \<subseteq> carrier (alt_group n)"
+ proof
+ fix p assume "p \<in> derived_set (alt_group n) (carrier (alt_group n))"
+ then obtain q r
+ where q: "q \<in> carrier (alt_group n)"
+ and r: "r \<in> carrier (alt_group n)"
+ and "p = q \<otimes>\<^bsub>(alt_group n)\<^esub> r \<otimes>\<^bsub>(alt_group n)\<^esub> (inv' q) \<otimes>\<^bsub>(alt_group n)\<^esub> (inv' r)"
+ using alt_group_inv_equality by auto
+ hence p: "p = q \<circ> r \<circ> (inv' q) \<circ> (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} \<and> 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 \<in> carrier (alt_group n)" by (simp add: alt_group_def)
+ qed
+
+ have "three_cycles n \<subseteq> derived_set (alt_group n) (carrier (alt_group n))"
+ proof
+ fix p :: "nat \<Rightarrow> nat" assume "p \<in> three_cycles n"
+ then obtain cs
+ where cs: "cycle cs" "length cs = 3" "set cs \<subseteq> {1..n}" and p: "p = cycle_of_list cs" 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 "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 \<notin> {i, j, k}"
+ hence "l \<notin> set cs \<and> l \<notin> 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 "\<not> l \<notin> {i, j, k}" hence "l \<in> {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 \<open>l \<in> {i, j, k}\<close> by auto
+ qed
+ qed
+ hence "p ^^ 2 = (Fun.swap j k id) \<circ> (cycle_of_list [i, j, k]) \<circ> (inv' (Fun.swap j k id))"
+ using conjugation_of_cycle[of "[i, j, k]" "Fun.swap j k id"] cs(1) ijk by auto
+
+ moreover have "card ({1..n} - {i, j, k}) \<ge> 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}) \<ge> 2"
+ using assms cs ijk by simp
+ then obtain f :: "nat \<Rightarrow> nat" where f: "inj_on f {1..2}" "f ` {1..2} \<subseteq> {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 \<noteq> 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 \<in> {1..n} - {i, j, k}" using f(2) gh(2) by force
+ hence h: "h \<in> {1..n} - {i, j, k}" using f(2) gh(2) by force
+ hence gh_simps: "g \<noteq> h \<and> g \<in> {1..n} \<and> h \<in> {1..n} \<and> g \<notin> {i, j, k} \<and> h \<notin> {i, j, k}"
+ using g gh(3) by blast
+
+ ultimately have final_step:
+ "p ^^ 2 = ((Fun.swap j k id) \<circ> (Fun.swap g h id)) \<circ>
+ (cycle_of_list [i, j, k]) \<circ>
+ (inv' ((Fun.swap j k id) \<circ> (Fun.swap g h id)))"
+ by (smt bij_id bij_swap_iff comp_id cycle_of_list.simps(1) cycle_of_list.simps(3)
+ fun.map_comp insertCI inv_swap_id o_inv_distrib o_inv_o_cancel surj_id
+ surj_imp_inj_inv surj_imp_surj_swap swap_id_independent)
+
+ define q where "q = (Fun.swap j k id) \<circ> (Fun.swap g h id)"
+
+ hence "(p \<circ> p) = q \<circ> p \<circ> (inv' q)"
+ by (metis final_step One_nat_def Suc_1 comp_id funpow.simps(2) funpow_simps_right(1) ijk p)
+ hence "(p \<circ> p) \<circ> (inv' p) = q \<circ> p \<circ> (inv' q) \<circ> (inv' p)" by simp
+ hence 1: "p = q \<circ> p \<circ> (inv' q) \<circ> (inv' p)"
+ using p cycle_permutes[OF cs(1)] o_assoc[of p p "inv' p"]
+ by (simp add: permutation_inverse_works(2))
+
+ have "(Fun.swap j k id) \<circ> (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) \<circ> (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 \<in> carrier (alt_group n)"
+ by (simp add: alt_group_def q_def)
+
+ have 3: "p \<in> carrier (alt_group n)"
+ using alt_group_as_three_cycles generate.incl[OF \<open>p \<in> three_cycles n\<close>] by simp
+
+ from 1 2 3 show "p \<in> 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) \<subseteq> 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
+qed
+
+corollary alt_group_not_solvable:
+ assumes "n \<ge> 5"
+ shows "\<not> solvable (alt_group n)"
+proof (rule ccontr)
+ assume "\<not> \<not> 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)) = { \<one>\<^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) = { \<one>\<^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] \<subseteq> {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] \<in> three_cycles n" by blast
+ hence "cycle_of_list [1 :: nat, 2, 3] \<in> 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 \<open>cycle [1 :: nat, 2, 3]\<close>, of 1] by simp
+ hence "cycle_of_list [1 :: nat, 2, 3] \<noteq> id"
+ by (metis list.map_id list.sel(1) numeral_One numeral_eq_iff semiring_norm(85))
+
+ ultimately show False using singleton by blast
+qed
+
+corollary sym_group_not_solvable:
+ assumes "n \<ge> 5"
+ shows "\<not> 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
+qed
+
+end
\ No newline at end of file
--- a/src/HOL/ROOT Mon Jul 02 16:26:11 2018 +0200
+++ b/src/HOL/ROOT Mon Jul 02 15:43:22 2018 +0100
@@ -288,7 +288,7 @@
session "HOL-Algebra" (main timing) in Algebra = "HOL-Computational_Algebra" +
description {*
- Author: Clemens Ballarin, started 24 September 1999
+ Author: Clemens Ballarin, started 24 September 1999, and many others
The Isabelle Algebraic Library.
*}
@@ -305,6 +305,8 @@
Divisibility (* Rings *)
IntRing (* Ideals and residue classes *)
UnivPoly (* Polynomials *)
+ (* Main theory *)
+ Algebra
document_files "root.bib" "root.tex"
session "HOL-Auth" (timing) in Auth = "HOL-Library" +