| author | wenzelm | 
| Sun, 10 Feb 2019 19:07:53 +0100 | |
| changeset 69798 | f610115ca3d0 | 
| parent 69597 | ff784d5a5bfb | 
| child 70160 | 8e9100dcde52 | 
| permissions | -rw-r--r-- | 
| 68582 | 1 | (* Title: HOL/Algebra/Ring_Divisibility.thy | 
| 2 | Author: Paulo EmÃlio de Vilhena | |
| 3 | *) | |
| 68578 | 4 | |
| 5 | theory Ring_Divisibility | |
| 69070 | 6 | imports Ideal Divisibility QuotRing Multiplicative_Group | 
| 68578 | 7 | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 8 | begin | 
| 68578 | 9 | |
| 10 | section \<open>The Arithmetic of Rings\<close> | |
| 11 | ||
| 12 | text \<open>In this section we study the links between the divisibility theory and that of rings\<close> | |
| 13 | ||
| 14 | ||
| 15 | subsection \<open>Definitions\<close> | |
| 16 | ||
| 17 | locale factorial_domain = domain + factorial_monoid "mult_of R" | |
| 18 | ||
| 19 | locale noetherian_ring = ring + | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 20 | assumes finetely_gen: "ideal I R \<Longrightarrow> \<exists>A \<subseteq> carrier R. finite A \<and> I = Idl A" | 
| 68578 | 21 | |
| 22 | locale noetherian_domain = noetherian_ring + domain | |
| 23 | ||
| 24 | locale principal_domain = domain + | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 25 | assumes exists_gen: "ideal I R \<Longrightarrow> \<exists>a \<in> carrier R. I = PIdl a" | 
| 68578 | 26 | |
| 69070 | 27 | locale euclidean_domain = | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 28 | domain R for R (structure) + fixes \<phi> :: "'a \<Rightarrow> nat" | 
| 68578 | 29 | assumes euclidean_function: | 
| 30 |   " \<lbrakk> a \<in> carrier R - { \<zero> }; b \<in> carrier R - { \<zero> } \<rbrakk> \<Longrightarrow>
 | |
| 31 | \<exists>q r. q \<in> carrier R \<and> r \<in> carrier R \<and> a = (b \<otimes> q) \<oplus> r \<and> ((r = \<zero>) \<or> (\<phi> r < \<phi> b))" | |
| 32 | ||
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 33 | definition ring_irreducible :: "('a, 'b) ring_scheme \<Rightarrow> 'a \<Rightarrow> bool" ("ring'_irreducible\<index>")
 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 34 | where "ring_irreducible\<^bsub>R\<^esub> a \<longleftrightarrow> (a \<noteq> \<zero>\<^bsub>R\<^esub>) \<and> (irreducible R a)" | 
| 68578 | 35 | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 36 | definition ring_prime :: "('a, 'b) ring_scheme \<Rightarrow> 'a \<Rightarrow> bool" ("ring'_prime\<index>")
 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 37 | where "ring_prime\<^bsub>R\<^esub> a \<longleftrightarrow> (a \<noteq> \<zero>\<^bsub>R\<^esub>) \<and> (prime R a)" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 38 | |
| 68578 | 39 | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 40 | subsection \<open>The cancellative monoid of a domain. \<close> | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 41 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 42 | sublocale domain < mult_of: comm_monoid_cancel "mult_of R" | 
| 68578 | 43 | rewrites "mult (mult_of R) = mult R" | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 44 | and "one (mult_of R) = one R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 45 | using m_comm m_assoc | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 46 | by (auto intro!: comm_monoid_cancelI comm_monoidI | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 47 | simp add: m_rcancel integral_iff) | 
| 68578 | 48 | |
| 49 | sublocale factorial_domain < mult_of: factorial_monoid "mult_of R" | |
| 50 | rewrites "mult (mult_of R) = mult R" | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 51 | and "one (mult_of R) = one R" | 
| 68578 | 52 | using factorial_monoid_axioms by auto | 
| 53 | ||
| 54 | lemma (in domain) euclidean_domainI: | |
| 55 |   assumes "\<And>a b. \<lbrakk> a \<in> carrier R - { \<zero> }; b \<in> carrier R - { \<zero> } \<rbrakk> \<Longrightarrow>
 | |
| 56 | \<exists>q r. q \<in> carrier R \<and> r \<in> carrier R \<and> a = (b \<otimes> q) \<oplus> r \<and> ((r = \<zero>) \<or> (\<phi> r < \<phi> b))" | |
| 57 | shows "euclidean_domain R \<phi>" | |
| 58 | using assms by unfold_locales auto | |
| 59 | ||
| 60 | ||
| 69597 | 61 | subsection \<open>Passing from \<^term>\<open>R\<close> to \<^term>\<open>mult_of R\<close> and vice-versa. \<close> | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 62 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 63 | lemma divides_mult_imp_divides [simp]: "a divides\<^bsub>(mult_of R)\<^esub> b \<Longrightarrow> a divides\<^bsub>R\<^esub> b" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 64 | unfolding factor_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 65 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 66 | lemma (in domain) divides_imp_divides_mult [simp]: | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 67 |   "\<lbrakk> a \<in> carrier R; b \<in> carrier R - { \<zero> } \<rbrakk> \<Longrightarrow> a divides\<^bsub>R\<^esub> b \<Longrightarrow> a divides\<^bsub>(mult_of R)\<^esub> b"
 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 68 | unfolding factor_def using integral_iff by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 69 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 70 | lemma (in cring) divides_one: | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 71 | assumes "a \<in> carrier R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 72 | shows "a divides \<one> \<longleftrightarrow> a \<in> Units R" | 
