Theory QuotRing

(*  Title:      HOL/Algebra/QuotRing.thy
    Author:     Stephan Hohe
    Author:     Paulo Emílio de Vilhena
*)

theory QuotRing
imports RingHom
begin

section ‹Quotient Rings›

subsection ‹Multiplication on Cosets›

definition rcoset_mult :: "[('a, _) ring_scheme, 'a set, 'a set, 'a set]  'a set"
    ("[mod _:] _ ı _" [81,81,81] 80)
  where "rcoset_mult R I A B = (aA. bB. I +>R(a Rb))"


text constrcoset_mult fulfils the properties required by congruences›
lemma (in ideal) rcoset_mult_add:
  assumes "x  carrier R" "y  carrier R"
  shows "[mod I:] (I +> x)  (I +> y) = I +> (x  y)"
proof -
  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
    obtain hx hy hz where hxI: "hx  I" and x': "x' = hx  x" and hyI: "hy  I" and y': "y' = hy  y"
      and hzI: "hz  I" and z: "z = hz  (x'  y')"
      by (auto simp: a_r_coset_def r_coset_def)
    note carr = assms hxI[THEN a_Hcarr] hyI[THEN a_Hcarr] hzI[THEN a_Hcarr]
    from z  x' y' have "z = hz  ((hx  x)  (hy  y))" by simp
    also from carr have " = (hz  (hx  (hy  y))  x  hy)  x  y" by algebra
    finally have z2: "z = (hz  (hx  (hy  y))  x  hy)  x  y" .
    from hxI hyI hzI carr have "hz  (hx  (hy  y))  x  hy  I"
      by (simp add: I_l_closed I_r_closed)
    with z2 show ?thesis
      by (auto simp add: a_r_coset_def r_coset_def)
  qed
  have 2: "aI +> x. bI +> y. z  I +> a  b" if "z  I +> x  y" for z
    using assms a_rcos_self that by blast
  show ?thesis
    unfolding rcoset_mult_def using assms
    by (auto simp: intro!: 1 2)
qed

subsection ‹Quotient Ring Definition›

definition FactRing :: "[('a,'b) ring_scheme, 'a set]  ('a set) ring"
    (infixl "Quot" 65)
  where "FactRing R I =
    carrier = a_rcosetsRI, mult = rcoset_mult R I,
      one = (I +>R𝟭R), zero = I, add = set_add R"

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
  show "Group.monoid (R Quot I)"
    by (rule monoidI)
      (auto simp add: FactRing_simps rcoset_mult_add m_assoc)
qed (auto simp: FactRing_simps rcoset_mult_add a_rcos_sum l_distr r_distr)


text ‹This is a ring homomorphism›

lemma (in ideal) rcos_ring_hom: "((+>) I)  ring_hom R (R Quot I)"
  by (simp add: ring_hom_memI FactRing_def a_rcosetsI[OF a_subset] rcoset_mult_add a_rcos_sum)

lemma (in ideal) rcos_ring_hom_ring: "ring_hom_ring R (R Quot I) ((+>) I)"
  by (simp add: local.ring_axioms quotient_is_ring rcos_ring_hom ring_hom_ringI2)

text ‹The quotient of a cring is also commutative›
lemma (in ideal) quotient_is_cring:
  assumes "cring R"
  shows "cring (R Quot I)"
proof -
  interpret cring R by fact
  show ?thesis
    apply (intro cring.intro comm_monoid.intro comm_monoid_axioms.intro quotient_is_ring)
     apply (rule ring.axioms[OF quotient_is_ring])
    apply (auto simp add: FactRing_simps rcoset_mult_add m_comm)
    done
qed

text ‹Cosets as a ring homomorphism on crings›
lemma (in ideal) rcos_ring_hom_cring:
  assumes "cring R"
  shows "ring_hom_cring R (R Quot I) ((+>) I)"
proof -
  interpret cring R by fact
  show ?thesis
    apply (rule ring_hom_cringI)
      apply (rule rcos_ring_hom_ring)
     apply (rule is_cring)
    apply (rule quotient_is_cring)
   apply (rule is_cring)
   done
qed


subsection ‹Factorization over Prime Ideals›

