| author | huffman | 
| Wed, 12 Oct 2005 01:43:37 +0200 | |
| changeset 17837 | 2922be3544f8 | 
| parent 16637 | 62dff56b43d3 | 
| child 19233 | 77ca20b0ed77 | 
| permissions | -rw-r--r-- | 
| 13835 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 1 | (* | 
| 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 2 | Title: The algebraic hierarchy of rings | 
| 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 3 | Id: $Id$ | 
| 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 4 | Author: Clemens Ballarin, started 9 December 1996 | 
| 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 5 | Copyright: Clemens Ballarin | 
| 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 6 | *) | 
| 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 7 | |
| 14577 | 8 | header {* Abelian Groups *}
 | 
| 9 | ||
| 16417 | 10 | theory CRing imports FiniteProduct | 
| 11 | uses ("ringsimp.ML") begin
 | |
| 13835 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 12 | |
| 13936 | 13 | record 'a ring = "'a monoid" + | 
| 14 |   zero :: 'a ("\<zero>\<index>")
 | |
| 15 | add :: "['a, 'a] => 'a" (infixl "\<oplus>\<index>" 65) | |
| 16 | ||
| 17 | text {* Derived operations. *}
 | |
| 18 | ||
| 14651 | 19 | constdefs (structure R) | 
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 20 |   a_inv :: "[('a, 'm) ring_scheme, 'a ] => 'a" ("\<ominus>\<index> _" [81] 80)
 | 
| 13936 | 21 | "a_inv R == m_inv (| carrier = carrier R, mult = add R, one = zero R |)" | 
| 22 | ||
| 23 |   minus :: "[('a, 'm) ring_scheme, 'a, 'a] => 'a" (infixl "\<ominus>\<index>" 65)
 | |
| 14651 | 24 | "[| x \<in> carrier R; y \<in> carrier R |] ==> x \<ominus> y == x \<oplus> (\<ominus> y)" | 
| 13936 | 25 | |
| 26 | locale abelian_monoid = struct G + | |
| 14963 | 27 | assumes a_comm_monoid: | 
| 28 | "comm_monoid (| carrier = carrier G, mult = add G, one = zero G |)" | |
| 29 | ||
| 13936 | 30 | |
| 31 | text {*
 | |
| 32 | The following definition is redundant but simple to use. | |
| 33 | *} | |
| 34 | ||
| 35 | locale abelian_group = abelian_monoid + | |
| 14963 | 36 | assumes a_comm_group: | 
| 37 | "comm_group (| carrier = carrier G, mult = add G, one = zero G |)" | |
| 38 | ||
| 13936 | 39 | |
| 40 | subsection {* Basic Properties *}
 | |
