| author | Fabian Huch <huch@in.tum.de> | 
| Wed, 05 Jun 2024 17:27:13 +0200 | |
| changeset 80254 | 6b3374d208b8 | 
| parent 76215 | a642599ffdea | 
| child 80917 | 2a77bc3b4eac | 
| permissions | -rw-r--r-- | 
| 14883 | 1 | (* Title: ZF/ex/Ring.thy | 
| 2 | ||
| 3 | *) | |
| 4 | ||
| 60770 | 5 | section \<open>Rings\<close> | 
| 14883 | 6 | |
| 16417 | 7 | theory Ring imports Group begin | 
| 14883 | 8 | |
| 46821 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 paulson parents: 
41524diff
changeset | 9 | no_notation | 
| 69587 | 10 | cadd (infixl \<open>\<oplus>\<close> 65) and | 
| 11 | cmult (infixl \<open>\<otimes>\<close> 70) | |
| 46821 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 paulson parents: 
41524diff
changeset | 12 | |
| 14883 | 13 | (*First, we must simulate a record declaration: | 
| 14 | record ring = monoid + | |
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 15 | add :: "[i, i] \<Rightarrow> i" (infixl "\<oplus>\<index>" 65) | 
| 14883 | 16 |   zero :: i ("\<zero>\<index>")
 | 
| 17 | *) | |
| 18 | ||
| 21233 | 19 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 20 | add_field :: "i \<Rightarrow> i" where | 
| 21233 | 21 | "add_field(M) = fst(snd(snd(snd(M))))" | 
| 14883 | 22 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 23 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 24 | ring_add :: "[i, i, i] \<Rightarrow> i" (infixl \<open>\<oplus>\<index>\<close> 65) where | 
| 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 25 | "ring_add(M,x,y) = add_field(M) ` \<langle>x,y\<rangle>" | 
| 14883 | 26 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 27 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 28 | zero :: "i \<Rightarrow> i" (\<open>\<zero>\<index>\<close>) where | 
| 21233 | 29 | "zero(M) = fst(snd(snd(snd(snd(M)))))" | 
| 14883 | 30 | |
| 31 | ||
| 32 | lemma add_field_eq [simp]: "add_field(<C,M,I,A,z>) = A" | |
| 33 | by (simp add: add_field_def) | |
| 34 | ||
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 35 | lemma add_eq [simp]: "ring_add(<C,M,I,A,z>, x, y) = A ` \<langle>x,y\<rangle>" | 
| 14883 | 36 | by (simp add: ring_add_def) | 
| 37 | ||
| 38 | lemma zero_eq [simp]: "zero(<C,M,I,A,Z,z>) = Z" | |
| 39 | by (simp add: zero_def) | |
| 40 | ||
| 41 | ||
| 60770 | 42 | text \<open>Derived operations.\<close> | 
| 14883 | 43 | |
| 21233 | 44 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 45 | a_inv :: "[i,i] \<Rightarrow> i" (\<open>\<ominus>\<index> _\<close> [81] 80) where | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 46 | "a_inv(R) \<equiv> m_inv (<carrier(R), add_field(R), zero(R), 0>)" | 
| 14883 | 47 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 48 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 49 | minus :: "[i,i,i] \<Rightarrow> i" (\<open>(_ \<ominus>\<index> _)\<close> [65,66] 65) where | 
| 21233 | 50 | "\<lbrakk>x \<in> carrier(R); y \<in> carrier(R)\<rbrakk> \<Longrightarrow> x \<ominus>\<^bsub>R\<^esub> y = x \<oplus>\<^bsub>R\<^esub> (\<ominus>\<^bsub>R\<^esub> y)" | 
| 14883 | 51 | |
| 29223 | 52 | locale abelian_monoid = fixes G (structure) | 
| 46821 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 paulson parents: 
41524diff
changeset | 53 | assumes a_comm_monoid: | 
| 14883 | 54 | "comm_monoid (<carrier(G), add_field(G), zero(G), 0>)" | 
| 55 | ||
| 60770 | 56 | text \<open> | 
| 14883 | 57 | The following definition is redundant but simple to use. | 
| 60770 | 58 | \<close> | 
| 14883 | 59 | |
| 60 | locale abelian_group = abelian_monoid + | |
| 46821 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 paulson parents: 
41524diff
changeset | 61 | assumes a_comm_group: | 
| 14883 | 62 | "comm_group (<carrier(G), add_field(G), zero(G), 0>)" | 
| 63 | ||
| 29223 | 64 | locale ring = abelian_group R + monoid R for R (structure) + | 
| 14883 | 65 | assumes l_distr: "\<lbrakk>x \<in> carrier(R); y \<in> carrier(R); z \<in> carrier(R)\<rbrakk> | 
| 66 | \<Longrightarrow> (x \<oplus> y) \<cdot> z = x \<cdot> z \<oplus> y \<cdot> z" | |
| 67 | and r_distr: "\<lbrakk>x \<in> carrier(R); y \<in> carrier(R); z \<in> carrier(R)\<rbrakk> | |
| 68 | \<Longrightarrow> z \<cdot> (x \<oplus> y) = z \<cdot> x \<oplus> z \<cdot> y" | |
| 69 | ||
| 70 | locale cring = ring + comm_monoid R | |
| 71 | ||
| 72 | locale "domain" = cring + | |
| 73 | assumes one_not_zero [simp]: "\<one> \<noteq> \<zero>" | |
| 74 | and integral: "\<lbrakk>a \<cdot> b = \<zero>; a \<in> carrier(R); b \<in> carrier(R)\<rbrakk> \<Longrightarrow> | |
| 75 | a = \<zero> | b = \<zero>" | |
| 76 | ||
| 77 | ||
| 60770 | 78 | subsection \<open>Basic Properties\<close> | 
| 14883 | 79 | |
| 80 | lemma (in abelian_monoid) a_monoid: | |
| 81 | "monoid (<carrier(G), add_field(G), zero(G), 0>)" | |
| 46821 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 paulson parents: 
41524diff
changeset | 82 | apply (insert a_comm_monoid) | 
| 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 paulson parents: 
41524diff
changeset | 83 | apply (simp add: comm_monoid_def) | 
| 14883 | 84 | done | 
| 85 | ||
| 86 | lemma (in abelian_group) a_group: | |
| 87 | "group (<carrier(G), add_field(G), zero(G), 0>)" | |
| 46821 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 paulson parents: 
41524diff
changeset | 88 | apply (insert a_comm_group) | 
| 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 paulson parents: 
41524diff
changeset | 89 | apply (simp add: comm_group_def group_def) | 
| 14883 | 90 | done | 
| 91 | ||
| 92 | ||
| 93 | lemma (in abelian_monoid) l_zero [simp]: | |
| 94 | "x \<in> carrier(G) \<Longrightarrow> \<zero> \<oplus> x = x" | |
| 95 | apply (insert monoid.l_one [OF a_monoid]) | |
| 46821 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 paulson parents: 
41524diff
changeset | 96 | apply (simp add: ring_add_def) | 
| 14883 | 97 | done | 
| 98 | ||
| 99 | lemma (in abelian_monoid) zero_closed [intro, simp]: | |
| 100 | "\<zero> \<in> carrier(G)" | |
| 101 | by (rule monoid.one_closed [OF a_monoid, simplified]) | |
| 102 | ||
| 103 | lemma (in abelian_group) a_inv_closed [intro, simp]: | |
| 104 | "x \<in> carrier(G) \<Longrightarrow> \<ominus> x \<in> carrier(G)" | |
| 105 | by (simp add: a_inv_def group.inv_closed [OF a_group, simplified]) | |
| 106 | ||
| 107 | lemma (in abelian_monoid) a_closed [intro, simp]: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 108 | "\<lbrakk>x \<in> carrier(G); y \<in> carrier(G)\<rbrakk> \<Longrightarrow> x \<oplus> y \<in> carrier(G)" | 
| 46821 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 paulson parents: 
41524diff
changeset | 109 | by (rule monoid.m_closed [OF a_monoid, | 
| 14883 | 110 | simplified, simplified ring_add_def [symmetric]]) | 
| 111 | ||
| 112 | lemma (in abelian_group) minus_closed [intro, simp]: | |
| 113 | "\<lbrakk>x \<in> carrier(G); y \<in> carrier(G)\<rbrakk> \<Longrightarrow> x \<ominus> y \<in> carrier(G)" | |
| 114 | by (simp add: minus_def) | |
| 115 | ||
| 116 | lemma (in abelian_group) a_l_cancel [simp]: | |
| 46821 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 paulson parents: 
41524diff
changeset | 117 | "\<lbrakk>x \<in> carrier(G); y \<in> carrier(G); z \<in> carrier(G)\<rbrakk> | 
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
46821diff
changeset | 118 | \<Longrightarrow> (x \<oplus> y = x \<oplus> z) \<longleftrightarrow> (y = z)" | 
| 46821 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 paulson parents: 
41524diff
changeset | 119 | by (rule group.l_cancel [OF a_group, | 
| 14883 | 120 | simplified, simplified ring_add_def [symmetric]]) | 
| 121 | ||
| 122 | lemma (in abelian_group) a_r_cancel [simp]: | |
| 46821 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 paulson parents: 
41524diff
changeset | 123 | "\<lbrakk>x \<in> carrier(G); y \<in> carrier(G); z \<in> carrier(G)\<rbrakk> | 
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
46821diff
changeset | 124 | \<Longrightarrow> (y \<oplus> x = z \<oplus> x) \<longleftrightarrow> (y = z)" | 
| 14883 | 125 | by (rule group.r_cancel [OF a_group, simplified, simplified ring_add_def [symmetric]]) | 
| 126 | ||
| 127 | lemma (in abelian_monoid) a_assoc: | |
| 46821 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 paulson parents: 
41524diff
changeset | 128 | "\<lbrakk>x \<in> carrier(G); y \<in> carrier(G); z \<in> carrier(G)\<rbrakk> | 
| 14883 | 129 | \<Longrightarrow> (x \<oplus> y) \<oplus> z = x \<oplus> (y \<oplus> z)" | 
| 130 | by (rule monoid.m_assoc [OF a_monoid, simplified, simplified ring_add_def [symmetric]]) | |
| 131 | ||
| 132 | lemma (in abelian_group) l_neg: | |
| 133 | "x \<in> carrier(G) \<Longrightarrow> \<ominus> x \<oplus> x = \<zero>" | |
| 134 | by (simp add: a_inv_def | |
| 135 | group.l_inv [OF a_group, simplified, simplified ring_add_def [symmetric]]) | |
| 136 | ||
| 137 | lemma (in abelian_monoid) a_comm: | |
| 138 | "\<lbrakk>x \<in> carrier(G); y \<in> carrier(G)\<rbrakk> \<Longrightarrow> x \<oplus> y = y \<oplus> x" | |
| 139 | by (rule comm_monoid.m_comm [OF a_comm_monoid, | |
| 140 | simplified, simplified ring_add_def [symmetric]]) | |
| 141 | ||
| 142 | lemma (in abelian_monoid) a_lcomm: | |
| 46821 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 paulson parents: 
41524diff
changeset | 143 | "\<lbrakk>x \<in> carrier(G); y \<in> carrier(G); z \<in> carrier(G)\<rbrakk> | 
| 14883 | 144 | \<Longrightarrow> x \<oplus> (y \<oplus> z) = y \<oplus> (x \<oplus> z)" | 
| 145 | by (rule comm_monoid.m_lcomm [OF a_comm_monoid, | |
| 146 | simplified, simplified ring_add_def [symmetric]]) | |
| 147 | ||
| 148 | lemma (in abelian_monoid) r_zero [simp]: | |
| 149 | "x \<in> carrier(G) \<Longrightarrow> x \<oplus> \<zero> = x" | |
| 150 | using monoid.r_one [OF a_monoid] | |
| 151 | by (simp add: ring_add_def [symmetric]) | |
| 152 | ||
| 153 | lemma (in abelian_group) r_neg: | |
| 154 | "x \<in> carrier(G) \<Longrightarrow> x \<oplus> (\<ominus> x) = \<zero>" | |
| 155 | using group.r_inv [OF a_group] | |
| 156 | by (simp add: a_inv_def ring_add_def [symmetric]) | |
| 157 | ||
| 158 | lemma (in abelian_group) minus_zero [simp]: | |
| 159 | "\<ominus> \<zero> = \<zero>" | |
| 160 | by (simp add: a_inv_def | |
| 161 | group.inv_one [OF a_group, simplified ]) | |
| 162 | ||
| 163 | lemma (in abelian_group) minus_minus [simp]: | |
| 164 | "x \<in> carrier(G) \<Longrightarrow> \<ominus> (\<ominus> x) = x" | |
| 165 | using group.inv_inv [OF a_group, simplified] | |
| 166 | by (simp add: a_inv_def) | |
| 167 | ||
| 168 | lemma (in abelian_group) minus_add: | |
| 169 | "\<lbrakk>x \<in> carrier(G); y \<in> carrier(G)\<rbrakk> \<Longrightarrow> \<ominus> (x \<oplus> y) = \<ominus> x \<oplus> \<ominus> y" | |
| 170 | using comm_group.inv_mult [OF a_comm_group] | |
| 171 | by (simp add: a_inv_def ring_add_def [symmetric]) | |
| 172 | ||
| 173 | lemmas (in abelian_monoid) a_ac = a_assoc a_comm a_lcomm | |
| 174 | ||
| 60770 | 175 | text \<open> | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 176 | The following proofs are from Jacobson, Basic Algebra I, pp.\<not>88--89 | 
| 60770 | 177 | \<close> | 
| 14883 | 178 | |
| 21233 | 179 | context ring | 
| 180 | begin | |
| 181 | ||
| 182 | lemma l_null [simp]: "x \<in> carrier(R) \<Longrightarrow> \<zero> \<cdot> x = \<zero>" | |
| 14883 | 183 | proof - | 
| 184 | assume R: "x \<in> carrier(R)" | |
| 185 | then have "\<zero> \<cdot> x \<oplus> \<zero> \<cdot> x = (\<zero> \<oplus> \<zero>) \<cdot> x" | |
| 46821 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 paulson parents: 
41524diff
changeset | 186 | by (blast intro: l_distr [THEN sym]) | 
| 14883 | 187 | also from R have "... = \<zero> \<cdot> x \<oplus> \<zero>" by simp | 
| 188 | finally have "\<zero> \<cdot> x \<oplus> \<zero> \<cdot> x = \<zero> \<cdot> x \<oplus> \<zero>" . | |
| 189 | with R show ?thesis by (simp del: r_zero) | |
| 190 | qed | |
| 191 | ||
| 21233 | 192 | lemma r_null [simp]: "x \<in> carrier(R) \<Longrightarrow> x \<cdot> \<zero> = \<zero>" | 
| 14883 | 193 | proof - | 
| 194 | assume R: "x \<in> carrier(R)" | |
| 195 | then have "x \<cdot> \<zero> \<oplus> x \<cdot> \<zero> = x \<cdot> (\<zero> \<oplus> \<zero>)" | |
| 196 | by (simp add: r_distr del: l_zero r_zero) | |
| 197 | also from R have "... = x \<cdot> \<zero> \<oplus> \<zero>" by simp | |
| 198 | finally have "x \<cdot> \<zero> \<oplus> x \<cdot> \<zero> = x \<cdot> \<zero> \<oplus> \<zero>" . | |
| 199 | with R show ?thesis by (simp del: r_zero) | |
| 200 | qed | |
| 201 | ||
| 21233 | 202 | lemma l_minus: | 
| 14883 | 203 | "\<lbrakk>x \<in> carrier(R); y \<in> carrier(R)\<rbrakk> \<Longrightarrow> \<ominus> x \<cdot> y = \<ominus> (x \<cdot> y)" | 
| 204 | proof - | |
| 205 | assume R: "x \<in> carrier(R)" "y \<in> carrier(R)" | |
| 206 | then have "(\<ominus> x) \<cdot> y \<oplus> x \<cdot> y = (\<ominus> x \<oplus> x) \<cdot> y" by (simp add: l_distr) | |
| 41524 | 207 | also from R have "... = \<zero>" by (simp add: l_neg) | 
| 14883 | 208 | finally have "(\<ominus> x) \<cdot> y \<oplus> x \<cdot> y = \<zero>" . | 
| 209 | with R have "(\<ominus> x) \<cdot> y \<oplus> x \<cdot> y \<oplus> \<ominus> (x \<cdot> y) = \<zero> \<oplus> \<ominus> (x \<cdot> y)" by simp | |
| 210 | with R show ?thesis by (simp add: a_assoc r_neg) | |
| 211 | qed | |
| 212 | ||
| 21233 | 213 | lemma r_minus: | 
| 14883 | 214 | "\<lbrakk>x \<in> carrier(R); y \<in> carrier(R)\<rbrakk> \<Longrightarrow> x \<cdot> \<ominus> y = \<ominus> (x \<cdot> y)" | 
| 215 | proof - | |
| 216 | assume R: "x \<in> carrier(R)" "y \<in> carrier(R)" | |
| 217 | then have "x \<cdot> (\<ominus> y) \<oplus> x \<cdot> y = x \<cdot> (\<ominus> y \<oplus> y)" by (simp add: r_distr) | |
| 41524 | 218 | also from R have "... = \<zero>" by (simp add: l_neg) | 
| 14883 | 219 | finally have "x \<cdot> (\<ominus> y) \<oplus> x \<cdot> y = \<zero>" . | 
| 220 | with R have "x \<cdot> (\<ominus> y) \<oplus> x \<cdot> y \<oplus> \<ominus> (x \<cdot> y) = \<zero> \<oplus> \<ominus> (x \<cdot> y)" by simp | |
| 221 | with R show ?thesis by (simp add: a_assoc r_neg) | |
| 222 | qed | |
| 223 | ||
| 21233 | 224 | lemma minus_eq: | 
| 14883 | 225 | "\<lbrakk>x \<in> carrier(R); y \<in> carrier(R)\<rbrakk> \<Longrightarrow> x \<ominus> y = x \<oplus> \<ominus> y" | 
| 226 | by (simp only: minus_def) | |
| 227 | ||
| 21233 | 228 | end | 
| 229 | ||
| 14883 | 230 | |
| 60770 | 231 | subsection \<open>Morphisms\<close> | 
| 14883 | 232 | |
| 21233 | 233 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 234 | ring_hom :: "[i,i] \<Rightarrow> i" where | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 235 | "ring_hom(R,S) \<equiv> | 
| 14883 | 236 |     {h \<in> carrier(R) -> carrier(S).
 | 
| 76214 | 237 | (\<forall>x y. x \<in> carrier(R) \<and> y \<in> carrier(R) \<longrightarrow> | 
| 238 | h ` (x \<cdot>\<^bsub>R\<^esub> y) = (h ` x) \<cdot>\<^bsub>S\<^esub> (h ` y) \<and> | |
| 239 | h ` (x \<oplus>\<^bsub>R\<^esub> y) = (h ` x) \<oplus>\<^bsub>S\<^esub> (h ` y)) \<and> | |
| 14883 | 240 | h ` \<one>\<^bsub>R\<^esub> = \<one>\<^bsub>S\<^esub>}" | 
| 241 | ||
| 242 | lemma ring_hom_memI: | |
| 243 | assumes hom_type: "h \<in> carrier(R) \<rightarrow> carrier(S)" | |
| 244 | and hom_mult: "\<And>x y. \<lbrakk>x \<in> carrier(R); y \<in> carrier(R)\<rbrakk> \<Longrightarrow> | |
| 245 | h ` (x \<cdot>\<^bsub>R\<^esub> y) = (h ` x) \<cdot>\<^bsub>S\<^esub> (h ` y)" | |
| 246 | and hom_add: "\<And>x y. \<lbrakk>x \<in> carrier(R); y \<in> carrier(R)\<rbrakk> \<Longrightarrow> | |
| 247 | h ` (x \<oplus>\<^bsub>R\<^esub> y) = (h ` x) \<oplus>\<^bsub>S\<^esub> (h ` y)" | |
| 248 | and hom_one: "h ` \<one>\<^bsub>R\<^esub> = \<one>\<^bsub>S\<^esub>" | |
| 249 | shows "h \<in> ring_hom(R,S)" | |
| 41524 | 250 | by (auto simp add: ring_hom_def assms) | 
| 14883 | 251 | |
| 252 | lemma ring_hom_closed: | |
| 253 | "\<lbrakk>h \<in> ring_hom(R,S); x \<in> carrier(R)\<rbrakk> \<Longrightarrow> h ` x \<in> carrier(S)" | |
| 254 | by (auto simp add: ring_hom_def) | |
| 255 | ||
| 256 | lemma ring_hom_mult: | |
| 46821 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 paulson parents: 
41524diff
changeset | 257 | "\<lbrakk>h \<in> ring_hom(R,S); x \<in> carrier(R); y \<in> carrier(R)\<rbrakk> | 
| 14883 | 258 | \<Longrightarrow> h ` (x \<cdot>\<^bsub>R\<^esub> y) = (h ` x) \<cdot>\<^bsub>S\<^esub> (h ` y)" | 
| 259 | by (simp add: ring_hom_def) | |
| 260 | ||
| 261 | lemma ring_hom_add: | |
| 46821 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 paulson parents: 
41524diff
changeset | 262 | "\<lbrakk>h \<in> ring_hom(R,S); x \<in> carrier(R); y \<in> carrier(R)\<rbrakk> | 
| 14883 | 263 | \<Longrightarrow> h ` (x \<oplus>\<^bsub>R\<^esub> y) = (h ` x) \<oplus>\<^bsub>S\<^esub> (h ` y)" | 
| 264 | by (simp add: ring_hom_def) | |
| 265 | ||
| 266 | lemma ring_hom_one: "h \<in> ring_hom(R,S) \<Longrightarrow> h ` \<one>\<^bsub>R\<^esub> = \<one>\<^bsub>S\<^esub>" | |
| 267 | by (simp add: ring_hom_def) | |
| 268 | ||
| 29223 | 269 | locale ring_hom_cring = R: cring R + S: cring S | 
| 270 | for R (structure) and S (structure) and h + | |
| 14883 | 271 | assumes homh [simp, intro]: "h \<in> ring_hom(R,S)" | 
| 272 | notes hom_closed [simp, intro] = ring_hom_closed [OF homh] | |
| 273 | and hom_mult [simp] = ring_hom_mult [OF homh] | |
| 274 | and hom_add [simp] = ring_hom_add [OF homh] | |
| 275 | and hom_one [simp] = ring_hom_one [OF homh] | |
| 276 | ||
| 277 | lemma (in ring_hom_cring) hom_zero [simp]: | |
| 278 | "h ` \<zero>\<^bsub>R\<^esub> = \<zero>\<^bsub>S\<^esub>" | |
| 279 | proof - | |
| 280 | have "h ` \<zero>\<^bsub>R\<^esub> \<oplus>\<^bsub>S\<^esub> h ` \<zero> = h ` \<zero>\<^bsub>R\<^esub> \<oplus>\<^bsub>S\<^esub> \<zero>\<^bsub>S\<^esub>" | |
| 281 | by (simp add: hom_add [symmetric] del: hom_add) | |
| 282 | then show ?thesis by (simp del: S.r_zero) | |
| 283 | qed | |
| 284 | ||
| 285 | lemma (in ring_hom_cring) hom_a_inv [simp]: | |
| 286 | "x \<in> carrier(R) \<Longrightarrow> h ` (\<ominus>\<^bsub>R\<^esub> x) = \<ominus>\<^bsub>S\<^esub> h ` x" | |
| 287 | proof - | |
| 288 | assume R: "x \<in> carrier(R)" | |
| 289 | then have "h ` x \<oplus>\<^bsub>S\<^esub> h ` (\<ominus> x) = h ` x \<oplus>\<^bsub>S\<^esub> (\<ominus>\<^bsub>S\<^esub> (h ` x))" | |
| 290 | by (simp add: hom_add [symmetric] R.r_neg S.r_neg del: hom_add) | |
| 291 | with R show ?thesis by simp | |
| 292 | qed | |
| 293 | ||
| 294 | lemma (in ring) id_ring_hom [simp]: "id(carrier(R)) \<in> ring_hom(R,R)" | |
| 46821 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 paulson parents: 
41524diff
changeset | 295 | apply (rule ring_hom_memI) | 
| 
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
 paulson parents: 
41524diff
changeset | 296 | apply (auto simp add: id_type) | 
| 14883 | 297 | done | 
| 298 | ||
| 299 | end |