Theory QuotRing

Up to index of Isabelle/HOL/HOL-Algebra

theory QuotRing
imports RingHom

(*  Title:      HOL/Algebra/QuotRing.thy
Author: Stephan Hohe
*)


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 _:] _ \<Otimes>\<index> _" [81,81,81] 80)
where "rcoset_mult R I A B = (\<Union>a∈A. \<Union>b∈B. I +>R (a ⊗R b))"



text {* @{const "rcoset_mult"} fulfils the properties required by
congruences *}

lemma (in ideal) rcoset_mult_add:
"[|x ∈ carrier R; y ∈ carrier R|] ==> [mod I:] (I +> x) \<Otimes> (I +> y) = I +> (x ⊗ y)"

apply rule
apply (rule, simp add: rcoset_mult_def, clarsimp)
defer 1
apply (rule, simp add: rcoset_mult_def)
defer 1
proof -
fix z x' y'
assume carr: "x ∈ carrier R" "y ∈ carrier R"
and x'rcos: "x' ∈ I +> x"
and y'rcos: "y' ∈ I +> y"
and zrcos: "z ∈ I +> x' ⊗ y'"


from x'rcos
have "∃h∈I. x' = h ⊕ x" by (simp add: a_r_coset_def r_coset_def)
from this obtain hx
where hxI: "hx ∈ I"
and x': "x' = hx ⊕ x"

by fast+

from y'rcos
have "∃h∈I. y' = h ⊕ y" by (simp add: a_r_coset_def r_coset_def)
from this
obtain hy
where hyI: "hy ∈ I"
and y': "y' = hy ⊕ y"

by fast+

from zrcos
have "∃h∈I. z = h ⊕ (x' ⊗ y')" by (simp add: a_r_coset_def r_coset_def)
from this
obtain hz
where hzI: "hz ∈ I"
and z: "z = hz ⊕ (x' ⊗ y')"

by fast+

note carr = carr hxI[THEN a_Hcarr] hyI[THEN a_Hcarr] hzI[THEN a_Hcarr]

from z have "z = hz ⊕ (x' ⊗ y')" .
also from x' y'
have "… = 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)

from this and z2
have "∃h∈I. z = h ⊕ x ⊗ y" by fast
thus "z ∈ I +> x ⊗ y" by (simp add: a_r_coset_def r_coset_def)
next
fix z
assume xcarr: "x ∈ carrier R"
and ycarr: "y ∈ carrier R"
and zrcos: "z ∈ I +> x ⊗ y"

from xcarr
have xself: "x ∈ I +> x" by (intro a_rcos_self)
from ycarr
have yself: "y ∈ I +> y" by (intro a_rcos_self)

from xself and yself and zrcos
show "∃a∈I +> x. ∃b∈I +> y. z ∈ I +> a ⊗ b" by fast
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_rcosetsR I, mult = rcoset_mult R I, one = (I +>R \<one>R), zero = I, add = set_add R|)),"



subsection {* Factorization over General Ideals *}

text {* The quotient is a ring *}
lemma (in ideal) quotient_is_ring:
shows "ring (R Quot I)"

apply (rule ringI)
--{* abelian group *}