| 41 | ||
| 42 | lemma abelian_monoidI: | |
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 43 | includes struct R | 
| 13936 | 44 | assumes a_closed: | 
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 45 | "!!x y. [| x \<in> carrier R; y \<in> carrier R |] ==> x \<oplus> y \<in> carrier R" | 
| 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 46 | and zero_closed: "\<zero> \<in> carrier R" | 
| 13936 | 47 | and a_assoc: | 
| 48 | "!!x y z. [| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |] ==> | |
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 49 | (x \<oplus> y) \<oplus> z = x \<oplus> (y \<oplus> z)" | 
| 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 50 | and l_zero: "!!x. x \<in> carrier R ==> \<zero> \<oplus> x = x" | 
| 13936 | 51 | and a_comm: | 
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 52 | "!!x y. [| x \<in> carrier R; y \<in> carrier R |] ==> x \<oplus> y = y \<oplus> x" | 
| 13936 | 53 | shows "abelian_monoid R" | 
| 54 | by (auto intro!: abelian_monoid.intro comm_monoidI intro: prems) | |
| 55 | ||
| 56 | lemma abelian_groupI: | |
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 57 | includes struct R | 
| 13936 | 58 | assumes a_closed: | 
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 59 | "!!x y. [| x \<in> carrier R; y \<in> carrier R |] ==> x \<oplus> y \<in> carrier R" | 
| 13936 | 60 | and zero_closed: "zero R \<in> carrier R" | 
| 61 | and a_assoc: | |
| 62 | "!!x y z. [| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |] ==> | |
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 63 | (x \<oplus> y) \<oplus> z = x \<oplus> (y \<oplus> z)" | 
| 13936 | 64 | and a_comm: | 
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 65 | "!!x y. [| x \<in> carrier R; y \<in> carrier R |] ==> x \<oplus> y = y \<oplus> x" | 
| 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 66 | and l_zero: "!!x. x \<in> carrier R ==> \<zero> \<oplus> x = x" | 
| 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 67 | and l_inv_ex: "!!x. x \<in> carrier R ==> EX y : carrier R. y \<oplus> x = \<zero>" | 
| 13936 | 68 | shows "abelian_group R" | 
| 69 | by (auto intro!: abelian_group.intro abelian_monoidI | |
| 70 | abelian_group_axioms.intro comm_monoidI comm_groupI | |
| 71 | intro: prems) | |
| 72 | ||
| 73 | lemma (in abelian_monoid) a_monoid: | |
| 74 | "monoid (| carrier = carrier G, mult = add G, one = zero G |)" | |
| 14963 | 75 | by (rule comm_monoid.axioms, rule a_comm_monoid) | 
| 13936 | 76 | |
| 77 | lemma (in abelian_group) a_group: | |
| 78 | "group (| carrier = carrier G, mult = add G, one = zero G |)" | |
| 14963 | 79 | by (simp add: group_def a_monoid comm_group.axioms a_comm_group) | 
| 13936 | 80 | |
| 14963 | 81 | lemmas monoid_record_simps = partial_object.simps monoid.simps | 
| 13936 | 82 | |
| 83 | lemma (in abelian_monoid) a_closed [intro, simp]: | |
| 14963 | 84 | "\<lbrakk> x \<in> carrier G; y \<in> carrier G \<rbrakk> \<Longrightarrow> x \<oplus> y \<in> carrier G" | 
| 85 | by (rule monoid.m_closed [OF a_monoid, simplified monoid_record_simps]) | |
| 13936 | 86 | |
| 87 | lemma (in abelian_monoid) zero_closed [intro, simp]: | |
| 88 | "\<zero> \<in> carrier G" | |
| 89 | by (rule monoid.one_closed [OF a_monoid, simplified monoid_record_simps]) | |
| 90 | ||
| 91 | lemma (in abelian_group) a_inv_closed [intro, simp]: | |
| 92 | "x \<in> carrier G ==> \<ominus> x \<in> carrier G" | |
| 93 | by (simp add: a_inv_def | |
| 94 | group.inv_closed [OF a_group, simplified monoid_record_simps]) | |
| 95 | ||
| 96 | lemma (in abelian_group) minus_closed [intro, simp]: | |
| 97 | "[| x \<in> carrier G; y \<in> carrier G |] ==> x \<ominus> y \<in> carrier G" | |
| 98 | by (simp add: minus_def) | |
| 99 | ||
| 100 | lemma (in abelian_group) a_l_cancel [simp]: | |
| 101 | "[| x \<in> carrier G; y \<in> carrier G; z \<in> carrier G |] ==> | |
| 102 | (x \<oplus> y = x \<oplus> z) = (y = z)" | |
| 103 | by (rule group.l_cancel [OF a_group, simplified monoid_record_simps]) | |
| 104 | ||
| 105 | lemma (in abelian_group) a_r_cancel [simp]: | |
| 106 | "[| x \<in> carrier G; y \<in> carrier G; z \<in> carrier G |] ==> | |
| 107 | (y \<oplus> x = z \<oplus> x) = (y = z)" | |
| 108 | by (rule group.r_cancel [OF a_group, simplified monoid_record_simps]) | |
| 109 | ||
| 110 | lemma (in abelian_monoid) a_assoc: | |
| 14963 | 111 | "\<lbrakk>x \<in> carrier G; y \<in> carrier G; z \<in> carrier G\<rbrakk> \<Longrightarrow> | 
| 13936 | 112 | (x \<oplus> y) \<oplus> z = x \<oplus> (y \<oplus> z)" | 
| 14963 | 113 | by (rule monoid.m_assoc [OF a_monoid, simplified monoid_record_simps]) | 
| 13936 | 114 | |
| 115 | lemma (in abelian_monoid) l_zero [simp]: | |
| 116 | "x \<in> carrier G ==> \<zero> \<oplus> x = x" | |
| 117 | by (rule monoid.l_one [OF a_monoid, simplified monoid_record_simps]) | |
| 118 | ||
| 119 | lemma (in abelian_group) l_neg: | |
| 120 | "x \<in> carrier G ==> \<ominus> x \<oplus> x = \<zero>" | |
| 121 | by (simp add: a_inv_def | |
| 122 | group.l_inv [OF a_group, simplified monoid_record_simps]) | |
| 123 | ||
| 124 | lemma (in abelian_monoid) a_comm: | |
| 14963 | 125 | "\<lbrakk>x \<in> carrier G; y \<in> carrier G\<rbrakk> \<Longrightarrow> x \<oplus> y = y \<oplus> x" | 
| 126 | by (rule comm_monoid.m_comm [OF a_comm_monoid, | |
| 13936 | 127 | simplified monoid_record_simps]) | 
| 128 | ||
| 129 | lemma (in abelian_monoid) a_lcomm: | |
| 14963 | 130 | "\<lbrakk>x \<in> carrier G; y \<in> carrier G; z \<in> carrier G\<rbrakk> \<Longrightarrow> | 
| 13936 | 131 | x \<oplus> (y \<oplus> z) = y \<oplus> (x \<oplus> z)" | 
| 14963 | 132 | by (rule comm_monoid.m_lcomm [OF a_comm_monoid, | 
| 133 | simplified monoid_record_simps]) | |
| 13936 | 134 | |
| 135 | lemma (in abelian_monoid) r_zero [simp]: | |
| 136 | "x \<in> carrier G ==> x \<oplus> \<zero> = x" | |
| 137 | using monoid.r_one [OF a_monoid] | |
| 138 | by simp | |
| 139 | ||
| 140 | lemma (in abelian_group) r_neg: | |
| 141 | "x \<in> carrier G ==> x \<oplus> (\<ominus> x) = \<zero>" | |
| 142 | using group.r_inv [OF a_group] | |
| 143 | by (simp add: a_inv_def) | |
| 144 | ||
| 145 | lemma (in abelian_group) minus_zero [simp]: | |
| 146 | "\<ominus> \<zero> = \<zero>" | |
| 147 | by (simp add: a_inv_def | |
| 148 | group.inv_one [OF a_group, simplified monoid_record_simps]) | |
| 149 | ||
| 150 | lemma (in abelian_group) minus_minus [simp]: | |
| 151 | "x \<in> carrier G ==> \<ominus> (\<ominus> x) = x" | |
| 152 | using group.inv_inv [OF a_group, simplified monoid_record_simps] | |
| 153 | by (simp add: a_inv_def) | |
| 154 | ||
| 155 | lemma (in abelian_group) a_inv_inj: | |
| 156 | "inj_on (a_inv G) (carrier G)" | |
| 157 | using group.inv_inj [OF a_group, simplified monoid_record_simps] | |
| 158 | by (simp add: a_inv_def) | |
| 159 | ||
| 160 | lemma (in abelian_group) minus_add: | |
| 161 | "[| x \<in> carrier G; y \<in> carrier G |] ==> \<ominus> (x \<oplus> y) = \<ominus> x \<oplus> \<ominus> y" | |
| 162 | using comm_group.inv_mult [OF a_comm_group] | |
| 163 | by (simp add: a_inv_def) | |
| 164 | ||
| 165 | lemmas (in abelian_monoid) a_ac = a_assoc a_comm a_lcomm | |
| 166 | ||
| 167 | subsection {* Sums over Finite Sets *}
 | |