| 69070 | 73 | using assms m_comm unfolding factor_def Units_def by force | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 74 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 75 | lemma (in ring) one_divides: | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 76 | assumes "a \<in> carrier R" shows "\<one> divides a" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 77 | using assms unfolding factor_def by simp | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 78 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 79 | lemma (in ring) divides_zero: | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 80 | assumes "a \<in> carrier R" shows "a divides \<zero>" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 81 | using r_null[OF assms] unfolding factor_def by force | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 82 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 83 | lemma (in ring) zero_divides: | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 84 | shows "\<zero> divides a \<longleftrightarrow> a = \<zero>" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 85 | unfolding factor_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 86 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 87 | lemma (in domain) divides_mult_zero: | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 88 | assumes "a \<in> carrier R" shows "a divides\<^bsub>(mult_of R)\<^esub> \<zero> \<Longrightarrow> a = \<zero>" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 89 | using integral[OF _ assms] unfolding factor_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 90 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 91 | lemma (in ring) divides_mult: | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 92 | assumes "a \<in> carrier R" "c \<in> carrier R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 93 | shows "a divides b \<Longrightarrow> (c \<otimes> a) divides (c \<otimes> b)" | 
| 69070 | 94 | using m_assoc[OF assms(2,1)] unfolding factor_def by auto | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 95 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 96 | lemma (in domain) mult_divides: | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 97 |   assumes "a \<in> carrier R" "b \<in> carrier R" "c \<in> carrier R - { \<zero> }"
 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 98 | shows "(c \<otimes> a) divides (c \<otimes> b) \<Longrightarrow> a divides b" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 99 | using assms m_assoc[of c] unfolding factor_def by (simp add: m_lcancel) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 100 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 101 | lemma (in domain) assoc_iff_assoc_mult: | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 102 | assumes "a \<in> carrier R" and "b \<in> carrier R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 103 | shows "a \<sim> b \<longleftrightarrow> a \<sim>\<^bsub>(mult_of R)\<^esub> b" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 104 | proof | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 105 | assume "a \<sim>\<^bsub>(mult_of R)\<^esub> b" thus "a \<sim> b" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 106 | unfolding associated_def factor_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 107 | next | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 108 | assume A: "a \<sim> b" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 109 | then obtain c1 c2 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 110 | where c1: "c1 \<in> carrier R" "a = b \<otimes> c1" and c2: "c2 \<in> carrier R" "b = a \<otimes> c2" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 111 | unfolding associated_def factor_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 112 | show "a \<sim>\<^bsub>(mult_of R)\<^esub> b" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 113 | proof (cases "a = \<zero>") | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 114 | assume a: "a = \<zero>" then have b: "b = \<zero>" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 115 | using c2 by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 116 | show ?thesis | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 117 | unfolding associated_def factor_def a b by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 118 | next | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 119 | assume a: "a \<noteq> \<zero>" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 120 | hence b: "b \<noteq> \<zero>" and c1_not_zero: "c1 \<noteq> \<zero>" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 121 | using c1 assms by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 122 | hence c2_not_zero: "c2 \<noteq> \<zero>" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 123 | using c2 assms by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 124 | show ?thesis | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 125 | unfolding associated_def factor_def using c1 c2 c1_not_zero c2_not_zero a b by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 126 | qed | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 127 | qed | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 128 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 129 | lemma (in domain) Units_mult_eq_Units [simp]: "Units (mult_of R) = Units R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 130 | unfolding Units_def using insert_Diff integral_iff by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 131 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 132 | lemma (in domain) ring_associated_iff: | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 133 | assumes "a \<in> carrier R" "b \<in> carrier R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 134 | shows "a \<sim> b \<longleftrightarrow> (\<exists>u \<in> Units R. a = u \<otimes> b)" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 135 | proof (cases "a = \<zero>") | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 136 | assume [simp]: "a = \<zero>" show ?thesis | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 137 | proof | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 138 | assume "a \<sim> b" thus "\<exists>u \<in> Units R. a = u \<otimes> b" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 139 | using zero_divides unfolding associated_def by force | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 140 | next | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 141 | assume "\<exists>u \<in> Units R. a = u \<otimes> b" then have "b = \<zero>" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 142 | by (metis Units_closed Units_l_cancel \<open>a = \<zero>\<close> assms r_null) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 143 | thus "a \<sim> b" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 144 | using zero_divides[of \<zero>] by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 145 | qed | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 146 | next | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 147 | assume a: "a \<noteq> \<zero>" show ?thesis | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 148 | proof (cases "b = \<zero>") | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 149 | assume "b = \<zero>" thus ?thesis | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 150 | using assms a zero_divides[of a] r_null unfolding associated_def by blast | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 151 | next | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 152 | assume b: "b \<noteq> \<zero>" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 153 | have "(\<exists>u \<in> Units R. a = u \<otimes> b) \<longleftrightarrow> (\<exists>u \<in> Units R. a = b \<otimes> u)" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 154 | using m_comm[OF assms(2)] Units_closed by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 155 | thus ?thesis | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 156 | using mult_of.associated_iff[of a b] a b assms | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 157 | unfolding assoc_iff_assoc_mult[OF assms] Units_mult_eq_Units | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 158 | by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 159 | qed | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 160 | qed | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 161 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 162 | lemma (in domain) properfactor_mult_imp_properfactor: | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 163 | "\<lbrakk> a \<in> carrier R; b \<in> carrier R \<rbrakk> \<Longrightarrow> properfactor (mult_of R) b a \<Longrightarrow> properfactor R b a" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 164 | proof - | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 165 | assume A: "a \<in> carrier R" "b \<in> carrier R" "properfactor (mult_of R) b a" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 166 | then obtain c where c: "c \<in> carrier (mult_of R)" "a = b \<otimes> c" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 167 | unfolding properfactor_def factor_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 168 | have "a \<noteq> \<zero>" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 169 | proof (rule ccontr) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 170 | assume a: "\<not> a \<noteq> \<zero>" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 171 | hence "b = \<zero>" using c A integral[of b c] by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 172 | hence "b = a \<otimes> \<one>" using a A by simp | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 173 | hence "a divides\<^bsub>(mult_of R)\<^esub> b" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 174 | unfolding factor_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 175 | thus False using A unfolding properfactor_def by simp | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 176 | qed | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 177 | hence "b \<noteq> \<zero>" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 178 | using c A integral_iff by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 179 | thus "properfactor R b a" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 180 | using A divides_imp_divides_mult[of a b] unfolding properfactor_def | 
| 69070 | 181 | by (meson DiffI divides_mult_imp_divides empty_iff insert_iff) | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 182 | qed | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 183 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 184 | lemma (in domain) properfactor_imp_properfactor_mult: | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 185 |   "\<lbrakk> a \<in> carrier R - { \<zero> }; b \<in> carrier R \<rbrakk> \<Longrightarrow> properfactor R b a \<Longrightarrow> properfactor (mult_of R) b a"
 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 186 | unfolding properfactor_def factor_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 187 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 188 | lemma (in domain) properfactor_of_zero: | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 189 | assumes "b \<in> carrier R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 190 | shows "\<not> properfactor (mult_of R) b \<zero>" and "properfactor R b \<zero> \<longleftrightarrow> b \<noteq> \<zero>" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 191 | using divides_mult_zero[OF assms] divides_zero[OF assms] unfolding properfactor_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 192 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 193 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 194 | subsection \<open>Irreducible\<close> | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 195 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 196 | text \<open>The following lemmas justify the need for a definition of irreducible specific to rings: | 
