author | wenzelm |
Wed, 12 Mar 2025 11:39:00 +0100 | |
changeset 82265 | 4b875a4c83b0 |
parent 81600 | b1772698bd78 |
permissions | -rw-r--r-- |
35849 | 1 |
(* Title: HOL/Algebra/QuotRing.thy |
2 |
Author: Stephan Hohe |
|
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
3 |
Author: Paulo EmÃlio de Vilhena |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
4 |
*) |
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
5 |
|
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
6 |
theory QuotRing |
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
7 |
imports RingHom |
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
8 |
begin |
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
9 |
|
61382 | 10 |
section \<open>Quotient Rings\<close> |
27717
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
ballarin
parents:
27611
diff
changeset
|
11 |
|
61382 | 12 |
subsection \<open>Multiplication on Cosets\<close> |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
13 |
|
45005 | 14 |
definition rcoset_mult :: "[('a, _) ring_scheme, 'a set, 'a set, 'a set] \<Rightarrow> 'a set" |
81142 | 15 |
(\<open>(\<open>open_block notation=\<open>mixfix rcoset_mult\<close>\<close>[mod _:] _ \<Otimes>\<index> _)\<close> [81,81,81] 80) |
35848
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents:
35847
diff
changeset
|
16 |
where "rcoset_mult R I A B = (\<Union>a\<in>A. \<Union>b\<in>B. I +>\<^bsub>R\<^esub> (a \<otimes>\<^bsub>R\<^esub> b))" |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
17 |
|
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
18 |
|
69597 | 19 |
text \<open>\<^const>\<open>rcoset_mult\<close> fulfils the properties required by congruences\<close> |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
20 |
lemma (in ideal) rcoset_mult_add: |
68673 | 21 |
assumes "x \<in> carrier R" "y \<in> carrier R" |
22 |
shows "[mod I:] (I +> x) \<Otimes> (I +> y) = I +> (x \<otimes> y)" |
|
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
23 |
proof - |
81600 | 24 |
have 1: "z \<in> I +> x \<otimes> y" |
68673 | 25 |
if x'rcos: "x' \<in> I +> x" and y'rcos: "y' \<in> I +> y" and zrcos: "z \<in> I +> x' \<otimes> y'" for z x' y' |
26 |
proof - |
|
27 |
from that |
|
28 |
obtain hx hy hz where hxI: "hx \<in> I" and x': "x' = hx \<oplus> x" and hyI: "hy \<in> I" and y': "y' = hy \<oplus> y" |
|
29 |
and hzI: "hz \<in> I" and z: "z = hz \<oplus> (x' \<otimes> y')" |
|
30 |
by (auto simp: a_r_coset_def r_coset_def) |
|
31 |
note carr = assms hxI[THEN a_Hcarr] hyI[THEN a_Hcarr] hzI[THEN a_Hcarr] |
|
32 |
from z x' y' have "z = hz \<oplus> ((hx \<oplus> x) \<otimes> (hy \<oplus> y))" by simp |
|
33 |
also from carr have "\<dots> = (hz \<oplus> (hx \<otimes> (hy \<oplus> y)) \<oplus> x \<otimes> hy) \<oplus> x \<otimes> y" by algebra |
|
34 |
finally have z2: "z = (hz \<oplus> (hx \<otimes> (hy \<oplus> y)) \<oplus> x \<otimes> hy) \<oplus> x \<otimes> y" . |
|
35 |
from hxI hyI hzI carr have "hz \<oplus> (hx \<otimes> (hy \<oplus> y)) \<oplus> x \<otimes> hy \<in> I" |
|
36 |
by (simp add: I_l_closed I_r_closed) |
|
37 |
with z2 show ?thesis |
|
38 |
by (auto simp add: a_r_coset_def r_coset_def) |
|
39 |
qed |
|
40 |
have 2: "\<exists>a\<in>I +> x. \<exists>b\<in>I +> y. z \<in> I +> a \<otimes> b" if "z \<in> I +> x \<otimes> y" for z |
|
41 |
using assms a_rcos_self that by blast |
|
42 |
show ?thesis |
|
43 |
unfolding rcoset_mult_def using assms |
|
44 |
by (auto simp: intro!: 1 2) |
|
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
45 |
qed |
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
46 |
|
81600 | 47 |
|
61382 | 48 |
subsection \<open>Quotient Ring Definition\<close> |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
49 |
|
45005 | 50 |
definition FactRing :: "[('a,'b) ring_scheme, 'a set] \<Rightarrow> ('a set) ring" |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
77138
diff
changeset
|
51 |
(infixl \<open>Quot\<close> 65) |
35848
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents:
35847
diff
changeset
|
52 |
where "FactRing R I = |
45005 | 53 |
\<lparr>carrier = a_rcosets\<^bsub>R\<^esub> I, mult = rcoset_mult R I, |
54 |
one = (I +>\<^bsub>R\<^esub> \<one>\<^bsub>R\<^esub>), zero = I, add = set_add R\<rparr>" |
|
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
55 |
|
68673 | 56 |
lemmas FactRing_simps = FactRing_def A_RCOSETS_defs a_r_coset_def[symmetric] |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
57 |
|
81600 | 58 |
|
61382 | 59 |
subsection \<open>Factorization over General Ideals\<close> |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
60 |
|
61382 | 61 |
text \<open>The quotient is a ring\<close> |
45005 | 62 |
lemma (in ideal) quotient_is_ring: "ring (R Quot I)" |
68673 | 63 |
proof (rule ringI) |
64 |
show "abelian_group (R Quot I)" |
|
81600 | 65 |
by (rule comm_group_abelian_groupI) |
66 |
(simp add: FactRing_def a_factorgroup_is_comm_group[unfolded A_FactGroup_def']) |
|
68673 | 67 |
show "Group.monoid (R Quot I)" |
68 |
by (rule monoidI) |
|
69 |
(auto simp add: FactRing_simps rcoset_mult_add m_assoc) |
|
70 |
qed (auto simp: FactRing_simps rcoset_mult_add a_rcos_sum l_distr r_distr) |
|
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
71 |
|
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
72 |
|
61382 | 73 |
text \<open>This is a ring homomorphism\<close> |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
74 |
|
67399 | 75 |
lemma (in ideal) rcos_ring_hom: "((+>) I) \<in> ring_hom R (R Quot I)" |
68673 | 76 |
by (simp add: ring_hom_memI FactRing_def a_rcosetsI[OF a_subset] rcoset_mult_add a_rcos_sum) |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
77 |
|
67399 | 78 |
lemma (in ideal) rcos_ring_hom_ring: "ring_hom_ring R (R Quot I) ((+>) I)" |
68673 | 79 |
by (simp add: local.ring_axioms quotient_is_ring rcos_ring_hom ring_hom_ringI2) |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
80 |
|
61382 | 81 |
text \<open>The quotient of a cring is also commutative\<close> |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
82 |
lemma (in ideal) quotient_is_cring: |
27611 | 83 |
assumes "cring R" |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
84 |
shows "cring (R Quot I)" |
27611 | 85 |
proof - |
29237 | 86 |
interpret cring R by fact |
45005 | 87 |
show ?thesis |
68673 | 88 |
apply (intro cring.intro comm_monoid.intro comm_monoid_axioms.intro quotient_is_ring) |
45005 | 89 |
apply (rule ring.axioms[OF quotient_is_ring]) |
68673 | 90 |
apply (auto simp add: FactRing_simps rcoset_mult_add m_comm) |
45005 | 91 |
done |
27611 | 92 |
qed |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
93 |
|
61382 | 94 |
text \<open>Cosets as a ring homomorphism on crings\<close> |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
95 |
lemma (in ideal) rcos_ring_hom_cring: |
27611 | 96 |
assumes "cring R" |
67399 | 97 |
shows "ring_hom_cring R (R Quot I) ((+>) I)" |
27611 | 98 |
proof - |
29237 | 99 |
interpret cring R by fact |
45005 | 100 |
show ?thesis |
101 |
apply (rule ring_hom_cringI) |
|
102 |
apply (rule rcos_ring_hom_ring) |
|
103 |
apply (rule is_cring) |
|
104 |
apply (rule quotient_is_cring) |
|
81600 | 105 |
apply (rule is_cring) |
106 |
done |
|
27611 | 107 |
qed |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
108 |
|
35849 | 109 |
|
61382 | 110 |
subsection \<open>Factorization over Prime Ideals\<close> |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
111 |
|
61382 | 112 |
text \<open>The quotient ring generated by a prime ideal is a domain\<close> |
45005 | 113 |
lemma (in primeideal) quotient_is_domain: "domain (R Quot I)" |
68673 | 114 |
proof - |
115 |
have 1: "I +> \<one> = I \<Longrightarrow> False" |
|
116 |
using I_notcarr a_rcos_self one_imp_carrier by blast |
|
117 |
have 2: "I +> x = I" |
|
118 |
if carr: "x \<in> carrier R" "y \<in> carrier R" |
|
45005 | 119 |
and a: "I +> x \<otimes> y = I" |
68673 | 120 |
and b: "I +> y \<noteq> I" for x y |
121 |
by (metis I_prime a a_rcos_const a_rcos_self b m_closed that) |
|
122 |
show ?thesis |
|
123 |
apply (intro domain.intro quotient_is_cring is_cring domain_axioms.intro) |
|
124 |
apply (metis "1" FactRing_def monoid.simps(2) ring.simps(1)) |
|
125 |
apply (simp add: FactRing_simps) |
|
81600 | 126 |
apply (metis "2" rcoset_mult_add) |
127 |
done |
|
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
128 |
qed |
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
129 |
|
61382 | 130 |
text \<open>Generating right cosets of a prime ideal is a homomorphism |
131 |
on commutative rings\<close> |
|
67399 | 132 |
lemma (in primeideal) rcos_ring_hom_cring: "ring_hom_cring R (R Quot I) ((+>) I)" |
45005 | 133 |
by (rule rcos_ring_hom_cring) (rule is_cring) |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
134 |
|
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
135 |
|
61382 | 136 |
subsection \<open>Factorization over Maximal Ideals\<close> |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
137 |
|
68673 | 138 |
text \<open>In a commutative ring, the quotient ring over a maximal ideal is a field. |
139 |
The proof follows ``W. Adkins, S. Weintraub: Algebra -- An Approach via Module Theory''\<close> |
|
140 |
proposition (in maximalideal) quotient_is_field: |
|
27611 | 141 |
assumes "cring R" |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
142 |
shows "field (R Quot I)" |
27611 | 143 |
proof - |
29237 | 144 |
interpret cring R by fact |
68673 | 145 |
have 1: "\<zero>\<^bsub>R Quot I\<^esub> \<noteq> \<one>\<^bsub>R Quot I\<^esub>" \<comment> \<open>Quotient is not empty\<close> |
146 |
proof |
|
45005 | 147 |
assume "\<zero>\<^bsub>R Quot I\<^esub> = \<one>\<^bsub>R Quot I\<^esub>" |
148 |
then have II1: "I = I +> \<one>" by (simp add: FactRing_def) |
|
68673 | 149 |
then have "I = carrier R" |
81600 | 150 |
using a_rcos_self one_imp_carrier by blast |
45005 | 151 |
with I_notcarr show False by simp |
68673 | 152 |
qed |
153 |
have 2: "\<exists>y\<in>carrier R. I +> a \<otimes> y = I +> \<one>" if IanI: "I +> a \<noteq> I" and acarr: "a \<in> carrier R" for a |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67399
diff
changeset
|
154 |
\<comment> \<open>Existence of Inverse\<close> |
68673 | 155 |
proof - |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67399
diff
changeset
|
156 |
\<comment> \<open>Helper ideal \<open>J\<close>\<close> |
63040 | 157 |
define J :: "'a set" where "J = (carrier R #> a) <+> I" |
45005 | 158 |
have idealJ: "ideal J R" |
68673 | 159 |
using J_def acarr add_ideals cgenideal_eq_rcos cgenideal_ideal is_ideal by auto |
45005 | 160 |
have IinJ: "I \<subseteq> J" |
68673 | 161 |
proof (clarsimp simp: J_def r_coset_def set_add_defs) |
45005 | 162 |
fix x |
163 |
assume xI: "x \<in> I" |
|
68673 | 164 |
have "x = \<zero> \<otimes> a \<oplus> x" |
165 |
by (simp add: acarr xI) |
|
166 |
with xI show "\<exists>xa\<in>carrier R. \<exists>k\<in>I. x = xa \<otimes> a \<oplus> k" by fast |
|
45005 | 167 |
qed |
81600 | 168 |
have JnI: "J \<noteq> I" |
68673 | 169 |
proof - |
170 |
have "a \<notin> I" |
|
171 |
using IanI a_rcos_const by blast |
|
172 |
moreover have "a \<in> J" |
|
173 |
proof (simp add: J_def r_coset_def set_add_defs) |
|
174 |
from acarr |
|
175 |
have "a = \<one> \<otimes> a \<oplus> \<zero>" by algebra |
|
176 |
with one_closed and additive_subgroup.zero_closed[OF is_additive_subgroup] |
|
177 |
show "\<exists>x\<in>carrier R. \<exists>k\<in>I. a = x \<otimes> a \<oplus> k" by fast |
|
178 |
qed |
|
179 |
ultimately show ?thesis by blast |
|
45005 | 180 |
qed |
68673 | 181 |
then have Jcarr: "J = carrier R" |
182 |
using I_maximal IinJ additive_subgroup.a_subset idealJ ideal_def by blast |
|
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
183 |
|
69597 | 184 |
\<comment> \<open>Calculating an inverse for \<^term>\<open>a\<close>\<close> |
45005 | 185 |
from one_closed[folded Jcarr] |
68673 | 186 |
obtain r i where rcarr: "r \<in> carrier R" |
81600 | 187 |
and iI: "i \<in> I" and one: "\<one> = r \<otimes> a \<oplus> i" |
68673 | 188 |
by (auto simp add: J_def r_coset_def set_add_defs) |
189 |
||
45005 | 190 |
from one and rcarr and acarr and iI[THEN a_Hcarr] |
191 |
have rai1: "a \<otimes> r = \<ominus>i \<oplus> \<one>" by algebra |
|
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
192 |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67399
diff
changeset
|
193 |
\<comment> \<open>Lifting to cosets\<close> |
45005 | 194 |
from iI have "\<ominus>i \<oplus> \<one> \<in> I +> \<one>" |
195 |
by (intro a_rcosI, simp, intro a_subset, simp) |
|
196 |
with rai1 have "a \<otimes> r \<in> I +> \<one>" by simp |
|
197 |
then have "I +> \<one> = I +> a \<otimes> r" |
|
198 |
by (rule a_repr_independence, simp) (rule a_subgroup) |
|
199 |
||
200 |
from rcarr and this[symmetric] |
|
201 |
show "\<exists>r\<in>carrier R. I +> a \<otimes> r = I +> \<one>" by fast |
|
202 |
qed |
|
68673 | 203 |
show ?thesis |
204 |
apply (intro cring.cring_fieldI2 quotient_is_cring is_cring 1) |
|
205 |
apply (clarsimp simp add: FactRing_simps rcoset_mult_add 2) |
|
206 |
done |
|
27611 | 207 |
qed |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
208 |
|
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
209 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
210 |
lemma (in ring_hom_ring) trivial_hom_iff: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
211 |
"(h ` (carrier R) = { \<zero>\<^bsub>S\<^esub> }) = (a_kernel R S h = carrier R)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
212 |
using group_hom.trivial_hom_iff[OF a_group_hom] by (simp add: a_kernel_def) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
213 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
214 |
lemma (in ring_hom_ring) trivial_ker_imp_inj: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
215 |
assumes "a_kernel R S h = { \<zero> }" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
216 |
shows "inj_on h (carrier R)" |
81600 | 217 |
using group_hom.trivial_ker_imp_inj[OF a_group_hom] assms a_kernel_def[of R S h] by simp |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
218 |
|
69122
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
219 |
(* NEW ========================================================================== *) |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
220 |
lemma (in ring_hom_ring) inj_iff_trivial_ker: |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
221 |
shows "inj_on h (carrier R) \<longleftrightarrow> a_kernel R S h = { \<zero> }" |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
222 |
using group_hom.inj_iff_trivial_ker[OF a_group_hom] a_kernel_def[of R S h] by simp |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
223 |
|
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
224 |
(* NEW ========================================================================== *) |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
225 |
corollary ring_hom_in_hom: |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
226 |
assumes "h \<in> ring_hom R S" shows "h \<in> hom R S" and "h \<in> hom (add_monoid R) (add_monoid S)" |
81600 | 227 |
using assms unfolding ring_hom_def hom_def by auto |
69122
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
228 |
|
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
229 |
(* NEW ========================================================================== *) |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
230 |
corollary set_add_hom: |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
231 |
assumes "h \<in> ring_hom R S" "I \<subseteq> carrier R" and "J \<subseteq> carrier R" |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
232 |
shows "h ` (I <+>\<^bsub>R\<^esub> J) = h ` I <+>\<^bsub>S\<^esub> h ` J" |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
233 |
using set_mult_hom[OF ring_hom_in_hom(2)[OF assms(1)]] assms(2-3) |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
234 |
unfolding a_kernel_def[of R S h] set_add_def by simp |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
235 |
|
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
236 |
(* NEW ========================================================================== *) |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
237 |
corollary a_coset_hom: |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
238 |
assumes "h \<in> ring_hom R S" "I \<subseteq> carrier R" "a \<in> carrier R" |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
239 |
shows "h ` (a <+\<^bsub>R\<^esub> I) = h a <+\<^bsub>S\<^esub> (h ` I)" and "h ` (I +>\<^bsub>R\<^esub> a) = (h ` I) +>\<^bsub>S\<^esub> h a" |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
240 |
using assms coset_hom[OF ring_hom_in_hom(2)[OF assms(1)], of I a] |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
241 |
unfolding a_l_coset_def l_coset_eq_set_mult |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
242 |
a_r_coset_def r_coset_eq_set_mult |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
243 |
by simp_all |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
244 |
|
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
245 |
(* NEW ========================================================================== *) |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
246 |
corollary (in ring_hom_ring) set_add_ker_hom: |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
247 |
assumes "I \<subseteq> carrier R" |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
248 |
shows "h ` (I <+> (a_kernel R S h)) = h ` I" and "h ` ((a_kernel R S h) <+> I) = h ` I" |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
249 |
using group_hom.set_mult_ker_hom[OF a_group_hom] assms |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
250 |
unfolding a_kernel_def[of R S h] set_add_def by simp+ |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
251 |
|
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
252 |
lemma (in ring_hom_ring) non_trivial_field_hom_imp_inj: |
81600 | 253 |
assumes R: "field R" |
254 |
and h: "h ` (carrier R) \<noteq> { \<zero>\<^bsub>S\<^esub> }" |
|
255 |
shows "inj_on h (carrier R)" |
|
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
256 |
proof - |
81600 | 257 |
from h have "a_kernel R S h \<noteq> carrier R" |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
258 |
using trivial_hom_iff by linarith |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
259 |
hence "a_kernel R S h = { \<zero> }" |
81600 | 260 |
using field.all_ideals[OF R] kernel_is_ideal by blast |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
261 |
thus "inj_on h (carrier R)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
262 |
using trivial_ker_imp_inj by blast |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
263 |
qed |
81600 | 264 |
|
69122
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
265 |
lemma "field R \<Longrightarrow> cring R" |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
266 |
using fieldE(1) by blast |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
267 |
|
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
268 |
lemma non_trivial_field_hom_is_inj: |
81600 | 269 |
assumes "h \<in> ring_hom R S" and "field R" and "field S" |
270 |
shows "inj_on h (carrier R)" |
|
69122
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
271 |
proof - |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
272 |
interpret ring_hom_cring R S h |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
273 |
using assms(1) ring_hom_cring.intro[OF assms(2-3)[THEN fieldE(1)]] |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
274 |
unfolding symmetric[OF ring_hom_cring_axioms_def] by simp |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
275 |
|
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
276 |
have "h \<one>\<^bsub>R\<^esub> = \<one>\<^bsub>S\<^esub>" and "\<one>\<^bsub>S\<^esub> \<noteq> \<zero>\<^bsub>S\<^esub>" |
81600 | 277 |
using domain.one_not_zero[OF field.axioms(1)[OF assms(3)]] by auto |
69122
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
278 |
hence "h ` (carrier R) \<noteq> { \<zero>\<^bsub>S\<^esub> }" |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
279 |
using ring.kernel_zero ring.trivial_hom_iff by fastforce |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
280 |
thus ?thesis |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
281 |
using ring.non_trivial_field_hom_imp_inj[OF assms(2)] by simp |
81600 | 282 |
qed |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
283 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
284 |
lemma (in ring_hom_ring) img_is_add_subgroup: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
285 |
assumes "subgroup H (add_monoid R)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
286 |
shows "subgroup (h ` H) (add_monoid S)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
287 |
proof - |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
288 |
have "group ((add_monoid R) \<lparr> carrier := H \<rparr>)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
289 |
using assms R.add.subgroup_imp_group by blast |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
290 |
moreover have "H \<subseteq> carrier R" by (simp add: R.add.subgroupE(1) assms) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
291 |
hence "h \<in> hom ((add_monoid R) \<lparr> carrier := H \<rparr>) (add_monoid S)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
292 |
unfolding hom_def by (auto simp add: subsetD) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
293 |
ultimately have "subgroup (h ` carrier ((add_monoid R) \<lparr> carrier := H \<rparr>)) (add_monoid S)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
294 |
using group_hom.img_is_subgroup[of "(add_monoid R) \<lparr> carrier := H \<rparr>" "add_monoid S" h] |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
295 |
using a_group_hom group_hom_axioms.intro group_hom_def by blast |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
296 |
thus "subgroup (h ` H) (add_monoid S)" by simp |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
297 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
298 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
299 |
lemma (in ring) ring_ideal_imp_quot_ideal: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
300 |
assumes "ideal I R" |
81600 | 301 |
and A: "ideal J R" |
302 |
shows "ideal ((+>) I ` J) (R Quot I)" |
|
303 |
proof (rule idealI) |
|
304 |
show "ring (R Quot I)" |
|
305 |
by (simp add: assms(1) ideal.quotient_is_ring) |
|
306 |
next |
|
307 |
have "subgroup J (add_monoid R)" |
|
308 |
by (simp add: additive_subgroup.a_subgroup A ideal.axioms(1)) |
|
309 |
moreover have "((+>) I) \<in> ring_hom R (R Quot I)" |
|
310 |
by (simp add: assms(1) ideal.rcos_ring_hom) |
|
311 |
ultimately show "subgroup ((+>) I ` J) (add_monoid (R Quot I))" |
|
312 |
using assms(1) ideal.rcos_ring_hom_ring ring_hom_ring.img_is_add_subgroup by blast |
|
313 |
next |
|
314 |
fix a x assume "a \<in> (+>) I ` J" "x \<in> carrier (R Quot I)" |
|
315 |
then obtain i j where i: "i \<in> carrier R" "x = I +> i" |
|
316 |
and j: "j \<in> J" "a = I +> j" |
|
317 |
unfolding FactRing_def using A_RCOSETS_def'[of R I] by auto |
|
318 |
hence "a \<otimes>\<^bsub>R Quot I\<^esub> x = [mod I:] (I +> j) \<Otimes> (I +> i)" |
|
319 |
unfolding FactRing_def by simp |
|
320 |
hence "a \<otimes>\<^bsub>R Quot I\<^esub> x = I +> (j \<otimes> i)" |
|
321 |
using ideal.rcoset_mult_add[OF assms(1), of j i] i(1) j(1) A ideal.Icarr by force |
|
322 |
thus "a \<otimes>\<^bsub>R Quot I\<^esub> x \<in> (+>) I ` J" |
|
323 |
using A i(1) j(1) by (simp add: ideal.I_r_closed) |
|
324 |
||
325 |
have "x \<otimes>\<^bsub>R Quot I\<^esub> a = [mod I:] (I +> i) \<Otimes> (I +> j)" |
|
326 |
unfolding FactRing_def i j by simp |
|
327 |
hence "x \<otimes>\<^bsub>R Quot I\<^esub> a = I +> (i \<otimes> j)" |
|
328 |
using ideal.rcoset_mult_add[OF assms(1), of i j] i(1) j(1) A ideal.Icarr by force |
|
329 |
thus "x \<otimes>\<^bsub>R Quot I\<^esub> a \<in> (+>) I ` J" |
|
330 |
using A i(1) j(1) by (simp add: ideal.I_l_closed) |
|
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
331 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
332 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
333 |
lemma (in ring_hom_ring) ideal_vimage: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
334 |
assumes "ideal I S" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
335 |
shows "ideal { r \<in> carrier R. h r \<in> I } R" (* or (carrier R) \<inter> (h -` I) *) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
336 |
proof |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
337 |
show "{ r \<in> carrier R. h r \<in> I } \<subseteq> carrier (add_monoid R)" by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
338 |
show "\<one>\<^bsub>add_monoid R\<^esub> \<in> { r \<in> carrier R. h r \<in> I }" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
339 |
by (simp add: additive_subgroup.zero_closed assms ideal.axioms(1)) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
340 |
next |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
341 |
fix a b |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
342 |
assume "a \<in> { r \<in> carrier R. h r \<in> I }" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
343 |
and "b \<in> { r \<in> carrier R. h r \<in> I }" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
344 |
hence a: "a \<in> carrier R" "h a \<in> I" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
345 |
and b: "b \<in> carrier R" "h b \<in> I" by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
346 |
hence "h (a \<oplus> b) = (h a) \<oplus>\<^bsub>S\<^esub> (h b)" using hom_add by blast |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
347 |
moreover have "(h a) \<oplus>\<^bsub>S\<^esub> (h b) \<in> I" using a b assms |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
348 |
by (simp add: additive_subgroup.a_closed ideal.axioms(1)) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
349 |
ultimately show "a \<otimes>\<^bsub>add_monoid R\<^esub> b \<in> { r \<in> carrier R. h r \<in> I }" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
350 |
using a(1) b (1) by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
351 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
352 |
have "h (\<ominus> a) = \<ominus>\<^bsub>S\<^esub> (h a)" by (simp add: a) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
353 |
moreover have "\<ominus>\<^bsub>S\<^esub> (h a) \<in> I" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
354 |
by (simp add: a(2) additive_subgroup.a_inv_closed assms ideal.axioms(1)) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
355 |
ultimately show "inv\<^bsub>add_monoid R\<^esub> a \<in> { r \<in> carrier R. h r \<in> I }" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
356 |
using a by (simp add: a_inv_def) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
357 |
next |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
358 |
fix a r |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
359 |
assume "a \<in> { r \<in> carrier R. h r \<in> I }" and r: "r \<in> carrier R" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
360 |
hence a: "a \<in> carrier R" "h a \<in> I" by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
361 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
362 |
have "h a \<otimes>\<^bsub>S\<^esub> h r \<in> I" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
363 |
using assms a r by (simp add: ideal.I_r_closed) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
364 |
thus "a \<otimes> r \<in> { r \<in> carrier R. h r \<in> I }" by (simp add: a(1) r) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
365 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
366 |
have "h r \<otimes>\<^bsub>S\<^esub> h a \<in> I" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
367 |
using assms a r by (simp add: ideal.I_l_closed) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
368 |
thus "r \<otimes> a \<in> { r \<in> carrier R. h r \<in> I }" by (simp add: a(1) r) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
369 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
370 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
371 |
lemma (in ring) canonical_proj_vimage_in_carrier: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
372 |
assumes "ideal I R" |
81600 | 373 |
and A: "J \<subseteq> carrier (R Quot I)" |
374 |
shows "\<Union> J \<subseteq> carrier R" |
|
375 |
proof |
|
376 |
fix j assume j: "j \<in> \<Union> J" |
|
377 |
then obtain j' where j': "j' \<in> J" "j \<in> j'" |
|
378 |
by blast |
|
379 |
then obtain r where r: "r \<in> carrier R" "j' = I +> r" |
|
380 |
using A j' unfolding FactRing_def using A_RCOSETS_def'[of R I] by auto |
|
381 |
thus "j \<in> carrier R" |
|
382 |
using j' assms |
|
383 |
by (meson a_r_coset_subset_G additive_subgroup.a_subset contra_subsetD ideal.axioms(1)) |
|
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
384 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
385 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
386 |
lemma (in ring) canonical_proj_vimage_mem_iff: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
387 |
assumes "ideal I R" "J \<subseteq> carrier (R Quot I)" |
81600 | 388 |
and a: "a \<in> carrier R" |
389 |
shows "(a \<in> \<Union> J) = (I +> a \<in> J)" |
|
390 |
proof |
|
391 |
assume "a \<in> \<Union> J" |
|
392 |
then obtain j where j: "j \<in> J" "a \<in> j" by blast |
|
393 |
then obtain r where r: "r \<in> carrier R" "j = I +> r" |
|
394 |
using assms j unfolding FactRing_def using A_RCOSETS_def'[of R I] by auto |
|
395 |
hence "I +> r = I +> a" |
|
396 |
using add.repr_independence[of a I r] j r |
|
397 |
by (metis a_r_coset_def additive_subgroup.a_subgroup assms(1) ideal.axioms(1)) |
|
398 |
thus "I +> a \<in> J" using r j by simp |
|
399 |
next |
|
400 |
assume "I +> a \<in> J" |
|
401 |
hence "\<zero> \<oplus> a \<in> I +> a" |
|
402 |
using additive_subgroup.zero_closed[OF ideal.axioms(1)[OF assms(1)]] |
|
403 |
a_r_coset_def'[of R I a] by blast |
|
404 |
thus "a \<in> \<Union> J" using a \<open>I +> a \<in> J\<close> by auto |
|
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
405 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
406 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
407 |
corollary (in ring) quot_ideal_imp_ring_ideal: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
408 |
assumes "ideal I R" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
409 |
shows "ideal J (R Quot I) \<Longrightarrow> ideal (\<Union> J) R" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
410 |
proof - |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
411 |
assume A: "ideal J (R Quot I)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
412 |
have "\<Union> J = { r \<in> carrier R. I +> r \<in> J }" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
413 |
using canonical_proj_vimage_in_carrier[OF assms, of J] |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
414 |
canonical_proj_vimage_mem_iff[OF assms, of J] |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
415 |
additive_subgroup.a_subset[OF ideal.axioms(1)[OF A]] by blast |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
416 |
thus "ideal (\<Union> J) R" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
417 |
using ring_hom_ring.ideal_vimage[OF ideal.rcos_ring_hom_ring[OF assms] A] by simp |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
418 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
419 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
420 |
lemma (in ring) ideal_incl_iff: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
421 |
assumes "ideal I R" "ideal J R" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
422 |
shows "(I \<subseteq> J) = (J = (\<Union> j \<in> J. I +> j))" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
423 |
proof |
81600 | 424 |
assume "J = (\<Union> j \<in> J. I +> j)" hence "I +> \<zero> \<subseteq> J" |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
425 |
using additive_subgroup.zero_closed[OF ideal.axioms(1)[OF assms(2)]] by blast |
81600 | 426 |
thus "I \<subseteq> J" using additive_subgroup.a_subset[OF ideal.axioms(1)[OF assms(1)]] by simp |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
427 |
next |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
428 |
assume A: "I \<subseteq> J" show "J = (\<Union>j\<in>J. I +> j)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
429 |
proof |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
430 |
show "J \<subseteq> (\<Union> j \<in> J. I +> j)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
431 |
proof |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
432 |
fix j assume j: "j \<in> J" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
433 |
have "\<zero> \<in> I" by (simp add: additive_subgroup.zero_closed assms(1) ideal.axioms(1)) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
434 |
hence "\<zero> \<oplus> j \<in> I +> j" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
435 |
using a_r_coset_def'[of R I j] by blast |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
436 |
thus "j \<in> (\<Union>j\<in>J. I +> j)" |
81600 | 437 |
using assms(2) j additive_subgroup.a_Hcarr ideal.axioms(1) by fastforce |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
438 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
439 |
show "(\<Union> j \<in> J. I +> j) \<subseteq> J" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
440 |
proof |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
441 |
fix x assume "x \<in> (\<Union> j \<in> J. I +> j)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
442 |
then obtain j where j: "j \<in> J" "x \<in> I +> j" by blast |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
443 |
then obtain i where i: "i \<in> I" "x = i \<oplus> j" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
444 |
using a_r_coset_def'[of R I j] by blast |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
445 |
thus "x \<in> J" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
446 |
using assms(2) j A additive_subgroup.a_closed[OF ideal.axioms(1)[OF assms(2)]] by blast |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
447 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
448 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
449 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
450 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
451 |
theorem (in ring) quot_ideal_correspondence: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
452 |
assumes "ideal I R" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
453 |
shows "bij_betw (\<lambda>J. (+>) I ` J) { J. ideal J R \<and> I \<subseteq> J } { J . ideal J (R Quot I) }" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
454 |
proof (rule bij_betw_byWitness[where ?f' = "\<lambda>X. \<Union> X"]) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
455 |
show "\<forall>J \<in> { J. ideal J R \<and> I \<subseteq> J }. (\<lambda>X. \<Union> X) ((+>) I ` J) = J" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
456 |
using assms ideal_incl_iff by blast |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
457 |
show "(\<lambda>J. (+>) I ` J) ` { J. ideal J R \<and> I \<subseteq> J } \<subseteq> { J. ideal J (R Quot I) }" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
458 |
using assms ring_ideal_imp_quot_ideal by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
459 |
show "(\<lambda>X. \<Union> X) ` { J. ideal J (R Quot I) } \<subseteq> { J. ideal J R \<and> I \<subseteq> J }" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
460 |
proof |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
461 |
fix J assume "J \<in> ((\<lambda>X. \<Union> X) ` { J. ideal J (R Quot I) })" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
462 |
then obtain J' where J': "ideal J' (R Quot I)" "J = \<Union> J'" by blast |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
463 |
hence "ideal J R" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
464 |
using assms quot_ideal_imp_ring_ideal by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
465 |
moreover have "I \<in> J'" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
466 |
using additive_subgroup.zero_closed[OF ideal.axioms(1)[OF J'(1)]] unfolding FactRing_def by simp |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
467 |
ultimately show "J \<in> { J. ideal J R \<and> I \<subseteq> J }" using J'(2) by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
468 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
469 |
show "\<forall>J' \<in> { J. ideal J (R Quot I) }. ((+>) I ` (\<Union> J')) = J'" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
470 |
proof |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
471 |
fix J' assume "J' \<in> { J. ideal J (R Quot I) }" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
472 |
hence subset: "J' \<subseteq> carrier (R Quot I) \<and> ideal J' (R Quot I)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
473 |
using additive_subgroup.a_subset ideal_def by blast |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
474 |
hence "((+>) I ` (\<Union> J')) \<subseteq> J'" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
475 |
using canonical_proj_vimage_in_carrier canonical_proj_vimage_mem_iff |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
476 |
by (meson assms contra_subsetD image_subsetI) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
477 |
moreover have "J' \<subseteq> ((+>) I ` (\<Union> J'))" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
478 |
proof |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
479 |
fix x assume "x \<in> J'" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
480 |
then obtain r where r: "r \<in> carrier R" "x = I +> r" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
481 |
using subset unfolding FactRing_def A_RCOSETS_def'[of R I] by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
482 |
hence "r \<in> (\<Union> J')" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
483 |
using \<open>x \<in> J'\<close> assms canonical_proj_vimage_mem_iff subset by blast |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
484 |
thus "x \<in> ((+>) I ` (\<Union> J'))" using r(2) by blast |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
485 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
486 |
ultimately show "((+>) I ` (\<Union> J')) = J'" by blast |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
487 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
488 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
489 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
490 |
lemma (in cring) quot_domain_imp_primeideal: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
491 |
assumes "ideal P R" |
81600 | 492 |
and A: "domain (R Quot P)" |
493 |
shows "primeideal P R" |
|
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
494 |
proof - |
81600 | 495 |
show "primeideal P R" |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
496 |
proof (rule primeidealI) |
81600 | 497 |
show "ideal P R" using assms(1) . |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
498 |
show "cring R" using is_cring . |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
499 |
next |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
500 |
show "carrier R \<noteq> P" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
501 |
proof (rule ccontr) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
502 |
assume "\<not> carrier R \<noteq> P" hence "carrier R = P" by simp |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
503 |
hence "\<And>I. I \<in> carrier (R Quot P) \<Longrightarrow> I = P" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
504 |
unfolding FactRing_def A_RCOSETS_def' apply simp |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
505 |
using a_coset_join2 additive_subgroup.a_subgroup assms ideal.axioms(1) by blast |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
506 |
hence "\<one>\<^bsub>(R Quot P)\<^esub> = \<zero>\<^bsub>(R Quot P)\<^esub>" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
507 |
by (metis assms ideal.quotient_is_ring ring.ring_simprules(2) ring.ring_simprules(6)) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
508 |
thus False using domain.one_not_zero[OF A] by simp |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
509 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
510 |
next |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
511 |
fix a b assume a: "a \<in> carrier R" and b: "b \<in> carrier R" and ab: "a \<otimes> b \<in> P" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
512 |
hence "P +> (a \<otimes> b) = \<zero>\<^bsub>(R Quot P)\<^esub>" unfolding FactRing_def |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
513 |
by (simp add: a_coset_join2 additive_subgroup.a_subgroup assms ideal.axioms(1)) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
514 |
moreover have "(P +> a) \<otimes>\<^bsub>(R Quot P)\<^esub> (P +> b) = P +> (a \<otimes> b)" unfolding FactRing_def |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
515 |
using a b by (simp add: assms ideal.rcoset_mult_add) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
516 |
moreover have "P +> a \<in> carrier (R Quot P) \<and> P +> b \<in> carrier (R Quot P)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
517 |
by (simp add: a b FactRing_def a_rcosetsI additive_subgroup.a_subset assms ideal.axioms(1)) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
518 |
ultimately have "P +> a = \<zero>\<^bsub>(R Quot P)\<^esub> \<or> P +> b = \<zero>\<^bsub>(R Quot P)\<^esub>" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
519 |
using domain.integral[OF A, of "P +> a" "P +> b"] by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
520 |
thus "a \<in> P \<or> b \<in> P" unfolding FactRing_def apply simp |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
521 |
using a b assms a_coset_join1 additive_subgroup.a_subgroup ideal.axioms(1) by blast |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
522 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
523 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
524 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
525 |
lemma (in cring) quot_domain_iff_primeideal: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
526 |
assumes "ideal P R" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
527 |
shows "domain (R Quot P) = primeideal P R" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
528 |
using quot_domain_imp_primeideal[OF assms] primeideal.quotient_is_domain[of P R] by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
529 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
530 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
531 |
subsection \<open>Isomorphism\<close> |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
532 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
533 |
definition |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
534 |
ring_iso :: "_ \<Rightarrow> _ \<Rightarrow> ('a \<Rightarrow> 'b) set" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
535 |
where "ring_iso R S = { h. h \<in> ring_hom R S \<and> bij_betw h (carrier R) (carrier S) }" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
536 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
537 |
definition |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
77138
diff
changeset
|
538 |
is_ring_iso :: "_ \<Rightarrow> _ \<Rightarrow> bool" (infixr \<open>\<simeq>\<close> 60) |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
539 |
where "R \<simeq> S = (ring_iso R S \<noteq> {})" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
540 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
541 |
definition |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
542 |
morphic_prop :: "_ \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
543 |
where "morphic_prop R P = |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
544 |
((P \<one>\<^bsub>R\<^esub>) \<and> |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
545 |
(\<forall>r \<in> carrier R. P r) \<and> |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
546 |
(\<forall>r1 \<in> carrier R. \<forall>r2 \<in> carrier R. P (r1 \<otimes>\<^bsub>R\<^esub> r2)) \<and> |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
547 |
(\<forall>r1 \<in> carrier R. \<forall>r2 \<in> carrier R. P (r1 \<oplus>\<^bsub>R\<^esub> r2)))" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
548 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
549 |
lemma ring_iso_memI: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
550 |
fixes R (structure) and S (structure) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
551 |
assumes "\<And>x. x \<in> carrier R \<Longrightarrow> h x \<in> carrier S" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
552 |
and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> h (x \<otimes> y) = h x \<otimes>\<^bsub>S\<^esub> h y" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
553 |
and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> h (x \<oplus> y) = h x \<oplus>\<^bsub>S\<^esub> h y" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
554 |
and "h \<one> = \<one>\<^bsub>S\<^esub>" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
555 |
and "bij_betw h (carrier R) (carrier S)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
556 |
shows "h \<in> ring_iso R S" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
557 |
by (auto simp add: ring_hom_memI assms ring_iso_def) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
558 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
559 |
lemma ring_iso_memE: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
560 |
fixes R (structure) and S (structure) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
561 |
assumes "h \<in> ring_iso R S" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
562 |
shows "\<And>x. x \<in> carrier R \<Longrightarrow> h x \<in> carrier S" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
563 |
and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> h (x \<otimes> y) = h x \<otimes>\<^bsub>S\<^esub> h y" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
564 |
and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> h (x \<oplus> y) = h x \<oplus>\<^bsub>S\<^esub> h y" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
565 |
and "h \<one> = \<one>\<^bsub>S\<^esub>" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
566 |
and "bij_betw h (carrier R) (carrier S)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
567 |
using assms unfolding ring_iso_def ring_hom_def by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
568 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
569 |
lemma morphic_propI: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
570 |
fixes R (structure) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
571 |
assumes "P \<one>" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
572 |
and "\<And>r. r \<in> carrier R \<Longrightarrow> P r" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
573 |
and "\<And>r1 r2. \<lbrakk> r1 \<in> carrier R; r2 \<in> carrier R \<rbrakk> \<Longrightarrow> P (r1 \<otimes> r2)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
574 |
and "\<And>r1 r2. \<lbrakk> r1 \<in> carrier R; r2 \<in> carrier R \<rbrakk> \<Longrightarrow> P (r1 \<oplus> r2)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
575 |
shows "morphic_prop R P" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
576 |
unfolding morphic_prop_def using assms by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
577 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
578 |
lemma morphic_propE: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
579 |
fixes R (structure) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
580 |
assumes "morphic_prop R P" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
581 |
shows "P \<one>" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
582 |
and "\<And>r. r \<in> carrier R \<Longrightarrow> P r" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
583 |
and "\<And>r1 r2. \<lbrakk> r1 \<in> carrier R; r2 \<in> carrier R \<rbrakk> \<Longrightarrow> P (r1 \<otimes> r2)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
584 |
and "\<And>r1 r2. \<lbrakk> r1 \<in> carrier R; r2 \<in> carrier R \<rbrakk> \<Longrightarrow> P (r1 \<oplus> r2)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
585 |
using assms unfolding morphic_prop_def by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
586 |
|
69122
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
587 |
(* NEW ============================================================================ *) |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
588 |
lemma (in ring) ring_hom_restrict: |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
589 |
assumes "f \<in> ring_hom R S" and "\<And>r. r \<in> carrier R \<Longrightarrow> f r = g r" shows "g \<in> ring_hom R S" |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
590 |
using assms(2) ring_hom_memE[OF assms(1)] by (auto intro: ring_hom_memI) |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
591 |
|
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
592 |
(* PROOF ========================================================================== *) |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
593 |
lemma (in ring) ring_iso_restrict: |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
594 |
assumes "f \<in> ring_iso R S" and "\<And>r. r \<in> carrier R \<Longrightarrow> f r = g r" shows "g \<in> ring_iso R S" |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
595 |
proof - |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
596 |
have hom: "g \<in> ring_hom R S" |
81600 | 597 |
using ring_hom_restrict assms unfolding ring_iso_def by auto |
69122
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
598 |
have "bij_betw g (carrier R) (carrier S)" |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
599 |
using bij_betw_cong[of "carrier R" f g] ring_iso_memE(5)[OF assms(1)] assms(2) by simp |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
600 |
thus ?thesis |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
601 |
using ring_hom_memE[OF hom] by (auto intro!: ring_iso_memI) |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
602 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
603 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
604 |
lemma ring_iso_morphic_prop: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
605 |
assumes "f \<in> ring_iso R S" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
606 |
and "morphic_prop R P" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
607 |
and "\<And>r. P r \<Longrightarrow> f r = g r" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
608 |
shows "g \<in> ring_iso R S" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
609 |
proof - |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
610 |
have eq0: "\<And>r. r \<in> carrier R \<Longrightarrow> f r = g r" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
611 |
and eq1: "f \<one>\<^bsub>R\<^esub> = g \<one>\<^bsub>R\<^esub>" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
612 |
and eq2: "\<And>r1 r2. \<lbrakk> r1 \<in> carrier R; r2 \<in> carrier R \<rbrakk> \<Longrightarrow> f (r1 \<otimes>\<^bsub>R\<^esub> r2) = g (r1 \<otimes>\<^bsub>R\<^esub> r2)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
613 |
and eq3: "\<And>r1 r2. \<lbrakk> r1 \<in> carrier R; r2 \<in> carrier R \<rbrakk> \<Longrightarrow> f (r1 \<oplus>\<^bsub>R\<^esub> r2) = g (r1 \<oplus>\<^bsub>R\<^esub> r2)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
614 |
using assms(2-3) unfolding morphic_prop_def by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
615 |
show ?thesis |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
616 |
apply (rule ring_iso_memI) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
617 |
using assms(1) eq0 ring_iso_memE(1) apply fastforce |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
618 |
apply (metis assms(1) eq0 eq2 ring_iso_memE(2)) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
619 |
apply (metis assms(1) eq0 eq3 ring_iso_memE(3)) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
620 |
using assms(1) eq1 ring_iso_memE(4) apply fastforce |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
621 |
using assms(1) bij_betw_cong eq0 ring_iso_memE(5) by blast |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
622 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
623 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
624 |
lemma (in ring) ring_hom_imp_img_ring: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
625 |
assumes "h \<in> ring_hom R S" |
69122
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
626 |
shows "ring (S \<lparr> carrier := h ` (carrier R), zero := h \<zero> \<rparr>)" (is "ring ?h_img") |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
627 |
proof - |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
628 |
have "h \<in> hom (add_monoid R) (add_monoid S)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
629 |
using assms unfolding hom_def ring_hom_def by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
630 |
hence "comm_group ((add_monoid S) \<lparr> carrier := h ` (carrier R), one := h \<zero> \<rparr>)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
631 |
using add.hom_imp_img_comm_group[of h "add_monoid S"] by simp |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
632 |
hence comm_group: "comm_group (add_monoid ?h_img)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
633 |
by (auto intro: comm_monoidI simp add: monoid.defs) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
634 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
635 |
moreover have "h \<in> hom R S" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
636 |
using assms unfolding ring_hom_def hom_def by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
637 |
hence "monoid (S \<lparr> carrier := h ` (carrier R), one := h \<one> \<rparr>)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
638 |
using hom_imp_img_monoid[of h S] by simp |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
639 |
hence monoid: "monoid ?h_img" |
69122
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
640 |
using ring_hom_memE(4)[OF assms] unfolding monoid_def by (simp add: monoid.defs) |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
641 |
show ?thesis |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
642 |
proof (rule ringI, simp_all add: comm_group_abelian_groupI[OF comm_group] monoid) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
643 |
fix x y z assume "x \<in> h ` carrier R" "y \<in> h ` carrier R" "z \<in> h ` carrier R" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
644 |
then obtain r1 r2 r3 |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
645 |
where r1: "r1 \<in> carrier R" "x = h r1" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
646 |
and r2: "r2 \<in> carrier R" "y = h r2" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
647 |
and r3: "r3 \<in> carrier R" "z = h r3" by blast |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
648 |
hence "(x \<oplus>\<^bsub>S\<^esub> y) \<otimes>\<^bsub>S\<^esub> z = h ((r1 \<oplus> r2) \<otimes> r3)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
649 |
using ring_hom_memE[OF assms] by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
650 |
also have " ... = h ((r1 \<otimes> r3) \<oplus> (r2 \<otimes> r3))" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
651 |
using l_distr[OF r1(1) r2(1) r3(1)] by simp |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
652 |
also have " ... = (x \<otimes>\<^bsub>S\<^esub> z) \<oplus>\<^bsub>S\<^esub> (y \<otimes>\<^bsub>S\<^esub> z)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
653 |
using ring_hom_memE[OF assms] r1 r2 r3 by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
654 |
finally show "(x \<oplus>\<^bsub>S\<^esub> y) \<otimes>\<^bsub>S\<^esub> z = (x \<otimes>\<^bsub>S\<^esub> z) \<oplus>\<^bsub>S\<^esub> (y \<otimes>\<^bsub>S\<^esub> z)" . |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
655 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
656 |
have "z \<otimes>\<^bsub>S\<^esub> (x \<oplus>\<^bsub>S\<^esub> y) = h (r3 \<otimes> (r1 \<oplus> r2))" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
657 |
using ring_hom_memE[OF assms] r1 r2 r3 by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
658 |
also have " ... = h ((r3 \<otimes> r1) \<oplus> (r3 \<otimes> r2))" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
659 |
using r_distr[OF r1(1) r2(1) r3(1)] by simp |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
660 |
also have " ... = (z \<otimes>\<^bsub>S\<^esub> x) \<oplus>\<^bsub>S\<^esub> (z \<otimes>\<^bsub>S\<^esub> y)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
661 |
using ring_hom_memE[OF assms] r1 r2 r3 by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
662 |
finally show "z \<otimes>\<^bsub>S\<^esub> (x \<oplus>\<^bsub>S\<^esub> y) = (z \<otimes>\<^bsub>S\<^esub> x) \<oplus>\<^bsub>S\<^esub> (z \<otimes>\<^bsub>S\<^esub> y)" . |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
663 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
664 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
665 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
666 |
lemma (in ring) ring_iso_imp_img_ring: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
667 |
assumes "h \<in> ring_iso R S" |
69122
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
668 |
shows "ring (S \<lparr> zero := h \<zero> \<rparr>)" |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
669 |
proof - |
69122
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
670 |
have "ring (S \<lparr> carrier := h ` (carrier R), zero := h \<zero> \<rparr>)" |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
671 |
using ring_hom_imp_img_ring[of h S] assms unfolding ring_iso_def by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
672 |
moreover have "h ` (carrier R) = carrier S" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
673 |
using assms unfolding ring_iso_def bij_betw_def by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
674 |
ultimately show ?thesis by simp |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
675 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
676 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
677 |
lemma (in cring) ring_iso_imp_img_cring: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
678 |
assumes "h \<in> ring_iso R S" |
69122
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
679 |
shows "cring (S \<lparr> zero := h \<zero> \<rparr>)" (is "cring ?h_img") |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
680 |
proof - |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
681 |
note m_comm |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
682 |
interpret h_img?: ring ?h_img |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
683 |
using ring_iso_imp_img_ring[OF assms] . |
81600 | 684 |
show ?thesis |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
685 |
proof (unfold_locales) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
686 |
fix x y assume "x \<in> carrier ?h_img" "y \<in> carrier ?h_img" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
687 |
then obtain r1 r2 |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
688 |
where r1: "r1 \<in> carrier R" "x = h r1" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
689 |
and r2: "r2 \<in> carrier R" "y = h r2" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
690 |
using assms image_iff[where ?f = h and ?A = "carrier R"] |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
691 |
unfolding ring_iso_def bij_betw_def by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
692 |
have "x \<otimes>\<^bsub>(?h_img)\<^esub> y = h (r1 \<otimes> r2)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
693 |
using assms r1 r2 unfolding ring_iso_def ring_hom_def by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
694 |
also have " ... = h (r2 \<otimes> r1)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
695 |
using m_comm[OF r1(1) r2(1)] by simp |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
696 |
also have " ... = y \<otimes>\<^bsub>(?h_img)\<^esub> x" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
697 |
using assms r1 r2 unfolding ring_iso_def ring_hom_def by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
698 |
finally show "x \<otimes>\<^bsub>(?h_img)\<^esub> y = y \<otimes>\<^bsub>(?h_img)\<^esub> x" . |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
699 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
700 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
701 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
702 |
lemma (in domain) ring_iso_imp_img_domain: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
703 |
assumes "h \<in> ring_iso R S" |
69122
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
704 |
shows "domain (S \<lparr> zero := h \<zero> \<rparr>)" (is "domain ?h_img") |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
705 |
proof - |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
706 |
note aux = m_closed integral one_not_zero one_closed zero_closed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
707 |
interpret h_img?: cring ?h_img |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
708 |
using ring_iso_imp_img_cring[OF assms] . |
81600 | 709 |
show ?thesis |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
710 |
proof (unfold_locales) |
69122
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
711 |
have "\<one>\<^bsub>?h_img\<^esub> = \<zero>\<^bsub>?h_img\<^esub> \<Longrightarrow> h \<one> = h \<zero>" |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
712 |
using ring_iso_memE(4)[OF assms] by simp |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
713 |
moreover have "h \<one> \<noteq> h \<zero>" |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
714 |
using ring_iso_memE(5)[OF assms] aux(3-4) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
715 |
unfolding bij_betw_def inj_on_def by force |
69122
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
716 |
ultimately show "\<one>\<^bsub>?h_img\<^esub> \<noteq> \<zero>\<^bsub>?h_img\<^esub>" |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
717 |
by auto |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
718 |
next |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
719 |
fix a b |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
720 |
assume A: "a \<otimes>\<^bsub>?h_img\<^esub> b = \<zero>\<^bsub>?h_img\<^esub>" "a \<in> carrier ?h_img" "b \<in> carrier ?h_img" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
721 |
then obtain r1 r2 |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
722 |
where r1: "r1 \<in> carrier R" "a = h r1" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
723 |
and r2: "r2 \<in> carrier R" "b = h r2" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
724 |
using assms image_iff[where ?f = h and ?A = "carrier R"] |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
725 |
unfolding ring_iso_def bij_betw_def by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
726 |
hence "a \<otimes>\<^bsub>?h_img\<^esub> b = h (r1 \<otimes> r2)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
727 |
using assms r1 r2 unfolding ring_iso_def ring_hom_def by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
728 |
hence "h (r1 \<otimes> r2) = h \<zero>" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
729 |
using A(1) by simp |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
730 |
hence "r1 \<otimes> r2 = \<zero>" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
731 |
using ring_iso_memE(5)[OF assms] aux(1)[OF r1(1) r2(1)] aux(5) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
732 |
unfolding bij_betw_def inj_on_def by force |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
733 |
hence "r1 = \<zero> \<or> r2 = \<zero>" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
734 |
using aux(2)[OF _ r1(1) r2(1)] by simp |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
735 |
thus "a = \<zero>\<^bsub>?h_img\<^esub> \<or> b = \<zero>\<^bsub>?h_img\<^esub>" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
736 |
unfolding r1 r2 by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
737 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
738 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
739 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
740 |
lemma (in field) ring_iso_imp_img_field: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
741 |
assumes "h \<in> ring_iso R S" |
69122
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
742 |
shows "field (S \<lparr> zero := h \<zero> \<rparr>)" (is "field ?h_img") |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
743 |
proof - |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
744 |
interpret h_img?: domain ?h_img |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
745 |
using ring_iso_imp_img_domain[OF assms] . |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
746 |
show ?thesis |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
747 |
proof (unfold_locales, auto simp add: Units_def) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
748 |
interpret field R using field_axioms . |
69122
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
749 |
fix a assume a: "a \<in> carrier S" "a \<otimes>\<^bsub>S\<^esub> h \<zero> = \<one>\<^bsub>S\<^esub>" |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
750 |
then obtain r where r: "r \<in> carrier R" "a = h r" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
751 |
using assms image_iff[where ?f = h and ?A = "carrier R"] |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
752 |
unfolding ring_iso_def bij_betw_def by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
753 |
have "a \<otimes>\<^bsub>S\<^esub> h \<zero> = h (r \<otimes> \<zero>)" unfolding r(2) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
754 |
using ring_iso_memE(2)[OF assms r(1)] by simp |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
755 |
hence "h \<one> = h \<zero>" |
69122
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
756 |
using ring_iso_memE(4)[OF assms] r(1) a(2) by simp |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
757 |
thus False |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
758 |
using ring_iso_memE(5)[OF assms] |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
759 |
unfolding bij_betw_def inj_on_def by force |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
760 |
next |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
761 |
interpret field R using field_axioms . |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
762 |
fix s assume s: "s \<in> carrier S" "s \<noteq> h \<zero>" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
763 |
then obtain r where r: "r \<in> carrier R" "s = h r" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
764 |
using assms image_iff[where ?f = h and ?A = "carrier R"] |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
765 |
unfolding ring_iso_def bij_betw_def by auto |
81600 | 766 |
hence "r \<noteq> \<zero>" using s(2) by auto |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
767 |
hence inv_r: "inv r \<in> carrier R" "inv r \<noteq> \<zero>" "r \<otimes> inv r = \<one>" "inv r \<otimes> r = \<one>" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
768 |
using field_Units r(1) by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
769 |
have "h (inv r) \<otimes>\<^bsub>S\<^esub> h r = h \<one>" and "h r \<otimes>\<^bsub>S\<^esub> h (inv r) = h \<one>" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
770 |
using ring_iso_memE(2)[OF assms inv_r(1) r(1)] inv_r(3-4) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
771 |
ring_iso_memE(2)[OF assms r(1) inv_r(1)] by auto |
69122
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
772 |
thus "\<exists>s' \<in> carrier S. s' \<otimes>\<^bsub>S\<^esub> s = \<one>\<^bsub>S\<^esub> \<and> s \<otimes>\<^bsub>S\<^esub> s' = \<one>\<^bsub>S\<^esub>" |
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
773 |
using ring_iso_memE(1,4)[OF assms] inv_r(1) r(2) by auto |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
774 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
775 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
776 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
777 |
lemma ring_iso_same_card: "R \<simeq> S \<Longrightarrow> card (carrier R) = card (carrier S)" |
81600 | 778 |
using bij_betw_same_card unfolding is_ring_iso_def ring_iso_def by auto |
69122
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
779 |
(* ========================================================================== *) |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
780 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
781 |
lemma ring_iso_set_refl: "id \<in> ring_iso R R" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
782 |
by (rule ring_iso_memI) (auto) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
783 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
784 |
corollary ring_iso_refl: "R \<simeq> R" |
81600 | 785 |
using is_ring_iso_def ring_iso_set_refl by auto |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
786 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
787 |
lemma ring_iso_set_trans: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
788 |
"\<lbrakk> f \<in> ring_iso R S; g \<in> ring_iso S Q \<rbrakk> \<Longrightarrow> (g \<circ> f) \<in> ring_iso R Q" |
81600 | 789 |
unfolding ring_iso_def using bij_betw_trans ring_hom_trans by fastforce |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
790 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
791 |
corollary ring_iso_trans: "\<lbrakk> R \<simeq> S; S \<simeq> Q \<rbrakk> \<Longrightarrow> R \<simeq> Q" |
81600 | 792 |
using ring_iso_set_trans unfolding is_ring_iso_def by blast |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
793 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
794 |
lemma ring_iso_set_sym: |
68604 | 795 |
assumes "ring R" and h: "h \<in> ring_iso R S" |
796 |
shows "(inv_into (carrier R) h) \<in> ring_iso S R" |
|
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
797 |
proof - |
68604 | 798 |
have h_hom: "h \<in> ring_hom R S" |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
799 |
and h_surj: "h ` (carrier R) = (carrier S)" |
81600 | 800 |
and h_inj: "\<And>x1 x2. \<lbrakk> x1 \<in> carrier R; x2 \<in> carrier R \<rbrakk> \<Longrightarrow> h x1 = h x2 \<Longrightarrow> x1 = x2" |
68604 | 801 |
using h unfolding ring_iso_def bij_betw_def inj_on_def by auto |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
802 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
803 |
have h_inv_bij: "bij_betw (inv_into (carrier R) h) (carrier S) (carrier R)" |
77138
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
804 |
by (simp add: bij_betw_inv_into h ring_iso_memE(5)) |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
805 |
|
77138
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
806 |
have "inv_into (carrier R) h \<in> ring_hom S R" |
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
807 |
using ring_iso_memE [OF h] bij_betwE [OF h_inv_bij] \<open>ring R\<close> |
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
808 |
by (simp add: bij_betw_imp_inj_on bij_betw_inv_into_right inv_into_f_eq ring.ring_simprules ring_hom_memI) |
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
809 |
moreover have "bij_betw (inv_into (carrier R) h) (carrier S) (carrier R)" |
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
810 |
using h_inv_bij by force |
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
811 |
ultimately show "inv_into (carrier R) h \<in> ring_iso S R" |
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
812 |
by (simp add: ring_iso_def) |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
813 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
814 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
815 |
corollary ring_iso_sym: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
816 |
assumes "ring R" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
817 |
shows "R \<simeq> S \<Longrightarrow> S \<simeq> R" |
81600 | 818 |
using assms ring_iso_set_sym unfolding is_ring_iso_def by auto |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
819 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
820 |
lemma (in ring_hom_ring) the_elem_simp [simp]: |
81600 | 821 |
assumes x: "x \<in> carrier R" |
822 |
shows "the_elem (h ` ((a_kernel R S h) +> x)) = h x" |
|
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
823 |
proof - |
81600 | 824 |
from x have "h x \<in> h ` ((a_kernel R S h) +> x)" |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
825 |
using homeq_imp_rcos by blast |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
826 |
thus "the_elem (h ` ((a_kernel R S h) +> x)) = h x" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
827 |
by (metis (no_types, lifting) x empty_iff homeq_imp_rcos rcos_imp_homeq the_elem_image_unique) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
828 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
829 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
830 |
lemma (in ring_hom_ring) the_elem_inj: |
81600 | 831 |
assumes "X \<in> carrier (R Quot (a_kernel R S h))" |
832 |
and "Y \<in> carrier (R Quot (a_kernel R S h))" |
|
833 |
and Eq: "the_elem (h ` X) = the_elem (h ` Y)" |
|
834 |
shows "X = Y" |
|
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
835 |
proof - |
81600 | 836 |
from assms obtain x y where x: "x \<in> carrier R" "X = (a_kernel R S h) +> x" |
837 |
and y: "y \<in> carrier R" "Y = (a_kernel R S h) +> y" |
|
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
838 |
unfolding FactRing_def A_RCOSETS_def' by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
839 |
hence "h x = h y" using Eq by simp |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
840 |
hence "x \<ominus> y \<in> (a_kernel R S h)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
841 |
by (simp add: a_minus_def abelian_subgroup.a_rcos_module_imp |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
842 |
abelian_subgroup_a_kernel homeq_imp_rcos x(1) y(1)) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
843 |
thus "X = Y" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
844 |
by (metis R.a_coset_add_inv1 R.minus_eq abelian_subgroup.a_rcos_const |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
845 |
abelian_subgroup_a_kernel additive_subgroup.a_subset additive_subgroup_a_kernel x y) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
846 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
847 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
848 |
lemma (in ring_hom_ring) quot_mem: |
81600 | 849 |
"X \<in> carrier (R Quot (a_kernel R S h)) \<Longrightarrow> \<exists>x \<in> carrier R. X = (a_kernel R S h) +> x" |
850 |
unfolding FactRing_simps by (simp add: a_r_coset_def) |
|
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
851 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
852 |
lemma (in ring_hom_ring) the_elem_wf: |
81600 | 853 |
assumes "X \<in> carrier (R Quot (a_kernel R S h))" |
854 |
shows "\<exists>y \<in> carrier S. (h ` X) = { y }" |
|
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
855 |
proof - |
81600 | 856 |
from assms obtain x where x: "x \<in> carrier R" and X: "X = (a_kernel R S h) +> x" |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
857 |
using quot_mem by blast |
81600 | 858 |
have "h x' = h x" if "x' \<in> X" for x' |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
859 |
proof - |
81600 | 860 |
from X that have "x' \<in> (a_kernel R S h) +> x" by simp |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
861 |
then obtain k where k: "k \<in> a_kernel R S h" "x' = k \<oplus> x" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
862 |
by (metis R.add.inv_closed R.add.m_assoc R.l_neg R.r_zero |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
863 |
abelian_subgroup.a_elemrcos_carrier |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
864 |
abelian_subgroup.a_rcos_module_imp abelian_subgroup_a_kernel x) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
865 |
hence "h x' = h k \<oplus>\<^bsub>S\<^esub> h x" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
866 |
by (meson additive_subgroup.a_Hcarr additive_subgroup_a_kernel hom_add x) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
867 |
also have " ... = h x" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
868 |
using k by (auto simp add: x) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
869 |
finally show "h x' = h x" . |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
870 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
871 |
moreover have "h x \<in> h ` X" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
872 |
by (simp add: X homeq_imp_rcos x) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
873 |
ultimately have "(h ` X) = { h x }" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
874 |
by blast |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
875 |
thus "\<exists>y \<in> carrier S. (h ` X) = { y }" using x by simp |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
876 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
877 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
878 |
corollary (in ring_hom_ring) the_elem_wf': |
81600 | 879 |
"X \<in> carrier (R Quot (a_kernel R S h)) \<Longrightarrow> \<exists>r \<in> carrier R. (h ` X) = { h r }" |
880 |
using the_elem_wf by (metis quot_mem the_elem_eq the_elem_simp) |
|
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
881 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
882 |
lemma (in ring_hom_ring) the_elem_hom: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
883 |
"(\<lambda>X. the_elem (h ` X)) \<in> ring_hom (R Quot (a_kernel R S h)) S" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
884 |
proof (rule ring_hom_memI) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
885 |
show "\<And>x. x \<in> carrier (R Quot a_kernel R S h) \<Longrightarrow> the_elem (h ` x) \<in> carrier S" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
886 |
using the_elem_wf by fastforce |
81600 | 887 |
|
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
888 |
show "the_elem (h ` \<one>\<^bsub>R Quot a_kernel R S h\<^esub>) = \<one>\<^bsub>S\<^esub>" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
889 |
unfolding FactRing_def using the_elem_simp[of "\<one>\<^bsub>R\<^esub>"] by simp |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
890 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
891 |
fix X Y |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
892 |
assume "X \<in> carrier (R Quot a_kernel R S h)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
893 |
and "Y \<in> carrier (R Quot a_kernel R S h)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
894 |
then obtain x y where x: "x \<in> carrier R" "X = (a_kernel R S h) +> x" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
895 |
and y: "y \<in> carrier R" "Y = (a_kernel R S h) +> y" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
896 |
using quot_mem by blast |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
897 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
898 |
have "X \<otimes>\<^bsub>R Quot a_kernel R S h\<^esub> Y = (a_kernel R S h) +> (x \<otimes> y)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
899 |
by (simp add: FactRing_def ideal.rcoset_mult_add kernel_is_ideal x y) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
900 |
thus "the_elem (h ` (X \<otimes>\<^bsub>R Quot a_kernel R S h\<^esub> Y)) = the_elem (h ` X) \<otimes>\<^bsub>S\<^esub> the_elem (h ` Y)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
901 |
by (simp add: x y) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
902 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
903 |
have "X \<oplus>\<^bsub>R Quot a_kernel R S h\<^esub> Y = (a_kernel R S h) +> (x \<oplus> y)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
904 |
using ideal.rcos_ring_hom kernel_is_ideal ring_hom_add x y by fastforce |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
905 |
thus "the_elem (h ` (X \<oplus>\<^bsub>R Quot a_kernel R S h\<^esub> Y)) = the_elem (h ` X) \<oplus>\<^bsub>S\<^esub> the_elem (h ` Y)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
906 |
by (simp add: x y) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
907 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
908 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
909 |
lemma (in ring_hom_ring) the_elem_surj: |
81600 | 910 |
"(\<lambda>X. (the_elem (h ` X))) ` carrier (R Quot (a_kernel R S h)) = (h ` (carrier R))" |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
911 |
proof |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
912 |
show "(\<lambda>X. the_elem (h ` X)) ` carrier (R Quot a_kernel R S h) \<subseteq> h ` carrier R" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
913 |
using the_elem_wf' by fastforce |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
914 |
show "h ` carrier R \<subseteq> (\<lambda>X. the_elem (h ` X)) ` carrier (R Quot a_kernel R S h)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
915 |
proof |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
916 |
fix y assume "y \<in> h ` carrier R" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
917 |
then obtain x where x: "x \<in> carrier R" "h x = y" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
918 |
by (metis image_iff) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
919 |
hence "the_elem (h ` ((a_kernel R S h) +> x)) = y" by simp |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
920 |
moreover have "(a_kernel R S h) +> x \<in> carrier (R Quot (a_kernel R S h))" |
68673 | 921 |
unfolding FactRing_simps by (auto simp add: x a_r_coset_def) |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
922 |
ultimately show "y \<in> (\<lambda>X. (the_elem (h ` X))) ` carrier (R Quot (a_kernel R S h))" by blast |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
923 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
924 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
925 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
926 |
proposition (in ring_hom_ring) FactRing_iso_set_aux: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
927 |
"(\<lambda>X. the_elem (h ` X)) \<in> ring_iso (R Quot (a_kernel R S h)) (S \<lparr> carrier := h ` (carrier R) \<rparr>)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
928 |
proof - |
81600 | 929 |
have *: "bij_betw (\<lambda>X. the_elem (h ` X)) (carrier (R Quot a_kernel R S h)) (h ` (carrier R))" |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
930 |
unfolding bij_betw_def inj_on_def using the_elem_surj the_elem_inj by simp |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
931 |
have "(\<lambda>X. the_elem (h ` X)): carrier (R Quot (a_kernel R S h)) \<rightarrow> h ` (carrier R)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
932 |
using the_elem_wf' by fastforce |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
933 |
hence "(\<lambda>X. the_elem (h ` X)) \<in> ring_hom (R Quot (a_kernel R S h)) (S \<lparr> carrier := h ` (carrier R) \<rparr>)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
934 |
using the_elem_hom the_elem_wf' unfolding ring_hom_def by simp |
81600 | 935 |
with * show ?thesis unfolding ring_iso_def using the_elem_hom by simp |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
936 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
937 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
938 |
theorem (in ring_hom_ring) FactRing_iso_set: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
939 |
assumes "h ` carrier R = carrier S" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
940 |
shows "(\<lambda>X. the_elem (h ` X)) \<in> ring_iso (R Quot (a_kernel R S h)) S" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
941 |
using FactRing_iso_set_aux assms by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
942 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
943 |
corollary (in ring_hom_ring) FactRing_iso: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
944 |
assumes "h ` carrier R = carrier S" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
945 |
shows "R Quot (a_kernel R S h) \<simeq> S" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
946 |
using FactRing_iso_set assms is_ring_iso_def by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
947 |
|
68583 | 948 |
corollary (in ring) FactRing_zeroideal: |
949 |
shows "R Quot { \<zero> } \<simeq> R" and "R \<simeq> R Quot { \<zero> }" |
|
950 |
proof - |
|
951 |
have "ring_hom_ring R R id" |
|
952 |
using ring_axioms by (auto intro: ring_hom_ringI) |
|
953 |
moreover have "a_kernel R R id = { \<zero> }" |
|
954 |
unfolding a_kernel_def' by auto |
|
955 |
ultimately show "R Quot { \<zero> } \<simeq> R" and "R \<simeq> R Quot { \<zero> }" |
|
956 |
using ring_hom_ring.FactRing_iso[of R R id] |
|
957 |
ring_iso_sym[OF ideal.quotient_is_ring[OF zeroideal], of R] by auto |
|
958 |
qed |
|
959 |
||
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
960 |
lemma (in ring_hom_ring) img_is_ring: "ring (S \<lparr> carrier := h ` (carrier R) \<rparr>)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
961 |
proof - |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
962 |
let ?the_elem = "\<lambda>X. the_elem (h ` X)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
963 |
have FactRing_is_ring: "ring (R Quot (a_kernel R S h))" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
964 |
by (simp add: ideal.quotient_is_ring kernel_is_ideal) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
965 |
have "ring ((S \<lparr> carrier := ?the_elem ` (carrier (R Quot (a_kernel R S h))) \<rparr>) |
69122
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68673
diff
changeset
|
966 |
\<lparr> zero := ?the_elem \<zero>\<^bsub>(R Quot (a_kernel R S h))\<^esub> \<rparr>)" |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
967 |
using ring.ring_iso_imp_img_ring[OF FactRing_is_ring, of ?the_elem |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
968 |
"S \<lparr> carrier := ?the_elem ` (carrier (R Quot (a_kernel R S h))) \<rparr>"] |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
969 |
FactRing_iso_set_aux the_elem_surj by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
970 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
971 |
moreover |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
972 |
have "\<zero> \<in> (a_kernel R S h)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
973 |
using a_kernel_def'[of R S h] by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
974 |
hence "\<one> \<in> (a_kernel R S h) +> \<one>" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
975 |
using a_r_coset_def'[of R "a_kernel R S h" \<one>] by force |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
976 |
hence "\<one>\<^bsub>S\<^esub> \<in> (h ` ((a_kernel R S h) +> \<one>))" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
977 |
using hom_one by force |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
978 |
hence "?the_elem \<one>\<^bsub>(R Quot (a_kernel R S h))\<^esub> = \<one>\<^bsub>S\<^esub>" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
979 |
using the_elem_wf[of "(a_kernel R S h) +> \<one>"] by (simp add: FactRing_def) |
81600 | 980 |
|
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
981 |
moreover |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
982 |
have "\<zero>\<^bsub>S\<^esub> \<in> (h ` (a_kernel R S h))" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
983 |
using a_kernel_def'[of R S h] hom_zero by force |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
984 |
hence "\<zero>\<^bsub>S\<^esub> \<in> (h ` \<zero>\<^bsub>(R Quot (a_kernel R S h))\<^esub>)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
985 |
by (simp add: FactRing_def) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
986 |
hence "?the_elem \<zero>\<^bsub>(R Quot (a_kernel R S h))\<^esub> = \<zero>\<^bsub>S\<^esub>" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
987 |
using the_elem_wf[OF ring.ring_simprules(2)[OF FactRing_is_ring]] |
81600 | 988 |
by (metis singletonD the_elem_eq) |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
989 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
990 |
ultimately |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
991 |
have "ring ((S \<lparr> carrier := h ` (carrier R) \<rparr>) \<lparr> one := \<one>\<^bsub>S\<^esub>, zero := \<zero>\<^bsub>S\<^esub> \<rparr>)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
992 |
using the_elem_surj by simp |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
993 |
thus ?thesis |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
994 |
by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
995 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
996 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
997 |
lemma (in ring_hom_ring) img_is_cring: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
998 |
assumes "cring S" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
999 |
shows "cring (S \<lparr> carrier := h ` (carrier R) \<rparr>)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1000 |
proof - |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1001 |
interpret ring "S \<lparr> carrier := h ` (carrier R) \<rparr>" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1002 |
using img_is_ring . |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1003 |
show ?thesis |
81600 | 1004 |
by unfold_locales (use assms in \<open>auto simp: cring_def comm_monoid_def comm_monoid_axioms_def\<close>) |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1005 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1006 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1007 |
lemma (in ring_hom_ring) img_is_domain: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1008 |
assumes "domain S" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1009 |
shows "domain (S \<lparr> carrier := h ` (carrier R) \<rparr>)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1010 |
proof - |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1011 |
interpret cring "S \<lparr> carrier := h ` (carrier R) \<rparr>" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1012 |
using img_is_cring assms unfolding domain_def by simp |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1013 |
show ?thesis |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1014 |
apply unfold_locales |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1015 |
using assms unfolding domain_def domain_axioms_def apply auto |
81600 | 1016 |
using hom_closed by blast |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1017 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1018 |
|
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1019 |
proposition (in ring_hom_ring) primeideal_vimage: |
81600 | 1020 |
assumes R: "cring R" |
1021 |
and A: "primeideal P S" |
|
1022 |
shows "primeideal { r \<in> carrier R. h r \<in> P } R" |
|
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1023 |
proof - |
81600 | 1024 |
from A have is_ideal: "ideal P S" unfolding primeideal_def by simp |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1025 |
have "ring_hom_ring R (S Quot P) (((+>\<^bsub>S\<^esub>) P) \<circ> h)" (is "ring_hom_ring ?A ?B ?h") |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1026 |
using ring_hom_trans[OF homh, of "(+>\<^bsub>S\<^esub>) P" "S Quot P"] |
81600 | 1027 |
ideal.rcos_ring_hom_ring[OF is_ideal] R |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1028 |
unfolding ring_hom_ring_def ring_hom_ring_axioms_def cring_def by simp |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1029 |
then interpret hom: ring_hom_ring R "S Quot P" "((+>\<^bsub>S\<^esub>) P) \<circ> h" by simp |
81600 | 1030 |
|
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1031 |
have "inj_on (\<lambda>X. the_elem (?h ` X)) (carrier (R Quot (a_kernel R (S Quot P) ?h)))" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1032 |
using hom.the_elem_inj unfolding inj_on_def by simp |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1033 |
moreover |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1034 |
have "ideal (a_kernel R (S Quot P) ?h) R" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1035 |
using hom.kernel_is_ideal by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1036 |
have hom': "ring_hom_ring (R Quot (a_kernel R (S Quot P) ?h)) (S Quot P) (\<lambda>X. the_elem (?h ` X))" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1037 |
using hom.the_elem_hom hom.kernel_is_ideal |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1038 |
by (meson hom.ring_hom_ring_axioms ideal.rcos_ring_hom_ring ring_hom_ring_axioms_def ring_hom_ring_def) |
81600 | 1039 |
|
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1040 |
ultimately |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1041 |
have "primeideal (a_kernel R (S Quot P) ?h) R" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1042 |
using ring_hom_ring.inj_on_domain[OF hom'] primeideal.quotient_is_domain[OF A] |
81600 | 1043 |
cring.quot_domain_imp_primeideal[OF R hom.kernel_is_ideal] by simp |
1044 |
||
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1045 |
moreover have "a_kernel R (S Quot P) ?h = { r \<in> carrier R. h r \<in> P }" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1046 |
proof |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1047 |
show "a_kernel R (S Quot P) ?h \<subseteq> { r \<in> carrier R. h r \<in> P }" |
81600 | 1048 |
proof |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1049 |
fix r assume "r \<in> a_kernel R (S Quot P) ?h" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1050 |
hence r: "r \<in> carrier R" "P +>\<^bsub>S\<^esub> (h r) = P" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1051 |
unfolding a_kernel_def kernel_def FactRing_def by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1052 |
hence "h r \<in> P" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1053 |
using S.a_rcosI R.l_zero S.l_zero additive_subgroup.a_subset[OF ideal.axioms(1)[OF is_ideal]] |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1054 |
additive_subgroup.zero_closed[OF ideal.axioms(1)[OF is_ideal]] hom_closed by metis |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1055 |
thus "r \<in> { r \<in> carrier R. h r \<in> P }" using r by simp |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1056 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1057 |
next |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1058 |
show "{ r \<in> carrier R. h r \<in> P } \<subseteq> a_kernel R (S Quot P) ?h" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1059 |
proof |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1060 |
fix r assume "r \<in> { r \<in> carrier R. h r \<in> P }" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1061 |
hence r: "r \<in> carrier R" "h r \<in> P" by simp_all |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1062 |
hence "?h r = P" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1063 |
by (simp add: S.a_coset_join2 additive_subgroup.a_subgroup ideal.axioms(1) is_ideal) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1064 |
thus "r \<in> a_kernel R (S Quot P) ?h" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1065 |
unfolding a_kernel_def kernel_def FactRing_def using r(1) by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1066 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1067 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1068 |
ultimately show "primeideal { r \<in> carrier R. h r \<in> P } R" by simp |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1069 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1070 |
|
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff
changeset
|
1071 |
end |