| 168 | ||
| 169 | text {*
 | |
| 170 |   This definition makes it easy to lift lemmas from @{term finprod}.
 | |
| 171 | *} | |
| 172 | ||
| 173 | constdefs | |
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 174 |   finsum :: "[('b, 'm) ring_scheme, 'a => 'b, 'a set] => 'b"
 | 
| 13936 | 175 | "finsum G f A == finprod (| carrier = carrier G, | 
| 176 | mult = add G, one = zero G |) f A" | |
| 177 | ||
| 14651 | 178 | syntax | 
| 179 | "_finsum" :: "index => idt => 'a set => 'b => 'b" | |
| 14666 | 180 |       ("(3\<Oplus>__:_. _)" [1000, 0, 51, 10] 10)
 | 
| 14651 | 181 | syntax (xsymbols) | 
| 182 | "_finsum" :: "index => idt => 'a set => 'b => 'b" | |
| 14666 | 183 |       ("(3\<Oplus>__\<in>_. _)" [1000, 0, 51, 10] 10)
 | 
| 14651 | 184 | syntax (HTML output) | 
| 185 | "_finsum" :: "index => idt => 'a set => 'b => 'b" | |
| 14666 | 186 |       ("(3\<Oplus>__\<in>_. _)" [1000, 0, 51, 10] 10)
 | 
| 14651 | 187 | translations | 
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 188 | "\<Oplus>\<index>i:A. b" == "finsum \<struct>\<index> (%i. b) A" | 
| 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 189 |   -- {* Beware of argument permutation! *}
 | 
| 14651 | 190 | |
| 13936 | 191 | (* | 
| 192 | lemmas (in abelian_monoid) finsum_empty [simp] = | |
| 193 | comm_monoid.finprod_empty [OF a_comm_monoid, simplified] | |
| 194 | is dangeous, because attributes (like simplified) are applied upon opening | |
| 195 | the locale, simplified refers to the simpset at that time!!! | |
| 196 | ||
| 197 | lemmas (in abelian_monoid) finsum_empty [simp] = | |
| 198 | abelian_monoid.finprod_empty [OF a_abelian_monoid, folded finsum_def, | |
| 199 | simplified monoid_record_simps] | |
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 200 | makes the locale slow, because proofs are repeated for every | 
| 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 201 | "lemma (in abelian_monoid)" command. | 
| 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 202 | When lemma is used time in UnivPoly.thy from beginning to UP_cring goes down | 
| 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 203 | from 110 secs to 60 secs. | 
| 13936 | 204 | *) | 
| 205 | ||
| 206 | lemma (in abelian_monoid) finsum_empty [simp]: | |
| 207 |   "finsum G f {} = \<zero>"
 | |
| 208 | by (rule comm_monoid.finprod_empty [OF a_comm_monoid, | |
| 209 | folded finsum_def, simplified monoid_record_simps]) | |
| 210 | ||
| 211 | lemma (in abelian_monoid) finsum_insert [simp]: | |
| 212 | "[| finite F; a \<notin> F; f \<in> F -> carrier G; f a \<in> carrier G |] | |
| 213 | ==> finsum G f (insert a F) = f a \<oplus> finsum G f F" | |
| 214 | by (rule comm_monoid.finprod_insert [OF a_comm_monoid, | |
| 215 | folded finsum_def, simplified monoid_record_simps]) | |
| 216 | ||
| 217 | lemma (in abelian_monoid) finsum_zero [simp]: | |
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 218 | "finite A ==> (\<Oplus>i\<in>A. \<zero>) = \<zero>" | 
| 13936 | 219 | by (rule comm_monoid.finprod_one [OF a_comm_monoid, folded finsum_def, | 
| 220 | simplified monoid_record_simps]) | |
| 221 | ||
| 222 | lemma (in abelian_monoid) finsum_closed [simp]: | |
| 223 | fixes A | |
| 224 | assumes fin: "finite A" and f: "f \<in> A -> carrier G" | |
| 225 | shows "finsum G f A \<in> carrier G" | |
| 226 | by (rule comm_monoid.finprod_closed [OF a_comm_monoid, | |
| 227 | folded finsum_def, simplified monoid_record_simps]) | |
| 228 | ||
| 229 | lemma (in abelian_monoid) finsum_Un_Int: | |
| 230 | "[| finite A; finite B; g \<in> A -> carrier G; g \<in> B -> carrier G |] ==> | |
| 231 | finsum G g (A Un B) \<oplus> finsum G g (A Int B) = | |
| 232 | finsum G g A \<oplus> finsum G g B" | |
| 233 | by (rule comm_monoid.finprod_Un_Int [OF a_comm_monoid, | |
| 234 | folded finsum_def, simplified monoid_record_simps]) | |
| 235 | ||
| 236 | lemma (in abelian_monoid) finsum_Un_disjoint: | |
| 237 |   "[| finite A; finite B; A Int B = {};
 | |