text ‹The quotient ring generated by a prime ideal is a domain›
lemma (in primeideal) quotient_is_domain: "domain (R Quot I)"
proof -
  have 1: "I +> 𝟭 = I  False"
    using I_notcarr a_rcos_self one_imp_carrier by blast
  have 2: "I +> x = I"
    if  carr: "x  carrier R" "y  carrier R"
    and a: "I +> x  y = I"
    and b: "I +> y  I" for x y
    by (metis I_prime a a_rcos_const a_rcos_self b m_closed that)
  show ?thesis
    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)
qed

text ‹Generating right cosets of a prime ideal is a homomorphism
        on commutative rings›
lemma (in primeideal) rcos_ring_hom_cring: "ring_hom_cring R (R Quot I) ((+>) I)"
  by (rule rcos_ring_hom_cring) (rule is_cring)


subsection ‹Factorization over Maximal Ideals›

text ‹In a commutative ring, the quotient ring over a maximal ideal is a field.
        The proof follows ``W. Adkins, S. Weintraub: Algebra -- An Approach via Module Theory''›
proposition (in maximalideal) quotient_is_field:
  assumes "cring R"
  shows "field (R Quot I)"
proof -
  interpret cring R by fact
  have 1: "𝟬R Quot I 𝟭R Quot I⇙"  ― ‹Quotient is not empty›
  proof
    assume "𝟬R Quot I= 𝟭R Quot I⇙"
    then have II1: "I = I +> 𝟭" by (simp add: FactRing_def)
    then have "I = carrier R"
      using a_rcos_self one_imp_carrier by blast 
    with I_notcarr show False by simp
  qed
  have 2: "ycarrier R. I +> a  y = I +> 𝟭" if IanI: "I +> a  I" and acarr: "a  carrier R" for a
    ― ‹Existence of Inverse›
  proof -
    ― ‹Helper ideal J›
    define J :: "'a set" where "J = (carrier R #> a) <+> I"
    have idealJ: "ideal J R"
      using J_def acarr add_ideals cgenideal_eq_rcos cgenideal_ideal is_ideal by auto
    have IinJ: "I  J"
    proof (clarsimp simp: J_def r_coset_def set_add_defs)
      fix x
      assume xI: "x  I"
      have "x = 𝟬  a  x"
        by (simp add: acarr xI)
      with xI show "xacarrier R. kI. x = xa  a  k" by fast
    qed
    have JnI: "J  I" 
    proof -
      have "a  I"
        using IanI a_rcos_const by blast
      moreover have "a  J"
      proof (simp add: J_def r_coset_def set_add_defs)
        from acarr
        have "a = 𝟭  a  𝟬" by algebra
        with one_closed and additive_subgroup.zero_closed[OF is_additive_subgroup]
        show "xcarrier R. kI. a = x  a  k" by fast
      qed
      ultimately show ?thesis by blast
    qed
    then have Jcarr: "J = carrier R"
      using I_maximal IinJ additive_subgroup.a_subset idealJ ideal_def by blast

    ― ‹Calculating an inverse for terma
    from one_closed[folded Jcarr]
    obtain r i where rcarr: "r  carrier R"
      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]
    have rai1: "a  r = i  𝟭" by algebra

    ― ‹Lifting to cosets›
    from iI have "i  𝟭  I +> 𝟭"
      by (intro a_rcosI, simp, intro a_subset, simp)
    with rai1 have "a  r  I +> 𝟭" by simp
    then have "I +> 𝟭 = I +> a  r"
      by (rule a_repr_independence, simp) (rule a_subgroup)

    from rcarr and this[symmetric]
    show "rcarrier R. I +> a  r = I +> 𝟭" by fast
  qed
  show ?thesis
    apply (intro cring.cring_fieldI2 quotient_is_cring is_cring 1)
     apply (clarsimp simp add: FactRing_simps rcoset_mult_add 2)
    done
qed


lemma (in ring_hom_ring) trivial_hom_iff:
  "(h ` (carrier R) = { 𝟬S}) = (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 = { 𝟬 }"
  shows "inj_on h (carrier R)"
  using group_hom.trivial_ker_imp_inj[OF a_group_hom] assms a_kernel_def[of R S h] by simp 

(* NEW ========================================================================== *)
lemma (in ring_hom_ring) inj_iff_trivial_ker:
  shows "inj_on h (carrier R)  a_kernel R S h = { 𝟬 }"
  using group_hom.inj_iff_trivial_ker[OF a_group_hom] a_kernel_def[of R S h] by simp

