More on Algebra by Paulo and Martin
authorpaulson <lp15@cam.ac.uk>
Sat, 30 Jun 2018 15:44:04 +0100
changeset 68551 b680e74eb6f2
parent 68550 516e81f75957
child 68552 391e89e03eef
More on Algebra by Paulo and Martin
src/HOL/Algebra/Divisibility.thy
src/HOL/Algebra/Group.thy
src/HOL/Algebra/Module.thy
src/HOL/Algebra/Multiplicative_Group.thy
src/HOL/Algebra/QuotRing.thy
src/HOL/Algebra/Ring.thy
src/HOL/Algebra/RingHom.thy
--- a/src/HOL/Algebra/Divisibility.thy	Fri Jun 29 23:04:36 2018 +0200
+++ b/src/HOL/Algebra/Divisibility.thy	Sat Jun 30 15:44:04 2018 +0100
@@ -763,6 +763,27 @@
   apply (metis pp' associated_sym divides_cong_l)
   done
 
+(*by Paulo Emílio de Vilhena*)
+lemma (in comm_monoid_cancel) prime_irreducible:
+  assumes "prime G p"
+  shows "irreducible G p"
+proof (rule irreducibleI)
+  show "p \<notin> Units G"
+    using assms unfolding prime_def by simp
+next
+  fix b assume A: "b \<in> carrier G" "properfactor G b p"
+  then obtain c where c: "c \<in> carrier G" "p = b \<otimes> c"
+    unfolding properfactor_def factor_def by auto
+  hence "p divides c"
+    using A assms unfolding prime_def properfactor_def by auto
+  then obtain b' where b': "b' \<in> carrier G" "c = p \<otimes> b'"
+    unfolding factor_def by auto
+  hence "\<one> = b \<otimes> b'"
+    by (metis A(1) l_cancel m_closed m_lcomm one_closed r_one c)
+  thus "b \<in> Units G"
+    using A(1) Units_one_closed b'(1) unit_factor by presburger
+qed
+
 
 subsection \<open>Factorization and Factorial Monoids\<close>
 
--- a/src/HOL/Algebra/Group.thy	Fri Jun 29 23:04:36 2018 +0200
+++ b/src/HOL/Algebra/Group.thy	Sat Jun 30 15:44:04 2018 +0100
@@ -1438,6 +1438,9 @@
   shows "m_inv (units_of G) x = m_inv G x"
   by (simp add: assms group.inv_equality units_group units_of_carrier units_of_mult units_of_one)
 
+lemma units_of_units [simp] : "Units (units_of G) = Units G"
+  unfolding units_of_def Units_def by force
+
 lemma (in group) surj_const_mult: "a \<in> carrier G \<Longrightarrow> (\<lambda>x. a \<otimes> x) ` carrier G = carrier G"
   apply (auto simp add: image_def)
   by (metis inv_closed inv_solve_left m_closed)
--- a/src/HOL/Algebra/Module.thy	Fri Jun 29 23:04:36 2018 +0200
+++ b/src/HOL/Algebra/Module.thy	Sat Jun 30 15:44:04 2018 +0100
@@ -76,87 +76,95 @@
       "!!a x y. [| a \<in> carrier R; x \<in> carrier M; y \<in> carrier M |] ==>
       (a \<odot>\<^bsub>M\<^esub> x) \<otimes>\<^bsub>M\<^esub> y = a \<odot>\<^bsub>M\<^esub> (x \<otimes>\<^bsub>M\<^esub> y)"
   shows "algebra R M"
-apply intro_locales
-apply (rule cring.axioms ring.axioms abelian_group.axioms comm_monoid.axioms assms)+
-apply (rule module_axioms.intro)
- apply (simp add: smult_closed)
- apply (simp add: smult_l_distr)
- apply (simp add: smult_r_distr)
- apply (simp add: smult_assoc1)
- apply (simp add: smult_one)
-apply (rule cring.axioms ring.axioms abelian_group.axioms comm_monoid.axioms assms)+
-apply (rule algebra_axioms.intro)
- apply (simp add: smult_assoc2)
-done
+  apply intro_locales
+             apply (rule cring.axioms ring.axioms abelian_group.axioms comm_monoid.axioms assms)+
+      apply (rule module_axioms.intro)
+          apply (simp add: smult_closed)
+         apply (simp add: smult_l_distr)
+        apply (simp add: smult_r_distr)
+       apply (simp add: smult_assoc1)
+      apply (simp add: smult_one)
+     apply (rule cring.axioms ring.axioms abelian_group.axioms comm_monoid.axioms assms)+
+  apply (rule algebra_axioms.intro)
+  apply (simp add: smult_assoc2)
+  done
 
-lemma (in algebra) R_cring:
-  "cring R"
-  ..
+lemma (in algebra) R_cring: "cring R" ..
 
-lemma (in algebra) M_cring:
-  "cring M"
-  ..
+lemma (in algebra) M_cring: "cring M" ..
 
-lemma (in algebra) module:
-  "module R M"
-  by (auto intro: moduleI R_cring is_abelian_group
-    smult_l_distr smult_r_distr smult_assoc1)
+lemma (in algebra) module: "module R M"
+  by (auto intro: moduleI R_cring is_abelian_group smult_l_distr smult_r_distr smult_assoc1)
 
 
-subsection \<open>Basic Properties of Algebras\<close>
+subsection \<open>Basic Properties of Modules\<close>
 
-lemma (in algebra) smult_l_null [simp]:
-  "x \<in> carrier M ==> \<zero> \<odot>\<^bsub>M\<^esub> x = \<zero>\<^bsub>M\<^esub>"
-proof -
-  assume M: "x \<in> carrier M"
+lemma (in module) smult_l_null [simp]:
+ "x \<in> carrier M ==> \<zero> \<odot>\<^bsub>M\<^esub> x = \<zero>\<^bsub>M\<^esub>"
+proof-
+  assume M : "x \<in> carrier M"
   note facts = M smult_closed [OF R.zero_closed]
-  from facts have "\<zero> \<odot>\<^bsub>M\<^esub> x = (\<zero> \<odot>\<^bsub>M\<^esub> x \<oplus>\<^bsub>M\<^esub> \<zero> \<odot>\<^bsub>M\<^esub> x) \<oplus>\<^bsub>M\<^esub> \<ominus>\<^bsub>M\<^esub> (\<zero> \<odot>\<^bsub>M\<^esub> x)" by algebra
-  also from M have "... = (\<zero> \<oplus> \<zero>) \<odot>\<^bsub>M\<^esub> x \<oplus>\<^bsub>M\<^esub> \<ominus>\<^bsub>M\<^esub> (\<zero> \<odot>\<^bsub>M\<^esub> x)"
-    by (simp add: smult_l_distr del: R.l_zero R.r_zero)
-  also from facts have "... = \<zero>\<^bsub>M\<^esub>" apply algebra apply algebra done
-  finally show ?thesis .
+  from facts have "\<zero> \<odot>\<^bsub>M\<^esub> x = (\<zero> \<oplus> \<zero>) \<odot>\<^bsub>M\<^esub> x "
+    using smult_l_distr by auto
+  also have "... = \<zero> \<odot>\<^bsub>M\<^esub> x \<oplus>\<^bsub>M\<^esub> \<zero> \<odot>\<^bsub>M\<^esub> x"
+    using smult_l_distr[of \<zero> \<zero> x] facts by auto
+  finally show "\<zero> \<odot>\<^bsub>M\<^esub> x = \<zero>\<^bsub>M\<^esub>" using facts
+    by (metis M.add.r_cancel_one')
 qed
 
-lemma (in algebra) smult_r_null [simp]:
+lemma (in module) smult_r_null [simp]:
   "a \<in> carrier R ==> a \<odot>\<^bsub>M\<^esub> \<zero>\<^bsub>M\<^esub> = \<zero>\<^bsub>M\<^esub>"
 proof -
   assume R: "a \<in> carrier R"
   note facts = R smult_closed
   from facts have "a \<odot>\<^bsub>M\<^esub> \<zero>\<^bsub>M\<^esub> = (a \<odot>\<^bsub>M\<^esub> \<zero>\<^bsub>M\<^esub> \<oplus>\<^bsub>M\<^esub> a \<odot>\<^bsub>M\<^esub> \<zero>\<^bsub>M\<^esub>) \<oplus>\<^bsub>M\<^esub> \<ominus>\<^bsub>M\<^esub> (a \<odot>\<^bsub>M\<^esub> \<zero>\<^bsub>M\<^esub>)"
-    by algebra
+    by (simp add: M.add.inv_solve_right)
   also from R have "... = a \<odot>\<^bsub>M\<^esub> (\<zero>\<^bsub>M\<^esub> \<oplus>\<^bsub>M\<^esub> \<zero>\<^bsub>M\<^esub>) \<oplus>\<^bsub>M\<^esub> \<ominus>\<^bsub>M\<^esub> (a \<odot>\<^bsub>M\<^esub> \<zero>\<^bsub>M\<^esub>)"
     by (simp add: smult_r_distr del: M.l_zero M.r_zero)
-  also from facts have "... = \<zero>\<^bsub>M\<^esub>" by algebra
+  also from facts have "... = \<zero>\<^bsub>M\<^esub>"
+    by (simp add: M.r_neg)
   finally show ?thesis .
 qed
 
-lemma (in algebra) smult_l_minus:
-  "[| a \<in> carrier R; x \<in> carrier M |] ==> (\<ominus>a) \<odot>\<^bsub>M\<^esub> x = \<ominus>\<^bsub>M\<^esub> (a \<odot>\<^bsub>M\<^esub> x)"
-proof -
+lemma (in module) smult_l_minus:
+"\<lbrakk> a \<in> carrier R; x \<in> carrier M \<rbrakk> \<Longrightarrow> (\<ominus>a) \<odot>\<^bsub>M\<^esub> x = \<ominus>\<^bsub>M\<^esub> (a \<odot>\<^bsub>M\<^esub> x)"
+proof-
   assume RM: "a \<in> carrier R" "x \<in> carrier M"
   from RM have a_smult: "a \<odot>\<^bsub>M\<^esub> x \<in> carrier M" by simp
   from RM have ma_smult: "\<ominus>a \<odot>\<^bsub>M\<^esub> x \<in> carrier M" by simp
   note facts = RM a_smult ma_smult
   from facts have "(\<ominus>a) \<odot>\<^bsub>M\<^esub> x = (\<ominus>a \<odot>\<^bsub>M\<^esub> x \<oplus>\<^bsub>M\<^esub> a \<odot>\<^bsub>M\<^esub> x) \<oplus>\<^bsub>M\<^esub> \<ominus>\<^bsub>M\<^esub>(a \<odot>\<^bsub>M\<^esub> x)"
-    by algebra
+    by (simp add: M.add.inv_solve_right)
   also from RM have "... = (\<ominus>a \<oplus> a) \<odot>\<^bsub>M\<^esub> x \<oplus>\<^bsub>M\<^esub> \<ominus>\<^bsub>M\<^esub>(a \<odot>\<^bsub>M\<^esub> x)"
     by (simp add: smult_l_distr)
   also from facts smult_l_null have "... = \<ominus>\<^bsub>M\<^esub>(a \<odot>\<^bsub>M\<^esub> x)"
-    apply algebra apply algebra done
+    by (simp add: R.l_neg)
   finally show ?thesis .
 qed
 
-lemma (in algebra) smult_r_minus:
+lemma (in module) smult_r_minus:
   "[| a \<in> carrier R; x \<in> carrier M |] ==> a \<odot>\<^bsub>M\<^esub> (\<ominus>\<^bsub>M\<^esub>x) = \<ominus>\<^bsub>M\<^esub> (a \<odot>\<^bsub>M\<^esub> x)"
 proof -
   assume RM: "a \<in> carrier R" "x \<in> carrier M"
   note facts = RM smult_closed
   from facts have "a \<odot>\<^bsub>M\<^esub> (\<ominus>\<^bsub>M\<^esub>x) = (a \<odot>\<^bsub>M\<^esub> \<ominus>\<^bsub>M\<^esub>x \<oplus>\<^bsub>M\<^esub> a \<odot>\<^bsub>M\<^esub> x) \<oplus>\<^bsub>M\<^esub> \<ominus>\<^bsub>M\<^esub>(a \<odot>\<^bsub>M\<^esub> x)"
-    by algebra
+    by (simp add: M.add.inv_solve_right)
   also from RM have "... = a \<odot>\<^bsub>M\<^esub> (\<ominus>\<^bsub>M\<^esub>x \<oplus>\<^bsub>M\<^esub> x) \<oplus>\<^bsub>M\<^esub> \<ominus>\<^bsub>M\<^esub>(a \<odot>\<^bsub>M\<^esub> x)"
     by (simp add: smult_r_distr)
-  also from facts smult_r_null have "... = \<ominus>\<^bsub>M\<^esub>(a \<odot>\<^bsub>M\<^esub> x)" by algebra
+  also from facts smult_l_null have "... = \<ominus>\<^bsub>M\<^esub>(a \<odot>\<^bsub>M\<^esub> x)"
+    by (metis M.add.inv_closed M.add.inv_solve_right M.l_neg R.zero_closed r_null smult_assoc1)
   finally show ?thesis .
 qed
 
+lemma (in module) finsum_smult_ldistr:
+  "\<lbrakk> finite A; a \<in> carrier R; f: A \<rightarrow> carrier M \<rbrakk> \<Longrightarrow>
+     a \<odot>\<^bsub>M\<^esub> (\<Oplus>\<^bsub>M\<^esub> i \<in> A. (f i)) = (\<Oplus>\<^bsub>M\<^esub> i \<in> A. ( a \<odot>\<^bsub>M\<^esub> (f i)))"
+proof (induct set: finite)
+  case empty then show ?case
+    by (metis M.finsum_empty M.zero_closed R.zero_closed r_null smult_assoc1 smult_l_null)
+next
+  case (insert x F) then show ?case
+    by (simp add: Pi_def smult_r_distr)
+qed
+
 end
--- a/src/HOL/Algebra/Multiplicative_Group.thy	Fri Jun 29 23:04:36 2018 +0200
+++ b/src/HOL/Algebra/Multiplicative_Group.thy	Sat Jun 30 15:44:04 2018 +0100
@@ -590,7 +590,11 @@
 
 lemmas mult_of_simps = carrier_mult_of mult_mult_of nat_pow_mult_of one_mult_of
 
-context field begin
+context field 
+begin
+
+lemma mult_of_is_Units: "mult_of R = units_of R" 
+  unfolding mult_of_def units_of_def using field_Units by auto
 
 lemma field_mult_group :
   shows "group (mult_of R)"
--- a/src/HOL/Algebra/QuotRing.thy	Fri Jun 29 23:04:36 2018 +0200
+++ b/src/HOL/Algebra/QuotRing.thy	Sat Jun 30 15:44:04 2018 +0100
@@ -1,5 +1,7 @@
 (*  Title:      HOL/Algebra/QuotRing.thy
     Author:     Stephan Hohe
+    Author:     Paulo Emílio de Vilhena
+
 *)
 
 theory QuotRing
@@ -306,4 +308,831 @@
   qed
 qed
 
+
+lemma (in ring_hom_ring) trivial_hom_iff:
+  "(h ` (carrier R) = { \<zero>\<^bsub>S\<^esub> }) = (a_kernel R S h = carrier R)"
+  using group_hom.trivial_hom_iff[OF a_group_hom] by (simp add: a_kernel_def)
+
+lemma (in ring_hom_ring) trivial_ker_imp_inj:
+  assumes "a_kernel R S h = { \<zero> }"
+  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 
+
+lemma (in ring_hom_ring) non_trivial_field_hom_imp_inj:
+  assumes "field R"
+  shows "h ` (carrier R) \<noteq> { \<zero>\<^bsub>S\<^esub> } \<Longrightarrow> inj_on h (carrier R)"
+proof -
+  assume "h ` (carrier R) \<noteq> { \<zero>\<^bsub>S\<^esub> }"
+  hence "a_kernel R S h \<noteq> carrier R"
+    using trivial_hom_iff by linarith
+  hence "a_kernel R S h = { \<zero> }"
+    using field.all_ideals[OF assms] kernel_is_ideal by blast
+  thus "inj_on h (carrier R)"
+    using trivial_ker_imp_inj by blast
+qed
+
+lemma (in ring_hom_ring) img_is_add_subgroup:
+  assumes "subgroup H (add_monoid R)"
+  shows "subgroup (h ` H) (add_monoid S)"
+proof -
+  have "group ((add_monoid R) \<lparr> carrier := H \<rparr>)"
+    using assms R.add.subgroup_imp_group by blast
+  moreover have "H \<subseteq> carrier R" by (simp add: R.add.subgroupE(1) assms)
+  hence "h \<in> hom ((add_monoid R) \<lparr> carrier := H \<rparr>) (add_monoid S)"
+    unfolding hom_def by (auto simp add: subsetD)
+  ultimately have "subgroup (h ` carrier ((add_monoid R) \<lparr> carrier := H \<rparr>)) (add_monoid S)"
+    using group_hom.img_is_subgroup[of "(add_monoid R) \<lparr> carrier := H \<rparr>" "add_monoid S" h]
+    using a_group_hom group_hom_axioms.intro group_hom_def by blast
+  thus "subgroup (h ` H) (add_monoid S)" by simp
+qed
+
+lemma (in ring) ring_ideal_imp_quot_ideal:
+  assumes "ideal I R"
+  shows "ideal J R \<Longrightarrow> 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) \<in> 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 \<in> (+>) I ` J" "x \<in> carrier (R Quot I)"
+    then obtain i j where i: "i \<in> carrier R" "x = I +> i"
+                      and j: "j \<in> J" "a = I +> j"
+      unfolding FactRing_def using A_RCOSETS_def'[of R I] by auto
+    hence "a \<otimes>\<^bsub>R Quot I\<^esub> x = [mod I:] (I +> j) \<Otimes> (I +> i)"
+      unfolding FactRing_def by simp
+    hence "a \<otimes>\<^bsub>R Quot I\<^esub> x = I +> (j \<otimes> i)"
+      using ideal.rcoset_mult_add[OF assms(1), of j i] i(1) j(1) A ideal.Icarr by force
+    thus "a \<otimes>\<^bsub>R Quot I\<^esub> x \<in> (+>) I ` J"
+      using A i(1) j(1) by (simp add: ideal.I_r_closed)
+  
+    have "x \<otimes>\<^bsub>R Quot I\<^esub> a = [mod I:] (I +> i) \<Otimes> (I +> j)"
+      unfolding FactRing_def i j by simp
+    hence "x \<otimes>\<^bsub>R Quot I\<^esub> a = I +> (i \<otimes> j)"
+      using ideal.rcoset_mult_add[OF assms(1), of i j] i(1) j(1) A ideal.Icarr by force
+    thus "x \<otimes>\<^bsub>R Quot I\<^esub> a \<in> (+>) I ` J"
+      using A i(1) j(1) by (simp add: ideal.I_l_closed)
+  qed
+qed
+
+lemma (in ring_hom_ring) ideal_vimage:
+  assumes "ideal I S"
+  shows "ideal { r \<in> carrier R. h r \<in> I } R" (* or (carrier R) \<inter> (h -` I) *)
+proof
+  show "{ r \<in> carrier R. h r \<in> I } \<subseteq> carrier (add_monoid R)" by auto
+next
+  show "\<one>\<^bsub>add_monoid R\<^esub> \<in> { r \<in> carrier R. h r \<in> I }"
+    by (simp add: additive_subgroup.zero_closed assms ideal.axioms(1))
+next
+  fix a b
+  assume "a \<in> { r \<in> carrier R. h r \<in> I }"
+     and "b \<in> { r \<in> carrier R. h r \<in> I }"
+  hence a: "a \<in> carrier R" "h a \<in> I"
+    and b: "b \<in> carrier R" "h b \<in> I" by auto
+  hence "h (a \<oplus> b) = (h a) \<oplus>\<^bsub>S\<^esub> (h b)" using hom_add by blast
+  moreover have "(h a) \<oplus>\<^bsub>S\<^esub> (h b) \<in> I" using a b assms
+    by (simp add: additive_subgroup.a_closed ideal.axioms(1))
+  ultimately show "a \<otimes>\<^bsub>add_monoid R\<^esub> b \<in> { r \<in> carrier R. h r \<in> I }"
+    using a(1) b (1) by auto
+
+  have "h (\<ominus> a) = \<ominus>\<^bsub>S\<^esub> (h a)" by (simp add: a)
+  moreover have "\<ominus>\<^bsub>S\<^esub> (h a) \<in> I"
+    by (simp add: a(2) additive_subgroup.a_inv_closed assms ideal.axioms(1))
+  ultimately show "inv\<^bsub>add_monoid R\<^esub> a \<in> { r \<in> carrier R. h r \<in> I }"
+    using a by (simp add: a_inv_def)
+next
+  fix a r
+  assume "a \<in> { r \<in> carrier R. h r \<in> I }" and r: "r \<in> carrier R"
+  hence a: "a \<in> carrier R" "h a \<in> I" by auto
+
+  have "h a \<otimes>\<^bsub>S\<^esub> h r \<in> I"
+    using assms a r by (simp add: ideal.I_r_closed)
+  thus "a \<otimes> r \<in> { r \<in> carrier R. h r \<in> I }" by (simp add: a(1) r)
+
+  have "h r \<otimes>\<^bsub>S\<^esub> h a \<in> I"
+    using assms a r by (simp add: ideal.I_l_closed)
+  thus "r \<otimes> a \<in> { r \<in> carrier R. h r \<in> I }" by (simp add: a(1) r)
+qed
+
+lemma (in ring) canonical_proj_vimage_in_carrier:
+  assumes "ideal I R"
+  shows "J \<subseteq> carrier (R Quot I) \<Longrightarrow> \<Union> J \<subseteq> carrier R"
+proof -
+  assume A: "J \<subseteq> carrier (R Quot I)" show "\<Union> J \<subseteq> carrier R"
+  proof
+    fix j assume j: "j \<in> \<Union> J"
+    then obtain j' where j': "j' \<in> J" "j \<in> j'" by blast
+    then obtain r where r: "r \<in> carrier R" "j' = I +> r"
+      using A j' unfolding FactRing_def using A_RCOSETS_def'[of R I] by auto
+    thus "j \<in> carrier R" using j' assms
+      by (meson a_r_coset_subset_G additive_subgroup.a_subset contra_subsetD ideal.axioms(1)) 
+  qed
+qed
+
+lemma (in ring) canonical_proj_vimage_mem_iff:
+  assumes "ideal I R" "J \<subseteq> carrier (R Quot I)"
+  shows "\<And>a. a \<in> carrier R \<Longrightarrow> (a \<in> (\<Union> J)) = (I +> a \<in> J)"
+proof -
+  fix a assume a: "a \<in> carrier R" show "(a \<in> (\<Union> J)) = (I +> a \<in> J)"
+  proof
+    assume "a \<in> \<Union> J"
+    then obtain j where j: "j \<in> J" "a \<in> j" by blast
+    then obtain r where r: "r \<in> 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 \<in> J" using r j by simp
+  next
+    assume "I +> a \<in> J"
+    hence "\<zero> \<oplus> a \<in> 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 \<in> \<Union> J" using a \<open>I +> a \<in> J\<close> by auto 
+  qed
+qed
+
+corollary (in ring) quot_ideal_imp_ring_ideal:
+  assumes "ideal I R"
+  shows "ideal J (R Quot I) \<Longrightarrow> ideal (\<Union> J) R"
+proof -
+  assume A: "ideal J (R Quot I)"
+  have "\<Union> J = { r \<in> carrier R. I +> r \<in> J }"
+    using canonical_proj_vimage_in_carrier[OF assms, of J]
+          canonical_proj_vimage_mem_iff[OF assms, of J]
+          additive_subgroup.a_subset[OF ideal.axioms(1)[OF A]] by blast
+  thus "ideal (\<Union> J) R"
+    using ring_hom_ring.ideal_vimage[OF ideal.rcos_ring_hom_ring[OF assms] A] by simp
+qed
+
+lemma (in ring) ideal_incl_iff:
+  assumes "ideal I R" "ideal J R"
+  shows "(I \<subseteq> J) = (J = (\<Union> j \<in> J. I +> j))"
+proof
+  assume A: "J = (\<Union> j \<in> J. I +> j)" hence "I +> \<zero> \<subseteq> J"
+    using additive_subgroup.zero_closed[OF ideal.axioms(1)[OF assms(2)]] by blast
+  thus "I \<subseteq> J" using additive_subgroup.a_subset[OF ideal.axioms(1)[OF assms(1)]] by simp 
+next
+  assume A: "I \<subseteq> J" show "J = (\<Union>j\<in>J. I +> j)"
+  proof
+    show "J \<subseteq> (\<Union> j \<in> J. I +> j)"
+    proof
+      fix j assume j: "j \<in> J"
+      have "\<zero> \<in> I" by (simp add: additive_subgroup.zero_closed assms(1) ideal.axioms(1))
+      hence "\<zero> \<oplus> j \<in> I +> j"
+        using a_r_coset_def'[of R I j] by blast
+      thus "j \<in> (\<Union>j\<in>J. I +> j)"
+        using assms(2) j additive_subgroup.a_Hcarr ideal.axioms(1) by fastforce 
+    qed
+  next
+    show "(\<Union> j \<in> J. I +> j) \<subseteq> J"
+    proof
+      fix x assume "x \<in> (\<Union> j \<in> J. I +> j)"
+      then obtain j where j: "j \<in> J" "x \<in> I +> j" by blast
+      then obtain i where i: "i \<in> I" "x = i \<oplus> j"
+        using a_r_coset_def'[of R I j] by blast
+      thus "x \<in> J"
+        using assms(2) j A additive_subgroup.a_closed[OF ideal.axioms(1)[OF assms(2)]] by blast
+    qed
+  qed
+qed
+
+theorem (in ring) quot_ideal_correspondence:
+  assumes "ideal I R"
+  shows "bij_betw (\<lambda>J. (+>) I ` J) { J. ideal J R \<and> I \<subseteq> J } { J . ideal J (R Quot I) }"
+proof (rule bij_betw_byWitness[where ?f' = "\<lambda>X. \<Union> X"])
+  show "\<forall>J \<in> { J. ideal J R \<and> I \<subseteq> J }. (\<lambda>X. \<Union> X) ((+>) I ` J) = J"
+    using assms ideal_incl_iff by blast
+next
+  show "(\<lambda>J. (+>) I ` J) ` { J. ideal J R \<and> I \<subseteq> J } \<subseteq> { J. ideal J (R Quot I) }"
+    using assms ring_ideal_imp_quot_ideal by auto
+next
+  show "(\<lambda>X. \<Union> X) ` { J. ideal J (R Quot I) } \<subseteq> { J. ideal J R \<and> I \<subseteq> J }"
+  proof
+    fix J assume "J \<in> ((\<lambda>X. \<Union> X) ` { J. ideal J (R Quot I) })"
+    then obtain J' where J': "ideal J' (R Quot I)" "J = \<Union> J'" by blast
+    hence "ideal J R"
+      using assms quot_ideal_imp_ring_ideal by auto
+    moreover have "I \<in> J'"
+      using additive_subgroup.zero_closed[OF ideal.axioms(1)[OF J'(1)]] unfolding FactRing_def by simp
+    ultimately show "J \<in> { J. ideal J R \<and> I \<subseteq> J }" using J'(2) by auto
+  qed
+next
+  show "\<forall>J' \<in> { J. ideal J (R Quot I) }. ((+>) I ` (\<Union> J')) = J'"
+  proof
+    fix J' assume "J' \<in> { J. ideal J (R Quot I) }"
+    hence subset: "J' \<subseteq> carrier (R Quot I) \<and> ideal J' (R Quot I)"
+      using additive_subgroup.a_subset ideal_def by blast
+    hence "((+>) I ` (\<Union> J')) \<subseteq> J'"
+      using canonical_proj_vimage_in_carrier canonical_proj_vimage_mem_iff
+      by (meson assms contra_subsetD image_subsetI)
+    moreover have "J' \<subseteq> ((+>) I ` (\<Union> J'))"
+    proof
+      fix x assume "x \<in> J'"
+      then obtain r where r: "r \<in> carrier R" "x = I +> r"
+        using subset unfolding FactRing_def A_RCOSETS_def'[of R I] by auto
+      hence "r \<in> (\<Union> J')"
+        using \<open>x \<in> J'\<close> assms canonical_proj_vimage_mem_iff subset by blast
+      thus "x \<in> ((+>) I ` (\<Union> J'))" using r(2) by blast
+    qed
+    ultimately show "((+>) I ` (\<Union> J')) = J'" by blast
+  qed
+qed
+
+lemma (in cring) quot_domain_imp_primeideal:
+  assumes "ideal P R"
+  shows "domain (R Quot P) \<Longrightarrow> primeideal P R"
+proof -
+  assume A: "domain (R Quot P)" show "primeideal P R"
+  proof (rule primeidealI)
+    show "ideal P R" using assms .
+    show "cring R" using is_cring .
+  next
+    show "carrier R \<noteq> P"
+    proof (rule ccontr)
+      assume "\<not> carrier R \<noteq> P" hence "carrier R = P" by simp
+      hence "\<And>I. I \<in> carrier (R Quot P) \<Longrightarrow> I = P"
+        unfolding FactRing_def A_RCOSETS_def' apply simp
+        using a_coset_join2 additive_subgroup.a_subgroup assms ideal.axioms(1) by blast
+      hence "\<one>\<^bsub>(R Quot P)\<^esub> = \<zero>\<^bsub>(R Quot P)\<^esub>"
+        by (metis assms ideal.quotient_is_ring ring.ring_simprules(2) ring.ring_simprules(6))
+      thus False using domain.one_not_zero[OF A] by simp
+    qed
+  next
+    fix a b assume a: "a \<in> carrier R" and b: "b \<in> carrier R" and ab: "a \<otimes> b \<in> P"
+    hence "P +> (a \<otimes> b) = \<zero>\<^bsub>(R Quot P)\<^esub>" unfolding FactRing_def
+      by (simp add: a_coset_join2 additive_subgroup.a_subgroup assms ideal.axioms(1))
+    moreover have "(P +> a) \<otimes>\<^bsub>(R Quot P)\<^esub> (P +> b) = P +> (a \<otimes> b)" unfolding FactRing_def
+      using a b by (simp add: assms ideal.rcoset_mult_add)
+    moreover have "P +> a \<in> carrier (R Quot P) \<and> P +> b \<in> carrier (R Quot P)"
+      by (simp add: a b FactRing_def a_rcosetsI additive_subgroup.a_subset assms ideal.axioms(1))
+    ultimately have "P +> a = \<zero>\<^bsub>(R Quot P)\<^esub> \<or> P +> b = \<zero>\<^bsub>(R Quot P)\<^esub>"
+      using domain.integral[OF A, of "P +> a" "P +> b"] by auto
+    thus "a \<in> P \<or> b \<in> P" unfolding FactRing_def apply simp
+      using a b assms a_coset_join1 additive_subgroup.a_subgroup ideal.axioms(1) by blast
+  qed
+qed
+
+lemma (in cring) quot_domain_iff_primeideal:
+  assumes "ideal P R"
+  shows "domain (R Quot P) = primeideal P R"
+  using quot_domain_imp_primeideal[OF assms] primeideal.quotient_is_domain[of P R] by auto
+
+
+subsection \<open>Isomorphism\<close>
+
+definition
+  ring_iso :: "_ \<Rightarrow> _ \<Rightarrow> ('a \<Rightarrow> 'b) set"
+  where "ring_iso R S = { h. h \<in> ring_hom R S \<and> bij_betw h (carrier R) (carrier S) }"
+
+definition
+  is_ring_iso :: "_ \<Rightarrow> _ \<Rightarrow> bool" (infixr "\<simeq>" 60)
+  where "R \<simeq> S = (ring_iso R S \<noteq> {})"
+
+definition
+  morphic_prop :: "_ \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
+  where "morphic_prop R P =
+           ((P \<one>\<^bsub>R\<^esub>) \<and>
+            (\<forall>r \<in> carrier R. P r) \<and>
+            (\<forall>r1 \<in> carrier R. \<forall>r2 \<in> carrier R. P (r1 \<otimes>\<^bsub>R\<^esub> r2)) \<and>
+            (\<forall>r1 \<in> carrier R. \<forall>r2 \<in> carrier R. P (r1 \<oplus>\<^bsub>R\<^esub> r2)))"
+
+lemma ring_iso_memI:
+  fixes R (structure) and S (structure)
+  assumes "\<And>x. x \<in> carrier R \<Longrightarrow> h x \<in> carrier S"
+      and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> h (x \<otimes> y) = h x \<otimes>\<^bsub>S\<^esub> h y"
+      and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> h (x \<oplus> y) = h x \<oplus>\<^bsub>S\<^esub> h y"
+      and "h \<one> = \<one>\<^bsub>S\<^esub>"
+      and "bij_betw h (carrier R) (carrier S)"
+  shows "h \<in> ring_iso R S"
+  by (auto simp add: ring_hom_memI assms ring_iso_def)
+
+lemma ring_iso_memE:
+  fixes R (structure) and S (structure)
+  assumes "h \<in> ring_iso R S"
+  shows "\<And>x. x \<in> carrier R \<Longrightarrow> h x \<in> carrier S"
+   and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> h (x \<otimes> y) = h x \<otimes>\<^bsub>S\<^esub> h y"
+   and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> h (x \<oplus> y) = h x \<oplus>\<^bsub>S\<^esub> h y"
+   and "h \<one> = \<one>\<^bsub>S\<^esub>"
+   and "bij_betw h (carrier R) (carrier S)"
+  using assms unfolding ring_iso_def ring_hom_def by auto
+
+lemma morphic_propI:
+  fixes R (structure)
+  assumes "P \<one>"
+    and "\<And>r. r \<in> carrier R \<Longrightarrow> P r"
+    and "\<And>r1 r2. \<lbrakk> r1 \<in> carrier R; r2 \<in> carrier R \<rbrakk> \<Longrightarrow> P (r1 \<otimes> r2)"
+    and "\<And>r1 r2. \<lbrakk> r1 \<in> carrier R; r2 \<in> carrier R \<rbrakk> \<Longrightarrow> P (r1 \<oplus> r2)"
+  shows "morphic_prop R P"
+  unfolding morphic_prop_def using assms by auto
+
+lemma morphic_propE:
+  fixes R (structure)
+  assumes "morphic_prop R P"
+  shows "P \<one>"
+    and "\<And>r. r \<in> carrier R \<Longrightarrow> P r"
+    and "\<And>r1 r2. \<lbrakk> r1 \<in> carrier R; r2 \<in> carrier R \<rbrakk> \<Longrightarrow> P (r1 \<otimes> r2)"
+    and "\<And>r1 r2. \<lbrakk> r1 \<in> carrier R; r2 \<in> carrier R \<rbrakk> \<Longrightarrow> P (r1 \<oplus> r2)"
+  using assms unfolding morphic_prop_def by auto
+
+lemma ring_iso_restrict:
+  assumes "f \<in> ring_iso R S"
+    and "\<And>r. r \<in> carrier R \<Longrightarrow> f r = g r"
+    and "ring R"
+  shows "g \<in> ring_iso R S"
+proof (rule ring_iso_memI)
+  show "bij_betw g (carrier R) (carrier S)"
+    using assms(1-2) bij_betw_cong ring_iso_memE(5) by blast
+  show "g \<one>\<^bsub>R\<^esub> = \<one>\<^bsub>S\<^esub>"
+    using assms ring.ring_simprules(6) ring_iso_memE(4) by force
+next
+  fix x y assume x: "x \<in> carrier R" and y: "y \<in> carrier R"
+  show "g x \<in> carrier S"
+    using assms(1-2) ring_iso_memE(1) x by fastforce
+  show "g (x \<otimes>\<^bsub>R\<^esub> y) = g x \<otimes>\<^bsub>S\<^esub> g y"
+    by (metis assms ring.ring_simprules(5) ring_iso_memE(2) x y)
+  show "g (x \<oplus>\<^bsub>R\<^esub> y) = g x \<oplus>\<^bsub>S\<^esub> g y"
+    by (metis assms ring.ring_simprules(1) ring_iso_memE(3) x y)
+qed
+
+lemma ring_iso_morphic_prop:
+  assumes "f \<in> ring_iso R S"
+    and "morphic_prop R P"
+    and "\<And>r. P r \<Longrightarrow> f r = g r"
+  shows "g \<in> ring_iso R S"
+proof -
+  have eq0: "\<And>r. r \<in> carrier R \<Longrightarrow> f r = g r"
+   and eq1: "f \<one>\<^bsub>R\<^esub> = g \<one>\<^bsub>R\<^esub>"
+   and eq2: "\<And>r1 r2. \<lbrakk> r1 \<in> carrier R; r2 \<in> carrier R \<rbrakk> \<Longrightarrow> f (r1 \<otimes>\<^bsub>R\<^esub> r2) = g (r1 \<otimes>\<^bsub>R\<^esub> r2)"
+   and eq3: "\<And>r1 r2. \<lbrakk> r1 \<in> carrier R; r2 \<in> carrier R \<rbrakk> \<Longrightarrow> f (r1 \<oplus>\<^bsub>R\<^esub> r2) = g (r1 \<oplus>\<^bsub>R\<^esub> r2)"
+    using assms(2-3) unfolding morphic_prop_def by auto
+  show ?thesis
+    apply (rule ring_iso_memI)
+    using assms(1) eq0 ring_iso_memE(1) apply fastforce
+    apply (metis assms(1) eq0 eq2 ring_iso_memE(2))
+    apply (metis assms(1) eq0 eq3 ring_iso_memE(3))
+    using assms(1) eq1 ring_iso_memE(4) apply fastforce
+    using assms(1) bij_betw_cong eq0 ring_iso_memE(5) by blast
+qed
+
+lemma (in ring) ring_hom_imp_img_ring:
+  assumes "h \<in> ring_hom R S"
+  shows "ring (S \<lparr> carrier := h ` (carrier R), one := h \<one>, zero := h \<zero> \<rparr>)" (is "ring ?h_img")
+proof -
+  have "h \<in> hom (add_monoid R) (add_monoid S)"
+    using assms unfolding hom_def ring_hom_def by auto
+  hence "comm_group ((add_monoid S) \<lparr>  carrier := h ` (carrier R), one := h \<zero> \<rparr>)"
+    using add.hom_imp_img_comm_group[of h "add_monoid S"] by simp
+  hence comm_group: "comm_group (add_monoid ?h_img)"
+    by (auto intro: comm_monoidI simp add: monoid.defs)
+
+  moreover have "h \<in> hom R S"
+    using assms unfolding ring_hom_def hom_def by auto
+  hence "monoid (S \<lparr>  carrier := h ` (carrier R), one := h \<one> \<rparr>)"
+    using hom_imp_img_monoid[of h S] by simp
+  hence monoid: "monoid ?h_img"
+    unfolding monoid_def by (simp add: monoid.defs)
+
+  show ?thesis
+  proof (rule ringI, simp_all add: comm_group_abelian_groupI[OF comm_group] monoid)
+    fix x y z assume "x \<in> h ` carrier R" "y \<in> h ` carrier R" "z \<in> h ` carrier R"
+    then obtain r1 r2 r3
+      where r1: "r1 \<in> carrier R" "x = h r1"
+        and r2: "r2 \<in> carrier R" "y = h r2"
+        and r3: "r3 \<in> carrier R" "z = h r3" by blast
+    hence "(x \<oplus>\<^bsub>S\<^esub> y) \<otimes>\<^bsub>S\<^esub> z = h ((r1 \<oplus> r2) \<otimes> r3)"
+      using ring_hom_memE[OF assms] by auto
+    also have " ... = h ((r1 \<otimes> r3) \<oplus> (r2 \<otimes> r3))"
+      using l_distr[OF r1(1) r2(1) r3(1)] by simp
+    also have " ... = (x \<otimes>\<^bsub>S\<^esub> z) \<oplus>\<^bsub>S\<^esub> (y \<otimes>\<^bsub>S\<^esub> z)"
+      using ring_hom_memE[OF assms] r1 r2 r3 by auto
+    finally show "(x \<oplus>\<^bsub>S\<^esub> y) \<otimes>\<^bsub>S\<^esub> z = (x \<otimes>\<^bsub>S\<^esub> z) \<oplus>\<^bsub>S\<^esub> (y \<otimes>\<^bsub>S\<^esub> z)" .
+
+    have "z \<otimes>\<^bsub>S\<^esub> (x \<oplus>\<^bsub>S\<^esub> y) = h (r3 \<otimes> (r1 \<oplus> r2))"
+      using ring_hom_memE[OF assms] r1 r2 r3 by auto
+    also have " ... =  h ((r3 \<otimes> r1) \<oplus> (r3 \<otimes> r2))"
+      using r_distr[OF r1(1) r2(1) r3(1)] by simp
+    also have " ... = (z \<otimes>\<^bsub>S\<^esub> x) \<oplus>\<^bsub>S\<^esub> (z \<otimes>\<^bsub>S\<^esub> y)"
+      using ring_hom_memE[OF assms] r1 r2 r3 by auto
+    finally show "z \<otimes>\<^bsub>S\<^esub> (x \<oplus>\<^bsub>S\<^esub> y) = (z \<otimes>\<^bsub>S\<^esub> x) \<oplus>\<^bsub>S\<^esub> (z \<otimes>\<^bsub>S\<^esub> y)" .
+  qed
+qed
+
+lemma (in ring) ring_iso_imp_img_ring:
+  assumes "h \<in> ring_iso R S"
+  shows "ring (S \<lparr> one := h \<one>, zero := h \<zero> \<rparr>)"
+proof -
+  have "ring (S \<lparr> carrier := h ` (carrier R), one := h \<one>, zero := h \<zero> \<rparr>)"
+    using ring_hom_imp_img_ring[of h S] assms unfolding ring_iso_def by auto
+  moreover have "h ` (carrier R) = carrier S"
+    using assms unfolding ring_iso_def bij_betw_def by auto
+  ultimately show ?thesis by simp
+qed
+
+lemma (in cring) ring_iso_imp_img_cring:
+  assumes "h \<in> ring_iso R S"
+  shows "cring (S \<lparr> one := h \<one>, zero := h \<zero> \<rparr>)" (is "cring ?h_img")
+proof -
+  note m_comm
+  interpret h_img?: ring ?h_img
+    using ring_iso_imp_img_ring[OF assms] .
+  show ?thesis 
+  proof (unfold_locales)
+    fix x y assume "x \<in> carrier ?h_img" "y \<in> carrier ?h_img"
+    then obtain r1 r2
+      where r1: "r1 \<in> carrier R" "x = h r1"
+        and r2: "r2 \<in> carrier R" "y = h r2"
+      using assms image_iff[where ?f = h and ?A = "carrier R"]
+      unfolding ring_iso_def bij_betw_def by auto
+    have "x \<otimes>\<^bsub>(?h_img)\<^esub> y = h (r1 \<otimes> r2)"
+      using assms r1 r2 unfolding ring_iso_def ring_hom_def by auto
+    also have " ... = h (r2 \<otimes> r1)"
+      using m_comm[OF r1(1) r2(1)] by simp
+    also have " ... = y \<otimes>\<^bsub>(?h_img)\<^esub> x"
+      using assms r1 r2 unfolding ring_iso_def ring_hom_def by auto
+    finally show "x \<otimes>\<^bsub>(?h_img)\<^esub> y = y \<otimes>\<^bsub>(?h_img)\<^esub> x" .
+  qed
+qed
+
+lemma (in domain) ring_iso_imp_img_domain:
+  assumes "h \<in> ring_iso R S"
+  shows "domain (S \<lparr> one := h \<one>, zero := h \<zero> \<rparr>)" (is "domain ?h_img")
+proof -
+  note aux = m_closed integral one_not_zero one_closed zero_closed
+  interpret h_img?: cring ?h_img
+    using ring_iso_imp_img_cring[OF assms] .
+  show ?thesis 
+  proof (unfold_locales)
+    show "\<one>\<^bsub>?h_img\<^esub> \<noteq> \<zero>\<^bsub>?h_img\<^esub>"
+      using ring_iso_memE(5)[OF assms] aux(3-4)
+      unfolding bij_betw_def inj_on_def by force
+  next
+    fix a b
+    assume A: "a \<otimes>\<^bsub>?h_img\<^esub> b = \<zero>\<^bsub>?h_img\<^esub>" "a \<in> carrier ?h_img" "b \<in> carrier ?h_img"
+    then obtain r1 r2
+      where r1: "r1 \<in> carrier R" "a = h r1"
+        and r2: "r2 \<in> carrier R" "b = h r2"
+      using assms image_iff[where ?f = h and ?A = "carrier R"]
+      unfolding ring_iso_def bij_betw_def by auto
+    hence "a \<otimes>\<^bsub>?h_img\<^esub> b = h (r1 \<otimes> r2)"
+      using assms r1 r2 unfolding ring_iso_def ring_hom_def by auto
+    hence "h (r1 \<otimes> r2) = h \<zero>"
+      using A(1) by simp
+    hence "r1 \<otimes> r2 = \<zero>"
+      using ring_iso_memE(5)[OF assms] aux(1)[OF r1(1) r2(1)] aux(5)
+      unfolding bij_betw_def inj_on_def by force
+    hence "r1 = \<zero> \<or> r2 = \<zero>"
+      using aux(2)[OF _ r1(1) r2(1)] by simp
+    thus "a = \<zero>\<^bsub>?h_img\<^esub> \<or> b = \<zero>\<^bsub>?h_img\<^esub>"
+      unfolding r1 r2 by auto
+  qed
+qed
+
+lemma (in field) ring_iso_imp_img_field:
+  assumes "h \<in> ring_iso R S"
+  shows "field (S \<lparr> one := h \<one>, zero := h \<zero> \<rparr>)" (is "field ?h_img")
+proof -
+  interpret h_img?: domain ?h_img
+    using ring_iso_imp_img_domain[OF assms] .
+  show ?thesis
+  proof (unfold_locales, auto simp add: Units_def)
+    interpret field R using field_axioms .
+    fix a assume a: "a \<in> carrier S" "a \<otimes>\<^bsub>S\<^esub> h \<zero> = h \<one>"
+    then obtain r where r: "r \<in> carrier R" "a = h r"
+      using assms image_iff[where ?f = h and ?A = "carrier R"]
+      unfolding ring_iso_def bij_betw_def by auto
+    have "a \<otimes>\<^bsub>S\<^esub> h \<zero> = h (r \<otimes> \<zero>)" unfolding r(2)
+      using ring_iso_memE(2)[OF assms r(1)] by simp
+    hence "h \<one> = h \<zero>"
+      using r(1) a(2) by simp
+    thus False
+      using ring_iso_memE(5)[OF assms]
+      unfolding bij_betw_def inj_on_def by force
+  next
+    interpret field R using field_axioms .
+    fix s assume s: "s \<in> carrier S" "s \<noteq> h \<zero>"
+    then obtain r where r: "r \<in> 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 \<noteq> \<zero>" using s(2) by auto 
+    hence inv_r: "inv r \<in> carrier R" "inv r \<noteq> \<zero>" "r \<otimes> inv r = \<one>" "inv r \<otimes> r = \<one>"
+      using field_Units r(1) by auto
+    have "h (inv r) \<otimes>\<^bsub>S\<^esub> h r = h \<one>" and "h r \<otimes>\<^bsub>S\<^esub> h (inv r) = h \<one>"
+      using ring_iso_memE(2)[OF assms inv_r(1) r(1)] inv_r(3-4)
+            ring_iso_memE(2)[OF assms r(1) inv_r(1)] by auto
+    thus "\<exists>s' \<in> carrier S. s' \<otimes>\<^bsub>S\<^esub> s = h \<one> \<and> s \<otimes>\<^bsub>S\<^esub> s' = h \<one>"
+      using ring_iso_memE(1)[OF assms inv_r(1)] r(2) by auto
+  qed
+qed
+
+lemma ring_iso_same_card: "R \<simeq> S \<Longrightarrow> card (carrier R) = card (carrier S)"
+proof -
+  assume "R \<simeq> S"
+  then obtain h where "bij_betw h (carrier R) (carrier S)"
+    unfolding is_ring_iso_def ring_iso_def by auto
+  thus "card (carrier R) = card (carrier S)"
+    using bij_betw_same_card[of h "carrier R" "carrier S"] by simp
+qed
+
+lemma ring_iso_set_refl: "id \<in> ring_iso R R"
+  by (rule ring_iso_memI) (auto)
+
+corollary ring_iso_refl: "R \<simeq> R"
+  using is_ring_iso_def ring_iso_set_refl by auto 
+
+lemma ring_iso_set_trans:
+  "\<lbrakk> f \<in> ring_iso R S; g \<in> ring_iso S Q \<rbrakk> \<Longrightarrow> (g \<circ> f) \<in> ring_iso R Q"
+  unfolding ring_iso_def using bij_betw_trans ring_hom_trans by fastforce 
+
+corollary ring_iso_trans: "\<lbrakk> R \<simeq> S; S \<simeq> Q \<rbrakk> \<Longrightarrow> R \<simeq> Q"
+  using ring_iso_set_trans unfolding is_ring_iso_def by blast 
+
+lemma ring_iso_set_sym:
+  assumes "ring R"
+  shows "h \<in> ring_iso R S \<Longrightarrow> (inv_into (carrier R) h) \<in> ring_iso S R"
+proof -
+  assume h: "h \<in> ring_iso R S"
+  hence h_hom:  "h \<in> ring_hom R S"
+    and h_surj: "h ` (carrier R) = (carrier S)"
+    and h_inj:  "\<And> x1 x2. \<lbrakk> x1 \<in> carrier R; x2 \<in> carrier R \<rbrakk> \<Longrightarrow>  h x1 = h x2 \<Longrightarrow> x1 = x2"
+    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)"
+      using bij_betw_inv_into h ring_iso_def by fastforce
+
+  show "inv_into (carrier R) h \<in> ring_iso S R"
+    apply (rule ring_iso_memI)
+    apply (simp add: h_surj inv_into_into)
+    apply (auto simp add: h_inv_bij)
+    apply (smt assms f_inv_into_f h_hom h_inj h_surj inv_into_into
+           ring.ring_simprules(5) ring_hom_closed ring_hom_mult)
+    apply (smt assms f_inv_into_f h_hom h_inj h_surj inv_into_into
+           ring.ring_simprules(1) ring_hom_add ring_hom_closed)
+    by (metis (no_types, hide_lams) assms f_inv_into_f h_hom h_inj
+        imageI inv_into_into ring.ring_simprules(6) ring_hom_one)
+qed
+
+corollary ring_iso_sym:
+  assumes "ring R"
+  shows "R \<simeq> S \<Longrightarrow> S \<simeq> R"
+  using assms ring_iso_set_sym unfolding is_ring_iso_def by auto 
+
+lemma (in ring_hom_ring) the_elem_simp [simp]:
+  "\<And>x. x \<in> carrier R \<Longrightarrow> the_elem (h ` ((a_kernel R S h) +> x)) = h x"
+proof -
+  fix x assume x: "x \<in> carrier R"
+  hence "h x \<in> 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:
+  "\<And>X Y. \<lbrakk> X \<in> carrier (R Quot (a_kernel R S h)); Y \<in> carrier (R Quot (a_kernel R S h)) \<rbrakk> \<Longrightarrow>
+           the_elem (h ` X) = the_elem (h ` Y) \<Longrightarrow> X = Y"
+proof -
+  fix X Y
+  assume "X \<in> carrier (R Quot (a_kernel R S h))"
+     and "Y \<in> 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 \<in> carrier R" "X = (a_kernel R S h) +> x"
+                    and y: "y \<in> 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 \<ominus> y \<in> (a_kernel R S h)"
+    by (simp add: a_minus_def abelian_subgroup.a_rcos_module_imp
+                  abelian_subgroup_a_kernel homeq_imp_rcos x(1) y(1))
+  thus "X = Y"
+    by (metis R.a_coset_add_inv1 R.minus_eq abelian_subgroup.a_rcos_const
+        abelian_subgroup_a_kernel additive_subgroup.a_subset additive_subgroup_a_kernel x y)
+qed
+
+lemma (in ring_hom_ring) quot_mem:
+  "\<And>X. X \<in> carrier (R Quot (a_kernel R S h)) \<Longrightarrow> \<exists>x \<in> carrier R. X = (a_kernel R S h) +> x"
+proof -
+  fix X assume "X \<in> carrier (R Quot (a_kernel R S h))"
+  thus "\<exists>x \<in> carrier R. X = (a_kernel R S h) +> x"
+    unfolding FactRing_def RCOSETS_def A_RCOSETS_def by (simp add: a_r_coset_def)
+qed
+
+lemma (in ring_hom_ring) the_elem_wf:
+  "\<And>X. X \<in> carrier (R Quot (a_kernel R S h)) \<Longrightarrow> \<exists>y \<in> carrier S. (h ` X) = { y }"
+proof -
+  fix X assume "X \<in> carrier (R Quot (a_kernel R S h))"
+  then obtain x where x: "x \<in> carrier R" and X: "X = (a_kernel R S h) +> x"
+    using quot_mem by blast
+  hence "\<And>x'. x' \<in> X \<Longrightarrow> h x' = h x"
+  proof -
+    fix x' assume "x' \<in> X" hence "x' \<in> (a_kernel R S h) +> x" using X by simp
+    then obtain k where k: "k \<in> a_kernel R S h" "x' = k \<oplus> x"
+      by (metis R.add.inv_closed R.add.m_assoc R.l_neg R.r_zero
+          abelian_subgroup.a_elemrcos_carrier
+          abelian_subgroup.a_rcos_module_imp abelian_subgroup_a_kernel x)
+    hence "h x' = h k \<oplus>\<^bsub>S\<^esub> h x"
+      by (meson additive_subgroup.a_Hcarr additive_subgroup_a_kernel hom_add x)
+    also have " ... =  h x"
+      using k by (auto simp add: x)
+    finally show "h x' = h x" .
+  qed
+  moreover have "h x \<in> h ` X"
+    by (simp add: X homeq_imp_rcos x)
+  ultimately have "(h ` X) = { h x }"
+    by blast
+  thus "\<exists>y \<in> carrier S. (h ` X) = { y }" using x by simp
+qed
+
+corollary (in ring_hom_ring) the_elem_wf':
+  "\<And>X. X \<in> carrier (R Quot (a_kernel R S h)) \<Longrightarrow> \<exists>r \<in> 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:
+  "(\<lambda>X. the_elem (h ` X)) \<in> ring_hom (R Quot (a_kernel R S h)) S"
+proof (rule ring_hom_memI)
+  show "\<And>x. x \<in> carrier (R Quot a_kernel R S h) \<Longrightarrow> the_elem (h ` x) \<in> carrier S"
+    using the_elem_wf by fastforce
+  
+  show "the_elem (h ` \<one>\<^bsub>R Quot a_kernel R S h\<^esub>) = \<one>\<^bsub>S\<^esub>"
+    unfolding FactRing_def  using the_elem_simp[of "\<one>\<^bsub>R\<^esub>"] by simp
+
+  fix X Y
+  assume "X \<in> carrier (R Quot a_kernel R S h)"
+     and "Y \<in> carrier (R Quot a_kernel R S h)"
+  then obtain x y where x: "x \<in> carrier R" "X = (a_kernel R S h) +> x"
+                    and y: "y \<in> carrier R" "Y = (a_kernel R S h) +> y"
+    using quot_mem by blast
+
+  have "X \<otimes>\<^bsub>R Quot a_kernel R S h\<^esub> Y = (a_kernel R S h) +> (x \<otimes> y)"
+    by (simp add: FactRing_def ideal.rcoset_mult_add kernel_is_ideal x y)
+  thus "the_elem (h ` (X \<otimes>\<^bsub>R Quot a_kernel R S h\<^esub> Y)) = the_elem (h ` X) \<otimes>\<^bsub>S\<^esub> the_elem (h ` Y)"
+    by (simp add: x y)
+
+  have "X \<oplus>\<^bsub>R Quot a_kernel R S h\<^esub> Y = (a_kernel R S h) +> (x \<oplus> y)"
+    using ideal.rcos_ring_hom kernel_is_ideal ring_hom_add x y by fastforce
+  thus "the_elem (h ` (X \<oplus>\<^bsub>R Quot a_kernel R S h\<^esub> Y)) = the_elem (h ` X) \<oplus>\<^bsub>S\<^esub> the_elem (h ` Y)"
+    by (simp add: x y)
+qed
+
+lemma (in ring_hom_ring) the_elem_surj:
+    "(\<lambda>X. (the_elem (h ` X))) ` carrier (R Quot (a_kernel R S h)) = (h ` (carrier R))"
+proof
+  show "(\<lambda>X. the_elem (h ` X)) ` carrier (R Quot a_kernel R S h) \<subseteq> h ` carrier R"
+    using the_elem_wf' by fastforce
+next
+  show "h ` carrier R \<subseteq> (\<lambda>X. the_elem (h ` X)) ` carrier (R Quot a_kernel R S h)"
+  proof
+    fix y assume "y \<in> h ` carrier R"
+    then obtain x where x: "x \<in> carrier R" "h x = y"
+      by (metis image_iff)
+    hence "the_elem (h ` ((a_kernel R S h) +> x)) = y" by simp
+    moreover have "(a_kernel R S h) +> x \<in> carrier (R Quot (a_kernel R S h))"
+     unfolding FactRing_def RCOSETS_def A_RCOSETS_def by (auto simp add: x a_r_coset_def)
+    ultimately show "y \<in> (\<lambda>X. (the_elem (h ` X))) ` carrier (R Quot (a_kernel R S h))" by blast
+  qed
+qed
+
+proposition (in ring_hom_ring) FactRing_iso_set_aux:
+  "(\<lambda>X. the_elem (h ` X)) \<in> ring_iso (R Quot (a_kernel R S h)) (S \<lparr> carrier := h ` (carrier R) \<rparr>)"
+proof -
+  have "bij_betw (\<lambda>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 "(\<lambda>X. the_elem (h ` X)): carrier (R Quot (a_kernel R S h)) \<rightarrow> h ` (carrier R)"
+    using the_elem_wf' by fastforce
+  hence "(\<lambda>X. the_elem (h ` X)) \<in> ring_hom (R Quot (a_kernel R S h)) (S \<lparr> carrier := h ` (carrier R) \<rparr>)"
+    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
+qed
+
+theorem (in ring_hom_ring) FactRing_iso_set:
+  assumes "h ` carrier R = carrier S"
+  shows "(\<lambda>X. the_elem (h ` X)) \<in> ring_iso (R Quot (a_kernel R S h)) S"
+  using FactRing_iso_set_aux assms by auto
+
+corollary (in ring_hom_ring) FactRing_iso:
+  assumes "h ` carrier R = carrier S"
+  shows "R Quot (a_kernel R S h) \<simeq> S"
+  using FactRing_iso_set assms is_ring_iso_def by auto
+
+lemma (in ring_hom_ring) img_is_ring: "ring (S \<lparr> carrier := h ` (carrier R) \<rparr>)"
+proof -
+  let ?the_elem = "\<lambda>X. the_elem (h ` X)"
+  have FactRing_is_ring: "ring (R Quot (a_kernel R S h))"
+    by (simp add: ideal.quotient_is_ring kernel_is_ideal)
+  have "ring ((S \<lparr> carrier := ?the_elem ` (carrier (R Quot (a_kernel R S h))) \<rparr>)
+                 \<lparr>     one := ?the_elem \<one>\<^bsub>(R Quot (a_kernel R S h))\<^esub>,
+                      zero := ?the_elem \<zero>\<^bsub>(R Quot (a_kernel R S h))\<^esub> \<rparr>)"
+    using ring.ring_iso_imp_img_ring[OF FactRing_is_ring, of ?the_elem
+          "S \<lparr> carrier := ?the_elem ` (carrier (R Quot (a_kernel R S h))) \<rparr>"]
+          FactRing_iso_set_aux the_elem_surj by auto
+
+  moreover
+  have "\<zero> \<in> (a_kernel R S h)"
+    using a_kernel_def'[of R S h] by auto
+  hence "\<one> \<in> (a_kernel R S h) +> \<one>"
+    using a_r_coset_def'[of R "a_kernel R S h" \<one>] by force
+  hence "\<one>\<^bsub>S\<^esub> \<in> (h ` ((a_kernel R S h) +> \<one>))"
+    using hom_one by force
+  hence "?the_elem \<one>\<^bsub>(R Quot (a_kernel R S h))\<^esub> = \<one>\<^bsub>S\<^esub>"
+    using the_elem_wf[of "(a_kernel R S h) +> \<one>"] by (simp add: FactRing_def)
+  
+  moreover
+  have "\<zero>\<^bsub>S\<^esub> \<in> (h ` (a_kernel R S h))"
+    using a_kernel_def'[of R S h] hom_zero by force
+  hence "\<zero>\<^bsub>S\<^esub> \<in> (h ` \<zero>\<^bsub>(R Quot (a_kernel R S h))\<^esub>)"
+    by (simp add: FactRing_def)
+  hence "?the_elem \<zero>\<^bsub>(R Quot (a_kernel R S h))\<^esub> = \<zero>\<^bsub>S\<^esub>"
+    using the_elem_wf[OF ring.ring_simprules(2)[OF FactRing_is_ring]]
+    by (metis singletonD the_elem_eq) 
+
+  ultimately
+  have "ring ((S \<lparr> carrier := h ` (carrier R) \<rparr>) \<lparr> one := \<one>\<^bsub>S\<^esub>, zero := \<zero>\<^bsub>S\<^esub> \<rparr>)"
+    using the_elem_surj by simp
+  thus ?thesis
+    by auto
+qed
+
+lemma (in ring_hom_ring) img_is_cring:
+  assumes "cring S"
+  shows "cring (S \<lparr> carrier := h ` (carrier R) \<rparr>)"
+proof -
+  interpret ring "S \<lparr> carrier := h ` (carrier R) \<rparr>"
+    using img_is_ring .
+  show ?thesis
+    apply unfold_locales
+    using assms unfolding cring_def comm_monoid_def comm_monoid_axioms_def by auto
+qed
+
+lemma (in ring_hom_ring) img_is_domain:
+  assumes "domain S"
+  shows "domain (S \<lparr> carrier := h ` (carrier R) \<rparr>)"
+proof -
+  interpret cring "S \<lparr> carrier := h ` (carrier R) \<rparr>"
+    using img_is_cring assms unfolding domain_def by simp
+  show ?thesis
+    apply unfold_locales
+    using assms unfolding domain_def domain_axioms_def apply auto
+    using hom_closed by blast 
+qed
+
+proposition (in ring_hom_ring) primeideal_vimage:
+  assumes "cring R"
+  shows "primeideal P S \<Longrightarrow> primeideal { r \<in> carrier R. h r \<in> P } R"
+proof -
+  assume A: "primeideal P S"
+  hence is_ideal: "ideal P S" unfolding primeideal_def by simp
+  have "ring_hom_ring R (S Quot P) (((+>\<^bsub>S\<^esub>) P) \<circ> 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
+    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) \<circ> h" by simp
+  
+  have "inj_on (\<lambda>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
+  have "ideal (a_kernel R (S Quot P) ?h) R"
+    using hom.kernel_is_ideal by auto
+  have hom': "ring_hom_ring (R Quot (a_kernel R (S Quot P) ?h)) (S Quot P) (\<lambda>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
+  
+  moreover have "a_kernel R (S Quot P) ?h = { r \<in> carrier R. h r \<in> P }"
+  proof
+    show "a_kernel R (S Quot P) ?h \<subseteq> { r \<in> carrier R. h r \<in> P }"
+    proof 
+      fix r assume "r \<in> a_kernel R (S Quot P) ?h"
+      hence r: "r \<in> carrier R" "P +>\<^bsub>S\<^esub> (h r) = P"
+        unfolding a_kernel_def kernel_def FactRing_def by auto
+      hence "h r \<in> P"
+        using S.a_rcosI R.l_zero S.l_zero additive_subgroup.a_subset[OF ideal.axioms(1)[OF is_ideal]]
+              additive_subgroup.zero_closed[OF ideal.axioms(1)[OF is_ideal]] hom_closed by metis
+      thus "r \<in> { r \<in> carrier R. h r \<in> P }" using r by simp
+    qed
+  next
+    show "{ r \<in> carrier R. h r \<in> P } \<subseteq> a_kernel R (S Quot P) ?h"
+    proof
+      fix r assume "r \<in> { r \<in> carrier R. h r \<in> P }"
+      hence r: "r \<in> carrier R" "h r \<in> P" by simp_all
+      hence "?h r = P"
+        by (simp add: S.a_coset_join2 additive_subgroup.a_subgroup ideal.axioms(1) is_ideal)
+      thus "r \<in> a_kernel R (S Quot P) ?h"
+        unfolding a_kernel_def kernel_def FactRing_def using r(1) by auto
+    qed
+  qed
+  ultimately show "primeideal { r \<in> carrier R. h r \<in> P } R" by simp
+qed
+
 end
--- a/src/HOL/Algebra/Ring.thy	Fri Jun 29 23:04:36 2018 +0200
+++ b/src/HOL/Algebra/Ring.thy	Sat Jun 30 15:44:04 2018 +0100
@@ -333,11 +333,6 @@
     and "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z"
   using assms cring_def apply auto by (simp add: assms cring.axioms(1) ringE(3))
 
-(*
-lemma (in cring) is_comm_monoid:
-  "comm_monoid R"
-  by (auto intro!: comm_monoidI m_assoc m_comm)
-*)
 lemma (in cring) is_cring:
   "cring R" by (rule cring_axioms)
 
@@ -652,6 +647,15 @@
 text \<open>Field would not need to be derived from domain, the properties
   for domain follow from the assumptions of field\<close>
 
+lemma fieldE :
+  fixes R (structure)
+  assumes "field R"
+  shows "cring R"
+    and one_not_zero : "\<one> \<noteq> \<zero>"
+    and integral: "\<And>a b. \<lbrakk> a \<otimes> b = \<zero>; a \<in> carrier R; b \<in> carrier R \<rbrakk> \<Longrightarrow> a = \<zero> \<or> b = \<zero>"
+  and field_Units: "Units R = carrier R - {\<zero>}"
+  using assms unfolding field_def field_axioms_def domain_def domain_axioms_def by simp_all
+
 lemma (in cring) cring_fieldI:
   assumes field_Units: "Units R = carrier R - {\<zero>}"
   shows "field R"
--- a/src/HOL/Algebra/RingHom.thy	Fri Jun 29 23:04:36 2018 +0200
+++ b/src/HOL/Algebra/RingHom.thy	Sat Jun 30 15:44:04 2018 +0100
@@ -203,4 +203,37 @@
       show "x \<in> a_kernel R S h +> a" by (rule homeq_imp_rcos)
 qed
 
+(*contributed by Paulo Emílio de Vilhena*)
+lemma (in ring_hom_ring) inj_on_domain:
+  assumes "inj_on h (carrier R)"
+  shows "domain S \<Longrightarrow> domain R"
+proof -
+  assume A: "domain S" show "domain R"
+  proof
+    have "h \<one> = \<one>\<^bsub>S\<^esub> \<and> h \<zero> = \<zero>\<^bsub>S\<^esub>" by simp
+    hence "h \<one> \<noteq> h \<zero>"
+      using domain.one_not_zero[OF A] by simp
+    thus "\<one> \<noteq> \<zero>"
+      using assms unfolding inj_on_def by fastforce 
+  next
+    fix a b
+    assume a: "a \<in> carrier R"
+       and b: "b \<in> carrier R"
+    have "h (a \<otimes> b) = (h a) \<otimes>\<^bsub>S\<^esub> (h b)" by (simp add: a b)
+    also have " ... = (h b) \<otimes>\<^bsub>S\<^esub> (h a)" using a b A cringE(1)[of S]
+      by (simp add: cring.cring_simprules(14) domain_def)
+    also have " ... = h (b \<otimes> a)" by (simp add: a b)
+    finally have "h (a \<otimes> b) = h (b \<otimes> a)" .
+    thus "a \<otimes> b = b \<otimes> a"
+      using assms a b unfolding inj_on_def by simp 
+    
+    assume  ab: "a \<otimes> b = \<zero>"
+    hence "h (a \<otimes> b) = \<zero>\<^bsub>S\<^esub>" by simp
+    hence "(h a) \<otimes>\<^bsub>S\<^esub> (h b) = \<zero>\<^bsub>S\<^esub>" using a b by simp
+    hence "h a =  \<zero>\<^bsub>S\<^esub> \<or> h b =  \<zero>\<^bsub>S\<^esub>" using a b domain.integral[OF A] by simp
+    thus "a = \<zero> \<or> b = \<zero>"
+      using a b assms unfolding inj_on_def by force
+  qed
+qed
+
 end