| 238 | g \<in> A -> carrier G; g \<in> B -> carrier G |] | |
| 239 | ==> finsum G g (A Un B) = finsum G g A \<oplus> finsum G g B" | |
| 240 | by (rule comm_monoid.finprod_Un_disjoint [OF a_comm_monoid, | |
| 241 | folded finsum_def, simplified monoid_record_simps]) | |
| 242 | ||
| 243 | lemma (in abelian_monoid) finsum_addf: | |
| 244 | "[| finite A; f \<in> A -> carrier G; g \<in> A -> carrier G |] ==> | |
| 245 | finsum G (%x. f x \<oplus> g x) A = (finsum G f A \<oplus> finsum G g A)" | |
| 246 | by (rule comm_monoid.finprod_multf [OF a_comm_monoid, | |
| 247 | folded finsum_def, simplified monoid_record_simps]) | |
| 248 | ||
| 249 | lemma (in abelian_monoid) finsum_cong': | |
| 250 | "[| A = B; g : B -> carrier G; | |
| 251 | !!i. i : B ==> f i = g i |] ==> finsum G f A = finsum G g B" | |
| 252 | by (rule comm_monoid.finprod_cong' [OF a_comm_monoid, | |
| 253 | folded finsum_def, simplified monoid_record_simps]) auto | |
| 254 | ||
| 255 | lemma (in abelian_monoid) finsum_0 [simp]: | |
| 256 |   "f : {0::nat} -> carrier G ==> finsum G f {..0} = f 0"
 | |
| 257 | by (rule comm_monoid.finprod_0 [OF a_comm_monoid, folded finsum_def, | |
| 258 | simplified monoid_record_simps]) | |
| 259 | ||
| 260 | lemma (in abelian_monoid) finsum_Suc [simp]: | |
| 261 |   "f : {..Suc n} -> carrier G ==>
 | |
| 262 |    finsum G f {..Suc n} = (f (Suc n) \<oplus> finsum G f {..n})"
 | |
| 263 | by (rule comm_monoid.finprod_Suc [OF a_comm_monoid, folded finsum_def, | |
| 264 | simplified monoid_record_simps]) | |
| 265 | ||
| 266 | lemma (in abelian_monoid) finsum_Suc2: | |
| 267 |   "f : {..Suc n} -> carrier G ==>
 | |
| 268 |    finsum G f {..Suc n} = (finsum G (%i. f (Suc i)) {..n} \<oplus> f 0)"
 | |
| 269 | by (rule comm_monoid.finprod_Suc2 [OF a_comm_monoid, folded finsum_def, | |
| 270 | simplified monoid_record_simps]) | |
| 271 | ||
| 272 | lemma (in abelian_monoid) finsum_add [simp]: | |
| 273 |   "[| f : {..n} -> carrier G; g : {..n} -> carrier G |] ==>
 | |
| 274 |      finsum G (%i. f i \<oplus> g i) {..n::nat} =
 | |
| 275 |      finsum G f {..n} \<oplus> finsum G g {..n}"
 | |
| 276 | by (rule comm_monoid.finprod_mult [OF a_comm_monoid, folded finsum_def, | |
| 277 | simplified monoid_record_simps]) | |
| 278 | ||
| 279 | lemma (in abelian_monoid) finsum_cong: | |
| 16637 
62dff56b43d3
Tuned finsum_cong to allow that premises are simplified more eagerly.
 berghofe parents: 
16417diff
changeset | 280 | "[| A = B; f : B -> carrier G; | 
| 
62dff56b43d3
Tuned finsum_cong to allow that premises are simplified more eagerly.
 berghofe parents: 
16417diff
changeset | 281 | !!i. i : B =simp=> f i = g i |] ==> finsum G f A = finsum G g B" | 
| 13936 | 282 | by (rule comm_monoid.finprod_cong [OF a_comm_monoid, folded finsum_def, | 
| 16637 
62dff56b43d3
Tuned finsum_cong to allow that premises are simplified more eagerly.
 berghofe parents: 
16417diff
changeset | 283 | simplified monoid_record_simps]) (auto simp add: simp_implies_def) | 
| 13936 | 284 | |
| 285 | text {*Usually, if this rule causes a failed congruence proof error,
 | |
| 286 |    the reason is that the premise @{text "g \<in> B -> carrier G"} cannot be shown.
 | |
| 287 |    Adding @{thm [source] Pi_def} to the simpset is often useful. *}
 | |
| 288 | ||
| 13835 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 289 | section {* The Algebraic Hierarchy of Rings *}
 | 
| 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 290 | |
| 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 291 | subsection {* Basic Definitions *}
 | 
| 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 292 | |
| 14399 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 293 | locale ring = abelian_group R + monoid R + | 
| 13936 | 294 | assumes l_distr: "[| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |] | 
| 13835 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 295 | ==> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z" | 
| 14399 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 296 | and r_distr: "[| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |] | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 297 | ==> z \<otimes> (x \<oplus> y) = z \<otimes> x \<oplus> z \<otimes> y" | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 298 | |
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 299 | locale cring = ring + comm_monoid R | 
| 13835 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 300 | |
| 13864 | 301 | locale "domain" = cring + | 
| 302 | assumes one_not_zero [simp]: "\<one> ~= \<zero>" | |
| 303 | and integral: "[| a \<otimes> b = \<zero>; a \<in> carrier R; b \<in> carrier R |] ==> | |
| 304 | a = \<zero> | b = \<zero>" | |
| 305 | ||
| 14551 | 306 | locale field = "domain" + | 
| 307 |   assumes field_Units: "Units R = carrier R - {\<zero>}"
 | |