(* NEW ========================================================================== *)
corollary ring_hom_in_hom:
  assumes "h  ring_hom R S" shows "h  hom R S" and "h  hom (add_monoid R) (add_monoid S)"
  using assms unfolding ring_hom_def hom_def by auto 

(* NEW ========================================================================== *)
corollary set_add_hom:
  assumes "h  ring_hom R S" "I  carrier R" and "J  carrier R"
  shows "h ` (I <+>RJ) = h ` I <+>Sh ` J"
  using set_mult_hom[OF ring_hom_in_hom(2)[OF assms(1)]] assms(2-3)
  unfolding a_kernel_def[of R S h] set_add_def by simp

(* NEW ========================================================================== *)
corollary a_coset_hom:
  assumes "h  ring_hom R S" "I  carrier R" "a  carrier R"
  shows "h ` (a <+RI) = h a <+S(h ` I)" and "h ` (I +>Ra) = (h ` I) +>Sh a"
  using assms coset_hom[OF ring_hom_in_hom(2)[OF assms(1)], of I a]
  unfolding a_l_coset_def l_coset_eq_set_mult
            a_r_coset_def r_coset_eq_set_mult
  by simp_all

(* NEW ========================================================================== *)
corollary (in ring_hom_ring) set_add_ker_hom:
  assumes "I  carrier R"
  shows "h ` (I <+> (a_kernel R S h)) = h ` I" and "h ` ((a_kernel R S h) <+> I) = h ` I"
  using group_hom.set_mult_ker_hom[OF a_group_hom] assms
  unfolding a_kernel_def[of R S h] set_add_def by simp+

lemma (in ring_hom_ring) non_trivial_field_hom_imp_inj:
  assumes "field R"
  shows "h ` (carrier R)  { 𝟬S}  inj_on h (carrier R)"
proof -
  assume "h ` (carrier R)  { 𝟬S}"
  hence "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
  thus "inj_on h (carrier R)"
    using trivial_ker_imp_inj by blast
qed
lemma "field R  cring R"
  using fieldE(1) by blast

lemma non_trivial_field_hom_is_inj:
  assumes "h  ring_hom R S" and "field R" and "field S" shows "inj_on h (carrier R)"
proof -
  interpret ring_hom_cring R S h
    using assms(1) ring_hom_cring.intro[OF assms(2-3)[THEN fieldE(1)]]
    unfolding symmetric[OF ring_hom_cring_axioms_def] by simp

  have "h 𝟭R= 𝟭S⇙" and "𝟭S 𝟬S⇙"
    using domain.one_not_zero[OF field.axioms(1)[OF assms(3)]] by auto 
  hence "h ` (carrier R)  { 𝟬S}"
    using ring.kernel_zero ring.trivial_hom_iff by fastforce
  thus ?thesis
    using ring.non_trivial_field_hom_imp_inj[OF assms(2)] by simp
qed    

lemma (in ring_hom_ring) img_is_add_subgroup:
  assumes "subgroup H (add_monoid R)"
  shows "subgroup (h ` H) (add_monoid S)"
