| author | wenzelm | 
| Sun, 04 Sep 2011 19:36:19 +0200 | |
| changeset 44706 | fe319b45315c | 
| parent 44677 | 3fb27b19e058 | 
| child 47409 | c5be1120980d | 
| permissions | -rw-r--r-- | 
| 41959 | 1 | (* Title: HOL/Algebra/Ideal.thy | 
| 35849 | 2 | Author: Stephan Hohe, TU Muenchen | 
| 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 Ideal | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 6 | imports Ring AbelCoset | 
| 
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 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 9 | section {* Ideals *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 10 | |
| 27717 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 ballarin parents: 
27714diff
changeset | 11 | subsection {* Definitions *}
 | 
| 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 ballarin parents: 
27714diff
changeset | 12 | |
| 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 ballarin parents: 
27714diff
changeset | 13 | subsubsection {* General definition *}
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 14 | |
| 29240 | 15 | locale ideal = additive_subgroup I R + ring R for I and R (structure) + | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 16 | assumes I_l_closed: "\<lbrakk>a \<in> I; x \<in> carrier R\<rbrakk> \<Longrightarrow> x \<otimes> a \<in> I" | 
| 44677 | 17 | and I_r_closed: "\<lbrakk>a \<in> I; x \<in> carrier R\<rbrakk> \<Longrightarrow> a \<otimes> x \<in> I" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 18 | |
| 29237 | 19 | sublocale ideal \<subseteq> abelian_subgroup I R | 
| 44677 | 20 | apply (intro abelian_subgroupI3 abelian_group.intro) | 
| 21 | apply (rule ideal.axioms, rule ideal_axioms) | |
| 22 | apply (rule abelian_group.axioms, rule ring.axioms, rule ideal.axioms, rule ideal_axioms) | |
| 23 | apply (rule abelian_group.axioms, rule ring.axioms, rule ideal.axioms, rule ideal_axioms) | |
| 24 | done | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 25 | |
| 44677 | 26 | lemma (in ideal) is_ideal: "ideal I R" | 
| 27 | by (rule ideal_axioms) | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 28 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 29 | lemma idealI: | 
| 27611 | 30 | fixes R (structure) | 
| 31 | assumes "ring R" | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 32 | assumes a_subgroup: "subgroup I \<lparr>carrier = carrier R, mult = add R, one = zero R\<rparr>" | 
| 44677 | 33 | and I_l_closed: "\<And>a x. \<lbrakk>a \<in> I; x \<in> carrier R\<rbrakk> \<Longrightarrow> x \<otimes> a \<in> I" | 
| 34 | and I_r_closed: "\<And>a x. \<lbrakk>a \<in> I; x \<in> carrier R\<rbrakk> \<Longrightarrow> a \<otimes> x \<in> I" | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 35 | shows "ideal I R" | 
| 27611 | 36 | proof - | 
| 29237 | 37 | interpret ring R by fact | 
| 27611 | 38 | show ?thesis apply (intro ideal.intro ideal_axioms.intro additive_subgroupI) | 
| 23463 | 39 | apply (rule a_subgroup) | 
| 40 | apply (rule is_ring) | |
| 41 | apply (erule (1) I_l_closed) | |
| 42 | apply (erule (1) I_r_closed) | |
| 43 | done | |
| 27611 | 44 | qed | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 45 | |
| 35849 | 46 | |
| 30363 
9b8d9b6ef803
proper local context for text with antiquotations;
 wenzelm parents: 
29240diff
changeset | 47 | subsubsection (in ring) {* Ideals Generated by a Subset of @{term "carrier R"} *}
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 48 | |
| 44677 | 49 | definition genideal :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a set"  ("Idl\<index> _" [80] 79)
 | 
| 35848 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 wenzelm parents: 
35847diff
changeset | 50 |   where "genideal R S = Inter {I. ideal I R \<and> S \<subseteq> I}"
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 51 | |
| 27717 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 ballarin parents: 
27714diff
changeset | 52 | subsubsection {* Principal Ideals *}
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 53 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 54 | locale principalideal = ideal + | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 55 |   assumes generate: "\<exists>i \<in> carrier R. I = Idl {i}"
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 56 | |
| 44677 | 57 | lemma (in principalideal) is_principalideal: "principalideal I R" | 
| 58 | by (rule principalideal_axioms) | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 59 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 60 | lemma principalidealI: | 
| 27611 | 61 | fixes R (structure) | 
| 62 | assumes "ideal I R" | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 63 |   assumes generate: "\<exists>i \<in> carrier R. I = Idl {i}"
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 64 | shows "principalideal I R" | 
| 27611 | 65 | proof - | 
| 29237 | 66 | interpret ideal I R by fact | 
| 44677 | 67 | show ?thesis | 
| 68 | by (intro principalideal.intro principalideal_axioms.intro) | |
| 69 | (rule is_ideal, rule generate) | |
| 27611 | 70 | qed | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 71 | |
| 35849 | 72 | |
| 27717 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 ballarin parents: 
27714diff
changeset | 73 | subsubsection {* Maximal Ideals *}
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 74 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 75 | locale maximalideal = ideal + | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 76 | assumes I_notcarr: "carrier R \<noteq> I" | 
| 44677 | 77 | and I_maximal: "\<lbrakk>ideal J R; I \<subseteq> J; J \<subseteq> carrier R\<rbrakk> \<Longrightarrow> J = I \<or> J = carrier R" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 78 | |
| 44677 | 79 | lemma (in maximalideal) is_maximalideal: "maximalideal I R" | 
| 80 | by (rule maximalideal_axioms) | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 81 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 82 | lemma maximalidealI: | 
| 27611 | 83 | fixes R | 
| 84 | assumes "ideal I R" | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 85 | assumes I_notcarr: "carrier R \<noteq> I" | 
| 44677 | 86 | and I_maximal: "\<And>J. \<lbrakk>ideal J R; I \<subseteq> J; J \<subseteq> carrier R\<rbrakk> \<Longrightarrow> J = I \<or> J = carrier R" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 87 | shows "maximalideal I R" | 
| 27611 | 88 | proof - | 
| 29237 | 89 | interpret ideal I R by fact | 
| 44677 | 90 | show ?thesis | 
| 91 | by (intro maximalideal.intro maximalideal_axioms.intro) | |
| 92 | (rule is_ideal, rule I_notcarr, rule I_maximal) | |
| 27611 | 93 | qed | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 94 | |
| 35849 | 95 | |
| 27717 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 ballarin parents: 
27714diff
changeset | 96 | subsubsection {* Prime Ideals *}
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 97 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 98 | locale primeideal = ideal + cring + | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 99 | assumes I_notcarr: "carrier R \<noteq> I" | 
| 44677 | 100 | and I_prime: "\<lbrakk>a \<in> carrier R; b \<in> carrier R; a \<otimes> b \<in> I\<rbrakk> \<Longrightarrow> a \<in> I \<or> b \<in> I" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 101 | |
| 44677 | 102 | lemma (in primeideal) is_primeideal: "primeideal I R" | 
| 103 | by (rule primeideal_axioms) | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 104 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 105 | lemma primeidealI: | 
| 27611 | 106 | fixes R (structure) | 
| 107 | assumes "ideal I R" | |
| 108 | assumes "cring R" | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 109 | assumes I_notcarr: "carrier R \<noteq> I" | 
| 44677 | 110 | and I_prime: "\<And>a b. \<lbrakk>a \<in> carrier R; b \<in> carrier R; a \<otimes> b \<in> I\<rbrakk> \<Longrightarrow> a \<in> I \<or> b \<in> I" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 111 | shows "primeideal I R" | 
| 27611 | 112 | proof - | 
| 29237 | 113 | interpret ideal I R by fact | 
| 114 | interpret cring R by fact | |
| 44677 | 115 | show ?thesis | 
| 116 | by (intro primeideal.intro primeideal_axioms.intro) | |
| 117 | (rule is_ideal, rule is_cring, rule I_notcarr, rule I_prime) | |
| 27611 | 118 | qed | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 119 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 120 | lemma primeidealI2: | 
| 27611 | 121 | fixes R (structure) | 
| 122 | assumes "additive_subgroup I R" | |
| 123 | assumes "cring R" | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 124 | assumes I_l_closed: "\<And>a x. \<lbrakk>a \<in> I; x \<in> carrier R\<rbrakk> \<Longrightarrow> x \<otimes> a \<in> I" | 
| 44677 | 125 | and I_r_closed: "\<And>a x. \<lbrakk>a \<in> I; x \<in> carrier R\<rbrakk> \<Longrightarrow> a \<otimes> x \<in> I" | 
| 126 | and I_notcarr: "carrier R \<noteq> I" | |
| 127 | and I_prime: "\<And>a b. \<lbrakk>a \<in> carrier R; b \<in> carrier R; a \<otimes> b \<in> I\<rbrakk> \<Longrightarrow> a \<in> I \<or> b \<in> I" | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 128 | shows "primeideal I R" | 
| 27611 | 129 | proof - | 
| 29240 | 130 | interpret additive_subgroup I R by fact | 
| 29237 | 131 | interpret cring R by fact | 
| 27611 | 132 | show ?thesis apply (intro_locales) | 
| 133 | apply (intro ideal_axioms.intro) | |
| 134 | apply (erule (1) I_l_closed) | |
| 135 | apply (erule (1) I_r_closed) | |
| 136 | apply (intro primeideal_axioms.intro) | |
| 137 | apply (rule I_notcarr) | |
| 138 | apply (erule (2) I_prime) | |
| 139 | done | |
| 140 | qed | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 141 | |
| 35849 | 142 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 143 | subsection {* Special Ideals *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 144 | |
| 44677 | 145 | lemma (in ring) zeroideal: "ideal {\<zero>} R"
 | 
| 146 | apply (intro idealI subgroup.intro) | |
| 147 | apply (rule is_ring) | |
| 148 | apply simp+ | |
| 149 | apply (fold a_inv_def, simp) | |
| 150 | apply simp+ | |
| 151 | done | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 152 | |
| 44677 | 153 | lemma (in ring) oneideal: "ideal (carrier R) R" | 
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 154 | by (rule idealI) (auto intro: is_ring add.subgroupI) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 155 | |
| 44677 | 156 | lemma (in "domain") zeroprimeideal: "primeideal {\<zero>} R"
 | 
| 157 | apply (intro primeidealI) | |
| 158 | apply (rule zeroideal) | |
| 159 | apply (rule domain.axioms, rule domain_axioms) | |
| 160 | defer 1 | |
| 161 | apply (simp add: integral) | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 162 | proof (rule ccontr, simp) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 163 |   assume "carrier R = {\<zero>}"
 | 
| 44677 | 164 | then have "\<one> = \<zero>" by (rule one_zeroI) | 
| 165 | with one_not_zero show False by simp | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 166 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 167 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 168 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 169 | subsection {* General Ideal Properies *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 170 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 171 | lemma (in ideal) one_imp_carrier: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 172 | assumes I_one_closed: "\<one> \<in> I" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 173 | shows "I = carrier R" | 
| 44677 | 174 | apply (rule) | 
| 175 | apply (rule) | |
| 176 | apply (rule a_Hcarr, simp) | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 177 | proof | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 178 | fix x | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 179 | assume xcarr: "x \<in> carrier R" | 
| 44677 | 180 | with I_one_closed have "x \<otimes> \<one> \<in> I" by (intro I_l_closed) | 
| 181 | with xcarr show "x \<in> I" by simp | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 182 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 183 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 184 | lemma (in ideal) Icarr: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 185 | assumes iI: "i \<in> I" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 186 | shows "i \<in> carrier R" | 
| 44677 | 187 | using iI by (rule a_Hcarr) | 
| 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 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 190 | subsection {* Intersection of Ideals *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 191 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 192 | text {* \paragraph{Intersection of two ideals} The intersection of any
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 193 |   two ideals is again an ideal in @{term R} *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 194 | lemma (in ring) i_intersect: | 
| 27611 | 195 | assumes "ideal I R" | 
| 196 | assumes "ideal J R" | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 197 | shows "ideal (I \<inter> J) R" | 
| 27611 | 198 | proof - | 
| 29237 | 199 | interpret ideal I R by fact | 
| 29240 | 200 | interpret ideal J R by fact | 
| 27611 | 201 | show ?thesis | 
| 44677 | 202 | apply (intro idealI subgroup.intro) | 
| 203 | apply (rule is_ring) | |
| 204 | apply (force simp add: a_subset) | |
| 205 | apply (simp add: a_inv_def[symmetric]) | |
| 206 | apply simp | |
| 207 | apply (simp add: a_inv_def[symmetric]) | |
| 208 | apply (clarsimp, rule) | |
| 209 | apply (fast intro: ideal.I_l_closed ideal.intro assms)+ | |
| 210 | apply (clarsimp, rule) | |
| 211 | apply (fast intro: ideal.I_r_closed ideal.intro assms)+ | |
| 212 | done | |
| 27611 | 213 | qed | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 214 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 215 | text {* The intersection of any Number of Ideals is again
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 216 |         an Ideal in @{term R} *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 217 | lemma (in ring) i_Intersect: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 218 | assumes Sideals: "\<And>I. I \<in> S \<Longrightarrow> ideal I R" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 219 |     and notempty: "S \<noteq> {}"
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 220 | shows "ideal (Inter S) R" | 
| 44677 | 221 | apply (unfold_locales) | 
| 222 | apply (simp_all add: Inter_eq) | |
| 223 | apply rule unfolding mem_Collect_eq defer 1 | |
| 224 | apply rule defer 1 | |
| 225 | apply rule defer 1 | |
| 226 | apply (fold a_inv_def, rule) defer 1 | |
| 227 | apply rule defer 1 | |
| 228 | apply rule defer 1 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 229 | proof - | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 230 | fix x y | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 231 | assume "\<forall>I\<in>S. x \<in> I" | 
| 44677 | 232 | then have xI: "\<And>I. I \<in> S \<Longrightarrow> x \<in> I" by simp | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 233 | assume "\<forall>I\<in>S. y \<in> I" | 
| 44677 | 234 | then have yI: "\<And>I. I \<in> S \<Longrightarrow> y \<in> I" by simp | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 235 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 236 | fix J | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 237 | assume JS: "J \<in> S" | 
| 29237 | 238 | interpret ideal J R by (rule Sideals[OF JS]) | 
| 44677 | 239 | from xI[OF JS] and yI[OF JS] show "x \<oplus> y \<in> J" by (rule a_closed) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 240 | next | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 241 | fix J | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 242 | assume JS: "J \<in> S" | 
| 29237 | 243 | interpret ideal J R by (rule Sideals[OF JS]) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 244 | show "\<zero> \<in> J" by simp | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 245 | next | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 246 | fix x | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 247 | assume "\<forall>I\<in>S. x \<in> I" | 
| 44677 | 248 | then have xI: "\<And>I. I \<in> S \<Longrightarrow> x \<in> I" by simp | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 249 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 250 | fix J | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 251 | assume JS: "J \<in> S" | 
| 29237 | 252 | interpret ideal J R by (rule Sideals[OF JS]) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 253 | |
| 44677 | 254 | from xI[OF JS] show "\<ominus> x \<in> J" by (rule a_inv_closed) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 255 | next | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 256 | fix x y | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 257 | assume "\<forall>I\<in>S. x \<in> I" | 
| 44677 | 258 | then have xI: "\<And>I. I \<in> S \<Longrightarrow> x \<in> I" by simp | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 259 | assume ycarr: "y \<in> carrier R" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 260 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 261 | fix J | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 262 | assume JS: "J \<in> S" | 
| 29237 | 263 | interpret ideal J R by (rule Sideals[OF JS]) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 264 | |
| 44677 | 265 | from xI[OF JS] and ycarr show "y \<otimes> x \<in> J" by (rule I_l_closed) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 266 | next | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 267 | fix x y | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 268 | assume "\<forall>I\<in>S. x \<in> I" | 
| 44677 | 269 | then have xI: "\<And>I. I \<in> S \<Longrightarrow> x \<in> I" by simp | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 270 | assume ycarr: "y \<in> carrier R" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 271 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 272 | fix J | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 273 | assume JS: "J \<in> S" | 
| 29237 | 274 | interpret ideal J R by (rule Sideals[OF JS]) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 275 | |
| 44677 | 276 | from xI[OF JS] and ycarr show "x \<otimes> y \<in> J" by (rule I_r_closed) | 
| 44106 | 277 | next | 
| 278 | fix x | |
| 279 | assume "\<forall>I\<in>S. x \<in> I" | |
| 44677 | 280 | then have xI: "\<And>I. I \<in> S \<Longrightarrow> x \<in> I" by simp | 
| 44106 | 281 | |
| 282 | from notempty have "\<exists>I0. I0 \<in> S" by blast | |
| 44677 | 283 | then obtain I0 where I0S: "I0 \<in> S" by auto | 
| 44106 | 284 | |
| 285 | interpret ideal I0 R by (rule Sideals[OF I0S]) | |
| 286 | ||
| 287 | from xI[OF I0S] have "x \<in> I0" . | |
| 44677 | 288 | with a_subset show "x \<in> carrier R" by fast | 
| 44106 | 289 | next | 
| 290 | ||
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 291 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 292 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 293 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 294 | subsection {* Addition of Ideals *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 295 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 296 | lemma (in ring) add_ideals: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 297 | assumes idealI: "ideal I R" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 298 | and idealJ: "ideal J R" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 299 | shows "ideal (I <+> J) R" | 
| 44677 | 300 | apply (rule ideal.intro) | 
| 301 | apply (rule add_additive_subgroups) | |
| 302 | apply (intro ideal.axioms[OF idealI]) | |
| 303 | apply (intro ideal.axioms[OF idealJ]) | |
| 304 | apply (rule is_ring) | |
| 305 | apply (rule ideal_axioms.intro) | |
| 306 | apply (simp add: set_add_defs, clarsimp) defer 1 | |
| 307 | apply (simp add: set_add_defs, clarsimp) defer 1 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 308 | proof - | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 309 | fix x i j | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 310 | assume xcarr: "x \<in> carrier R" | 
| 44677 | 311 | and iI: "i \<in> I" | 
| 312 | and jJ: "j \<in> J" | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 313 | from xcarr ideal.Icarr[OF idealI iI] ideal.Icarr[OF idealJ jJ] | 
| 44677 | 314 | have c: "(i \<oplus> j) \<otimes> x = (i \<otimes> x) \<oplus> (j \<otimes> x)" | 
| 315 | by algebra | |
| 316 | from xcarr and iI have a: "i \<otimes> x \<in> I" | |
| 317 | by (simp add: ideal.I_r_closed[OF idealI]) | |
| 318 | from xcarr and jJ have b: "j \<otimes> x \<in> J" | |
| 319 | by (simp add: ideal.I_r_closed[OF idealJ]) | |
| 320 | from a b c show "\<exists>ha\<in>I. \<exists>ka\<in>J. (i \<oplus> j) \<otimes> x = ha \<oplus> ka" | |
| 321 | by fast | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 322 | next | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 323 | fix x i j | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 324 | assume xcarr: "x \<in> carrier R" | 
| 44677 | 325 | and iI: "i \<in> I" | 
| 326 | and jJ: "j \<in> J" | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 327 | from xcarr ideal.Icarr[OF idealI iI] ideal.Icarr[OF idealJ jJ] | 
| 44677 | 328 | have c: "x \<otimes> (i \<oplus> j) = (x \<otimes> i) \<oplus> (x \<otimes> j)" by algebra | 
| 329 | from xcarr and iI have a: "x \<otimes> i \<in> I" | |
| 330 | by (simp add: ideal.I_l_closed[OF idealI]) | |
| 331 | from xcarr and jJ have b: "x \<otimes> j \<in> J" | |
| 332 | by (simp add: ideal.I_l_closed[OF idealJ]) | |
| 333 | from a b c show "\<exists>ha\<in>I. \<exists>ka\<in>J. x \<otimes> (i \<oplus> j) = ha \<oplus> ka" | |
| 334 | by fast | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 335 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 336 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 337 | |
| 30363 
9b8d9b6ef803
proper local context for text with antiquotations;
 wenzelm parents: 
29240diff
changeset | 338 | subsection (in ring) {* Ideals generated by a subset of @{term "carrier R"} *}
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 339 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 340 | text {* @{term genideal} generates an ideal *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 341 | lemma (in ring) genideal_ideal: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 342 | assumes Scarr: "S \<subseteq> carrier R" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 343 | shows "ideal (Idl S) R" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 344 | unfolding genideal_def | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 345 | proof (rule i_Intersect, fast, simp) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 346 | from oneideal and Scarr | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 347 | show "\<exists>I. ideal I R \<and> S \<le> I" by fast | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 348 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 349 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 350 | lemma (in ring) genideal_self: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 351 | assumes "S \<subseteq> carrier R" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 352 | shows "S \<subseteq> Idl S" | 
| 44677 | 353 | unfolding genideal_def by fast | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 354 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 355 | lemma (in ring) genideal_self': | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 356 | assumes carr: "i \<in> carrier R" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 357 |   shows "i \<in> Idl {i}"
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 358 | proof - | 
| 44677 | 359 |   from carr have "{i} \<subseteq> Idl {i}" by (fast intro!: genideal_self)
 | 
| 360 |   then show "i \<in> Idl {i}" by fast
 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 361 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 362 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 363 | text {* @{term genideal} generates the minimal ideal *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 364 | lemma (in ring) genideal_minimal: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 365 | assumes a: "ideal I R" | 
| 44677 | 366 | and b: "S \<subseteq> I" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 367 | shows "Idl S \<subseteq> I" | 
| 44677 | 368 | unfolding genideal_def by rule (elim InterD, simp add: a b) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 369 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 370 | text {* Generated ideals and subsets *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 371 | lemma (in ring) Idl_subset_ideal: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 372 | assumes Iideal: "ideal I R" | 
| 44677 | 373 | and Hcarr: "H \<subseteq> carrier R" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 374 | shows "(Idl H \<subseteq> I) = (H \<subseteq> I)" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 375 | proof | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 376 | assume a: "Idl H \<subseteq> I" | 
| 23350 | 377 | from Hcarr have "H \<subseteq> Idl H" by (rule genideal_self) | 
| 44677 | 378 | with a show "H \<subseteq> I" by simp | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 379 | next | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 380 | fix x | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 381 | assume HI: "H \<subseteq> I" | 
| 44677 | 382 |   from Iideal and HI have "I \<in> {I. ideal I R \<and> H \<subseteq> I}" by fast
 | 
| 383 | then show "Idl H \<subseteq> I" unfolding genideal_def by fast | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 384 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 385 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 386 | lemma (in ring) subset_Idl_subset: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 387 | assumes Icarr: "I \<subseteq> carrier R" | 
| 44677 | 388 | and HI: "H \<subseteq> I" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 389 | shows "Idl H \<subseteq> Idl I" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 390 | proof - | 
| 44677 | 391 | from HI and genideal_self[OF Icarr] have HIdlI: "H \<subseteq> Idl I" | 
| 392 | by fast | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 393 | |
| 44677 | 394 | from Icarr have Iideal: "ideal (Idl I) R" | 
| 395 | by (rule genideal_ideal) | |
| 396 | from HI and Icarr have "H \<subseteq> carrier R" | |
| 397 | by fast | |
| 398 | with Iideal have "(H \<subseteq> Idl I) = (Idl H \<subseteq> Idl I)" | |
| 399 | by (rule Idl_subset_ideal[symmetric]) | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 400 | |
| 44677 | 401 | with HIdlI show "Idl H \<subseteq> Idl I" by simp | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 402 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 403 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 404 | lemma (in ring) Idl_subset_ideal': | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 405 | assumes acarr: "a \<in> carrier R" and bcarr: "b \<in> carrier R" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 406 |   shows "(Idl {a} \<subseteq> Idl {b}) = (a \<in> Idl {b})"
 | 
| 44677 | 407 |   apply (subst Idl_subset_ideal[OF genideal_ideal[of "{b}"], of "{a}"])
 | 
| 408 | apply (fast intro: bcarr, fast intro: acarr) | |
| 409 | apply fast | |
| 410 | done | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 411 | |
| 44677 | 412 | lemma (in ring) genideal_zero: "Idl {\<zero>} = {\<zero>}"
 | 
| 413 | apply rule | |
| 414 | apply (rule genideal_minimal[OF zeroideal], simp) | |
| 415 | apply (simp add: genideal_self') | |
| 416 | done | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 417 | |
| 44677 | 418 | lemma (in ring) genideal_one: "Idl {\<one>} = carrier R"
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 419 | proof - | 
| 44677 | 420 |   interpret ideal "Idl {\<one>}" "R" by (rule genideal_ideal) fast
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 421 |   show "Idl {\<one>} = carrier R"
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 422 | apply (rule, rule a_subset) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 423 | apply (simp add: one_imp_carrier genideal_self') | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 424 | done | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 425 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 426 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 427 | |
| 27717 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 ballarin parents: 
27714diff
changeset | 428 | text {* Generation of Principal Ideals in Commutative Rings *}
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 429 | |
| 44677 | 430 | definition cgenideal :: "_ \<Rightarrow> 'a \<Rightarrow> 'a set"  ("PIdl\<index> _" [80] 79)
 | 
| 35848 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 wenzelm parents: 
35847diff
changeset | 431 |   where "cgenideal R a = {x \<otimes>\<^bsub>R\<^esub> a | x. x \<in> carrier R}"
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 432 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 433 | text {* genhideal (?) really generates an ideal *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 434 | lemma (in cring) cgenideal_ideal: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 435 | assumes acarr: "a \<in> carrier R" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 436 | shows "ideal (PIdl a) R" | 
| 44677 | 437 | apply (unfold cgenideal_def) | 
| 438 | apply (rule idealI[OF is_ring]) | |
| 439 | apply (rule subgroup.intro) | |
| 440 | apply simp_all | |
| 441 | apply (blast intro: acarr) | |
| 442 | apply clarsimp defer 1 | |
| 443 | defer 1 | |
| 444 | apply (fold a_inv_def, clarsimp) defer 1 | |
| 445 | apply clarsimp defer 1 | |
| 446 | apply clarsimp defer 1 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 447 | proof - | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 448 | fix x y | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 449 | assume xcarr: "x \<in> carrier R" | 
| 44677 | 450 | and ycarr: "y \<in> carrier R" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 451 | note carr = acarr xcarr ycarr | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 452 | |
| 44677 | 453 | from carr have "x \<otimes> a \<oplus> y \<otimes> a = (x \<oplus> y) \<otimes> a" | 
| 454 | by (simp add: l_distr) | |
| 455 | with carr show "\<exists>z. x \<otimes> a \<oplus> y \<otimes> a = z \<otimes> a \<and> z \<in> carrier R" | |
| 456 | by fast | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 457 | next | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 458 | from l_null[OF acarr, symmetric] and zero_closed | 
| 44677 | 459 | show "\<exists>x. \<zero> = x \<otimes> a \<and> x \<in> carrier R" by fast | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 460 | next | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 461 | fix x | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 462 | assume xcarr: "x \<in> carrier R" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 463 | note carr = acarr xcarr | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 464 | |
| 44677 | 465 | from carr have "\<ominus> (x \<otimes> a) = (\<ominus> x) \<otimes> a" | 
| 466 | by (simp add: l_minus) | |
| 467 | with carr show "\<exists>z. \<ominus> (x \<otimes> a) = z \<otimes> a \<and> z \<in> carrier R" | |
| 468 | by fast | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 469 | next | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 470 | fix x y | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 471 | assume xcarr: "x \<in> carrier R" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 472 | and ycarr: "y \<in> carrier R" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 473 | note carr = acarr xcarr ycarr | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 474 | |
| 44677 | 475 | from carr have "y \<otimes> a \<otimes> x = (y \<otimes> x) \<otimes> a" | 
| 476 | by (simp add: m_assoc) (simp add: m_comm) | |
| 477 | with carr show "\<exists>z. y \<otimes> a \<otimes> x = z \<otimes> a \<and> z \<in> carrier R" | |
| 478 | by fast | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 479 | next | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 480 | fix x y | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 481 | assume xcarr: "x \<in> carrier R" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 482 | and ycarr: "y \<in> carrier R" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 483 | note carr = acarr xcarr ycarr | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 484 | |
| 44677 | 485 | from carr have "x \<otimes> (y \<otimes> a) = (x \<otimes> y) \<otimes> a" | 
| 486 | by (simp add: m_assoc) | |
| 487 | with carr show "\<exists>z. x \<otimes> (y \<otimes> a) = z \<otimes> a \<and> z \<in> carrier R" | |
| 488 | by fast | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 489 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 490 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 491 | lemma (in ring) cgenideal_self: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 492 | assumes icarr: "i \<in> carrier R" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 493 | shows "i \<in> PIdl i" | 
| 44677 | 494 | unfolding cgenideal_def | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 495 | proof simp | 
| 44677 | 496 | from icarr have "i = \<one> \<otimes> i" | 
| 497 | by simp | |
| 498 | with icarr show "\<exists>x. i = x \<otimes> i \<and> x \<in> carrier R" | |
| 499 | by fast | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 500 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 501 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 502 | text {* @{const "cgenideal"} is minimal *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 503 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 504 | lemma (in ring) cgenideal_minimal: | 
| 27611 | 505 | assumes "ideal J R" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 506 | assumes aJ: "a \<in> J" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 507 | shows "PIdl a \<subseteq> J" | 
| 27611 | 508 | proof - | 
| 29240 | 509 | interpret ideal J R by fact | 
| 44677 | 510 | show ?thesis | 
| 511 | unfolding cgenideal_def | |
| 27611 | 512 | apply rule | 
| 513 | apply clarify | |
| 514 | using aJ | |
| 515 | apply (erule I_l_closed) | |
| 516 | done | |
| 517 | qed | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 518 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 519 | lemma (in cring) cgenideal_eq_genideal: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 520 | assumes icarr: "i \<in> carrier R" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 521 |   shows "PIdl i = Idl {i}"
 | 
| 44677 | 522 | apply rule | 
| 523 | apply (intro cgenideal_minimal) | |
| 524 | apply (rule genideal_ideal, fast intro: icarr) | |
| 525 | apply (rule genideal_self', fast intro: icarr) | |
| 526 | apply (intro genideal_minimal) | |
| 527 | apply (rule cgenideal_ideal [OF icarr]) | |
| 528 | apply (simp, rule cgenideal_self [OF icarr]) | |
| 529 | done | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 530 | |
| 44677 | 531 | lemma (in cring) cgenideal_eq_rcos: "PIdl i = carrier R #> i" | 
| 532 | unfolding cgenideal_def r_coset_def by fast | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 533 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 534 | lemma (in cring) cgenideal_is_principalideal: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 535 | assumes icarr: "i \<in> carrier R" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 536 | shows "principalideal (PIdl i) R" | 
| 44677 | 537 | apply (rule principalidealI) | 
| 538 | apply (rule cgenideal_ideal [OF icarr]) | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 539 | proof - | 
| 44677 | 540 |   from icarr have "PIdl i = Idl {i}"
 | 
| 541 | by (rule cgenideal_eq_genideal) | |
| 542 |   with icarr show "\<exists>i'\<in>carrier R. PIdl i = Idl {i'}"
 | |
| 543 | by fast | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 544 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 545 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 546 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 547 | subsection {* Union of Ideals *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 548 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 549 | lemma (in ring) union_genideal: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 550 | assumes idealI: "ideal I R" | 
| 44677 | 551 | and idealJ: "ideal J R" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 552 | shows "Idl (I \<union> J) = I <+> J" | 
| 44677 | 553 | apply rule | 
| 554 | apply (rule ring.genideal_minimal) | |
| 555 | apply (rule is_ring) | |
| 556 | apply (rule add_ideals[OF idealI idealJ]) | |
| 557 | apply (rule) | |
| 558 | apply (simp add: set_add_defs) apply (elim disjE) defer 1 defer 1 | |
| 559 | apply (rule) apply (simp add: set_add_defs genideal_def) apply clarsimp defer 1 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 560 | proof - | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 561 | fix x | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 562 | assume xI: "x \<in> I" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 563 | have ZJ: "\<zero> \<in> J" | 
| 44677 | 564 | by (intro additive_subgroup.zero_closed) (rule ideal.axioms[OF idealJ]) | 
| 565 | from ideal.Icarr[OF idealI xI] have "x = x \<oplus> \<zero>" | |
| 566 | by algebra | |
| 567 | with xI and ZJ show "\<exists>h\<in>I. \<exists>k\<in>J. x = h \<oplus> k" | |
| 568 | by fast | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 569 | next | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 570 | fix x | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 571 | assume xJ: "x \<in> J" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 572 | have ZI: "\<zero> \<in> I" | 
| 44677 | 573 | by (intro additive_subgroup.zero_closed, rule ideal.axioms[OF idealI]) | 
| 574 | from ideal.Icarr[OF idealJ xJ] have "x = \<zero> \<oplus> x" | |
| 575 | by algebra | |
| 576 | with ZI and xJ show "\<exists>h\<in>I. \<exists>k\<in>J. x = h \<oplus> k" | |
| 577 | by fast | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 578 | next | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 579 | fix i j K | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 580 | assume iI: "i \<in> I" | 
| 44677 | 581 | and jJ: "j \<in> J" | 
| 582 | and idealK: "ideal K R" | |
| 583 | and IK: "I \<subseteq> K" | |
| 584 | and JK: "J \<subseteq> K" | |
| 585 | from iI and IK have iK: "i \<in> K" by fast | |
| 586 | from jJ and JK have jK: "j \<in> K" by fast | |
| 587 | from iK and jK show "i \<oplus> j \<in> K" | |
| 588 | by (intro additive_subgroup.a_closed) (rule ideal.axioms[OF idealK]) | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 589 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 590 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 591 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 592 | subsection {* Properties of Principal Ideals *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 593 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 594 | text {* @{text "\<zero>"} generates the zero ideal *}
 | 
| 44677 | 595 | lemma (in ring) zero_genideal: "Idl {\<zero>} = {\<zero>}"
 | 
| 596 | apply rule | |
| 597 | apply (simp add: genideal_minimal zeroideal) | |
| 598 | apply (fast intro!: genideal_self) | |
| 599 | done | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 600 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 601 | text {* @{text "\<one>"} generates the unit ideal *}
 | 
| 44677 | 602 | lemma (in ring) one_genideal: "Idl {\<one>} = carrier R"
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 603 | proof - | 
| 44677 | 604 |   have "\<one> \<in> Idl {\<one>}"
 | 
| 605 | by (simp add: genideal_self') | |
| 606 |   then show "Idl {\<one>} = carrier R"
 | |
| 607 | by (intro ideal.one_imp_carrier) (fast intro: genideal_ideal) | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 608 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 609 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 610 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 611 | text {* The zero ideal is a principal ideal *}
 | 
| 44677 | 612 | corollary (in ring) zeropideal: "principalideal {\<zero>} R"
 | 
| 613 | apply (rule principalidealI) | |
| 614 | apply (rule zeroideal) | |
| 615 | apply (blast intro!: zero_genideal[symmetric]) | |
| 616 | done | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 617 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 618 | text {* The unit ideal is a principal ideal *}
 | 
| 44677 | 619 | corollary (in ring) onepideal: "principalideal (carrier R) R" | 
| 620 | apply (rule principalidealI) | |
| 621 | apply (rule oneideal) | |
| 622 | apply (blast intro!: one_genideal[symmetric]) | |
| 623 | done | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 624 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 625 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 626 | text {* Every principal ideal is a right coset of the carrier *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 627 | lemma (in principalideal) rcos_generate: | 
| 27611 | 628 | assumes "cring R" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 629 | shows "\<exists>x\<in>I. I = carrier R #> x" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 630 | proof - | 
| 29237 | 631 | interpret cring R by fact | 
| 44677 | 632 |   from generate obtain i where icarr: "i \<in> carrier R" and I1: "I = Idl {i}"
 | 
| 633 | by fast+ | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 634 | |
| 44677 | 635 |   from icarr and genideal_self[of "{i}"] have "i \<in> Idl {i}"
 | 
| 636 | by fast | |
| 637 | then have iI: "i \<in> I" by (simp add: I1) | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 638 | |
| 44677 | 639 | from I1 icarr have I2: "I = PIdl i" | 
| 640 | by (simp add: cgenideal_eq_genideal) | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 641 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 642 | have "PIdl i = carrier R #> i" | 
| 44677 | 643 | unfolding cgenideal_def r_coset_def by fast | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 644 | |
| 44677 | 645 | with I2 have "I = carrier R #> i" | 
| 646 | by simp | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 647 | |
| 44677 | 648 | with iI show "\<exists>x\<in>I. I = carrier R #> x" | 
| 649 | by fast | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 650 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 651 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 652 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 653 | subsection {* Prime Ideals *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 654 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 655 | lemma (in ideal) primeidealCD: | 
| 27611 | 656 | assumes "cring R" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 657 | assumes notprime: "\<not> primeideal I R" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 658 | shows "carrier R = I \<or> (\<exists>a b. a \<in> carrier R \<and> b \<in> carrier R \<and> a \<otimes> b \<in> I \<and> a \<notin> I \<and> b \<notin> I)" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 659 | proof (rule ccontr, clarsimp) | 
| 29237 | 660 | interpret cring R by fact | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 661 | assume InR: "carrier R \<noteq> I" | 
| 44677 | 662 | and "\<forall>a. a \<in> carrier R \<longrightarrow> (\<forall>b. a \<otimes> b \<in> I \<longrightarrow> b \<in> carrier R \<longrightarrow> a \<in> I \<or> b \<in> I)" | 
| 663 | then have I_prime: "\<And> a b. \<lbrakk>a \<in> carrier R; b \<in> carrier R; a \<otimes> b \<in> I\<rbrakk> \<Longrightarrow> a \<in> I \<or> b \<in> I" | |
| 664 | by simp | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 665 | have "primeideal I R" | 
| 44677 | 666 | apply (rule primeideal.intro [OF is_ideal is_cring]) | 
| 667 | apply (rule primeideal_axioms.intro) | |
| 668 | apply (rule InR) | |
| 669 | apply (erule (2) I_prime) | |
| 670 | done | |
| 671 | with notprime show False by simp | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 672 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 673 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 674 | lemma (in ideal) primeidealCE: | 
| 27611 | 675 | assumes "cring R" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 676 | assumes notprime: "\<not> primeideal I R" | 
| 23383 | 677 | obtains "carrier R = I" | 
| 678 | | "\<exists>a b. a \<in> carrier R \<and> b \<in> carrier R \<and> a \<otimes> b \<in> I \<and> a \<notin> I \<and> b \<notin> I" | |
| 27611 | 679 | proof - | 
| 30729 
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
 wenzelm parents: 
30363diff
changeset | 680 | interpret R: cring R by fact | 
| 27611 | 681 | assume "carrier R = I ==> thesis" | 
| 682 | and "\<exists>a b. a \<in> carrier R \<and> b \<in> carrier R \<and> a \<otimes> b \<in> I \<and> a \<notin> I \<and> b \<notin> I \<Longrightarrow> thesis" | |
| 683 | then show thesis using primeidealCD [OF R.is_cring notprime] by blast | |
| 684 | qed | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 685 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 686 | text {* If @{text "{\<zero>}"} is a prime ideal of a commutative ring, the ring is a domain *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 687 | lemma (in cring) zeroprimeideal_domainI: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 688 |   assumes pi: "primeideal {\<zero>} R"
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 689 | shows "domain R" | 
| 44677 | 690 | apply (rule domain.intro, rule is_cring) | 
| 691 | apply (rule domain_axioms.intro) | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 692 | proof (rule ccontr, simp) | 
| 29237 | 693 |   interpret primeideal "{\<zero>}" "R" by (rule pi)
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 694 | assume "\<one> = \<zero>" | 
| 44677 | 695 |   then have "carrier R = {\<zero>}" by (rule one_zeroD)
 | 
| 696 | from this[symmetric] and I_notcarr show False | |
| 697 | by simp | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 698 | next | 
| 29237 | 699 |   interpret primeideal "{\<zero>}" "R" by (rule pi)
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 700 | fix a b | 
| 44677 | 701 | assume ab: "a \<otimes> b = \<zero>" and carr: "a \<in> carrier R" "b \<in> carrier R" | 
| 702 |   from ab have abI: "a \<otimes> b \<in> {\<zero>}"
 | |
| 703 | by fast | |
| 704 |   with carr have "a \<in> {\<zero>} \<or> b \<in> {\<zero>}"
 | |
| 705 | by (rule I_prime) | |
| 706 | then show "a = \<zero> \<or> b = \<zero>" by simp | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 707 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 708 | |
| 44677 | 709 | corollary (in cring) domain_eq_zeroprimeideal: "domain R = primeideal {\<zero>} R"
 | 
| 710 | apply rule | |
| 711 | apply (erule domain.zeroprimeideal) | |
| 712 | apply (erule zeroprimeideal_domainI) | |
| 713 | done | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 714 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 715 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 716 | subsection {* Maximal Ideals *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 717 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 718 | lemma (in ideal) helper_I_closed: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 719 | assumes carr: "a \<in> carrier R" "x \<in> carrier R" "y \<in> carrier R" | 
| 44677 | 720 | and axI: "a \<otimes> x \<in> I" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 721 | shows "a \<otimes> (x \<otimes> y) \<in> I" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 722 | proof - | 
| 44677 | 723 | from axI and carr have "(a \<otimes> x) \<otimes> y \<in> I" | 
| 724 | by (simp add: I_r_closed) | |
| 725 | also from carr have "(a \<otimes> x) \<otimes> y = a \<otimes> (x \<otimes> y)" | |
| 726 | by (simp add: m_assoc) | |
| 727 | finally show "a \<otimes> (x \<otimes> y) \<in> I" . | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 728 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 729 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 730 | lemma (in ideal) helper_max_prime: | 
| 27611 | 731 | assumes "cring R" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 732 | assumes acarr: "a \<in> carrier R" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 733 |   shows "ideal {x\<in>carrier R. a \<otimes> x \<in> I} R"
 | 
| 27611 | 734 | proof - | 
| 29237 | 735 | interpret cring R by fact | 
| 27611 | 736 | show ?thesis apply (rule idealI) | 
| 737 | apply (rule cring.axioms[OF is_cring]) | |
| 738 | apply (rule subgroup.intro) | |
| 739 | apply (simp, fast) | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 740 | apply clarsimp apply (simp add: r_distr acarr) | 
| 27611 | 741 | apply (simp add: acarr) | 
| 742 | apply (simp add: a_inv_def[symmetric], clarify) defer 1 | |
| 743 | apply clarsimp defer 1 | |
| 744 | apply (fast intro!: helper_I_closed acarr) | |
| 745 | proof - | |
| 746 | fix x | |
| 747 | assume xcarr: "x \<in> carrier R" | |
| 748 | and ax: "a \<otimes> x \<in> I" | |
| 749 | from ax and acarr xcarr | |
| 750 | have "\<ominus>(a \<otimes> x) \<in> I" by simp | |
| 751 | also from acarr xcarr | |
| 752 | have "\<ominus>(a \<otimes> x) = a \<otimes> (\<ominus>x)" by algebra | |
| 44677 | 753 | finally show "a \<otimes> (\<ominus>x) \<in> I" . | 
| 754 | from acarr have "a \<otimes> \<zero> = \<zero>" by simp | |
| 27611 | 755 | next | 
| 756 | fix x y | |
| 757 | assume xcarr: "x \<in> carrier R" | |
| 758 | and ycarr: "y \<in> carrier R" | |
| 759 | and ayI: "a \<otimes> y \<in> I" | |
| 44677 | 760 | from ayI and acarr xcarr ycarr have "a \<otimes> (y \<otimes> x) \<in> I" | 
| 761 | by (simp add: helper_I_closed) | |
| 762 | moreover | |
| 763 | from xcarr ycarr have "y \<otimes> x = x \<otimes> y" | |
| 764 | by (simp add: m_comm) | |
| 27611 | 765 | ultimately | 
| 766 | show "a \<otimes> (x \<otimes> y) \<in> I" by simp | |
| 767 | qed | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 768 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 769 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 770 | text {* In a cring every maximal ideal is prime *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 771 | lemma (in cring) maximalideal_is_prime: | 
| 27611 | 772 | assumes "maximalideal I R" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 773 | shows "primeideal I R" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 774 | proof - | 
| 29237 | 775 | interpret maximalideal I R by fact | 
| 27611 | 776 | show ?thesis apply (rule ccontr) | 
| 777 | apply (rule primeidealCE) | |
| 778 | apply (rule is_cring) | |
| 779 | apply assumption | |
| 780 | apply (simp add: I_notcarr) | |
| 781 | proof - | |
| 782 | assume "\<exists>a b. a \<in> carrier R \<and> b \<in> carrier R \<and> a \<otimes> b \<in> I \<and> a \<notin> I \<and> b \<notin> I" | |
| 44677 | 783 | then obtain a b where | 
| 784 | acarr: "a \<in> carrier R" and | |
| 785 | bcarr: "b \<in> carrier R" and | |
| 786 | abI: "a \<otimes> b \<in> I" and | |
| 787 | anI: "a \<notin> I" and | |
| 788 | bnI: "b \<notin> I" by fast | |
| 27611 | 789 |     def J \<equiv> "{x\<in>carrier R. a \<otimes> x \<in> I}"
 | 
| 790 | ||
| 44677 | 791 | from is_cring and acarr have idealJ: "ideal J R" | 
| 792 | unfolding J_def by (rule helper_max_prime) | |
| 27611 | 793 | |
| 794 | have IsubJ: "I \<subseteq> J" | |
| 795 | proof | |
| 796 | fix x | |
| 797 | assume xI: "x \<in> I" | |
| 44677 | 798 | with acarr have "a \<otimes> x \<in> I" | 
| 799 | by (intro I_l_closed) | |
| 800 | with xI[THEN a_Hcarr] show "x \<in> J" | |
| 801 | unfolding J_def by fast | |
| 27611 | 802 | qed | 
| 803 | ||
| 44677 | 804 | from abI and acarr bcarr have "b \<in> J" | 
| 805 | unfolding J_def by fast | |
| 806 | with bnI have JnI: "J \<noteq> I" by fast | |
| 27611 | 807 | from acarr | 
| 808 | have "a = a \<otimes> \<one>" by algebra | |
| 44677 | 809 | with anI have "a \<otimes> \<one> \<notin> I" by simp | 
| 810 | with one_closed have "\<one> \<notin> J" | |
| 811 | unfolding J_def by fast | |
| 812 | then have Jncarr: "J \<noteq> carrier R" by fast | |
| 27611 | 813 | |
| 29237 | 814 | interpret ideal J R by (rule idealJ) | 
| 27611 | 815 | |
| 816 | have "J = I \<or> J = carrier R" | |
| 817 | apply (intro I_maximal) | |
| 818 | apply (rule idealJ) | |
| 819 | apply (rule IsubJ) | |
| 820 | apply (rule a_subset) | |
| 821 | done | |
| 822 | ||
| 44677 | 823 | with JnI and Jncarr show False by simp | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 824 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 825 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 826 | |
| 35849 | 827 | |
| 27717 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 ballarin parents: 
27714diff
changeset | 828 | subsection {* Derived Theorems *}
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 829 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 830 | --"A non-zero cring that has only the two trivial ideals is a field" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 831 | lemma (in cring) trivialideals_fieldI: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 832 |   assumes carrnzero: "carrier R \<noteq> {\<zero>}"
 | 
| 44677 | 833 |     and haveideals: "{I. ideal I R} = {{\<zero>}, carrier R}"
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 834 | shows "field R" | 
| 44677 | 835 | apply (rule cring_fieldI) | 
| 836 | apply (rule, rule, rule) | |
| 837 | apply (erule Units_closed) | |
| 838 | defer 1 | |
| 839 | apply rule | |
| 840 | defer 1 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 841 | proof (rule ccontr, simp) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 842 | assume zUnit: "\<zero> \<in> Units R" | 
| 44677 | 843 | then have a: "\<zero> \<otimes> inv \<zero> = \<one>" by (rule Units_r_inv) | 
| 844 | from zUnit have "\<zero> \<otimes> inv \<zero> = \<zero>" | |
| 845 | by (intro l_null) (rule Units_inv_closed) | |
| 846 | with a[symmetric] have "\<one> = \<zero>" by simp | |
| 847 |   then have "carrier R = {\<zero>}" by (rule one_zeroD)
 | |
| 848 | with carrnzero show False by simp | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 849 | next | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 850 | fix x | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 851 |   assume xcarr': "x \<in> carrier R - {\<zero>}"
 | 
| 44677 | 852 | then have xcarr: "x \<in> carrier R" by fast | 
| 853 | from xcarr' have xnZ: "x \<noteq> \<zero>" by fast | |
| 854 | from xcarr have xIdl: "ideal (PIdl x) R" | |
| 855 | by (intro cgenideal_ideal) fast | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 856 | |
| 44677 | 857 | from xcarr have "x \<in> PIdl x" | 
| 858 | by (intro cgenideal_self) fast | |
| 859 |   with xnZ have "PIdl x \<noteq> {\<zero>}" by fast
 | |
| 860 | with haveideals have "PIdl x = carrier R" | |
| 861 | by (blast intro!: xIdl) | |
| 862 | then have "\<one> \<in> PIdl x" by simp | |
| 863 | then have "\<exists>y. \<one> = y \<otimes> x \<and> y \<in> carrier R" | |
| 864 | unfolding cgenideal_def by blast | |
| 865 | then obtain y where ycarr: " y \<in> carrier R" and ylinv: "\<one> = y \<otimes> x" | |
| 866 | by fast+ | |
| 867 | from ylinv and xcarr ycarr have yrinv: "\<one> = x \<otimes> y" | |
| 868 | by (simp add: m_comm) | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 869 | from ycarr and ylinv[symmetric] and yrinv[symmetric] | 
| 44677 | 870 | have "\<exists>y \<in> carrier R. y \<otimes> x = \<one> \<and> x \<otimes> y = \<one>" by fast | 
| 871 | with xcarr show "x \<in> Units R" | |
| 872 | unfolding Units_def by fast | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 873 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 874 | |
| 44677 | 875 | lemma (in field) all_ideals: "{I. ideal I R} = {{\<zero>}, carrier R}"
 | 
| 876 | apply (rule, rule) | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 877 | proof - | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 878 | fix I | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 879 |   assume a: "I \<in> {I. ideal I R}"
 | 
| 44677 | 880 | then interpret ideal I R by simp | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 881 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 882 |   show "I \<in> {{\<zero>}, carrier R}"
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 883 |   proof (cases "\<exists>a. a \<in> I - {\<zero>}")
 | 
| 44677 | 884 | case True | 
| 885 | then obtain a where aI: "a \<in> I" and anZ: "a \<noteq> \<zero>" | |
| 886 | by fast+ | |
| 887 | from aI[THEN a_Hcarr] anZ have aUnit: "a \<in> Units R" | |
| 888 | by (simp add: field_Units) | |
| 889 | then have a: "a \<otimes> inv a = \<one>" by (rule Units_r_inv) | |
| 890 | from aI and aUnit have "a \<otimes> inv a \<in> I" | |
| 891 | by (simp add: I_r_closed del: Units_r_inv) | |
| 892 | then have oneI: "\<one> \<in> I" by (simp add: a[symmetric]) | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 893 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 894 | have "carrier R \<subseteq> I" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 895 | proof | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 896 | fix x | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 897 | assume xcarr: "x \<in> carrier R" | 
| 44677 | 898 | with oneI have "\<one> \<otimes> x \<in> I" by (rule I_r_closed) | 
| 899 | with xcarr show "x \<in> I" by simp | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 900 | qed | 
| 44677 | 901 | with a_subset have "I = carrier R" by fast | 
| 902 |     then show "I \<in> {{\<zero>}, carrier R}" by fast
 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 903 | next | 
| 44677 | 904 | case False | 
| 905 | then have IZ: "\<And>a. a \<in> I \<Longrightarrow> a = \<zero>" by simp | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 906 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 907 |     have a: "I \<subseteq> {\<zero>}"
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 908 | proof | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 909 | fix x | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 910 | assume "x \<in> I" | 
| 44677 | 911 | then have "x = \<zero>" by (rule IZ) | 
| 912 |       then show "x \<in> {\<zero>}" by fast
 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 913 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 914 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 915 | have "\<zero> \<in> I" by simp | 
| 44677 | 916 |     then have "{\<zero>} \<subseteq> I" by fast
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 917 | |
| 44677 | 918 |     with a have "I = {\<zero>}" by fast
 | 
| 919 |     then show "I \<in> {{\<zero>}, carrier R}" by fast
 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 920 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 921 | qed (simp add: zeroideal oneideal) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 922 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 923 | --"Jacobson Theorem 2.2" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 924 | lemma (in cring) trivialideals_eq_field: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 925 |   assumes carrnzero: "carrier R \<noteq> {\<zero>}"
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 926 |   shows "({I. ideal I R} = {{\<zero>}, carrier R}) = field R"
 | 
| 44677 | 927 | by (fast intro!: trivialideals_fieldI[OF carrnzero] field.all_ideals) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 928 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 929 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 930 | text {* Like zeroprimeideal for domains *}
 | 
| 44677 | 931 | lemma (in field) zeromaximalideal: "maximalideal {\<zero>} R"
 | 
| 932 | apply (rule maximalidealI) | |
| 933 | apply (rule zeroideal) | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 934 | proof- | 
| 44677 | 935 |   from one_not_zero have "\<one> \<notin> {\<zero>}" by simp
 | 
| 936 |   with one_closed show "carrier R \<noteq> {\<zero>}" by fast
 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 937 | next | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 938 | fix J | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 939 | assume Jideal: "ideal J R" | 
| 44677 | 940 |   then have "J \<in> {I. ideal I R}" by fast
 | 
| 941 |   with all_ideals show "J = {\<zero>} \<or> J = carrier R"
 | |
| 942 | by simp | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 943 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 944 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 945 | lemma (in cring) zeromaximalideal_fieldI: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 946 |   assumes zeromax: "maximalideal {\<zero>} R"
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 947 | shows "field R" | 
| 44677 | 948 | apply (rule trivialideals_fieldI, rule maximalideal.I_notcarr[OF zeromax]) | 
| 949 | apply rule apply clarsimp defer 1 | |
| 950 | apply (simp add: zeroideal oneideal) | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 951 | proof - | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 952 | fix J | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 953 |   assume Jn0: "J \<noteq> {\<zero>}"
 | 
| 44677 | 954 | and idealJ: "ideal J R" | 
| 29237 | 955 | interpret ideal J R by (rule idealJ) | 
| 44677 | 956 |   have "{\<zero>} \<subseteq> J" by (rule ccontr) simp
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 957 | from zeromax and idealJ and this and a_subset | 
| 44677 | 958 |   have "J = {\<zero>} \<or> J = carrier R"
 | 
| 959 | by (rule maximalideal.I_maximal) | |
| 960 | with Jn0 show "J = carrier R" | |
| 961 | by simp | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 962 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 963 | |
| 44677 | 964 | lemma (in cring) zeromaximalideal_eq_field: "maximalideal {\<zero>} R = field R"
 | 
| 965 | apply rule | |
| 966 | apply (erule zeromaximalideal_fieldI) | |
| 967 | apply (erule field.zeromaximalideal) | |
| 968 | done | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 969 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 970 | end |