| 308 | ||
| 13864 | 309 | subsection {* Basic Facts of Rings *}
 | 
| 13835 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 310 | |
| 14399 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 311 | lemma ringI: | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 312 | includes struct R | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 313 | assumes abelian_group: "abelian_group R" | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 314 | and monoid: "monoid R" | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 315 | and l_distr: "!!x y z. [| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |] | 
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 316 | ==> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z" | 
| 14399 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 317 | and r_distr: "!!x y z. [| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |] | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 318 | ==> z \<otimes> (x \<oplus> y) = z \<otimes> x \<oplus> z \<otimes> y" | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 319 | shows "ring R" | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 320 | by (auto intro: ring.intro | 
| 14963 | 321 | abelian_group.axioms ring_axioms.intro prems) | 
| 14399 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 322 | |
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 323 | lemma (in ring) is_abelian_group: | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 324 | "abelian_group R" | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 325 | by (auto intro!: abelian_groupI a_assoc a_comm l_neg) | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 326 | |
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 327 | lemma (in ring) is_monoid: | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 328 | "monoid R" | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 329 | by (auto intro!: monoidI m_assoc) | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 330 | |
| 13936 | 331 | lemma cringI: | 
| 14399 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 332 | includes struct R | 
| 13936 | 333 | assumes abelian_group: "abelian_group R" | 
| 334 | and comm_monoid: "comm_monoid R" | |
| 335 | and l_distr: "!!x y z. [| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |] | |
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 336 | ==> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z" | 
| 13936 | 337 | shows "cring R" | 
| 14399 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 338 | proof (rule cring.intro) | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 339 | show "ring_axioms R" | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 340 |     -- {* Right-distributivity follows from left-distributivity and
 | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 341 | commutativity. *} | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 342 | proof (rule ring_axioms.intro) | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 343 | fix x y z | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 344 | assume R: "x \<in> carrier R" "y \<in> carrier R" "z \<in> carrier R" | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 345 | note [simp]= comm_monoid.axioms [OF comm_monoid] | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 346 | abelian_group.axioms [OF abelian_group] | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 347 | abelian_monoid.a_closed | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 348 | |
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 349 | from R have "z \<otimes> (x \<oplus> y) = (x \<oplus> y) \<otimes> z" | 
| 14963 | 350 | by (simp add: comm_monoid.m_comm [OF comm_monoid.intro]) | 
| 14399 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 351 | also from R have "... = x \<otimes> z \<oplus> y \<otimes> z" by (simp add: l_distr) | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 352 | also from R have "... = z \<otimes> x \<oplus> z \<otimes> y" | 
| 14963 | 353 | by (simp add: comm_monoid.m_comm [OF comm_monoid.intro]) | 
| 14399 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 354 | finally show "z \<otimes> (x \<oplus> y) = z \<otimes> x \<oplus> z \<otimes> y" . | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 355 | qed | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 356 | qed (auto intro: cring.intro | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 357 | abelian_group.axioms comm_monoid.axioms ring_axioms.intro prems) | 
| 13854 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 358 | |
| 13936 | 359 | lemma (in cring) is_comm_monoid: | 
| 360 | "comm_monoid R" | |
| 361 | by (auto intro!: comm_monoidI m_assoc m_comm) | |
| 13835 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 362 | |
| 14551 | 363 | subsection {* Normaliser for Rings *}
 | 