proof -
  have "group ((add_monoid R)  carrier := H )"
    using assms R.add.subgroup_imp_group by blast
  moreover have "H  carrier R" by (simp add: R.add.subgroupE(1) assms)
  hence "h  hom ((add_monoid R)  carrier := H ) (add_monoid S)"
    unfolding hom_def by (auto simp add: subsetD)
  ultimately have "subgroup (h ` carrier ((add_monoid R)  carrier := H )) (add_monoid S)"
    using group_hom.img_is_subgroup[of "(add_monoid R)  carrier := H " "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  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 R Quot Ix = [mod I:] (I +> j)  (I +> i)"
      unfolding FactRing_def by simp
    hence "a R Quot Ix = 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 R Quot Ix  (+>) I ` J"
      using A i(1) j(1) by (simp add: ideal.I_r_closed)
  
    have "x R Quot Ia = [mod I:] (I +> i)  (I +> j)"
      unfolding FactRing_def i j by simp
    hence "x R Quot Ia = 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 R Quot Ia  (+>) 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  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 "𝟭add_monoid R { r  carrier R. h r  I }"
    by (simp add: additive_subgroup.zero_closed assms ideal.axioms(1))
next
  fix a b
  assume "a  { r  carrier R. h r  I }"
     and "b  { r  carrier R. h r  I }"
  hence a: "a  carrier R" "h a  I"
    and b: "b  carrier R" "h b  I" by auto
  hence "h (a  b) = (h a) S(h b)" using hom_add by blast
  moreover have "(h a) S(h b)  I" using a b assms
    by (simp add: additive_subgroup.a_closed ideal.axioms(1))
  ultimately show "a add_monoid Rb  { r  carrier R. h r  I }"
    using a(1) b (1) by auto

  have "h ( a) = S(h a)" by (simp add: a)
  moreover have "S(h a)  I"
    by (simp add: a(2) additive_subgroup.a_inv_closed assms ideal.axioms(1))
  ultimately show "invadd_monoid Ra  { r  carrier R. h r  I }"
    using a by (simp add: a_inv_def)
next
  fix a r
  assume "a  { r  carrier R. h r  I }" and r: "r  carrier R"
  hence a: "a  carrier R" "h a  I" by auto

  have "h a Sh r  I"
    using assms a r by (simp add: ideal.I_r_closed)
  thus "a  r  { r  carrier R. h r  I }" by (simp add: a(1) r)

  have "h r Sh a  I"
    using assms a r by (simp add: ideal.I_l_closed)
  thus "r  a  { r  carrier R. h r  I }" by (simp add: a(1) r)
qed

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
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
qed

corollary (in ring) quot_ideal_imp_ring_ideal:
  assumes "ideal I R"
  shows "ideal J (R Quot I)  ideal ( J) R"
proof -
  assume A: "ideal J (R Quot I)"
  have " J = { r  carrier R. I +> r  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 ( 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  J) = (J = ( j  J. I +> j))"
proof
  assume A: "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 
next
  assume A: "I  J" show "J = (jJ. I +> j)"
  proof
    show "J  ( j  J. I +> j)"
    proof
      fix j assume j: "j  J"
      have "𝟬  I" by (simp add: additive_subgroup.zero_closed assms(1) ideal.axioms(1))
      hence "𝟬  j  I +> j"
        using a_r_coset_def'[of R I j] by blast
      thus "j  (jJ. I +> j)"
        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)"
      then obtain j where j: "j  J" "x  I +> j" by blast
      then obtain i where i: "i  I" "x = i  j"
        using a_r_coset_def'[of R I j] by blast
      thus "x  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 (λJ. (+>) I ` J) { J. ideal J R  I  J } { J . ideal J (R Quot I) }"
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) })"
    then obtain J' where J': "ideal J' (R Quot I)" "J =  J'" by blast
    hence "ideal J R"
      using assms quot_ideal_imp_ring_ideal by auto
    moreover have "I  J'"
      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) }"
    hence subset: "J'  carrier (R Quot I)  ideal J' (R Quot I)"
      using additive_subgroup.a_subset ideal_def by blast
    hence "((+>) I ` ( J'))  J'"
      using canonical_proj_vimage_in_carrier canonical_proj_vimage_mem_iff
      by (meson assms contra_subsetD image_subsetI)
    moreover have "J'  ((+>) I ` ( J'))"
    proof
      fix x assume "x  J'"
      then obtain r where r: "r  carrier R" "x = I +> r"
        using subset unfolding FactRing_def A_RCOSETS_def'[of R I] by auto
      hence "r  ( J')"
        using x  J' assms canonical_proj_vimage_mem_iff subset by blast
      thus "x  ((+>) I ` ( J'))" using r(2) by blast
    qed
    ultimately show "((+>) I ` ( J')) = J'" by blast
  qed
qed

lemma (in cring) quot_domain_imp_primeideal:
  assumes "ideal P R"
  shows "domain (R Quot P)  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  P"
    proof (rule ccontr)
      assume "¬ carrier R  P" hence "carrier R = P" by simp
      hence "I. I  carrier (R Quot P)  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 "𝟭(R Quot P)= 𝟬(R Quot P)⇙"
        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  carrier R" and b: "b  carrier R" and ab: "a  b  P"
    hence "P +> (a  b) = 𝟬(R Quot P)⇙" unfolding FactRing_def
      by (simp add: a_coset_join2 additive_subgroup.a_subgroup assms ideal.axioms(1))
    moreover have "(P +> a) (R Quot P)(P +> b) = P +> (a  b)" unfolding FactRing_def
      using a b by (simp add: assms ideal.rcoset_mult_add)
    moreover have "P +> a  carrier (R Quot P)  P +> b  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 = 𝟬(R Quot P) P +> b = 𝟬(R Quot P)⇙"
      using domain.integral[OF A, of "P +> a" "P +> b"] by auto
    thus "a  P  b  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 ‹Isomorphism›