| 69597 | 197 | for \<^term>\<open>irreducible R\<close>, we need to suppose we are not in a field (which is plausible, | 
| 198 | but \<^term>\<open>\<not> field R\<close> is an assumption we want to avoid; for \<^term>\<open>irreducible (mult_of R)\<close>, zero | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 199 | is allowed. \<close> | 
| 68578 | 200 | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 201 | lemma (in domain) zero_is_irreducible_mult: | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 202 | shows "irreducible (mult_of R) \<zero>" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 203 | unfolding irreducible_def Units_def properfactor_def factor_def | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 204 | using integral_iff by fastforce | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 205 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 206 | lemma (in domain) zero_is_irreducible_iff_field: | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 207 | shows "irreducible R \<zero> \<longleftrightarrow> field R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 208 | proof | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 209 | assume irr: "irreducible R \<zero>" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 210 | have "\<And>a. \<lbrakk> a \<in> carrier R; a \<noteq> \<zero> \<rbrakk> \<Longrightarrow> properfactor R a \<zero>" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 211 | unfolding properfactor_def factor_def by (auto, metis r_null zero_closed) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 212 |   hence "carrier R - { \<zero> } = Units R"
 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 213 | using irr unfolding irreducible_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 214 | thus "field R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 215 | using field.intro[OF domain_axioms] unfolding field_axioms_def by simp | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 216 | next | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 217 | assume field: "field R" show "irreducible R \<zero>" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 218 | using field.field_Units[OF field] | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 219 | unfolding irreducible_def properfactor_def factor_def by blast | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 220 | qed | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 221 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 222 | lemma (in domain) irreducible_imp_irreducible_mult: | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 223 | "\<lbrakk> a \<in> carrier R; irreducible R a \<rbrakk> \<Longrightarrow> irreducible (mult_of R) a" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 224 | using zero_is_irreducible_mult Units_mult_eq_Units properfactor_mult_imp_properfactor | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 225 | by (cases "a = \<zero>") (auto simp add: irreducible_def) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 226 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 227 | lemma (in domain) irreducible_mult_imp_irreducible: | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 228 |   "\<lbrakk> a \<in> carrier R - { \<zero> }; irreducible (mult_of R) a \<rbrakk> \<Longrightarrow> irreducible R a"
 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 229 | unfolding irreducible_def using properfactor_imp_properfactor_mult properfactor_divides by fastforce | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 230 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 231 | lemma (in domain) ring_irreducibleE: | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 232 | assumes "r \<in> carrier R" and "ring_irreducible r" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 233 | shows "r \<noteq> \<zero>" "irreducible R r" "irreducible (mult_of R) r" "r \<notin> Units R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 234 | and "\<And>a b. \<lbrakk> a \<in> carrier R; b \<in> carrier R \<rbrakk> \<Longrightarrow> r = a \<otimes> b \<Longrightarrow> a \<in> Units R \<or> b \<in> Units R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 235 | proof - | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 236 | show "irreducible (mult_of R) r" "irreducible R r" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 237 | using assms irreducible_imp_irreducible_mult unfolding ring_irreducible_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 238 | show "r \<noteq> \<zero>" "r \<notin> Units R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 239 | using assms unfolding ring_irreducible_def irreducible_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 240 | next | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 241 | fix a b assume a: "a \<in> carrier R" and b: "b \<in> carrier R" and r: "r = a \<otimes> b" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 242 | show "a \<in> Units R \<or> b \<in> Units R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 243 | proof (cases "properfactor R a r") | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 244 | assume "properfactor R a r" thus ?thesis | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 245 | using a assms(2) unfolding ring_irreducible_def irreducible_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 246 | next | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 247 | assume not_proper: "\<not> properfactor R a r" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 248 | hence "r divides a" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 249 | using r b unfolding properfactor_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 250 | then obtain c where c: "c \<in> carrier R" "a = r \<otimes> c" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 251 | unfolding factor_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 252 | have "\<one> = c \<otimes> b" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 253 | using r c(1) b assms m_assoc m_lcancel[OF _ assms(1) one_closed m_closed[OF c(1) b]] | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 254 | unfolding c(2) ring_irreducible_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 255 | thus ?thesis | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 256 | using c(1) b m_comm unfolding Units_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 257 | qed | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 258 | qed | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 259 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 260 | lemma (in domain) ring_irreducibleI: | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 261 |   assumes "r \<in> carrier R - { \<zero> }" "r \<notin> Units R"
 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 262 | and "\<And>a b. \<lbrakk> a \<in> carrier R; b \<in> carrier R \<rbrakk> \<Longrightarrow> r = a \<otimes> b \<Longrightarrow> a \<in> Units R \<or> b \<in> Units R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 263 | shows "ring_irreducible r" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 264 | unfolding ring_irreducible_def | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 265 | proof | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 266 | show "r \<noteq> \<zero>" using assms(1) by blast | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 267 | next | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 268 | show "irreducible R r" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 269 | proof (rule irreducibleI[OF assms(2)]) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 270 | fix a assume a: "a \<in> carrier R" "properfactor R a r" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 271 | then obtain b where b: "b \<in> carrier R" "r = a \<otimes> b" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 272 | unfolding properfactor_def factor_def by blast | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 273 | hence "a \<in> Units R \<or> b \<in> Units R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 274 | using assms(3)[OF a(1) b(1)] by simp | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 275 | thus "a \<in> Units R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 276 | proof (auto) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 277 | assume "b \<in> Units R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 278 | hence "r \<otimes> inv b = a" and "inv b \<in> carrier R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 279 | using b a by (simp add: m_assoc)+ | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 280 | thus ?thesis | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 281 | using a(2) unfolding properfactor_def factor_def by blast | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 282 | qed | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 283 | qed | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 284 | qed | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 285 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 286 | lemma (in domain) ring_irreducibleI': | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 287 |   assumes "r \<in> carrier R - { \<zero> }" "irreducible (mult_of R) r"
 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 288 | shows "ring_irreducible r" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 289 | unfolding ring_irreducible_def | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 290 | using irreducible_mult_imp_irreducible[OF assms] assms(1) by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 291 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 292 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 293 | subsection \<open>Primes\<close> | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 294 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 295 | lemma (in domain) zero_is_prime: "prime R \<zero>" "prime (mult_of R) \<zero>" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 296 | using integral unfolding prime_def factor_def Units_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 297 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 298 | lemma (in domain) prime_eq_prime_mult: | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 299 | assumes "p \<in> carrier R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 300 | shows "prime R p \<longleftrightarrow> prime (mult_of R) p" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 301 | proof (cases "p = \<zero>", auto simp add: zero_is_prime) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 302 | assume "p \<noteq> \<zero>" "prime R p" thus "prime (mult_of R) p" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 303 | unfolding prime_def | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 304 | using divides_mult_imp_divides Units_mult_eq_Units mult_mult_of | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 305 | by (metis DiffD1 assms carrier_mult_of divides_imp_divides_mult) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 306 | next | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 307 | assume p: "p \<noteq> \<zero>" "prime (mult_of R) p" show "prime R p" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 308 | proof (rule primeI) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 309 | show "p \<notin> Units R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 310 | using p(2) Units_mult_eq_Units unfolding prime_def by simp | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 311 | next | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 312 | fix a b assume a: "a \<in> carrier R" and b: "b \<in> carrier R" and dvd: "p divides a \<otimes> b" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 313 | then obtain c where c: "c \<in> carrier R" "a \<otimes> b = p \<otimes> c" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 314 | unfolding factor_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 315 | show "p divides a \<or> p divides b" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 316 | proof (cases "a \<otimes> b = \<zero>") | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 317 | case True thus ?thesis | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 318 | using integral[OF _ a b] c unfolding factor_def by force | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 319 | next | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 320 | case False | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 321 | hence "p divides\<^bsub>(mult_of R)\<^esub> (a \<otimes> b)" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 322 | using divides_imp_divides_mult[OF assms _ dvd] m_closed[OF a b] by simp | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 323 | moreover have "a \<noteq> \<zero>" "b \<noteq> \<zero>" "c \<noteq> \<zero>" | 
| 69070 | 324 | using False a b c p l_null integral_iff by (auto, simp add: assms) | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 325 | ultimately show ?thesis | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 326 | using a b p unfolding prime_def | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 327 | by (auto, metis Diff_iff divides_mult_imp_divides singletonD) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 328 | qed | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 329 | qed | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 330 | qed | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 331 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 332 | lemma (in domain) ring_primeE: | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 333 | assumes "p \<in> carrier R" "ring_prime p" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 334 | shows "p \<noteq> \<zero>" "prime (mult_of R) p" "prime R p" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 335 | using assms prime_eq_prime_mult unfolding ring_prime_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 336 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 337 | lemma (in ring) ring_primeI: (* in ring only to avoid instantiating R *) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 338 | assumes "p \<noteq> \<zero>" "prime R p" shows "ring_prime p" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 339 | using assms unfolding ring_prime_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 340 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 341 | lemma (in domain) ring_primeI': | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 342 |   assumes "p \<in> carrier R - { \<zero> }" "prime (mult_of R) p"
 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 343 | shows "ring_prime p" | 
| 69070 | 344 | using assms prime_eq_prime_mult unfolding ring_prime_def by auto | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 345 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 346 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 347 | subsection \<open>Basic Properties\<close> | 
| 68578 | 348 | |
| 349 | lemma (in cring) to_contain_is_to_divide: | |
| 350 | assumes "a \<in> carrier R" "b \<in> carrier R" | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 351 | shows "PIdl b \<subseteq> PIdl a \<longleftrightarrow> a divides b" | 
| 69070 | 352 | proof | 
| 68578 | 353 | show "PIdl b \<subseteq> PIdl a \<Longrightarrow> a divides b" | 
| 354 | proof - | |
| 355 | assume "PIdl b \<subseteq> PIdl a" | |
| 356 | hence "b \<in> PIdl a" | |
| 357 | by (meson assms(2) local.ring_axioms ring.cgenideal_self subsetCE) | |
| 358 | thus ?thesis | |
| 69070 | 359 | unfolding factor_def cgenideal_def using m_comm assms(1) by blast | 
| 68578 | 360 | qed | 
| 361 | show "a divides b \<Longrightarrow> PIdl b \<subseteq> PIdl a" | |
| 362 | proof - | |
| 363 | assume "a divides b" then obtain c where c: "c \<in> carrier R" "b = c \<otimes> a" | |
| 364 | unfolding factor_def using m_comm[OF assms(1)] by blast | |
| 365 | show "PIdl b \<subseteq> PIdl a" | |
| 366 | proof | |
| 367 | fix x assume "x \<in> PIdl b" | |
| 368 | then obtain d where d: "d \<in> carrier R" "x = d \<otimes> b" | |
| 369 | unfolding cgenideal_def by blast | |
| 370 | hence "x = (d \<otimes> c) \<otimes> a" | |
| 371 | using c d m_assoc assms by simp | |
| 372 | thus "x \<in> PIdl a" | |
| 373 | unfolding cgenideal_def using m_assoc assms c d by blast | |
| 374 | qed | |
| 375 | qed | |
| 376 | qed | |
| 377 | ||
| 378 | lemma (in cring) associated_iff_same_ideal: | |
| 379 | assumes "a \<in> carrier R" "b \<in> carrier R" | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 380 | shows "a \<sim> b \<longleftrightarrow> PIdl a = PIdl b" | 
| 68578 | 381 | unfolding associated_def | 
| 382 | using to_contain_is_to_divide[OF assms] | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 383 | to_contain_is_to_divide[OF assms(2,1)] by auto | 
| 68578 | 384 | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 385 | lemma (in cring) ideal_eq_carrier_iff: | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 386 | assumes "a \<in> carrier R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 387 | shows "carrier R = PIdl a \<longleftrightarrow> a \<in> Units R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 388 | proof | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 389 | assume "carrier R = PIdl a" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 390 | hence "\<one> \<in> PIdl a" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 391 | by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 392 | then obtain b where "b \<in> carrier R" "\<one> = a \<otimes> b" "\<one> = b \<otimes> a" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 393 | unfolding cgenideal_def using m_comm[OF assms] by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 394 | thus "a \<in> Units R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 395 | using assms unfolding Units_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 396 | next | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 397 | assume "a \<in> Units R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 398 | then have inv_a: "inv a \<in> carrier R" "inv a \<otimes> a = \<one>" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 399 | by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 400 | have "carrier R \<subseteq> PIdl a" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 401 | proof | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 402 | fix b assume "b \<in> carrier R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 403 | hence "(b \<otimes> inv a) \<otimes> a = b" and "b \<otimes> inv a \<in> carrier R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 404 | using m_assoc[OF _ inv_a(1) assms] inv_a by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 405 | thus "b \<in> PIdl a" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 406 | unfolding cgenideal_def by force | 
| 68578 | 407 | qed | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 408 | thus "carrier R = PIdl a" | 
| 69070 | 409 | using assms by (simp add: cgenideal_eq_rcos r_coset_subset_G subset_antisym) | 
| 68578 | 410 | qed | 
| 411 | ||
| 412 | lemma (in domain) primeideal_iff_prime: | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 413 |   assumes "p \<in> carrier R - { \<zero> }"
 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 414 | shows "primeideal (PIdl p) R \<longleftrightarrow> ring_prime p" | 
| 69070 | 415 | proof | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 416 | assume PIdl: "primeideal (PIdl p) R" show "ring_prime p" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 417 | proof (rule ring_primeI) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 418 | show "p \<noteq> \<zero>" using assms by simp | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 419 | next | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 420 | show "prime R p" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 421 | proof (rule primeI) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 422 | show "p \<notin> Units R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 423 | using ideal_eq_carrier_iff[of p] assms primeideal.I_notcarr[OF PIdl] by auto | 
| 68578 | 424 | next | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 425 | fix a b assume a: "a \<in> carrier R" and b: "b \<in> carrier R" and dvd: "p divides a \<otimes> b" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 426 | hence "a \<otimes> b \<in> PIdl p" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 427 | by (meson assms DiffD1 cgenideal_self contra_subsetD m_closed to_contain_is_to_divide) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 428 | hence "a \<in> PIdl p \<or> b \<in> PIdl p" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 429 | using primeideal.I_prime[OF PIdl a b] by simp | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 430 | thus "p divides a \<or> p divides b" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 431 | using assms a b Idl_subset_ideal' cgenideal_eq_genideal to_contain_is_to_divide by auto | 
| 68578 | 432 | qed | 
| 433 | qed | |
| 434 | next | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 435 | assume prime: "ring_prime p" show "primeideal (PIdl p) R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 436 | proof (rule primeidealI[OF cgenideal_ideal cring_axioms]) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 437 | show "p \<in> carrier R" and "carrier R \<noteq> PIdl p" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 438 | using ideal_eq_carrier_iff[of p] assms prime | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 439 | unfolding ring_prime_def prime_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 440 | next | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 441 | fix a b assume a: "a \<in> carrier R" and b: "b \<in> carrier R" "a \<otimes> b \<in> PIdl p" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 442 | hence "p divides a \<otimes> b" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 443 | using assms Idl_subset_ideal' cgenideal_eq_genideal to_contain_is_to_divide by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 444 | hence "p divides a \<or> p divides b" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 445 | using a b assms primeE[OF ring_primeE(3)[OF _ prime]] by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 446 | thus "a \<in> PIdl p \<or> b \<in> PIdl p" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 447 | using a b assms Idl_subset_ideal' cgenideal_eq_genideal to_contain_is_to_divide by auto | 
| 68578 | 448 | qed | 
| 449 | qed | |
| 450 | ||
| 451 | ||
| 452 | subsection \<open>Noetherian Rings\<close> | |
| 453 | ||
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 454 | lemma (in ring) chain_Union_is_ideal: | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 455 |   assumes "subset.chain { I. ideal I R } C"
 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 456 |   shows "ideal (if C = {} then { \<zero> } else (\<Union>C)) R"
 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 457 | proof (cases "C = {}")
 | 
| 69070 | 458 | case True thus ?thesis by (simp add: zeroideal) | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 459 | next | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 460 | case False have "ideal (\<Union>C) R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 461 | proof (rule idealI[OF ring_axioms]) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 462 | show "subgroup (\<Union>C) (add_monoid R)" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 463 | proof | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 464 | show "\<Union>C \<subseteq> carrier (add_monoid R)" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 465 | using assms subgroup.subset[OF additive_subgroup.a_subgroup] | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 466 | unfolding pred_on.chain_def ideal_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 467 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 468 | obtain I where I: "I \<in> C" "ideal I R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 469 | using False assms unfolding pred_on.chain_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 470 | thus "\<one>\<^bsub>add_monoid R\<^esub> \<in> \<Union>C" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 471 | using additive_subgroup.zero_closed[OF ideal.axioms(1)[OF I(2)]] by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 472 | next | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 473 | fix x y assume "x \<in> \<Union>C" "y \<in> \<Union>C" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 474 | then obtain I where I: "I \<in> C" "x \<in> I" "y \<in> I" | 
| 69070 | 475 | using assms unfolding pred_on.chain_def by blast | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 476 | hence ideal: "ideal I R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 477 | using assms unfolding pred_on.chain_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 478 | thus "x \<otimes>\<^bsub>add_monoid R\<^esub> y \<in> \<Union>C" | 
| 69070 | 479 | using UnionI I additive_subgroup.a_closed unfolding ideal_def by fastforce | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 480 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 481 | show "inv\<^bsub>add_monoid R\<^esub> x \<in> \<Union>C" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 482 | using UnionI I additive_subgroup.a_inv_closed ideal unfolding ideal_def a_inv_def by metis | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 483 | qed | 
| 68578 | 484 | next | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 485 | fix a x assume a: "a \<in> \<Union>C" and x: "x \<in> carrier R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 486 | then obtain I where I: "ideal I R" "a \<in> I" and "I \<in> C" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 487 | using assms unfolding pred_on.chain_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 488 | thus "x \<otimes> a \<in> \<Union>C" and "a \<otimes> x \<in> \<Union>C" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 489 | using ideal.I_l_closed[OF I x] ideal.I_r_closed[OF I x] by auto | 
| 68578 | 490 | qed | 
| 491 | thus ?thesis | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 492 | using False by simp | 
| 68578 | 493 | qed | 
| 494 | ||
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 495 | lemma (in noetherian_ring) ideal_chain_is_trivial: | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 496 |   assumes "C \<noteq> {}" "subset.chain { I. ideal I R } C"
 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 497 | shows "\<Union>C \<in> C" | 
| 68578 | 498 | proof - | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 499 |   { fix S assume "finite S" "S \<subseteq> \<Union>C"
 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 500 | hence "\<exists>I. I \<in> C \<and> S \<subseteq> I" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 501 | proof (induct S) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 502 | case empty thus ?case | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 503 | using assms(1) by blast | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 504 | next | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 505 | case (insert s S) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 506 | then obtain I where I: "I \<in> C" "S \<subseteq> I" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 507 | by blast | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 508 | moreover obtain I' where I': "I' \<in> C" "s \<in> I'" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 509 | using insert(4) by blast | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 510 | ultimately have "I \<subseteq> I' \<or> I' \<subseteq> I" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 511 | using assms unfolding pred_on.chain_def by blast | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 512 | thus ?case | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 513 | using I I' by blast | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 514 | qed } note aux_lemma = this | 
| 68578 | 515 | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 516 | obtain S where S: "finite S" "S \<subseteq> carrier R" "\<Union>C = Idl S" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 517 | using finetely_gen[OF chain_Union_is_ideal[OF assms(2)]] assms(1) by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 518 | then obtain I where I: "I \<in> C" and "S \<subseteq> I" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 519 | using aux_lemma[OF S(1)] genideal_self[OF S(2)] by blast | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 520 | hence "Idl S \<subseteq> I" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 521 | using assms unfolding pred_on.chain_def | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 522 | by (metis genideal_minimal mem_Collect_eq rev_subsetD) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 523 | hence "\<Union>C = I" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 524 | using S(3) I by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 525 | thus ?thesis | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 526 | using I by simp | 
| 68578 | 527 | qed | 
| 528 | ||
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 529 | lemma (in ring) trivial_ideal_chain_imp_noetherian: | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 530 |   assumes "\<And>C. \<lbrakk> C \<noteq> {}; subset.chain { I. ideal I R } C \<rbrakk> \<Longrightarrow> \<Union>C \<in> C"
 | 
| 68578 | 531 | shows "noetherian_ring R" | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 532 | proof (auto simp add: noetherian_ring_def noetherian_ring_axioms_def ring_axioms) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 533 | fix I assume I: "ideal I R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 534 | have in_carrier: "I \<subseteq> carrier R" and add_subgroup: "additive_subgroup I R" | 
| 69070 | 535 | using ideal.axioms(1)[OF I] additive_subgroup.a_subset by auto | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 536 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 537 |   define S where "S = { Idl S' | S'. S' \<subseteq> I \<and> finite S' }"
 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 538 | have "\<exists>M \<in> S. \<forall>S' \<in> S. M \<subseteq> S' \<longrightarrow> S' = M" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 539 | proof (rule subset_Zorn) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 540 | fix C assume C: "subset.chain S C" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 541 | show "\<exists>U \<in> S. \<forall>S' \<in> C. S' \<subseteq> U" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 542 |     proof (cases "C = {}")
 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 543 | case True | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 544 |       have "{ \<zero> } \<in> S"
 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 545 | using additive_subgroup.zero_closed[OF add_subgroup] genideal_zero | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 546 | by (auto simp add: S_def) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 547 | thus ?thesis | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 548 | using True by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 549 | next | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 550 | case False | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 551 |       have "S \<subseteq> { I. ideal I R }"
 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 552 | using additive_subgroup.a_subset[OF add_subgroup] genideal_ideal | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 553 | by (auto simp add: S_def) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 554 |       hence "subset.chain { I. ideal I R } C"
 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 555 | using C unfolding pred_on.chain_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 556 | then have "\<Union>C \<in> C" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 557 | using assms False by simp | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 558 | thus ?thesis | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 559 | by (meson C Union_upper pred_on.chain_def subsetCE) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 560 | qed | 
| 68578 | 561 | qed | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 562 | then obtain M where M: "M \<in> S" "\<And>S'. \<lbrakk>S' \<in> S; M \<subseteq> S' \<rbrakk> \<Longrightarrow> S' = M" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 563 | by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 564 | then obtain S' where S': "S' \<subseteq> I" "finite S'" "M = Idl S'" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 565 | by (auto simp add: S_def) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 566 | hence "M \<subseteq> I" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 567 | using I genideal_minimal by (auto simp add: S_def) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 568 | moreover have "I \<subseteq> M" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 569 | proof (rule ccontr) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 570 | assume "\<not> I \<subseteq> M" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 571 | then obtain a where a: "a \<in> I" "a \<notin> M" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 572 | by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 573 | have "M \<subseteq> Idl (insert a S')" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 574 | using S' a(1) genideal_minimal[of "Idl (insert a S')" S'] | 
| 69070 | 575 | in_carrier genideal_ideal genideal_self | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 576 | by (meson insert_subset subset_trans) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 577 | moreover have "Idl (insert a S') \<in> S" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 578 | using a(1) S' by (auto simp add: S_def) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 579 | ultimately have "M = Idl (insert a S')" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 580 | using M(2) by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 581 | hence "a \<in> M" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 582 | using genideal_self S'(1) a (1) in_carrier by (meson insert_subset subset_trans) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 583 | from \<open>a \<in> M\<close> and \<open>a \<notin> M\<close> show False by simp | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 584 | qed | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 585 | ultimately have "M = I" by simp | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 586 | thus "\<exists>A \<subseteq> carrier R. finite A \<and> I = Idl A" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 587 | using S' in_carrier by blast | 
| 68578 | 588 | qed | 
| 589 | ||
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 590 | lemma (in noetherian_domain) factorization_property: | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 591 |   assumes "a \<in> carrier R - { \<zero> }" "a \<notin> Units R"
 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 592 | shows "\<exists>fs. set fs \<subseteq> carrier (mult_of R) \<and> wfactors (mult_of R) fs a" (is "?factorizable a") | 