| 13835 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 364 | |
| 13936 | 365 | lemma (in abelian_group) r_neg2: | 
| 366 | "[| x \<in> carrier G; y \<in> carrier G |] ==> x \<oplus> (\<ominus> x \<oplus> y) = y" | |
| 13854 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 367 | proof - | 
| 13936 | 368 | assume G: "x \<in> carrier G" "y \<in> carrier G" | 
| 369 | then have "(x \<oplus> \<ominus> x) \<oplus> y = y" | |
| 370 | by (simp only: r_neg l_zero) | |
| 371 | with G show ?thesis | |
| 372 | by (simp add: a_ac) | |
| 13835 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 373 | qed | 
| 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 374 | |
| 13936 | 375 | lemma (in abelian_group) r_neg1: | 
| 376 | "[| x \<in> carrier G; y \<in> carrier G |] ==> \<ominus> x \<oplus> (x \<oplus> y) = y" | |
| 13854 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 377 | proof - | 
| 13936 | 378 | assume G: "x \<in> carrier G" "y \<in> carrier G" | 
| 379 | then have "(\<ominus> x \<oplus> x) \<oplus> y = y" | |
| 380 | by (simp only: l_neg l_zero) | |
| 13854 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 381 | with G show ?thesis by (simp add: a_ac) | 
| 13835 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 382 | qed | 
| 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 383 | |
| 13854 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 384 | text {* 
 | 
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 385 | The following proofs are from Jacobson, Basic Algebra I, pp.~88--89 | 
| 13835 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 386 | *} | 
| 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 387 | |
| 14399 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 388 | lemma (in ring) l_null [simp]: | 
| 13854 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 389 | "x \<in> carrier R ==> \<zero> \<otimes> x = \<zero>" | 
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 390 | proof - | 
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 391 | assume R: "x \<in> carrier R" | 
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 392 | then have "\<zero> \<otimes> x \<oplus> \<zero> \<otimes> x = (\<zero> \<oplus> \<zero>) \<otimes> x" | 
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 393 | by (simp add: l_distr del: l_zero r_zero) | 
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 394 | also from R have "... = \<zero> \<otimes> x \<oplus> \<zero>" by simp | 
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 395 | finally have "\<zero> \<otimes> x \<oplus> \<zero> \<otimes> x = \<zero> \<otimes> x \<oplus> \<zero>" . | 
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 396 | with R show ?thesis by (simp del: r_zero) | 
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 397 | qed | 
| 13835 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 398 | |
| 14399 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 399 | lemma (in ring) r_null [simp]: | 
| 13854 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 400 | "x \<in> carrier R ==> x \<otimes> \<zero> = \<zero>" | 
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 401 | proof - | 
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 402 | assume R: "x \<in> carrier R" | 
| 14399 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 403 | then have "x \<otimes> \<zero> \<oplus> x \<otimes> \<zero> = x \<otimes> (\<zero> \<oplus> \<zero>)" | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 404 | by (simp add: r_distr del: l_zero r_zero) | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 405 | also from R have "... = x \<otimes> \<zero> \<oplus> \<zero>" by simp | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 406 | finally have "x \<otimes> \<zero> \<oplus> x \<otimes> \<zero> = x \<otimes> \<zero> \<oplus> \<zero>" . | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 407 | with R show ?thesis by (simp del: r_zero) | 
| 13854 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 408 | qed | 
| 13835 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 409 | |
| 14399 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 410 | lemma (in ring) l_minus: | 
| 13854 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 411 | "[| x \<in> carrier R; y \<in> carrier R |] ==> \<ominus> x \<otimes> y = \<ominus> (x \<otimes> y)" | 
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 412 | proof - | 
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 413 | assume R: "x \<in> carrier R" "y \<in> carrier R" | 
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 414 | then have "(\<ominus> x) \<otimes> y \<oplus> x \<otimes> y = (\<ominus> x \<oplus> x) \<otimes> y" by (simp add: l_distr) | 
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 415 | also from R have "... = \<zero>" by (simp add: l_neg l_null) | 
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 416 | finally have "(\<ominus> x) \<otimes> y \<oplus> x \<otimes> y = \<zero>" . | 
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 417 | with R have "(\<ominus> x) \<otimes> y \<oplus> x \<otimes> y \<oplus> \<ominus> (x \<otimes> y) = \<zero> \<oplus> \<ominus> (x \<otimes> y)" by simp | 
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 418 | with R show ?thesis by (simp add: a_assoc r_neg ) | 
| 13835 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 419 | qed | 
| 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 420 | |
| 14399 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 421 | lemma (in ring) r_minus: | 
| 13854 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 422 | "[| x \<in> carrier R; y \<in> carrier R |] ==> x \<otimes> \<ominus> y = \<ominus> (x \<otimes> y)" | 
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 423 | proof - | 
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 424 | assume R: "x \<in> carrier R" "y \<in> carrier R" | 
| 14399 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 425 | then have "x \<otimes> (\<ominus> y) \<oplus> x \<otimes> y = x \<otimes> (\<ominus> y \<oplus> y)" by (simp add: r_distr) | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 426 | also from R have "... = \<zero>" by (simp add: l_neg r_null) | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 427 | finally have "x \<otimes> (\<ominus> y) \<oplus> x \<otimes> y = \<zero>" . | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 428 | with R have "x \<otimes> (\<ominus> y) \<oplus> x \<otimes> y \<oplus> \<ominus> (x \<otimes> y) = \<zero> \<oplus> \<ominus> (x \<otimes> y)" by simp | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 429 | with R show ?thesis by (simp add: a_assoc r_neg ) | 
| 13835 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 430 | qed | 
| 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 431 | |
| 14399 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 432 | lemma (in ring) minus_eq: | 
| 13936 | 433 | "[| x \<in> carrier R; y \<in> carrier R |] ==> x \<ominus> y = x \<oplus> \<ominus> y" | 
| 434 | by (simp only: minus_def) | |
| 435 | ||
| 14399 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 436 | lemmas (in ring) ring_simprules = | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 437 | a_closed zero_closed a_inv_closed minus_closed m_closed one_closed | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 438 | a_assoc l_zero l_neg a_comm m_assoc l_one l_distr minus_eq | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 439 | r_zero r_neg r_neg2 r_neg1 minus_add minus_minus minus_zero | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 440 | a_lcomm r_distr l_null r_null l_minus r_minus | 
| 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 441 | |
| 13854 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 442 | lemmas (in cring) cring_simprules = | 
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 443 | a_closed zero_closed a_inv_closed minus_closed m_closed one_closed | 
| 13936 | 444 | a_assoc l_zero l_neg a_comm m_assoc l_one l_distr m_comm minus_eq | 
| 13854 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 445 | r_zero r_neg r_neg2 r_neg1 minus_add minus_minus minus_zero | 
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 446 | a_lcomm m_lcomm r_distr l_null r_null l_minus r_minus | 
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 447 | |
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 448 | use "ringsimp.ML" | 
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 449 | |
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 450 | method_setup algebra = | 
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 451 |   {* Method.ctxt_args cring_normalise *}
 | 
| 13936 | 452 |   {* computes distributive normal form in locale context cring *}
 | 
| 453 | ||
| 454 | lemma (in cring) nat_pow_zero: | |
| 455 | "(n::nat) ~= 0 ==> \<zero> (^) n = \<zero>" | |
| 456 | by (induct n) simp_all | |
| 13854 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 457 | |
| 13864 | 458 | text {* Two examples for use of method algebra *}
 | 
| 459 | ||
| 13854 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 460 | lemma | 
| 14399 
dc677b35e54f
New lemmas about inversion of restricted functions.
 ballarin parents: 
