# HG changeset patch # User wenzelm # Date 1734299928 -3600 # Node ID b1772698bd78bfc0c7015f5b8b6c7c95f19b4bd4 # Parent ca6b2e49424b16261aa8a321f7e62f22082a43c3 tuned proofs; tuned whitespace; diff -r ca6b2e49424b -r b1772698bd78 src/HOL/Algebra/QuotRing.thy --- a/src/HOL/Algebra/QuotRing.thy Sun Dec 15 21:39:43 2024 +0100 +++ b/src/HOL/Algebra/QuotRing.thy Sun Dec 15 22:58:48 2024 +0100 @@ -21,7 +21,7 @@ assumes "x \ carrier R" "y \ carrier R" shows "[mod I:] (I +> x) \ (I +> y) = I +> (x \ y)" proof - - have 1: "z \ I +> x \ y" + have 1: "z \ I +> x \ y" if x'rcos: "x' \ I +> x" and y'rcos: "y' \ I +> y" and zrcos: "z \ I +> x' \ y'" for z x' y' proof - from that @@ -44,6 +44,7 @@ by (auto simp: intro!: 1 2) qed + subsection \Quotient Ring Definition\ definition FactRing :: "[('a,'b) ring_scheme, 'a set] \ ('a set) ring" @@ -54,15 +55,15 @@ lemmas FactRing_simps = FactRing_def A_RCOSETS_defs a_r_coset_def[symmetric] + subsection \Factorization over General Ideals\ text \The quotient is a ring\ lemma (in ideal) quotient_is_ring: "ring (R Quot I)" proof (rule ringI) show "abelian_group (R Quot I)" - apply (rule comm_group_abelian_groupI) - apply (simp add: FactRing_def a_factorgroup_is_comm_group[unfolded A_FactGroup_def']) - done + by (rule comm_group_abelian_groupI) + (simp add: FactRing_def a_factorgroup_is_comm_group[unfolded A_FactGroup_def']) show "Group.monoid (R Quot I)" by (rule monoidI) (auto simp add: FactRing_simps rcoset_mult_add m_assoc) @@ -101,8 +102,8 @@ apply (rule rcos_ring_hom_ring) apply (rule is_cring) apply (rule quotient_is_cring) - apply (rule is_cring) - done + apply (rule is_cring) + done qed @@ -122,7 +123,8 @@ apply (intro domain.intro quotient_is_cring is_cring domain_axioms.intro) apply (metis "1" FactRing_def monoid.simps(2) ring.simps(1)) apply (simp add: FactRing_simps) - by (metis "2" rcoset_mult_add) + apply (metis "2" rcoset_mult_add) + done qed text \Generating right cosets of a prime ideal is a homomorphism @@ -145,7 +147,7 @@ assume "\\<^bsub>R Quot I\<^esub> = \\<^bsub>R Quot I\<^esub>" then have II1: "I = I +> \" by (simp add: FactRing_def) then have "I = carrier R" - using a_rcos_self one_imp_carrier by blast + using a_rcos_self one_imp_carrier by blast with I_notcarr show False by simp qed have 2: "\y\carrier R. I +> a \ y = I +> \" if IanI: "I +> a \ I" and acarr: "a \ carrier R" for a @@ -163,7 +165,7 @@ by (simp add: acarr xI) with xI show "\xa\carrier R. \k\I. x = xa \ a \ k" by fast qed - have JnI: "J \ I" + have JnI: "J \ I" proof - have "a \ I" using IanI a_rcos_const by blast @@ -182,7 +184,7 @@ \ \Calculating an inverse for \<^term>\a\\ from one_closed[folded Jcarr] obtain r i where rcarr: "r \ carrier R" - and iI: "i \ I" and one: "\ = r \ a \ i" + and iI: "i \ I" and one: "\ = r \ a \ i" by (auto simp add: J_def r_coset_def set_add_defs) from one and rcarr and acarr and iI[THEN a_Hcarr] @@ -212,7 +214,7 @@ lemma (in ring_hom_ring) trivial_ker_imp_inj: assumes "a_kernel R S h = { \ }" shows "inj_on h (carrier R)" - using group_hom.trivial_ker_imp_inj[OF a_group_hom] assms a_kernel_def[of R S h] by simp + using group_hom.trivial_ker_imp_inj[OF a_group_hom] assms a_kernel_def[of R S h] by simp (* NEW ========================================================================== *) lemma (in ring_hom_ring) inj_iff_trivial_ker: @@ -222,7 +224,7 @@ (* NEW ========================================================================== *) corollary ring_hom_in_hom: assumes "h \ ring_hom R S" shows "h \ hom R S" and "h \ hom (add_monoid R) (add_monoid S)" - using assms unfolding ring_hom_def hom_def by auto + using assms unfolding ring_hom_def hom_def by auto (* NEW ========================================================================== *) corollary set_add_hom: @@ -248,34 +250,36 @@ unfolding a_kernel_def[of R S h] set_add_def by simp+ lemma (in ring_hom_ring) non_trivial_field_hom_imp_inj: - assumes "field R" - shows "h ` (carrier R) \ { \\<^bsub>S\<^esub> } \ inj_on h (carrier R)" + assumes R: "field R" + and h: "h ` (carrier R) \ { \\<^bsub>S\<^esub> }" + shows "inj_on h (carrier R)" proof - - assume "h ` (carrier R) \ { \\<^bsub>S\<^esub> }" - hence "a_kernel R S h \ carrier R" + from h have "a_kernel R S h \ carrier R" using trivial_hom_iff by linarith hence "a_kernel R S h = { \ }" - using field.all_ideals[OF assms] kernel_is_ideal by blast + using field.all_ideals[OF R] kernel_is_ideal by blast thus "inj_on h (carrier R)" using trivial_ker_imp_inj by blast qed + lemma "field R \ cring R" using fieldE(1) by blast lemma non_trivial_field_hom_is_inj: - assumes "h \ ring_hom R S" and "field R" and "field S" shows "inj_on h (carrier R)" + assumes "h \ ring_hom R S" and "field R" and "field S" + shows "inj_on h (carrier R)" proof - interpret ring_hom_cring R S h using assms(1) ring_hom_cring.intro[OF assms(2-3)[THEN fieldE(1)]] unfolding symmetric[OF ring_hom_cring_axioms_def] by simp have "h \\<^bsub>R\<^esub> = \\<^bsub>S\<^esub>" and "\\<^bsub>S\<^esub> \ \\<^bsub>S\<^esub>" - using domain.one_not_zero[OF field.axioms(1)[OF assms(3)]] by auto + using domain.one_not_zero[OF field.axioms(1)[OF assms(3)]] by auto hence "h ` (carrier R) \ { \\<^bsub>S\<^esub> }" using ring.kernel_zero ring.trivial_hom_iff by fastforce thus ?thesis using ring.non_trivial_field_hom_imp_inj[OF assms(2)] by simp -qed +qed lemma (in ring_hom_ring) img_is_add_subgroup: assumes "subgroup H (add_monoid R)" @@ -294,38 +298,36 @@ lemma (in ring) ring_ideal_imp_quot_ideal: assumes "ideal I R" - shows "ideal J R \ ideal ((+>) I ` J) (R Quot I)" -proof - - assume A: "ideal J R" show "ideal (((+>) I) ` J) (R Quot I)" - proof (rule idealI) - show "ring (R Quot I)" - by (simp add: assms(1) ideal.quotient_is_ring) - next - have "subgroup J (add_monoid R)" - by (simp add: additive_subgroup.a_subgroup A ideal.axioms(1)) - moreover have "((+>) I) \ ring_hom R (R Quot I)" - by (simp add: assms(1) ideal.rcos_ring_hom) - ultimately show "subgroup ((+>) I ` J) (add_monoid (R Quot I))" - using assms(1) ideal.rcos_ring_hom_ring ring_hom_ring.img_is_add_subgroup by blast - next - fix a x assume "a \ (+>) I ` J" "x \ carrier (R Quot I)" - then obtain i j where i: "i \ carrier R" "x = I +> i" - and j: "j \ J" "a = I +> j" - unfolding FactRing_def using A_RCOSETS_def'[of R I] by auto - hence "a \\<^bsub>R Quot I\<^esub> x = [mod I:] (I +> j) \ (I +> i)" - unfolding FactRing_def by simp - hence "a \\<^bsub>R Quot I\<^esub> x = I +> (j \ i)" - using ideal.rcoset_mult_add[OF assms(1), of j i] i(1) j(1) A ideal.Icarr by force - thus "a \\<^bsub>R Quot I\<^esub> x \ (+>) I ` J" - using A i(1) j(1) by (simp add: ideal.I_r_closed) - - have "x \\<^bsub>R Quot I\<^esub> a = [mod I:] (I +> i) \ (I +> j)" - unfolding FactRing_def i j by simp - hence "x \\<^bsub>R Quot I\<^esub> a = I +> (i \ j)" - using ideal.rcoset_mult_add[OF assms(1), of i j] i(1) j(1) A ideal.Icarr by force - thus "x \\<^bsub>R Quot I\<^esub> a \ (+>) I ` J" - using A i(1) j(1) by (simp add: ideal.I_l_closed) - qed + and A: "ideal J R" + shows "ideal ((+>) I ` J) (R Quot I)" +proof (rule idealI) + show "ring (R Quot I)" + by (simp add: assms(1) ideal.quotient_is_ring) +next + have "subgroup J (add_monoid R)" + by (simp add: additive_subgroup.a_subgroup A ideal.axioms(1)) + moreover have "((+>) I) \ ring_hom R (R Quot I)" + by (simp add: assms(1) ideal.rcos_ring_hom) + ultimately show "subgroup ((+>) I ` J) (add_monoid (R Quot I))" + using assms(1) ideal.rcos_ring_hom_ring ring_hom_ring.img_is_add_subgroup by blast +next + fix a x assume "a \ (+>) I ` J" "x \ carrier (R Quot I)" + then obtain i j where i: "i \ carrier R" "x = I +> i" + and j: "j \ J" "a = I +> j" + unfolding FactRing_def using A_RCOSETS_def'[of R I] by auto + hence "a \\<^bsub>R Quot I\<^esub> x = [mod I:] (I +> j) \ (I +> i)" + unfolding FactRing_def by simp + hence "a \\<^bsub>R Quot I\<^esub> x = I +> (j \ i)" + using ideal.rcoset_mult_add[OF assms(1), of j i] i(1) j(1) A ideal.Icarr by force + thus "a \\<^bsub>R Quot I\<^esub> x \ (+>) I ` J" + using A i(1) j(1) by (simp add: ideal.I_r_closed) + + have "x \\<^bsub>R Quot I\<^esub> a = [mod I:] (I +> i) \ (I +> j)" + unfolding FactRing_def i j by simp + hence "x \\<^bsub>R Quot I\<^esub> a = I +> (i \ j)" + using ideal.rcoset_mult_add[OF assms(1), of i j] i(1) j(1) A ideal.Icarr by force + thus "x \\<^bsub>R Quot I\<^esub> a \ (+>) I ` J" + using A i(1) j(1) by (simp add: ideal.I_l_closed) qed lemma (in ring_hom_ring) ideal_vimage: @@ -333,7 +335,6 @@ shows "ideal { r \ carrier R. h r \ I } R" (* or (carrier R) \ (h -` I) *) proof show "{ r \ carrier R. h r \ I } \ carrier (add_monoid R)" by auto -next show "\\<^bsub>add_monoid R\<^esub> \ { r \ carrier R. h r \ I }" by (simp add: additive_subgroup.zero_closed assms ideal.axioms(1)) next @@ -369,40 +370,38 @@ lemma (in ring) canonical_proj_vimage_in_carrier: assumes "ideal I R" - shows "J \ carrier (R Quot I) \ \ J \ carrier R" -proof - - assume A: "J \ carrier (R Quot I)" show "\ J \ carrier R" - proof - fix j assume j: "j \ \ J" - then obtain j' where j': "j' \ J" "j \ j'" by blast - then obtain r where r: "r \ carrier R" "j' = I +> r" - using A j' unfolding FactRing_def using A_RCOSETS_def'[of R I] by auto - thus "j \ carrier R" using j' assms - by (meson a_r_coset_subset_G additive_subgroup.a_subset contra_subsetD ideal.axioms(1)) - qed + and A: "J \ carrier (R Quot I)" + shows "\ J \ carrier R" +proof + fix j assume j: "j \ \ J" + then obtain j' where j': "j' \ J" "j \ j'" + by blast + then obtain r where r: "r \ carrier R" "j' = I +> r" + using A j' unfolding FactRing_def using A_RCOSETS_def'[of R I] by auto + thus "j \ carrier R" + using j' assms + by (meson a_r_coset_subset_G additive_subgroup.a_subset contra_subsetD ideal.axioms(1)) qed lemma (in ring) canonical_proj_vimage_mem_iff: assumes "ideal I R" "J \ carrier (R Quot I)" - shows "\a. a \ carrier R \ (a \ (\ J)) = (I +> a \ J)" -proof - - fix a assume a: "a \ carrier R" show "(a \ (\ J)) = (I +> a \ J)" - proof - assume "a \ \ J" - then obtain j where j: "j \ J" "a \ j" by blast - then obtain r where r: "r \ carrier R" "j = I +> r" - using assms j unfolding FactRing_def using A_RCOSETS_def'[of R I] by auto - hence "I +> r = I +> a" - using add.repr_independence[of a I r] j r - by (metis a_r_coset_def additive_subgroup.a_subgroup assms(1) ideal.axioms(1)) - thus "I +> a \ J" using r j by simp - next - assume "I +> a \ J" - hence "\ \ a \ I +> a" - using additive_subgroup.zero_closed[OF ideal.axioms(1)[OF assms(1)]] - a_r_coset_def'[of R I a] by blast - thus "a \ \ J" using a \I +> a \ J\ by auto - qed + and a: "a \ carrier R" + shows "(a \ \ J) = (I +> a \ J)" +proof + assume "a \ \ J" + then obtain j where j: "j \ J" "a \ j" by blast + then obtain r where r: "r \ carrier R" "j = I +> r" + using assms j unfolding FactRing_def using A_RCOSETS_def'[of R I] by auto + hence "I +> r = I +> a" + using add.repr_independence[of a I r] j r + by (metis a_r_coset_def additive_subgroup.a_subgroup assms(1) ideal.axioms(1)) + thus "I +> a \ J" using r j by simp +next + assume "I +> a \ J" + hence "\ \ a \ I +> a" + using additive_subgroup.zero_closed[OF ideal.axioms(1)[OF assms(1)]] + a_r_coset_def'[of R I a] by blast + thus "a \ \ J" using a \I +> a \ J\ by auto qed corollary (in ring) quot_ideal_imp_ring_ideal: @@ -422,9 +421,9 @@ assumes "ideal I R" "ideal J R" shows "(I \ J) = (J = (\ j \ J. I +> j))" proof - assume A: "J = (\ j \ J. I +> j)" hence "I +> \ \ J" + assume "J = (\ j \ J. I +> j)" hence "I +> \ \ J" using additive_subgroup.zero_closed[OF ideal.axioms(1)[OF assms(2)]] by blast - thus "I \ J" using additive_subgroup.a_subset[OF ideal.axioms(1)[OF assms(1)]] by simp + thus "I \ J" using additive_subgroup.a_subset[OF ideal.axioms(1)[OF assms(1)]] by simp next assume A: "I \ J" show "J = (\j\J. I +> j)" proof @@ -435,9 +434,8 @@ hence "\ \ j \ I +> j" using a_r_coset_def'[of R I j] by blast thus "j \ (\j\J. I +> j)" - using assms(2) j additive_subgroup.a_Hcarr ideal.axioms(1) by fastforce + using assms(2) j additive_subgroup.a_Hcarr ideal.axioms(1) by fastforce qed - next show "(\ j \ J. I +> j) \ J" proof fix x assume "x \ (\ j \ J. I +> j)" @@ -456,10 +454,8 @@ proof (rule bij_betw_byWitness[where ?f' = "\X. \ X"]) show "\J \ { J. ideal J R \ I \ J }. (\X. \ X) ((+>) I ` J) = J" using assms ideal_incl_iff by blast -next show "(\J. (+>) I ` J) ` { J. ideal J R \ I \ J } \ { J. ideal J (R Quot I) }" using assms ring_ideal_imp_quot_ideal by auto -next show "(\X. \ X) ` { J. ideal J (R Quot I) } \ { J. ideal J R \ I \ J }" proof fix J assume "J \ ((\X. \ X) ` { J. ideal J (R Quot I) })" @@ -470,7 +466,6 @@ using additive_subgroup.zero_closed[OF ideal.axioms(1)[OF J'(1)]] unfolding FactRing_def by simp ultimately show "J \ { J. ideal J R \ I \ J }" using J'(2) by auto qed -next show "\J' \ { J. ideal J (R Quot I) }. ((+>) I ` (\ J')) = J'" proof fix J' assume "J' \ { J. ideal J (R Quot I) }" @@ -494,11 +489,12 @@ lemma (in cring) quot_domain_imp_primeideal: assumes "ideal P R" - shows "domain (R Quot P) \ primeideal P R" + and A: "domain (R Quot P)" + shows "primeideal P R" proof - - assume A: "domain (R Quot P)" show "primeideal P R" + show "primeideal P R" proof (rule primeidealI) - show "ideal P R" using assms . + show "ideal P R" using assms(1) . show "cring R" using is_cring . next show "carrier R \ P" @@ -598,7 +594,7 @@ assumes "f \ ring_iso R S" and "\r. r \ carrier R \ f r = g r" shows "g \ ring_iso R S" proof - have hom: "g \ ring_hom R S" - using ring_hom_restrict assms unfolding ring_iso_def by auto + using ring_hom_restrict assms unfolding ring_iso_def by auto have "bij_betw g (carrier R) (carrier S)" using bij_betw_cong[of "carrier R" f g] ring_iso_memE(5)[OF assms(1)] assms(2) by simp thus ?thesis @@ -685,7 +681,7 @@ note m_comm interpret h_img?: ring ?h_img using ring_iso_imp_img_ring[OF assms] . - show ?thesis + show ?thesis proof (unfold_locales) fix x y assume "x \ carrier ?h_img" "y \ carrier ?h_img" then obtain r1 r2 @@ -710,7 +706,7 @@ note aux = m_closed integral one_not_zero one_closed zero_closed interpret h_img?: cring ?h_img using ring_iso_imp_img_cring[OF assms] . - show ?thesis + show ?thesis proof (unfold_locales) have "\\<^bsub>?h_img\<^esub> = \\<^bsub>?h_img\<^esub> \ h \ = h \" using ring_iso_memE(4)[OF assms] by simp @@ -767,7 +763,7 @@ then obtain r where r: "r \ carrier R" "s = h r" using assms image_iff[where ?f = h and ?A = "carrier R"] unfolding ring_iso_def bij_betw_def by auto - hence "r \ \" using s(2) by auto + hence "r \ \" using s(2) by auto hence inv_r: "inv r \ carrier R" "inv r \ \" "r \ inv r = \" "inv r \ r = \" using field_Units r(1) by auto have "h (inv r) \\<^bsub>S\<^esub> h r = h \" and "h r \\<^bsub>S\<^esub> h (inv r) = h \" @@ -779,21 +775,21 @@ qed lemma ring_iso_same_card: "R \ S \ card (carrier R) = card (carrier S)" - using bij_betw_same_card unfolding is_ring_iso_def ring_iso_def by auto + using bij_betw_same_card unfolding is_ring_iso_def ring_iso_def by auto (* ========================================================================== *) lemma ring_iso_set_refl: "id \ ring_iso R R" by (rule ring_iso_memI) (auto) corollary ring_iso_refl: "R \ R" - using is_ring_iso_def ring_iso_set_refl by auto + using is_ring_iso_def ring_iso_set_refl by auto lemma ring_iso_set_trans: "\ f \ ring_iso R S; g \ ring_iso S Q \ \ (g \ f) \ ring_iso R Q" - unfolding ring_iso_def using bij_betw_trans ring_hom_trans by fastforce + unfolding ring_iso_def using bij_betw_trans ring_hom_trans by fastforce corollary ring_iso_trans: "\ R \ S; S \ Q \ \ R \ Q" - using ring_iso_set_trans unfolding is_ring_iso_def by blast + using ring_iso_set_trans unfolding is_ring_iso_def by blast lemma ring_iso_set_sym: assumes "ring R" and h: "h \ ring_iso R S" @@ -801,7 +797,7 @@ proof - have h_hom: "h \ ring_hom R S" and h_surj: "h ` (carrier R) = (carrier S)" - and h_inj: "\ x1 x2. \ x1 \ carrier R; x2 \ carrier R \ \ h x1 = h x2 \ x1 = x2" + and h_inj: "\x1 x2. \ x1 \ carrier R; x2 \ carrier R \ \ h x1 = h x2 \ x1 = x2" using h unfolding ring_iso_def bij_betw_def inj_on_def by auto have h_inv_bij: "bij_betw (inv_into (carrier R) h) (carrier S) (carrier R)" @@ -819,28 +815,26 @@ corollary ring_iso_sym: assumes "ring R" shows "R \ S \ S \ R" - using assms ring_iso_set_sym unfolding is_ring_iso_def by auto + using assms ring_iso_set_sym unfolding is_ring_iso_def by auto lemma (in ring_hom_ring) the_elem_simp [simp]: - "\x. x \ carrier R \ the_elem (h ` ((a_kernel R S h) +> x)) = h x" + assumes x: "x \ carrier R" + shows "the_elem (h ` ((a_kernel R S h) +> x)) = h x" proof - - fix x assume x: "x \ carrier R" - hence "h x \ h ` ((a_kernel R S h) +> x)" + from x have "h x \ h ` ((a_kernel R S h) +> x)" using homeq_imp_rcos by blast thus "the_elem (h ` ((a_kernel R S h) +> x)) = h x" by (metis (no_types, lifting) x empty_iff homeq_imp_rcos rcos_imp_homeq the_elem_image_unique) qed lemma (in ring_hom_ring) the_elem_inj: - "\X Y. \ X \ carrier (R Quot (a_kernel R S h)); Y \ carrier (R Quot (a_kernel R S h)) \ \ - the_elem (h ` X) = the_elem (h ` Y) \ X = Y" + assumes "X \ carrier (R Quot (a_kernel R S h))" + and "Y \ carrier (R Quot (a_kernel R S h))" + and Eq: "the_elem (h ` X) = the_elem (h ` Y)" + shows "X = Y" proof - - fix X Y - assume "X \ carrier (R Quot (a_kernel R S h))" - and "Y \ carrier (R Quot (a_kernel R S h))" - and Eq: "the_elem (h ` X) = the_elem (h ` Y)" - then obtain x y where x: "x \ carrier R" "X = (a_kernel R S h) +> x" - and y: "y \ carrier R" "Y = (a_kernel R S h) +> y" + from assms obtain x y where x: "x \ carrier R" "X = (a_kernel R S h) +> x" + and y: "y \ carrier R" "Y = (a_kernel R S h) +> y" unfolding FactRing_def A_RCOSETS_def' by auto hence "h x = h y" using Eq by simp hence "x \ y \ (a_kernel R S h)" @@ -852,22 +846,18 @@ qed lemma (in ring_hom_ring) quot_mem: - "\X. X \ carrier (R Quot (a_kernel R S h)) \ \x \ carrier R. X = (a_kernel R S h) +> x" -proof - - fix X assume "X \ carrier (R Quot (a_kernel R S h))" - thus "\x \ carrier R. X = (a_kernel R S h) +> x" - unfolding FactRing_simps by (simp add: a_r_coset_def) -qed + "X \ carrier (R Quot (a_kernel R S h)) \ \x \ carrier R. X = (a_kernel R S h) +> x" + unfolding FactRing_simps by (simp add: a_r_coset_def) lemma (in ring_hom_ring) the_elem_wf: - "\X. X \ carrier (R Quot (a_kernel R S h)) \ \y \ carrier S. (h ` X) = { y }" + assumes "X \ carrier (R Quot (a_kernel R S h))" + shows "\y \ carrier S. (h ` X) = { y }" proof - - fix X assume "X \ carrier (R Quot (a_kernel R S h))" - then obtain x where x: "x \ carrier R" and X: "X = (a_kernel R S h) +> x" + from assms obtain x where x: "x \ carrier R" and X: "X = (a_kernel R S h) +> x" using quot_mem by blast - hence "\x'. x' \ X \ h x' = h x" + have "h x' = h x" if "x' \ X" for x' proof - - fix x' assume "x' \ X" hence "x' \ (a_kernel R S h) +> x" using X by simp + from X that have "x' \ (a_kernel R S h) +> x" by simp then obtain k where k: "k \ a_kernel R S h" "x' = k \ x" by (metis R.add.inv_closed R.add.m_assoc R.l_neg R.r_zero abelian_subgroup.a_elemrcos_carrier @@ -886,15 +876,15 @@ qed corollary (in ring_hom_ring) the_elem_wf': - "\X. X \ carrier (R Quot (a_kernel R S h)) \ \r \ carrier R. (h ` X) = { h r }" - using the_elem_wf by (metis quot_mem the_elem_eq the_elem_simp) + "X \ carrier (R Quot (a_kernel R S h)) \ \r \ carrier R. (h ` X) = { h r }" + using the_elem_wf by (metis quot_mem the_elem_eq the_elem_simp) lemma (in ring_hom_ring) the_elem_hom: "(\X. the_elem (h ` X)) \ ring_hom (R Quot (a_kernel R S h)) S" proof (rule ring_hom_memI) show "\x. x \ carrier (R Quot a_kernel R S h) \ the_elem (h ` x) \ carrier S" using the_elem_wf by fastforce - + show "the_elem (h ` \\<^bsub>R Quot a_kernel R S h\<^esub>) = \\<^bsub>S\<^esub>" unfolding FactRing_def using the_elem_simp[of "\\<^bsub>R\<^esub>"] by simp @@ -917,11 +907,10 @@ qed lemma (in ring_hom_ring) the_elem_surj: - "(\X. (the_elem (h ` X))) ` carrier (R Quot (a_kernel R S h)) = (h ` (carrier R))" + "(\X. (the_elem (h ` X))) ` carrier (R Quot (a_kernel R S h)) = (h ` (carrier R))" proof show "(\X. the_elem (h ` X)) ` carrier (R Quot a_kernel R S h) \ h ` carrier R" using the_elem_wf' by fastforce -next show "h ` carrier R \ (\X. the_elem (h ` X)) ` carrier (R Quot a_kernel R S h)" proof fix y assume "y \ h ` carrier R" @@ -937,16 +926,13 @@ proposition (in ring_hom_ring) FactRing_iso_set_aux: "(\X. the_elem (h ` X)) \ ring_iso (R Quot (a_kernel R S h)) (S \ carrier := h ` (carrier R) \)" proof - - have "bij_betw (\X. the_elem (h ` X)) (carrier (R Quot a_kernel R S h)) (h ` (carrier R))" + have *: "bij_betw (\X. the_elem (h ` X)) (carrier (R Quot a_kernel R S h)) (h ` (carrier R))" unfolding bij_betw_def inj_on_def using the_elem_surj the_elem_inj by simp - - moreover have "(\X. the_elem (h ` X)): carrier (R Quot (a_kernel R S h)) \ h ` (carrier R)" using the_elem_wf' by fastforce hence "(\X. the_elem (h ` X)) \ ring_hom (R Quot (a_kernel R S h)) (S \ carrier := h ` (carrier R) \)" using the_elem_hom the_elem_wf' unfolding ring_hom_def by simp - - ultimately show ?thesis unfolding ring_iso_def using the_elem_hom by simp + with * show ?thesis unfolding ring_iso_def using the_elem_hom by simp qed theorem (in ring_hom_ring) FactRing_iso_set: @@ -991,7 +977,7 @@ using hom_one by force hence "?the_elem \\<^bsub>(R Quot (a_kernel R S h))\<^esub> = \\<^bsub>S\<^esub>" using the_elem_wf[of "(a_kernel R S h) +> \"] by (simp add: FactRing_def) - + moreover have "\\<^bsub>S\<^esub> \ (h ` (a_kernel R S h))" using a_kernel_def'[of R S h] hom_zero by force @@ -999,7 +985,7 @@ by (simp add: FactRing_def) hence "?the_elem \\<^bsub>(R Quot (a_kernel R S h))\<^esub> = \\<^bsub>S\<^esub>" using the_elem_wf[OF ring.ring_simprules(2)[OF FactRing_is_ring]] - by (metis singletonD the_elem_eq) + by (metis singletonD the_elem_eq) ultimately have "ring ((S \ carrier := h ` (carrier R) \) \ one := \\<^bsub>S\<^esub>, zero := \\<^bsub>S\<^esub> \)" @@ -1015,8 +1001,7 @@ interpret ring "S \ carrier := h ` (carrier R) \" using img_is_ring . show ?thesis - apply unfold_locales - using assms unfolding cring_def comm_monoid_def comm_monoid_axioms_def by auto + by unfold_locales (use assms in \auto simp: cring_def comm_monoid_def comm_monoid_axioms_def\) qed lemma (in ring_hom_ring) img_is_domain: @@ -1028,21 +1013,21 @@ show ?thesis apply unfold_locales using assms unfolding domain_def domain_axioms_def apply auto - using hom_closed by blast + using hom_closed by blast qed proposition (in ring_hom_ring) primeideal_vimage: - assumes "cring R" - shows "primeideal P S \ primeideal { r \ carrier R. h r \ P } R" + assumes R: "cring R" + and A: "primeideal P S" + shows "primeideal { r \ carrier R. h r \ P } R" proof - - assume A: "primeideal P S" - hence is_ideal: "ideal P S" unfolding primeideal_def by simp + from A have is_ideal: "ideal P S" unfolding primeideal_def by simp have "ring_hom_ring R (S Quot P) (((+>\<^bsub>S\<^esub>) P) \ h)" (is "ring_hom_ring ?A ?B ?h") using ring_hom_trans[OF homh, of "(+>\<^bsub>S\<^esub>) P" "S Quot P"] - ideal.rcos_ring_hom_ring[OF is_ideal] assms + ideal.rcos_ring_hom_ring[OF is_ideal] R unfolding ring_hom_ring_def ring_hom_ring_axioms_def cring_def by simp then interpret hom: ring_hom_ring R "S Quot P" "((+>\<^bsub>S\<^esub>) P) \ h" by simp - + have "inj_on (\X. the_elem (?h ` X)) (carrier (R Quot (a_kernel R (S Quot P) ?h)))" using hom.the_elem_inj unfolding inj_on_def by simp moreover @@ -1051,16 +1036,16 @@ have hom': "ring_hom_ring (R Quot (a_kernel R (S Quot P) ?h)) (S Quot P) (\X. the_elem (?h ` X))" using hom.the_elem_hom hom.kernel_is_ideal by (meson hom.ring_hom_ring_axioms ideal.rcos_ring_hom_ring ring_hom_ring_axioms_def ring_hom_ring_def) - + ultimately have "primeideal (a_kernel R (S Quot P) ?h) R" using ring_hom_ring.inj_on_domain[OF hom'] primeideal.quotient_is_domain[OF A] - cring.quot_domain_imp_primeideal[OF assms hom.kernel_is_ideal] by simp - + cring.quot_domain_imp_primeideal[OF R hom.kernel_is_ideal] by simp + moreover have "a_kernel R (S Quot P) ?h = { r \ carrier R. h r \ P }" proof show "a_kernel R (S Quot P) ?h \ { r \ carrier R. h r \ P }" - proof + proof fix r assume "r \ a_kernel R (S Quot P) ?h" hence r: "r \ carrier R" "P +>\<^bsub>S\<^esub> (h r) = P" unfolding a_kernel_def kernel_def FactRing_def by auto