| 68578 | 593 | proof (rule ccontr) | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 594 | assume a: "\<not> ?factorizable a" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 595 |   define S where "S = { PIdl r | r. r \<in> carrier R - { \<zero> } \<and> r \<notin> Units R \<and> \<not> ?factorizable r }"
 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 596 | then obtain C where C: "subset.maxchain S C" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 597 | using subset.Hausdorff by blast | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 598 | hence chain: "subset.chain S C" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 599 | using pred_on.maxchain_def by blast | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 600 |   moreover have "S \<subseteq> { I. ideal I R }"
 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 601 | using cgenideal_ideal by (auto simp add: S_def) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 602 |   ultimately have "subset.chain { I. ideal I R } C"
 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 603 | by (meson dual_order.trans pred_on.chain_def) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 604 | moreover have "PIdl a \<in> S" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 605 | using assms a by (auto simp add: S_def) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 606 |   hence "subset.chain S { PIdl a }"
 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 607 | unfolding pred_on.chain_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 608 |   hence "C \<noteq> {}"
 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 609 | using C unfolding pred_on.maxchain_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 610 | ultimately have "\<Union>C \<in> C" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 611 | using ideal_chain_is_trivial by simp | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 612 | hence "\<Union>C \<in> S" | 
| 69070 | 613 | using chain unfolding pred_on.chain_def by auto | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 614 |   then obtain r where r: "\<Union>C = PIdl r" "r \<in> carrier R - { \<zero> }" "r \<notin> Units R" "\<not> ?factorizable r"
 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 615 | by (auto simp add: S_def) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 616 |   have "\<exists>p. p \<in> carrier R - { \<zero> } \<and> p \<notin> Units R \<and> \<not> ?factorizable p \<and> p divides r \<and> \<not> r divides p"
 | 