14286diff
changeset | 461 | includes ring R + cring S | 
| 13854 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 462 | shows "[| a \<in> carrier R; b \<in> carrier R; c \<in> carrier S; d \<in> carrier S |] ==> | 
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 463 | a \<oplus> \<ominus> (a \<oplus> \<ominus> b) = b & c \<otimes>\<^bsub>S\<^esub> d = d \<otimes>\<^bsub>S\<^esub> c" | 
| 13854 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 464 | by algebra | 
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 465 | |
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 466 | lemma | 
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 467 | includes cring | 
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 468 | shows "[| a \<in> carrier R; b \<in> carrier R |] ==> a \<ominus> (a \<ominus> b) = b" | 
| 
91c9ab25fece
First distributed version of Group and Ring theory.
 ballarin parents: 
13835diff
changeset | 469 | by algebra | 
| 13835 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 470 | |
| 13864 | 471 | subsection {* Sums over Finite Sets *}
 | 
| 472 | ||
| 473 | lemma (in cring) finsum_ldistr: | |
| 474 | "[| finite A; a \<in> carrier R; f \<in> A -> carrier R |] ==> | |
| 475 | finsum R f A \<otimes> a = finsum R (%i. f i \<otimes> a) A" | |
| 476 | proof (induct set: Finites) | |
| 477 | case empty then show ?case by simp | |
| 478 | next | |
| 15328 | 479 | case (insert x F) then show ?case by (simp add: Pi_def l_distr) | 
| 13864 | 480 | qed | 
| 481 | ||
| 482 | lemma (in cring) finsum_rdistr: | |
| 483 | "[| finite A; a \<in> carrier R; f \<in> A -> carrier R |] ==> | |
| 484 | a \<otimes> finsum R f A = finsum R (%i. a \<otimes> f i) A" | |
| 485 | proof (induct set: Finites) | |
| 486 | case empty then show ?case by simp | |
| 487 | next | |
| 15328 | 488 | case (insert x F) then show ?case by (simp add: Pi_def r_distr) | 
| 13864 | 489 | qed | 
| 490 | ||
| 491 | subsection {* Facts of Integral Domains *}
 | |
| 492 | ||
| 493 | lemma (in "domain") zero_not_one [simp]: | |
| 494 | "\<zero> ~= \<one>" | |
| 495 | by (rule not_sym) simp | |
| 496 | ||
| 497 | lemma (in "domain") integral_iff: (* not by default a simp rule! *) | |
| 498 | "[| a \<in> carrier R; b \<in> carrier R |] ==> (a \<otimes> b = \<zero>) = (a = \<zero> | b = \<zero>)" | |
| 499 | proof | |
| 500 | assume "a \<in> carrier R" "b \<in> carrier R" "a \<otimes> b = \<zero>" | |
| 501 | then show "a = \<zero> | b = \<zero>" by (simp add: integral) | |
| 502 | next | |
| 503 | assume "a \<in> carrier R" "b \<in> carrier R" "a = \<zero> | b = \<zero>" | |
| 504 | then show "a \<otimes> b = \<zero>" by auto | |
| 505 | qed | |
| 506 | ||
| 507 | lemma (in "domain") m_lcancel: | |
| 508 | assumes prem: "a ~= \<zero>" | |
| 509 | and R: "a \<in> carrier R" "b \<in> carrier R" "c \<in> carrier R" | |
| 510 | shows "(a \<otimes> b = a \<otimes> c) = (b = c)" | |
| 511 | proof | |
| 512 | assume eq: "a \<otimes> b = a \<otimes> c" | |
| 513 | with R have "a \<otimes> (b \<ominus> c) = \<zero>" by algebra | |
| 514 | with R have "a = \<zero> | (b \<ominus> c) = \<zero>" by (simp add: integral_iff) | |
| 515 | with prem and R have "b \<ominus> c = \<zero>" by auto | |
| 516 | with R have "b = b \<ominus> (b \<ominus> c)" by algebra | |
| 517 | also from R have "b \<ominus> (b \<ominus> c) = c" by algebra | |
| 518 | finally show "b = c" . | |
| 519 | next | |
| 520 | assume "b = c" then show "a \<otimes> b = a \<otimes> c" by simp | |
| 521 | qed | |
| 522 | ||
| 523 | lemma (in "domain") m_rcancel: | |
| 524 | assumes prem: "a ~= \<zero>" | |
| 525 | and R: "a \<in> carrier R" "b \<in> carrier R" "c \<in> carrier R" | |
| 526 | shows conc: "(b \<otimes> a = c \<otimes> a) = (b = c)" | |
| 527 | proof - | |
| 528 | from prem and R have "(a \<otimes> b = a \<otimes> c) = (b = c)" by (rule m_lcancel) | |
| 529 | with R show ?thesis by algebra | |
| 530 | qed | |
| 531 | ||
| 13936 | 532 | subsection {* Morphisms *}
 | 
| 533 | ||
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 534 | constdefs (structure R S) | 
| 13936 | 535 |   ring_hom :: "[('a, 'm) ring_scheme, ('b, 'n) ring_scheme] => ('a => 'b) set"
 | 
| 536 |   "ring_hom R S == {h. h \<in> carrier R -> carrier S &
 | |
