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