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