| 537 | (ALL x y. x \<in> carrier R & y \<in> carrier R --> | |
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 538 | h (x \<otimes> y) = h x \<otimes>\<^bsub>S\<^esub> h y & h (x \<oplus> y) = h x \<oplus>\<^bsub>S\<^esub> h y) & | 
| 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 539 | h \<one> = \<one>\<^bsub>S\<^esub>}" | 
| 13936 | 540 | |
| 541 | lemma ring_hom_memI: | |
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 542 | includes struct R + struct S | 
| 13936 | 543 | assumes hom_closed: "!!x. x \<in> carrier R ==> h x \<in> carrier S" | 
| 544 | and hom_mult: "!!x y. [| x \<in> carrier R; y \<in> carrier R |] ==> | |
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 545 | h (x \<otimes> y) = h x \<otimes>\<^bsub>S\<^esub> h y" | 
| 13936 | 546 | and hom_add: "!!x y. [| x \<in> carrier R; y \<in> carrier R |] ==> | 
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 547 | h (x \<oplus> y) = h x \<oplus>\<^bsub>S\<^esub> h y" | 
| 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 548 | and hom_one: "h \<one> = \<one>\<^bsub>S\<^esub>" | 
| 13936 | 549 | shows "h \<in> ring_hom R S" | 
| 550 | by (auto simp add: ring_hom_def prems Pi_def) | |
| 551 | ||
| 552 | lemma ring_hom_closed: | |
| 553 | "[| h \<in> ring_hom R S; x \<in> carrier R |] ==> h x \<in> carrier S" | |
| 554 | by (auto simp add: ring_hom_def funcset_mem) | |
| 555 | ||
| 556 | lemma ring_hom_mult: | |
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 557 | includes struct R + struct S | 
| 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 558 | shows | 
| 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 559 | "[| h \<in> ring_hom R S; x \<in> carrier R; y \<in> carrier R |] ==> | 
| 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 560 | h (x \<otimes> y) = h x \<otimes>\<^bsub>S\<^esub> h y" | 
| 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 561 | by (simp add: ring_hom_def) | 
| 13936 | 562 | |
| 563 | lemma ring_hom_add: | |
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 564 | includes struct R + struct S | 
| 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 565 | shows | 
| 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 566 | "[| h \<in> ring_hom R S; x \<in> carrier R; y \<in> carrier R |] ==> | 
| 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 567 | h (x \<oplus> y) = h x \<oplus>\<^bsub>S\<^esub> h y" | 
| 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 568 | by (simp add: ring_hom_def) | 
| 13936 | 569 | |
| 570 | lemma ring_hom_one: | |
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 571 | includes struct R + struct S | 
| 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 572 | shows "h \<in> ring_hom R S ==> h \<one> = \<one>\<^bsub>S\<^esub>" | 
| 13936 | 573 | by (simp add: ring_hom_def) | 
| 574 | ||
| 575 | locale ring_hom_cring = cring R + cring S + var h + | |
| 576 | assumes homh [simp, intro]: "h \<in> ring_hom R S" | |
| 577 | notes hom_closed [simp, intro] = ring_hom_closed [OF homh] | |
| 578 | and hom_mult [simp] = ring_hom_mult [OF homh] | |
| 579 | and hom_add [simp] = ring_hom_add [OF homh] | |
| 580 | and hom_one [simp] = ring_hom_one [OF homh] | |
| 581 | ||
| 582 | lemma (in ring_hom_cring) hom_zero [simp]: | |
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 583 | "h \<zero> = \<zero>\<^bsub>S\<^esub>" | 
| 13936 | 584 | proof - | 
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 585 | have "h \<zero> \<oplus>\<^bsub>S\<^esub> h \<zero> = h \<zero> \<oplus>\<^bsub>S\<^esub> \<zero>\<^bsub>S\<^esub>" | 
| 13936 | 586 | by (simp add: hom_add [symmetric] del: hom_add) | 
| 587 | then show ?thesis by (simp del: S.r_zero) | |
| 588 | qed | |
| 589 | ||
| 590 | lemma (in ring_hom_cring) hom_a_inv [simp]: | |
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 591 | "x \<in> carrier R ==> h (\<ominus> x) = \<ominus>\<^bsub>S\<^esub> h x" | 
| 13936 | 592 | proof - | 
| 593 | assume R: "x \<in> carrier R" | |
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14963diff
changeset | 594 | then have "h x \<oplus>\<^bsub>S\<^esub> h (\<ominus> x) = h x \<oplus>\<^bsub>S\<^esub> (\<ominus>\<^bsub>S\<^esub> h x)" | 
| 13936 | 595 | by (simp add: hom_add [symmetric] R.r_neg S.r_neg del: hom_add) | 
| 596 | with R show ?thesis by simp | |
| 597 | qed | |
| 598 | ||
| 599 | lemma (in ring_hom_cring) hom_finsum [simp]: | |
| 600 | "[| finite A; f \<in> A -> carrier R |] ==> | |
| 601 | h (finsum R f A) = finsum S (h o f) A" | |
| 602 | proof (induct set: Finites) | |
| 603 | case empty then show ?case by simp | |
| 604 | next | |
| 605 | case insert then show ?case by (simp add: Pi_def) | |
| 606 | qed | |
| 607 | ||
| 608 | lemma (in ring_hom_cring) hom_finprod: | |
| 609 | "[| finite A; f \<in> A -> carrier R |] ==> | |
| 610 | h (finprod R f A) = finprod S (h o f) A" | |
| 611 | proof (induct set: Finites) | |
| 612 | case empty then show ?case by simp | |
| 613 | next | |
| 614 | case insert then show ?case by (simp add: Pi_def) | |
| 615 | qed | |
| 616 | ||
| 617 | declare ring_hom_cring.hom_finprod [simp] | |
| 618 | ||
| 619 | lemma id_ring_hom [simp]: | |
| 620 | "id \<in> ring_hom R R" | |
| 621 | by (auto intro!: ring_hom_memI) | |
| 622 | ||
| 13835 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: diff
changeset | 623 | end |