| author | wenzelm | 
| Sun, 13 Mar 2016 12:39:12 +0100 | |
| changeset 62611 | dc7cc407c911 | 
| parent 61382 | efac889fccbc | 
| child 63040 | eb4ddd18d635 | 
| permissions | -rw-r--r-- | 
| 35849 | 1 | (* Title: HOL/Algebra/QuotRing.thy | 
| 2 | Author: Stephan Hohe | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 3 | *) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 4 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 5 | theory QuotRing | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 6 | imports RingHom | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 7 | begin | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 8 | |
| 61382 | 9 | section \<open>Quotient Rings\<close> | 
| 27717 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 ballarin parents: 
27611diff
changeset | 10 | |
| 61382 | 11 | subsection \<open>Multiplication on Cosets\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 12 | |
| 45005 | 13 | definition rcoset_mult :: "[('a, _) ring_scheme, 'a set, 'a set, 'a set] \<Rightarrow> 'a set"
 | 
| 23463 | 14 |     ("[mod _:] _ \<Otimes>\<index> _" [81,81,81] 80)
 | 
| 35848 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 wenzelm parents: 
35847diff
changeset | 15 | where "rcoset_mult R I A B = (\<Union>a\<in>A. \<Union>b\<in>B. I +>\<^bsub>R\<^esub> (a \<otimes>\<^bsub>R\<^esub> b))" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 16 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 17 | |
| 61382 | 18 | text \<open>@{const "rcoset_mult"} fulfils the properties required by
 | 
| 19 | congruences\<close> | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 20 | lemma (in ideal) rcoset_mult_add: | 
| 45005 | 21 | "x \<in> carrier R \<Longrightarrow> y \<in> carrier R \<Longrightarrow> [mod I:] (I +> x) \<Otimes> (I +> y) = I +> (x \<otimes> y)" | 
| 22 | apply rule | |
| 23 | apply (rule, simp add: rcoset_mult_def, clarsimp) | |
| 24 | defer 1 | |
| 25 | apply (rule, simp add: rcoset_mult_def) | |
| 26 | defer 1 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 27 | proof - | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 28 | fix z x' y' | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 29 | assume carr: "x \<in> carrier R" "y \<in> carrier R" | 
| 45005 | 30 | and x'rcos: "x' \<in> I +> x" | 
| 31 | and y'rcos: "y' \<in> I +> y" | |
| 32 | and zrcos: "z \<in> I +> x' \<otimes> y'" | |
| 33 | ||
| 34 | from x'rcos have "\<exists>h\<in>I. x' = h \<oplus> x" | |
| 35 | by (simp add: a_r_coset_def r_coset_def) | |
| 36 | then obtain hx where hxI: "hx \<in> I" and x': "x' = hx \<oplus> x" | |
| 37 | by fast+ | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 38 | |
| 45005 | 39 | from y'rcos have "\<exists>h\<in>I. y' = h \<oplus> y" | 
| 40 | by (simp add: a_r_coset_def r_coset_def) | |
| 41 | then obtain hy where hyI: "hy \<in> I" and y': "y' = hy \<oplus> y" | |
| 42 | by fast+ | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 43 | |
| 45005 | 44 | from zrcos have "\<exists>h\<in>I. z = h \<oplus> (x' \<otimes> y')" | 
| 45 | by (simp add: a_r_coset_def r_coset_def) | |
| 46 | then obtain hz where hzI: "hz \<in> I" and z: "z = hz \<oplus> (x' \<otimes> y')" | |
| 47 | by fast+ | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 48 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 49 | note carr = carr hxI[THEN a_Hcarr] hyI[THEN a_Hcarr] hzI[THEN a_Hcarr] | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 50 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 51 | from z have "z = hz \<oplus> (x' \<otimes> y')" . | 
| 45005 | 52 | also from x' y' have "\<dots> = hz \<oplus> ((hx \<oplus> x) \<otimes> (hy \<oplus> y))" by simp | 
| 53 | also from carr have "\<dots> = (hz \<oplus> (hx \<otimes> (hy \<oplus> y)) \<oplus> x \<otimes> hy) \<oplus> x \<otimes> y" by algebra | |
| 54 | finally have z2: "z = (hz \<oplus> (hx \<otimes> (hy \<oplus> y)) \<oplus> x \<otimes> hy) \<oplus> x \<otimes> y" . | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 55 | |
| 45005 | 56 | from hxI hyI hzI carr have "hz \<oplus> (hx \<otimes> (hy \<oplus> y)) \<oplus> x \<otimes> hy \<in> I" | 
| 57 | by (simp add: I_l_closed I_r_closed) | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 58 | |
| 45005 | 59 | with z2 have "\<exists>h\<in>I. z = h \<oplus> x \<otimes> y" by fast | 
| 60 | then show "z \<in> I +> x \<otimes> y" by (simp add: a_r_coset_def r_coset_def) | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 61 | next | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 62 | fix z | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 63 | assume xcarr: "x \<in> carrier R" | 
| 45005 | 64 | and ycarr: "y \<in> carrier R" | 
| 65 | and zrcos: "z \<in> I +> x \<otimes> y" | |
| 66 | from xcarr have xself: "x \<in> I +> x" by (intro a_rcos_self) | |
| 67 | from ycarr have yself: "y \<in> I +> y" by (intro a_rcos_self) | |
| 68 | show "\<exists>a\<in>I +> x. \<exists>b\<in>I +> y. z \<in> I +> a \<otimes> b" | |
| 69 | using xself and yself and zrcos by fast | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 70 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 71 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 72 | |
| 61382 | 73 | subsection \<open>Quotient Ring Definition\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 74 | |
| 45005 | 75 | definition FactRing :: "[('a,'b) ring_scheme, 'a set] \<Rightarrow> ('a set) ring"
 | 