apply (rule comm_group_abelian_groupI)
apply (simp add: FactRing_def)
apply (rule a_factorgroup_is_comm_group[unfolded A_FactGroup_def'])
--{* mult monoid *}

apply (rule monoidI)
apply (simp_all add: FactRing_def A_RCOSETS_def RCOSETS_def
a_r_coset_def[symmetric])
--{* mult closed *}

apply (clarify)
apply (simp add: rcoset_mult_add, fast)
--{* mult @{text one_closed} *}

apply (force intro: one_closed)
--{* mult assoc *}

apply clarify
apply (simp add: rcoset_mult_add m_assoc)
--{* mult one *}

apply clarify
apply (simp add: rcoset_mult_add l_one)
apply clarify
apply (simp add: rcoset_mult_add r_one)
--{* distr *}

apply clarify
apply (simp add: rcoset_mult_add a_rcos_sum l_distr)
apply clarify
apply (simp add: rcoset_mult_add a_rcos_sum r_distr)
done


text {* This is a ring homomorphism *}

lemma (in ideal) rcos_ring_hom:
"(op +> I) ∈ ring_hom R (R Quot I)"

apply (rule ring_hom_memI)
apply (simp add: FactRing_def a_rcosetsI[OF a_subset])
apply (simp add: FactRing_def rcoset_mult_add)
apply (simp add: FactRing_def a_rcos_sum)
apply (simp add: FactRing_def)
done

lemma (in ideal) rcos_ring_hom_ring:
"ring_hom_ring R (R Quot I) (op +> I)"

apply (rule ring_hom_ringI)
apply (rule is_ring, rule quotient_is_ring)
apply (simp add: FactRing_def a_rcosetsI[OF a_subset])
apply (simp add: FactRing_def rcoset_mult_add)
apply (simp add: FactRing_def a_rcos_sum)
apply (simp add: FactRing_def)
done

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)
apply (rule quotient_is_ring)
apply (rule ring.axioms[OF quotient_is_ring])
apply (simp add: FactRing_def A_RCOSETS_defs a_r_coset_def[symmetric])
apply clarify
apply (simp add: 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) (op +> 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:
shows "domain (R Quot I)"

apply (rule domain.intro)
apply (rule quotient_is_cring, rule is_cring)
apply (rule domain_axioms.intro)
apply (simp add: FactRing_def) defer 1
apply (simp add: FactRing_def A_RCOSETS_defs a_r_coset_def[symmetric], clarify)
apply (simp add: rcoset_mult_add) defer 1
proof (rule ccontr, clarsimp)
assume "I +> \<one> = I"
hence "\<one> ∈ I" by (simp only: a_coset_join1 one_closed a_subgroup)
hence "carrier R ⊆ I" by (subst one_imp_carrier, simp, fast)
from this and a_subset
have "I = carrier R" by fast
from this and I_notcarr
show "False" by fast
next
fix x y
assume carr: "x ∈ carrier R" "y ∈ carrier R"
and a: "I +> x ⊗ y = I"
and b: "I +> y ≠ I"


have ynI: "y ∉ I"
proof (rule ccontr, simp)
assume "y ∈ I"
hence "I +> y = I" by (rule a_rcos_const)
from this and b
show "False" by simp
qed

from carr
have "x ⊗ y ∈ I +> x ⊗ y" by (simp add: a_rcos_self)
from this
have xyI: "x ⊗ y ∈ I" by (simp add: a)

from xyI and carr
have xI: "x ∈ I ∨ y ∈ I" by (simp add: I_prime)
from this and ynI
have "x ∈ I" by fast
thus "I +> x = I" by (rule a_rcos_const)
qed

text {* Generating right cosets of a prime ideal is a homomorphism
on commutative rings *}

lemma (in primeideal) rcos_ring_hom_cring:
shows "ring_hom_cring R (R Quot I) (op +> 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'' *}

lemma (in maximalideal) quotient_is_field:
assumes "cring R"
shows "field (R Quot I)"

proof -
interpret cring R by fact
show ?thesis apply (intro cring.cring_fieldI2)
apply (rule quotient_is_cring, rule is_cring)
defer 1
apply (simp add: FactRing_def A_RCOSETS_defs a_r_coset_def[symmetric], clarsimp)
apply (simp add: rcoset_mult_add) defer 1
proof (rule ccontr, simp)
--{* Quotient is not empty *}

assume "\<zero>R Quot I = \<one>R Quot I"
hence II1: "I = I +> \<one>" by (simp add: FactRing_def)
from a_rcos_self[OF one_closed]
have "\<one> ∈ I" by (simp add: II1[symmetric])
hence "I = carrier R" by (rule one_imp_carrier)
from this and I_notcarr
show "False" by simp
next
--{* Existence of Inverse *}

fix a
assume IanI: "I +> a ≠ I"
and acarr: "a ∈ carrier R"

--{* Helper ideal @{text "J"} *}

def J "(carrier R #> a) <+> I :: 'a set"
have idealJ: "ideal J R"
apply (unfold J_def, rule add_ideals)
apply (simp only: cgenideal_eq_rcos[symmetric], rule cgenideal_ideal, rule acarr)
apply (rule is_ideal)
done

--{* Showing @{term "J"} not smaller than @{term "I"} *}

have IinJ: "I ⊆ J"
proof (rule, simp add: J_def r_coset_def set_add_defs)
fix x
assume xI: "x ∈ I"
have Zcarr: "\<zero> ∈ carrier R" by fast
from xI[THEN a_Hcarr] acarr
have "x = \<zero> ⊗ a ⊕ x" by algebra

from Zcarr and xI and this
show "∃xa∈carrier R. ∃k∈I. x = xa ⊗ a ⊕ k" by fast
qed

--{* Showing @{term "J ≠ I"} *}

have anI: "a ∉ I"
proof (rule ccontr, simp)
assume "a ∈ I"
hence "I +> a = I" by (rule a_rcos_const)
from this and IanI
show "False" by simp
qed

have aJ: "a ∈ J"
proof (simp add: J_def r_coset_def set_add_defs)
from acarr
have "a = \<one> ⊗ a ⊕ \<zero>" by algebra
from one_closed and additive_subgroup.zero_closed[OF is_additive_subgroup] and this
show "∃x∈carrier R. ∃k∈I. a = x ⊗ a ⊕ k" by fast
qed

from aJ and anI
have JnI: "J ≠ I" by fast

--{* Deducing @{term "J = carrier R"} because @{term "I"} is maximal *}

from idealJ and IinJ
have "J = I ∨ J = carrier R"
proof (rule I_maximal, unfold J_def)
have "carrier R #> a ⊆ carrier R"
using subset_refl acarr
by (rule r_coset_subset_G)
from this and a_subset
show "carrier R #> a <+> I ⊆ carrier R" by (rule set_add_closed)
qed

from this and JnI
have Jcarr: "J = carrier R" by simp

--{* Calculating an inverse for @{term "a"} *}

from one_closed[folded Jcarr]
have "∃r∈carrier R. ∃i∈I. \<one> = r ⊗ a ⊕ i"
by (simp add: J_def r_coset_def set_add_defs)
from this
obtain r i
where rcarr: "r ∈ carrier R"
and iI: "i ∈ I"
and one: "\<one> = r ⊗ a ⊕ i"

by fast
from one and rcarr and acarr and iI[THEN a_Hcarr]
have rai1: "a ⊗ r = \<ominus>i ⊕ \<one>" by algebra

--{* Lifting to cosets *}

from iI
have "\<ominus>i ⊕ \<one> ∈ I +> \<one>"
by (intro a_rcosI, simp, intro a_subset, simp)
from this and rai1
have "a ⊗ r ∈ I +> \<one>" by simp
from this have "I +> \<one> = I +> a ⊗ r"
by (rule a_repr_independence, simp) (rule a_subgroup)

from rcarr and this[symmetric]
show "∃r∈carrier R. I +> a ⊗ r = I +> \<one>" by fast
qed
qed

end