| 68578 | 617 | proof - | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 618 | have "wfactors (mult_of R) [ r ] r" if "irreducible (mult_of R) r" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 619 | using r(2) that unfolding wfactors_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 620 | hence "\<not> irreducible (mult_of R) r" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 621 | using r(2,4) by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 622 | hence "\<not> ring_irreducible r" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 623 | using ring_irreducibleE(3) r(2) by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 624 | then obtain p1 p2 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 625 | where p1_p2: "p1 \<in> carrier R" "p2 \<in> carrier R" "r = p1 \<otimes> p2" "p1 \<notin> Units R" "p2 \<notin> Units R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 626 | using ring_irreducibleI[OF r(2-3)] by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 627 | hence in_carrier: "p1 \<in> carrier (mult_of R)" "p2 \<in> carrier (mult_of R)" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 628 | using r(2) by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 629 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 630 | have "\<lbrakk> ?factorizable p1; ?factorizable p2 \<rbrakk> \<Longrightarrow> ?factorizable r" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 631 | using mult_of.wfactors_mult[OF _ _ in_carrier] p1_p2(3) by (metis le_sup_iff set_append) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 632 | hence "\<not> ?factorizable p1 \<or> \<not> ?factorizable p2" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 633 | using r(4) by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 634 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 635 | moreover | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 636 | have "\<And>p1 p2. \<lbrakk> p1 \<in> carrier R; p2 \<in> carrier R; r = p1 \<otimes> p2; r divides p1 \<rbrakk> \<Longrightarrow> p2 \<in> Units R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 637 | proof - | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 638 | fix p1 p2 assume A: "p1 \<in> carrier R" "p2 \<in> carrier R" "r = p1 \<otimes> p2" "r divides p1" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 639 | then obtain c where c: "c \<in> carrier R" "p1 = r \<otimes> c" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 640 | unfolding factor_def by blast | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 641 | hence "\<one> = c \<otimes> p2" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 642 | using A m_lcancel[OF _ _ one_closed, of r "c \<otimes> p2"] r(2) by (auto, metis m_assoc m_closed) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 643 | thus "p2 \<in> Units R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 644 | unfolding Units_def using c A(2) m_comm[OF c(1) A(2)] by auto | 
| 68578 | 645 | qed | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 646 | hence "\<not> r divides p1" and "\<not> r divides p2" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 647 | using p1_p2 m_comm[OF p1_p2(1-2)] by blast+ | 
| 68578 | 648 | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 649 | ultimately show ?thesis | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 650 | using p1_p2 in_carrier by (metis carrier_mult_of dividesI' m_comm) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 651 | qed | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 652 | then obtain p | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 653 |     where p: "p \<in> carrier R - { \<zero> }" "p \<notin> Units R" "\<not> ?factorizable p" "p divides r" "\<not> r divides p"
 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 654 | by blast | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 655 | hence "PIdl p \<in> S" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 656 | unfolding S_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 657 | moreover have "\<Union>C \<subset> PIdl p" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 658 | using p r to_contain_is_to_divide unfolding r(1) by (metis Diff_iff psubsetI) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 659 | ultimately have "subset.chain S (insert (PIdl p) C)" and "C \<subset> (insert (PIdl p) C)" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 660 | unfolding pred_on.chain_def by (metis C psubsetE subset_maxchain_max, blast) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 661 | thus False | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 662 | using C unfolding pred_on.maxchain_def by blast | 
| 68578 | 663 | qed | 
| 664 | ||
| 665 | ||
| 666 | subsection \<open>Principal Domains\<close> | |
| 667 | ||
| 668 | sublocale principal_domain \<subseteq> noetherian_domain | |
| 669 | proof | |
| 670 | fix I assume "ideal I R" | |
| 671 |   then obtain i where "i \<in> carrier R" "I = Idl { i }"
 | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 672 | using exists_gen cgenideal_eq_genideal by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 673 | thus "\<exists>A \<subseteq> carrier R. finite A \<and> I = Idl A" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 674 | by blast | 
| 68578 | 675 | qed | 
| 676 | ||
| 677 | lemma (in principal_domain) irreducible_imp_maximalideal: | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 678 | assumes "p \<in> carrier R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 679 | and "ring_irreducible p" | 
| 68578 | 680 | shows "maximalideal (PIdl p) R" | 
| 681 | proof (rule maximalidealI) | |
| 682 | show "ideal (PIdl p) R" | |
| 683 | using assms(1) by (simp add: cgenideal_ideal) | |
| 684 | next | |
| 685 | show "carrier R \<noteq> PIdl p" | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 686 | using ideal_eq_carrier_iff[OF assms(1)] ring_irreducibleE(4)[OF assms] by auto | 
| 68578 | 687 | next | 
| 688 | fix J assume J: "ideal J R" "PIdl p \<subseteq> J" "J \<subseteq> carrier R" | |
| 689 | then obtain q where q: "q \<in> carrier R" "J = PIdl q" | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 690 | using exists_gen[OF J(1)] cgenideal_eq_rcos by metis | 
| 68578 | 691 | hence "q divides p" | 
| 692 | using to_contain_is_to_divide[of q p] using assms(1) J(1-2) by simp | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 693 | then obtain r where r: "r \<in> carrier R" "p = q \<otimes> r" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 694 | unfolding factor_def by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 695 | hence "q \<in> Units R \<or> r \<in> Units R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 696 | using ring_irreducibleE(5)[OF assms q(1)] by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 697 | thus "J = PIdl p \<or> J = carrier R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 698 | proof | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 699 | assume "q \<in> Units R" thus ?thesis | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 700 | using ideal_eq_carrier_iff[OF q(1)] q(2) by auto | 
| 68578 | 701 | next | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 702 | assume "r \<in> Units R" hence "p \<sim> q" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 703 | using assms(1) r q(1) associatedI2' by blast | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 704 | thus ?thesis | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 705 | unfolding associated_iff_same_ideal[OF assms(1) q(1)] q(2) by auto | 
| 68578 | 706 | qed | 
| 707 | qed | |
| 708 | ||
| 709 | corollary (in principal_domain) primeness_condition: | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 710 | assumes "p \<in> carrier R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 711 | shows "ring_irreducible p \<longleftrightarrow> ring_prime p" | 
| 68578 | 712 | proof | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 713 | show "ring_irreducible p \<Longrightarrow> ring_prime p" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 714 | using maximalideal_prime[OF irreducible_imp_maximalideal] ring_irreducibleE(1) | 
| 69070 | 715 | primeideal_iff_prime assms by auto | 
| 68578 | 716 | next | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 717 | show "ring_prime p \<Longrightarrow> ring_irreducible p" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 718 | using mult_of.prime_irreducible ring_primeI[of p] ring_irreducibleI' assms | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 719 | unfolding ring_prime_def prime_eq_prime_mult[OF assms] by auto | 
| 68578 | 720 | qed | 
| 721 | ||
| 722 | lemma (in principal_domain) domain_iff_prime: | |
| 723 |   assumes "a \<in> carrier R - { \<zero> }"
 | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 724 | shows "domain (R Quot (PIdl a)) \<longleftrightarrow> ring_prime a" | 
| 68578 | 725 | using quot_domain_iff_primeideal[of "PIdl a"] primeideal_iff_prime[of a] | 
| 726 | cgenideal_ideal[of a] assms by auto | |
| 727 | ||
| 728 | lemma (in principal_domain) field_iff_prime: | |
| 729 |   assumes "a \<in> carrier R - { \<zero> }"
 | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 730 | shows "field (R Quot (PIdl a)) \<longleftrightarrow> ring_prime a" | 
| 68578 | 731 | proof | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 732 | show "ring_prime a \<Longrightarrow> field (R Quot (PIdl a))" | 
| 68578 | 733 | using primeness_condition[of a] irreducible_imp_maximalideal[of a] | 
| 734 | maximalideal.quotient_is_field[of "PIdl a" R] is_cring assms by auto | |
| 735 | next | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 736 | show "field (R Quot (PIdl a)) \<Longrightarrow> ring_prime a" | 
| 68578 | 737 | unfolding field_def using domain_iff_prime[of a] assms by auto | 
| 738 | qed | |
| 739 | ||
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 740 | |
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 741 | sublocale principal_domain < mult_of: primeness_condition_monoid "mult_of R" | 
| 68578 | 742 | rewrites "mult (mult_of R) = mult R" | 
| 743 | and "one (mult_of R) = one R" | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 744 | unfolding primeness_condition_monoid_def | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 745 | primeness_condition_monoid_axioms_def | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 746 | proof (auto simp add: mult_of.is_comm_monoid_cancel) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 747 | fix a assume a: "a \<in> carrier R" "a \<noteq> \<zero>" "irreducible (mult_of R) a" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 748 | show "prime (mult_of R) a" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 749 | using primeness_condition[OF a(1)] irreducible_mult_imp_irreducible[OF _ a(3)] a(1-2) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 750 | unfolding ring_prime_def ring_irreducible_def prime_eq_prime_mult[OF a(1)] by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 751 | qed | 
| 68578 | 752 | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 753 | sublocale principal_domain < mult_of: factorial_monoid "mult_of R" | 
| 68578 | 754 | rewrites "mult (mult_of R) = mult R" | 
| 755 | and "one (mult_of R) = one R" | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 756 | using mult_of.wfactors_unique factorization_property mult_of.is_comm_monoid_cancel | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 757 | by (auto intro!: mult_of.factorial_monoidI) | 
| 68578 | 758 | |
| 759 | sublocale principal_domain \<subseteq> factorial_domain | |
| 69070 | 760 | unfolding factorial_domain_def using domain_axioms mult_of.factorial_monoid_axioms by simp | 
| 68578 | 761 | |
| 762 | lemma (in principal_domain) ideal_sum_iff_gcd: | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 763 | assumes "a \<in> carrier R" "b \<in> carrier R" "d \<in> carrier R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 764 | shows "PIdl d = PIdl a <+>\<^bsub>R\<^esub> PIdl b \<longleftrightarrow> d gcdof a b" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 765 | proof - | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 766 |   { fix a b d
 | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 767 | assume in_carrier: "a \<in> carrier R" "b \<in> carrier R" "d \<in> carrier R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 768 | and ideal_eq: "PIdl d = PIdl a <+>\<^bsub>R\<^esub> PIdl b" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 769 | have "d gcdof a b" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 770 | proof (auto simp add: isgcd_def) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 771 | have "a \<in> PIdl d" and "b \<in> PIdl d" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 772 | using in_carrier(1-2)[THEN cgenideal_ideal] additive_subgroup.zero_closed[OF ideal.axioms(1)] | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 773 | in_carrier(1-2)[THEN cgenideal_self] in_carrier(1-2) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 774 | unfolding ideal_eq set_add_def' by force+ | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 775 | thus "d divides a" and "d divides b" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 776 | using in_carrier(1,2)[THEN to_contain_is_to_divide[OF in_carrier(3)]] | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 777 | cgenideal_minimal[OF cgenideal_ideal[OF in_carrier(3)]] by simp+ | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 778 | next | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 779 | fix c assume c: "c \<in> carrier R" "c divides a" "c divides b" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 780 | hence "PIdl a \<subseteq> PIdl c" and "PIdl b \<subseteq> PIdl c" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 781 | using to_contain_is_to_divide in_carrier by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 782 | hence "PIdl a <+>\<^bsub>R\<^esub> PIdl b \<subseteq> PIdl c" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 783 | by (metis Un_subset_iff c(1) in_carrier(1-2) cgenideal_ideal genideal_minimal union_genideal) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 784 | thus "c divides d" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 785 | using ideal_eq to_contain_is_to_divide[OF c(1) in_carrier(3)] by simp | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 786 | qed } note aux_lemma = this | 
| 68578 | 787 | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 788 | have "PIdl d = PIdl a <+>\<^bsub>R\<^esub> PIdl b \<Longrightarrow> d gcdof a b" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 789 | using aux_lemma assms by simp | 
| 68578 | 790 | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 791 | moreover | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 792 | have "d gcdof a b \<Longrightarrow> PIdl d = PIdl a <+>\<^bsub>R\<^esub> PIdl b" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 793 | proof - | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 794 | assume d: "d gcdof a b" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 795 | obtain c where c: "c \<in> carrier R" "PIdl c = PIdl a <+>\<^bsub>R\<^esub> PIdl b" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 796 | using exists_gen[OF add_ideals[OF assms(1-2)[THEN cgenideal_ideal]]] by blast | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 797 | hence "c gcdof a b" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 798 | using aux_lemma assms by simp | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 799 | from \<open>d gcdof a b\<close> and \<open>c gcdof a b\<close> have "d \<sim> c" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 800 | using assms c(1) by (simp add: associated_def isgcd_def) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 801 | thus ?thesis | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 802 | using c(2) associated_iff_same_ideal[OF assms(3) c(1)] by simp | 
| 68578 | 803 | qed | 
| 804 | ||
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 805 | ultimately show ?thesis by auto | 
| 68578 | 806 | qed | 
| 807 | ||
| 808 | lemma (in principal_domain) bezout_identity: | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 809 | assumes "a \<in> carrier R" "b \<in> carrier R" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 810 | shows "PIdl a <+>\<^bsub>R\<^esub> PIdl b = PIdl (somegcd R a b)" | 
| 68578 | 811 | proof - | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 812 | have "\<exists>d \<in> carrier R. d gcdof a b" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 813 | using exists_gen[OF add_ideals[OF assms(1-2)[THEN cgenideal_ideal]]] | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 814 | ideal_sum_iff_gcd[OF assms(1-2)] by auto | 
| 68578 | 815 | thus ?thesis | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 816 | using ideal_sum_iff_gcd[OF assms(1-2)] somegcd_def | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 817 | by (metis (no_types, lifting) tfl_some) | 
| 68578 | 818 | qed | 
| 819 | ||
| 820 | subsection \<open>Euclidean Domains\<close> | |
| 821 | ||
| 822 | sublocale euclidean_domain \<subseteq> principal_domain | |
| 823 | unfolding principal_domain_def principal_domain_axioms_def | |
| 824 | proof (auto) | |
| 825 | show "domain R" by (simp add: domain_axioms) | |
| 826 | next | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 827 | fix I assume I: "ideal I R" have "principalideal I R" | 
| 68578 | 828 |   proof (cases "I = { \<zero> }")
 | 
| 69070 | 829 | case True thus ?thesis by (simp add: zeropideal) | 
| 68578 | 830 | next | 
| 831 |     case False hence A: "I - { \<zero> } \<noteq> {}"
 | |
| 69070 | 832 | using I additive_subgroup.zero_closed ideal.axioms(1) by auto | 
| 68578 | 833 |     define phi_img :: "nat set" where "phi_img = (\<phi> ` (I - { \<zero> }))"
 | 
| 69070 | 834 |     hence "phi_img \<noteq> {}" using A by simp
 | 
| 68578 | 835 | then obtain m where "m \<in> phi_img" "\<And>k. k \<in> phi_img \<Longrightarrow> m \<le> k" | 
| 836 | using exists_least_iff[of "\<lambda>n. n \<in> phi_img"] not_less by force | |
| 837 |     then obtain a where a: "a \<in> I - { \<zero> }" "\<And>b. b \<in> I - { \<zero> } \<Longrightarrow> \<phi> a \<le> \<phi> b"
 | |
| 838 | using phi_img_def by blast | |
| 839 | have "I = PIdl a" | |
| 840 | proof (rule ccontr) | |
| 841 | assume "I \<noteq> PIdl a" | |
| 842 | then obtain b where b: "b \<in> I" "b \<notin> PIdl a" | |
| 843 |         using I \<open>a \<in> I - {\<zero>}\<close> cgenideal_minimal by auto
 | |
| 844 | hence "b \<noteq> \<zero>" | |
| 845 | by (metis DiffD1 I a(1) additive_subgroup.zero_closed cgenideal_ideal ideal.Icarr ideal.axioms(1)) | |
| 846 | then obtain q r | |
| 847 | where eucl_div: "q \<in> carrier R" "r \<in> carrier R" "b = (a \<otimes> q) \<oplus> r" "r = \<zero> \<or> \<phi> r < \<phi> a" | |
| 848 | using euclidean_function[of b a] a(1) b(1) ideal.Icarr[OF I] by auto | |
| 849 | hence "r = \<zero> \<Longrightarrow> b \<in> PIdl a" | |
| 850 | unfolding cgenideal_def using m_comm[of a] ideal.Icarr[OF I] a(1) by auto | |
| 851 | hence 1: "\<phi> r < \<phi> a \<and> r \<noteq> \<zero>" | |
| 852 | using eucl_div(4) b(2) by auto | |
| 69070 | 853 | |
| 68578 | 854 | have "r = (\<ominus> (a \<otimes> q)) \<oplus> b" | 
| 855 | using eucl_div(1-3) a(1) b(1) ideal.Icarr[OF I] r_neg1 by auto | |
| 856 | moreover have "\<ominus> (a \<otimes> q) \<in> I" | |
| 857 | using eucl_div(1) a(1) I | |
| 858 | by (meson DiffD1 additive_subgroup.a_inv_closed ideal.I_r_closed ideal.axioms(1)) | |
| 859 | ultimately have 2: "r \<in> I" | |
| 860 | using b(1) additive_subgroup.a_closed[OF ideal.axioms(1)[OF I]] by auto | |
| 861 | ||
| 862 | from 1 and 2 show False | |
| 863 | using a(2) by fastforce | |
| 864 | qed | |
| 865 | thus ?thesis | |
| 866 | by (meson DiffD1 I cgenideal_is_principalideal ideal.Icarr local.a(1)) | |
| 867 | qed | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 868 | thus "\<exists>a \<in> carrier R. I = PIdl a" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 869 | by (simp add: cgenideal_eq_genideal principalideal.generate) | 
| 68578 | 870 | qed | 
| 871 | ||
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 872 | |
| 68578 | 873 | sublocale field \<subseteq> euclidean_domain R "\<lambda>_. 0" | 
| 874 | proof (rule euclidean_domainI) | |
| 875 | fix a b | |
| 876 | let ?eucl_div = "\<lambda>q r. q \<in> carrier R \<and> r \<in> carrier R \<and> a = b \<otimes> q \<oplus> r \<and> (r = \<zero> \<or> 0 < 0)" | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68604diff
changeset | 877 | |
| 68578 | 878 |   assume a: "a \<in> carrier R - { \<zero> }" and b: "b \<in> carrier R - { \<zero> }"
 | 
| 879 | hence "a = b \<otimes> ((inv b) \<otimes> a) \<oplus> \<zero>" | |
| 880 | by (metis DiffD1 Units_inv_closed Units_r_inv field_Units l_one m_assoc r_zero) | |
| 881 | hence "?eucl_div _ ((inv b) \<otimes> a) \<zero>" | |
| 882 | using a b field_Units by auto | |
| 883 | thus "\<exists>q r. ?eucl_div _ q r" | |
| 884 | by blast | |
| 885 | qed | |
| 886 | ||
| 69070 | 887 | end |