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