definition
  ring_iso :: "_  _  ('a  'b) set"
  where "ring_iso R S = { h. h  ring_hom R S  bij_betw h (carrier R) (carrier S) }"

definition
  is_ring_iso :: "_  _  bool" (infixr "" 60)
  where "R  S = (ring_iso R S  {})"

definition
  morphic_prop :: "_  ('a  bool)  bool"
  where "morphic_prop R P =
           ((P 𝟭R) 
            (r  carrier R. P r) 
            (r1  carrier R. r2  carrier R. P (r1 Rr2)) 
            (r1  carrier R. r2  carrier R. P (r1 Rr2)))"

lemma ring_iso_memI:
  fixes R (structure) and S (structure)
  assumes "x. x  carrier R  h x  carrier S"
      and "x y.  x  carrier R; y  carrier R   h (x  y) = h x Sh y"
      and "x y.  x  carrier R; y  carrier R   h (x  y) = h x Sh y"
      and "h 𝟭 = 𝟭S⇙"
      and "bij_betw h (carrier R) (carrier S)"
  shows "h  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  ring_iso R S"
  shows "x. x  carrier R  h x  carrier S"
   and "x y.  x  carrier R; y  carrier R   h (x  y) = h x Sh y"
   and "x y.  x  carrier R; y  carrier R   h (x  y) = h x Sh y"
   and "h 𝟭 = 𝟭S⇙"
   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 𝟭"
    and "r. r  carrier R  P r"
    and "r1 r2.  r1  carrier R; r2  carrier R   P (r1  r2)"
    and "r1 r2.  r1  carrier R; r2  carrier R   P (r1  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 𝟭"
    and "r. r  carrier R  P r"
    and "r1 r2.  r1  carrier R; r2  carrier R   P (r1  r2)"
    and "r1 r2.  r1  carrier R; r2  carrier R   P (r1  r2)"
  using assms unfolding morphic_prop_def by auto

(* NEW ============================================================================ *)
lemma (in ring) ring_hom_restrict:
  assumes "f  ring_hom R S" and "r. r  carrier R  f r = g r" shows "g  ring_hom R S"
  using assms(2) ring_hom_memE[OF assms(1)] by (auto intro: ring_hom_memI)

(* PROOF ========================================================================== *)
lemma (in ring) ring_iso_restrict:
  assumes "f  ring_iso R S" and "r. r  carrier R  f r = g r" shows "g  ring_iso R S"
proof -
  have hom: "g  ring_hom R S"
    using ring_hom_restrict assms unfolding ring_iso_def by auto 
  have "bij_betw g (carrier R) (carrier S)"
    using bij_betw_cong[of "carrier R" f g] ring_iso_memE(5)[OF assms(1)] assms(2) by simp
  thus ?thesis
    using ring_hom_memE[OF hom] by (auto intro!: ring_iso_memI)
qed

lemma ring_iso_morphic_prop:
  assumes "f  ring_iso R S"
    and "morphic_prop R P"
    and "r. P r  f r = g r"
  shows "g  ring_iso R S"
proof -
  have eq0: "r. r  carrier R  f r = g r"
   and eq1: "f 𝟭R= g 𝟭R⇙"
   and eq2: "r1 r2.  r1  carrier R; r2  carrier R   f (r1 Rr2) = g (r1 Rr2)"
   and eq3: "r1 r2.  r1  carrier R; r2  carrier R   f (r1 Rr2) = g (r1 Rr2)"
    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  ring_hom R S"
  shows "ring (S  carrier := h ` (carrier R), zero := h 𝟬 )" (is "ring ?h_img")