| 76 | (infixl "Quot" 65) | |
| 35848 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 wenzelm parents: 
35847diff
changeset | 77 | where "FactRing R I = | 
| 45005 | 78 | \<lparr>carrier = a_rcosets\<^bsub>R\<^esub> I, mult = rcoset_mult R I, | 
| 79 | one = (I +>\<^bsub>R\<^esub> \<one>\<^bsub>R\<^esub>), zero = I, add = set_add R\<rparr>" | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 80 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 81 | |
| 61382 | 82 | subsection \<open>Factorization over General Ideals\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 83 | |
| 61382 | 84 | text \<open>The quotient is a ring\<close> | 
| 45005 | 85 | lemma (in ideal) quotient_is_ring: "ring (R Quot I)" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 86 | apply (rule ringI) | 
| 61382 | 87 | --\<open>abelian group\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 88 | apply (rule comm_group_abelian_groupI) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 89 | apply (simp add: FactRing_def) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 90 | apply (rule a_factorgroup_is_comm_group[unfolded A_FactGroup_def']) | 
| 61382 | 91 | --\<open>mult monoid\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 92 | apply (rule monoidI) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 93 | apply (simp_all add: FactRing_def A_RCOSETS_def RCOSETS_def | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 94 | a_r_coset_def[symmetric]) | 
| 61382 | 95 | --\<open>mult closed\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 96 | apply (clarify) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 97 | apply (simp add: rcoset_mult_add, fast) | 
| 61382 | 98 |      --\<open>mult @{text one_closed}\<close>
 | 
| 45005 | 99 | apply force | 
| 61382 | 100 | --\<open>mult assoc\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 101 | apply clarify | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 102 | apply (simp add: rcoset_mult_add m_assoc) | 
| 61382 | 103 | --\<open>mult one\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 104 | apply clarify | 
| 45005 | 105 | apply (simp add: rcoset_mult_add) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 106 | apply clarify | 
| 45005 | 107 | apply (simp add: rcoset_mult_add) | 
| 61382 | 108 | --\<open>distr\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 109 | apply clarify | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 110 | apply (simp add: rcoset_mult_add a_rcos_sum l_distr) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 111 | apply clarify | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 112 | apply (simp add: rcoset_mult_add a_rcos_sum r_distr) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 113 | done | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 114 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 115 | |
| 61382 | 116 | text \<open>This is a ring homomorphism\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 117 | |
| 45005 | 118 | lemma (in ideal) rcos_ring_hom: "(op +> I) \<in> ring_hom R (R Quot I)" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 119 | apply (rule ring_hom_memI) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 120 | apply (simp add: FactRing_def a_rcosetsI[OF a_subset]) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 121 | apply (simp add: FactRing_def rcoset_mult_add) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 122 | apply (simp add: FactRing_def a_rcos_sum) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 123 | apply (simp add: FactRing_def) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 124 | done | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 125 | |
| 45005 | 126 | lemma (in ideal) rcos_ring_hom_ring: "ring_hom_ring R (R Quot I) (op +> I)" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 127 | apply (rule ring_hom_ringI) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 128 | apply (rule is_ring, rule quotient_is_ring) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 129 | apply (simp add: FactRing_def a_rcosetsI[OF a_subset]) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 130 | apply (simp add: FactRing_def rcoset_mult_add) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 131 | apply (simp add: FactRing_def a_rcos_sum) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 132 | apply (simp add: FactRing_def) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 133 | done | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 134 | |
| 61382 | 135 | text \<open>The quotient of a cring is also commutative\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 136 | lemma (in ideal) quotient_is_cring: | 
| 27611 | 137 | assumes "cring R" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 138 | shows "cring (R Quot I)" | 
| 27611 | 139 | proof - | 
| 29237 | 140 | interpret cring R by fact | 
| 45005 | 141 | show ?thesis | 
| 142 | apply (intro cring.intro comm_monoid.intro comm_monoid_axioms.intro) | |
| 143 | apply (rule quotient_is_ring) | |
| 144 | apply (rule ring.axioms[OF quotient_is_ring]) | |
| 145 | apply (simp add: FactRing_def A_RCOSETS_defs a_r_coset_def[symmetric]) | |
| 146 | apply clarify | |
| 147 | apply (simp add: rcoset_mult_add m_comm) | |
| 148 | done | |
| 27611 | 149 | qed | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 150 | |
| 61382 | 151 | text \<open>Cosets as a ring homomorphism on crings\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 152 | lemma (in ideal) rcos_ring_hom_cring: | 
| 27611 | 153 | assumes "cring R" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 154 | shows "ring_hom_cring R (R Quot I) (op +> I)" | 
| 27611 | 155 | proof - | 
| 29237 | 156 | interpret cring R by fact | 
| 45005 | 157 | show ?thesis | 
| 158 | apply (rule ring_hom_cringI) | |
| 159 | apply (rule rcos_ring_hom_ring) | |
| 160 | apply (rule is_cring) | |
| 161 | apply (rule quotient_is_cring) | |
| 162 | apply (rule is_cring) | |
| 163 | done | |
| 27611 | 164 | qed | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 165 | |
| 35849 | 166 | |
| 61382 | 167 | subsection \<open>Factorization over Prime Ideals\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 168 | |
| 61382 | 169 | text \<open>The quotient ring generated by a prime ideal is a domain\<close> | 
| 45005 | 170 | lemma (in primeideal) quotient_is_domain: "domain (R Quot I)" | 
| 171 | apply (rule domain.intro) | |
| 172 | apply (rule quotient_is_cring, rule is_cring) | |
| 173 | apply (rule domain_axioms.intro) | |
| 174 | apply (simp add: FactRing_def) defer 1 | |
| 175 | apply (simp add: FactRing_def A_RCOSETS_defs a_r_coset_def[symmetric], clarify) | |
| 176 | apply (simp add: rcoset_mult_add) defer 1 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 177 | proof (rule ccontr, clarsimp) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 178 | assume "I +> \<one> = I" | 
| 45005 | 179 | then have "\<one> \<in> I" by (simp only: a_coset_join1 one_closed a_subgroup) | 
| 180 | then have "carrier R \<subseteq> I" by (subst one_imp_carrier, simp, fast) | |
| 181 | with a_subset have "I = carrier R" by fast | |
| 182 | with I_notcarr show False by fast | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 183 | next | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 184 | fix x y | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 185 | assume carr: "x \<in> carrier R" "y \<in> carrier R" | 
| 45005 | 186 | and a: "I +> x \<otimes> y = I" | 
| 187 | and b: "I +> y \<noteq> I" | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 188 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 189 | have ynI: "y \<notin> I" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 190 | proof (rule ccontr, simp) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 191 | assume "y \<in> I" | 
| 45005 | 192 | then have "I +> y = I" by (rule a_rcos_const) | 
| 193 | with b show False by simp | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 194 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 195 | |
| 45005 | 196 | from carr have "x \<otimes> y \<in> I +> x \<otimes> y" by (simp add: a_rcos_self) | 
| 197 | then have xyI: "x \<otimes> y \<in> I" by (simp add: a) | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 198 | |
| 45005 | 199 | from xyI and carr have xI: "x \<in> I \<or> y \<in> I" by (simp add: I_prime) | 
| 200 | with ynI have "x \<in> I" by fast | |
| 201 | then show "I +> x = I" by (rule a_rcos_const) | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 202 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 203 | |
| 61382 | 204 | text \<open>Generating right cosets of a prime ideal is a homomorphism | 
| 205 | on commutative rings\<close> | |
| 45005 | 206 | lemma (in primeideal) rcos_ring_hom_cring: "ring_hom_cring R (R Quot I) (op +> I)" | 
| 207 | by (rule rcos_ring_hom_cring) (rule is_cring) | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 208 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 209 | |
| 61382 | 210 | subsection \<open>Factorization over Maximal Ideals\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 211 | |
| 61382 | 212 | text \<open>In a commutative ring, the quotient ring over a maximal ideal | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 213 | is a field. | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 214 | The proof follows ``W. Adkins, S. Weintraub: Algebra -- | 
| 61382 | 215 | An Approach via Module Theory''\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 216 | lemma (in maximalideal) quotient_is_field: | 
| 27611 | 217 | assumes "cring R" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 218 | shows "field (R Quot I)" | 
| 27611 | 219 | proof - | 
| 29237 | 220 | interpret cring R by fact | 
| 45005 | 221 | show ?thesis | 
| 222 | apply (intro cring.cring_fieldI2) | |
| 223 | apply (rule quotient_is_cring, rule is_cring) | |
| 224 | defer 1 | |
| 225 | apply (simp add: FactRing_def A_RCOSETS_defs a_r_coset_def[symmetric], clarsimp) | |
| 226 | apply (simp add: rcoset_mult_add) defer 1 | |
| 227 | proof (rule ccontr, simp) | |
| 61382 | 228 | --\<open>Quotient is not empty\<close> | 
| 45005 | 229 | assume "\<zero>\<^bsub>R Quot I\<^esub> = \<one>\<^bsub>R Quot I\<^esub>" | 
| 230 | then have II1: "I = I +> \<one>" by (simp add: FactRing_def) | |
| 231 | from a_rcos_self[OF one_closed] have "\<one> \<in> I" | |
| 232 | by (simp add: II1[symmetric]) | |
| 233 | then have "I = carrier R" by (rule one_imp_carrier) | |
| 234 | with I_notcarr show False by simp | |
| 235 | next | |
| 61382 | 236 | --\<open>Existence of Inverse\<close> | 
| 45005 | 237 | fix a | 
| 238 | assume IanI: "I +> a \<noteq> I" and acarr: "a \<in> carrier R" | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 239 | |
| 61382 | 240 |     --\<open>Helper ideal @{text "J"}\<close>
 | 
| 45005 | 241 | def J \<equiv> "(carrier R #> a) <+> I :: 'a set" | 
| 242 | have idealJ: "ideal J R" | |
| 243 | apply (unfold J_def, rule add_ideals) | |
| 244 | apply (simp only: cgenideal_eq_rcos[symmetric], rule cgenideal_ideal, rule acarr) | |
| 245 | apply (rule is_ideal) | |
| 246 | done | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 247 | |
| 61382 | 248 |     --\<open>Showing @{term "J"} not smaller than @{term "I"}\<close>
 | 
| 45005 | 249 | have IinJ: "I \<subseteq> J" | 
| 250 | proof (rule, simp add: J_def r_coset_def set_add_defs) | |
| 251 | fix x | |
| 252 | assume xI: "x \<in> I" | |
| 253 | have Zcarr: "\<zero> \<in> carrier R" by fast | |
| 254 | from xI[THEN a_Hcarr] acarr | |
| 255 | have "x = \<zero> \<otimes> a \<oplus> x" by algebra | |
| 256 | with Zcarr and xI show "\<exists>xa\<in>carrier R. \<exists>k\<in>I. x = xa \<otimes> a \<oplus> k" by fast | |
| 257 | qed | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 258 | |
| 61382 | 259 |     --\<open>Showing @{term "J \<noteq> I"}\<close>
 | 
| 45005 | 260 | have anI: "a \<notin> I" | 
| 261 | proof (rule ccontr, simp) | |
| 262 | assume "a \<in> I" | |
| 263 | then have "I +> a = I" by (rule a_rcos_const) | |
| 264 | with IanI show False by simp | |
| 265 | qed | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 266 | |
| 45005 | 267 | have aJ: "a \<in> J" | 
| 268 | proof (simp add: J_def r_coset_def set_add_defs) | |
| 269 | from acarr | |
| 270 | have "a = \<one> \<otimes> a \<oplus> \<zero>" by algebra | |
| 271 | with one_closed and additive_subgroup.zero_closed[OF is_additive_subgroup] | |
| 272 | show "\<exists>x\<in>carrier R. \<exists>k\<in>I. a = x \<otimes> a \<oplus> k" by fast | |
| 273 | qed | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 274 | |
| 45005 | 275 | from aJ and anI have JnI: "J \<noteq> I" by fast | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 276 | |
| 61382 | 277 |     --\<open>Deducing @{term "J = carrier R"} because @{term "I"} is maximal\<close>
 | 
| 45005 | 278 | from idealJ and IinJ have "J = I \<or> J = carrier R" | 
| 279 | proof (rule I_maximal, unfold J_def) | |
| 280 | have "carrier R #> a \<subseteq> carrier R" | |
| 281 | using subset_refl acarr by (rule r_coset_subset_G) | |
| 282 | then show "carrier R #> a <+> I \<subseteq> carrier R" | |
| 283 | using a_subset by (rule set_add_closed) | |
| 284 | qed | |
| 285 | ||
| 286 | with JnI have Jcarr: "J = carrier R" by simp | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 287 | |
| 61382 | 288 |     --\<open>Calculating an inverse for @{term "a"}\<close>
 | 
| 45005 | 289 | from one_closed[folded Jcarr] | 
| 290 | have "\<exists>r\<in>carrier R. \<exists>i\<in>I. \<one> = r \<otimes> a \<oplus> i" | |
| 291 | by (simp add: J_def r_coset_def set_add_defs) | |
| 292 | then obtain r i where rcarr: "r \<in> carrier R" | |
| 293 | and iI: "i \<in> I" and one: "\<one> = r \<otimes> a \<oplus> i" by fast | |
| 294 | from one and rcarr and acarr and iI[THEN a_Hcarr] | |
| 295 | have rai1: "a \<otimes> r = \<ominus>i \<oplus> \<one>" by algebra | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 296 | |
| 61382 | 297 | --\<open>Lifting to cosets\<close> | 
| 45005 | 298 | from iI have "\<ominus>i \<oplus> \<one> \<in> I +> \<one>" | 
| 299 | by (intro a_rcosI, simp, intro a_subset, simp) | |
| 300 | with rai1 have "a \<otimes> r \<in> I +> \<one>" by simp | |
| 301 | then have "I +> \<one> = I +> a \<otimes> r" | |
| 302 | by (rule a_repr_independence, simp) (rule a_subgroup) | |
| 303 | ||
| 304 | from rcarr and this[symmetric] | |
| 305 | show "\<exists>r\<in>carrier R. I +> a \<otimes> r = I +> \<one>" by fast | |
| 306 | qed | |
| 27611 | 307 | qed | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 308 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 309 | end |