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