proof -
  have "h  hom (add_monoid R) (add_monoid S)"
    using assms unfolding hom_def ring_hom_def by auto
  hence "comm_group ((add_monoid S)   carrier := h ` (carrier R), one := h 𝟬 )"
    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  hom R S"
    using assms unfolding ring_hom_def hom_def by auto
  hence "monoid (S   carrier := h ` (carrier R), one := h 𝟭 )"
    using hom_imp_img_monoid[of h S] by simp
  hence monoid: "monoid ?h_img"
    using ring_hom_memE(4)[OF assms] unfolding monoid_def by (simp add: monoid.defs)
  show ?thesis
  proof (rule ringI, simp_all add: comm_group_abelian_groupI[OF comm_group] monoid)
    fix x y z assume "x  h ` carrier R" "y  h ` carrier R" "z  h ` carrier R"
    then obtain r1 r2 r3
      where r1: "r1  carrier R" "x = h r1"
        and r2: "r2  carrier R" "y = h r2"
        and r3: "r3  carrier R" "z = h r3" by blast
    hence "(x Sy) Sz = h ((r1  r2)  r3)"
      using ring_hom_memE[OF assms] by auto
    also have " ... = h ((r1  r3)  (r2  r3))"
      using l_distr[OF r1(1) r2(1) r3(1)] by simp
    also have " ... = (x Sz) S(y Sz)"
      using ring_hom_memE[OF assms] r1 r2 r3 by auto
    finally show "(x Sy) Sz = (x Sz) S(y Sz)" .

    have "z S(x Sy) = h (r3  (r1  r2))"
      using ring_hom_memE[OF assms] r1 r2 r3 by auto
    also have " ... =  h ((r3  r1)  (r3  r2))"
      using r_distr[OF r1(1) r2(1) r3(1)] by simp
    also have " ... = (z Sx) S(z Sy)"
      using ring_hom_memE[OF assms] r1 r2 r3 by auto
    finally show "z S(x Sy) = (z Sx) S(z Sy)" .
  qed
qed

lemma (in ring) ring_iso_imp_img_ring:
  assumes "h  ring_iso R S"
  shows "ring (S  zero := h 𝟬 )"
proof -
  have "ring (S  carrier := h ` (carrier R), zero := h 𝟬 )"
    using ring_hom_imp_img_ring[of h S] assms unfolding ring_iso_def by auto
  moreover have "h ` (carrier R) = carrier S"
    using assms unfolding ring_iso_def bij_betw_def by auto
  ultimately show ?thesis by simp
qed

lemma (in cring) ring_iso_imp_img_cring:
  assumes "h  ring_iso R S"
  shows "cring (S  zero := h 𝟬 )" (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  carrier ?h_img" "y  carrier ?h_img"
    then obtain r1 r2
      where r1: "r1  carrier R" "x = h r1"
        and r2: "r2  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 (?h_img)y = h (r1  r2)"
      using assms r1 r2 unfolding ring_iso_def ring_hom_def by auto
    also have " ... = h (r2  r1)"
      using m_comm[OF r1(1) r2(1)] by simp
    also have " ... = y (?h_img)x"
      using assms r1 r2 unfolding ring_iso_def ring_hom_def by auto
    finally show "x (?h_img)y = y (?h_img)x" .
  qed
qed

lemma (in domain) ring_iso_imp_img_domain:
  assumes "h  ring_iso R S"
  shows "domain (S  zero := h 𝟬 )" (is "domain ?h_img")
proof -
  note aux = m_closed integral one_not_zero one_closed zero_closed
  interpret h_img?: cring ?h_img
    using ring_iso_imp_img_cring[OF assms] .
  show ?thesis 
  proof (unfold_locales)
    have "𝟭?h_img= 𝟬?h_img h 𝟭 = h 𝟬"
      using ring_iso_memE(4)[OF assms] by simp
    moreover have "h 𝟭  h 𝟬"
      using ring_iso_memE(5)[OF assms] aux(3-4)
      unfolding bij_betw_def inj_on_def by force
    ultimately show "𝟭?h_img 𝟬?h_img⇙"
      by auto
  next
    fix a b
    assume A: "a ?h_imgb = 𝟬?h_img⇙" "a  carrier ?h_img" "b  carrier ?h_img"
    then obtain r1 r2
      where r1: "r1  carrier R" "a = h r1"
        and r2: "r2  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 ?h_imgb = h (r1  r2)"
      using assms r1 r2 unfolding ring_iso_def ring_hom_def by auto
    hence "h (r1  r2) = h 𝟬"
      using A(1) by simp
    hence "r1  r2 = 𝟬"
      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 = 𝟬  r2 = 𝟬"
      using aux(2)[OF _ r1(1) r2(1)] by simp
    thus "a = 𝟬?h_img b = 𝟬?h_img⇙"
      unfolding r1 r2 by auto
  qed
qed

lemma (in field) ring_iso_imp_img_field:
  assumes "h  ring_iso R S"
  shows "field (S  zero := h 𝟬 )" (is "field ?h_img")
proof -
  interpret h_img?: domain ?h_img
    using ring_iso_imp_img_domain[OF assms] .
  show ?thesis
  proof (unfold_locales, auto simp add: Units_def)
    interpret field R using field_axioms .
    fix a assume a: "a  carrier S" "a Sh 𝟬 = 𝟭S⇙"
    then obtain r where r: "r  carrier R" "a = h r"
      using assms image_iff[where ?f = h and ?A = "carrier R"]
      unfolding ring_iso_def bij_betw_def by auto
    have "a Sh 𝟬 = h (r  𝟬)" unfolding r(2)
      using ring_iso_memE(2)[OF assms r(1)] by simp
    hence "h 𝟭 = h 𝟬"
      using ring_iso_memE(4)[OF assms] r(1) a(2) by simp
    thus False
      using ring_iso_memE(5)[OF assms]
      unfolding bij_betw_def inj_on_def by force
  next
    interpret field R using field_axioms .
    fix s assume s: "s  carrier S" "s  h 𝟬"
    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 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) Sh r = h 𝟭" and "h r Sh (inv r) = h 𝟭"
      using ring_iso_memE(2)[OF assms inv_r(1) r(1)] inv_r(3-4)
            ring_iso_memE(2)[OF assms r(1) inv_r(1)] by auto
    thus "s'  carrier S. s' Ss = 𝟭S s Ss' = 𝟭S⇙"
      using ring_iso_memE(1,4)[OF assms] inv_r(1) r(2) by auto
  qed
qed

lemma ring_iso_same_card: "R  S  card (carrier R) = card (carrier S)"
  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 

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 

corollary ring_iso_trans: " R  S; S  Q   R  Q"
  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"
  shows "(inv_into (carrier R) h)  ring_iso S R"
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"
    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)"
    by (simp add: bij_betw_inv_into h ring_iso_memE(5))

  have "inv_into (carrier R) h  ring_hom S R"
    using ring_iso_memE [OF h] bij_betwE [OF h_inv_bij] ring R
    by (simp add: bij_betw_imp_inj_on bij_betw_inv_into_right inv_into_f_eq ring.ring_simprules ring_hom_memI)
  moreover have "bij_betw (inv_into (carrier R) h) (carrier S) (carrier R)"
    using h_inv_bij by force
    ultimately show "inv_into (carrier R) h  ring_iso S R"
      by (simp add: ring_iso_def)
qed

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 

lemma (in ring_hom_ring) the_elem_simp [simp]:
  "x. x  carrier R  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)"
    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"
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"
    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)"
    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:
  "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

lemma (in ring_hom_ring) the_elem_wf:
  "X. X  carrier (R Quot (a_kernel R S h))  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"
    using quot_mem by blast
  hence "x'. x'  X  h x' = h x"
  proof -
    fix x' assume "x'  X" hence "x'  (a_kernel R S h) +> x" using 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
          abelian_subgroup.a_rcos_module_imp abelian_subgroup_a_kernel x)
    hence "h x' = h k Sh 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  h ` X"
    by (simp add: X homeq_imp_rcos x)
  ultimately have "(h ` X) = { h x }"
    by blast
  thus "y  carrier S. (h ` X) = { y }" using x by simp
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) 

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 ` 𝟭R Quot a_kernel R S h) = 𝟭S⇙"
    unfolding FactRing_def  using the_elem_simp[of "𝟭R⇙"] by simp

  fix X Y
  assume "X  carrier (R Quot a_kernel R S h)"
     and "Y  carrier (R Quot a_kernel R S h)"
  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"
    using quot_mem by blast

  have "X R Quot a_kernel R S hY = (a_kernel R S h) +> (x  y)"
    by (simp add: FactRing_def ideal.rcoset_mult_add kernel_is_ideal x y)
  thus "the_elem (h ` (X R Quot a_kernel R S hY)) = the_elem (h ` X) Sthe_elem (h ` Y)"
    by (simp add: x y)

  have "X R Quot a_kernel R S hY = (a_kernel R S h) +> (x  y)"
    using ideal.rcos_ring_hom kernel_is_ideal ring_hom_add x y by fastforce
  thus "the_elem (h ` (X R Quot a_kernel R S hY)) = the_elem (h ` X) Sthe_elem (h ` Y)"
    by (simp add: x y)
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))"
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"
    then obtain x where x: "x  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  carrier (R Quot (a_kernel R S h))"
     unfolding FactRing_simps by (auto simp add: x a_r_coset_def)
    ultimately show "y  (λ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:
  "(λ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))"
    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
qed

theorem (in ring_hom_ring) FactRing_iso_set:
  assumes "h ` carrier R = carrier S"
  shows "(λX. the_elem (h ` X))  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)  S"
  using FactRing_iso_set assms is_ring_iso_def by auto

corollary (in ring) FactRing_zeroideal:
  shows "R Quot { 𝟬 }  R" and "R  R Quot { 𝟬 }"
proof -
  have "ring_hom_ring R R id"
    using ring_axioms by (auto intro: ring_hom_ringI)
  moreover have "a_kernel R R id = { 𝟬 }"
    unfolding a_kernel_def' by auto
  ultimately show "R Quot { 𝟬 }  R" and "R  R Quot { 𝟬 }"
    using ring_hom_ring.FactRing_iso[of R R id]
          ring_iso_sym[OF ideal.quotient_is_ring[OF zeroideal], of R] by auto
qed

lemma (in ring_hom_ring) img_is_ring: "ring (S  carrier := h ` (carrier R) )"
proof -
  let ?the_elem = "λ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  carrier := ?the_elem ` (carrier (R Quot (a_kernel R S h))) )
                     zero := ?the_elem 𝟬(R Quot (a_kernel R S h)))"
    using ring.ring_iso_imp_img_ring[OF FactRing_is_ring, of ?the_elem
          "S  carrier := ?the_elem ` (carrier (R Quot (a_kernel R S h))) "]
          FactRing_iso_set_aux the_elem_surj by auto

  moreover
  have "𝟬  (a_kernel R S h)"
    using a_kernel_def'[of R S h] by auto
  hence "𝟭  (a_kernel R S h) +> 𝟭"
    using a_r_coset_def'[of R "a_kernel R S h" 𝟭] by force
  hence "𝟭S (h ` ((a_kernel R S h) +> 𝟭))"
    using hom_one by force
  hence "?the_elem 𝟭(R Quot (a_kernel R S h))= 𝟭S⇙"
    using the_elem_wf[of "(a_kernel R S h) +> 𝟭"] by (simp add: FactRing_def)
  
  moreover
  have "𝟬S (h ` (a_kernel R S h))"
    using a_kernel_def'[of R S h] hom_zero by force
  hence "𝟬S (h ` 𝟬(R Quot (a_kernel R S h)))"
    by (simp add: FactRing_def)
  hence "?the_elem 𝟬(R Quot (a_kernel R S h))= 𝟬S⇙"
    using the_elem_wf[OF ring.ring_simprules(2)[OF FactRing_is_ring]]
    by (metis singletonD the_elem_eq) 

  ultimately
  have "ring ((S  carrier := h ` (carrier R) )  one := 𝟭S, zero := 𝟬S)"
    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  carrier := h ` (carrier R) )"
proof -
  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
qed

lemma (in ring_hom_ring) img_is_domain:
  assumes "domain S"
  shows "domain (S  carrier := h ` (carrier R) )"
proof -
  interpret cring "S  carrier := h ` (carrier R) "
    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  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
  have "ring_hom_ring R (S Quot P) (((+>S) P)  h)" (is "ring_hom_ring ?A ?B ?h")
    using ring_hom_trans[OF homh, of "(+>S) 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" "((+>S) 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
  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) (λ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  carrier R. h r  P }"
  proof
    show "a_kernel R (S Quot P) ?h  { r  carrier R. h r  P }"
    proof 
      fix r assume "r  a_kernel R (S Quot P) ?h"
      hence r: "r  carrier R" "P +>S(h r) = P"
        unfolding a_kernel_def kernel_def FactRing_def by auto
      hence "h r  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  { r  carrier R. h r  P }" using r by simp
    qed
  next
    show "{ r  carrier R. h r  P }  a_kernel R (S Quot P) ?h"
    proof
      fix r assume "r  { r  carrier R. h r  P }"
      hence r: "r  carrier R" "h r  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  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  carrier R. h r  P } R" by simp
qed

end