| author | bulwahn | 
| Thu, 07 Apr 2011 14:51:26 +0200 | |
| changeset 42274 | 50850486f8dc | 
| parent 41959 | b460124855b8 | 
| child 44677 | 3fb27b19e058 | 
| permissions | -rw-r--r-- | 
| 41959 | 1  | 
(* Title: HOL/Algebra/Ring.thy  | 
| 35849 | 2  | 
Author: Clemens Ballarin, started 9 December 1996  | 
3  | 
Copyright: Clemens Ballarin  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
4  | 
*)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
5  | 
|
| 28823 | 6  | 
theory Ring  | 
7  | 
imports FiniteProduct  | 
|
| 35847 | 8  | 
uses ("ringsimp.ML")
 | 
9  | 
begin  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
10  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
11  | 
section {* The Algebraic Hierarchy of Rings *}
 | 
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
12  | 
|
| 
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
13  | 
subsection {* Abelian Groups *}
 | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
14  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
15  | 
record 'a ring = "'a monoid" +  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
16  | 
  zero :: 'a ("\<zero>\<index>")
 | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
17  | 
add :: "['a, 'a] => 'a" (infixl "\<oplus>\<index>" 65)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
18  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
19  | 
text {* Derived operations. *}
 | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
20  | 
|
| 35847 | 21  | 
definition  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
22  | 
  a_inv :: "[('a, 'm) ring_scheme, 'a ] => 'a" ("\<ominus>\<index> _" [81] 80)
 | 
| 
35848
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
35847 
diff
changeset
 | 
23  | 
where "a_inv R = m_inv (| carrier = carrier R, mult = add R, one = zero R |)"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
24  | 
|
| 35847 | 25  | 
definition  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
26  | 
  a_minus :: "[('a, 'm) ring_scheme, 'a, 'a] => 'a" (infixl "\<ominus>\<index>" 65)
 | 
| 
35848
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
35847 
diff
changeset
 | 
27  | 
where "[| x \<in> carrier R; y \<in> carrier R |] ==> x \<ominus>\<^bsub>R\<^esub> y = x \<oplus>\<^bsub>R\<^esub> (\<ominus>\<^bsub>R\<^esub> y)"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
28  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
29  | 
locale abelian_monoid =  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
30  | 
fixes G (structure)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
31  | 
assumes a_comm_monoid:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
32  | 
"comm_monoid (| carrier = carrier G, mult = add G, one = zero G |)"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
33  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
34  | 
definition  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
35  | 
  finsum :: "[('b, 'm) ring_scheme, 'a => 'b, 'a set] => 'b" where
 | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
36  | 
"finsum G = finprod (| carrier = carrier G, mult = add G, one = zero G |)"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
37  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
38  | 
syntax  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
39  | 
"_finsum" :: "index => idt => 'a set => 'b => 'b"  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
40  | 
      ("(3\<Oplus>__:_. _)" [1000, 0, 51, 10] 10)
 | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
41  | 
syntax (xsymbols)  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
42  | 
"_finsum" :: "index => idt => 'a set => 'b => 'b"  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
43  | 
      ("(3\<Oplus>__\<in>_. _)" [1000, 0, 51, 10] 10)
 | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
44  | 
syntax (HTML output)  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
45  | 
"_finsum" :: "index => idt => 'a set => 'b => 'b"  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
46  | 
      ("(3\<Oplus>__\<in>_. _)" [1000, 0, 51, 10] 10)
 | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
47  | 
translations  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
48  | 
"\<Oplus>\<index>i:A. b" == "CONST finsum \<struct>\<index> (%i. b) A"  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
49  | 
  -- {* Beware of argument permutation! *}
 | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
50  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
51  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
52  | 
locale abelian_group = abelian_monoid +  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
53  | 
assumes a_comm_group:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
54  | 
"comm_group (| carrier = carrier G, mult = add G, one = zero G |)"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
55  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
56  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
57  | 
subsection {* Basic Properties *}
 | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
58  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
59  | 
lemma abelian_monoidI:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
60  | 
fixes R (structure)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
61  | 
assumes a_closed:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
62  | 
"!!x y. [| x \<in> carrier R; y \<in> carrier R |] ==> x \<oplus> y \<in> carrier R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
63  | 
and zero_closed: "\<zero> \<in> carrier R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
64  | 
and a_assoc:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
65  | 
"!!x y z. [| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |] ==>  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
66  | 
(x \<oplus> y) \<oplus> z = x \<oplus> (y \<oplus> z)"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
67  | 
and l_zero: "!!x. x \<in> carrier R ==> \<zero> \<oplus> x = x"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
68  | 
and a_comm:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
69  | 
"!!x y. [| x \<in> carrier R; y \<in> carrier R |] ==> x \<oplus> y = y \<oplus> x"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
70  | 
shows "abelian_monoid R"  | 
| 
27714
 
27b4d7c01f8b
Tuned (for the sake of a meaningless log entry).
 
ballarin 
parents: 
27699 
diff
changeset
 | 
71  | 
by (auto intro!: abelian_monoid.intro comm_monoidI intro: assms)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
72  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
73  | 
lemma abelian_groupI:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
74  | 
fixes R (structure)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
75  | 
assumes a_closed:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
76  | 
"!!x y. [| x \<in> carrier R; y \<in> carrier R |] ==> x \<oplus> y \<in> carrier R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
77  | 
and zero_closed: "zero R \<in> carrier R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
78  | 
and a_assoc:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
79  | 
"!!x y z. [| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |] ==>  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
80  | 
(x \<oplus> y) \<oplus> z = x \<oplus> (y \<oplus> z)"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
81  | 
and a_comm:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
82  | 
"!!x y. [| x \<in> carrier R; y \<in> carrier R |] ==> x \<oplus> y = y \<oplus> x"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
83  | 
and l_zero: "!!x. x \<in> carrier R ==> \<zero> \<oplus> x = x"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
84  | 
and l_inv_ex: "!!x. x \<in> carrier R ==> EX y : carrier R. y \<oplus> x = \<zero>"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
85  | 
shows "abelian_group R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
86  | 
by (auto intro!: abelian_group.intro abelian_monoidI  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
87  | 
abelian_group_axioms.intro comm_monoidI comm_groupI  | 
| 
27714
 
27b4d7c01f8b
Tuned (for the sake of a meaningless log entry).
 
ballarin 
parents: 
27699 
diff
changeset
 | 
88  | 
intro: assms)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
89  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
90  | 
lemma (in abelian_monoid) a_monoid:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
91  | 
"monoid (| carrier = carrier G, mult = add G, one = zero G |)"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
92  | 
by (rule comm_monoid.axioms, rule a_comm_monoid)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
93  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
94  | 
lemma (in abelian_group) a_group:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
95  | 
"group (| carrier = carrier G, mult = add G, one = zero G |)"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
96  | 
by (simp add: group_def a_monoid)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
97  | 
(simp add: comm_group.axioms group.axioms a_comm_group)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
98  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
99  | 
lemmas monoid_record_simps = partial_object.simps monoid.simps  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
100  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
101  | 
text {* Transfer facts from multiplicative structures via interpretation. *}
 | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
102  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
103  | 
sublocale abelian_monoid <  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
104  | 
add!: monoid "(| carrier = carrier G, mult = add G, one = zero G |)"  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
105  | 
where "carrier (| carrier = carrier G, mult = add G, one = zero G |) = carrier G"  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
106  | 
and "mult (| carrier = carrier G, mult = add G, one = zero G |) = add G"  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
107  | 
and "one (| carrier = carrier G, mult = add G, one = zero G |) = zero G"  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
108  | 
by (rule a_monoid) auto  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
109  | 
|
| 27933 | 110  | 
context abelian_monoid begin  | 
111  | 
||
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
112  | 
lemmas a_closed = add.m_closed  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
113  | 
lemmas zero_closed = add.one_closed  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
114  | 
lemmas a_assoc = add.m_assoc  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
115  | 
lemmas l_zero = add.l_one  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
116  | 
lemmas r_zero = add.r_one  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
117  | 
lemmas minus_unique = add.inv_unique  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
118  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
119  | 
end  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
120  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
121  | 
sublocale abelian_monoid <  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
122  | 
add!: comm_monoid "(| carrier = carrier G, mult = add G, one = zero G |)"  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
123  | 
where "carrier (| carrier = carrier G, mult = add G, one = zero G |) = carrier G"  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
124  | 
and "mult (| carrier = carrier G, mult = add G, one = zero G |) = add G"  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
125  | 
and "one (| carrier = carrier G, mult = add G, one = zero G |) = zero G"  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
126  | 
and "finprod (| carrier = carrier G, mult = add G, one = zero G |) = finsum G"  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
127  | 
by (rule a_comm_monoid) (auto simp: finsum_def)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
128  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
129  | 
context abelian_monoid begin  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
130  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
131  | 
lemmas a_comm = add.m_comm  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
132  | 
lemmas a_lcomm = add.m_lcomm  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
133  | 
lemmas a_ac = a_assoc a_comm a_lcomm  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
134  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
135  | 
lemmas finsum_empty = add.finprod_empty  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
136  | 
lemmas finsum_insert = add.finprod_insert  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
137  | 
lemmas finsum_zero = add.finprod_one  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
138  | 
lemmas finsum_closed = add.finprod_closed  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
139  | 
lemmas finsum_Un_Int = add.finprod_Un_Int  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
140  | 
lemmas finsum_Un_disjoint = add.finprod_Un_disjoint  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
141  | 
lemmas finsum_addf = add.finprod_multf  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
142  | 
lemmas finsum_cong' = add.finprod_cong'  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
143  | 
lemmas finsum_0 = add.finprod_0  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
144  | 
lemmas finsum_Suc = add.finprod_Suc  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
145  | 
lemmas finsum_Suc2 = add.finprod_Suc2  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
146  | 
lemmas finsum_add = add.finprod_mult  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
147  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
148  | 
lemmas finsum_cong = add.finprod_cong  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
149  | 
text {*Usually, if this rule causes a failed congruence proof error,
 | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
150  | 
   the reason is that the premise @{text "g \<in> B -> carrier G"} cannot be shown.
 | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
151  | 
   Adding @{thm [source] Pi_def} to the simpset is often useful. *}
 | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
152  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
153  | 
lemmas finsum_reindex = add.finprod_reindex  | 
| 27699 | 154  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
155  | 
(* The following would be wrong. Needed is the equivalent of (^) for addition,  | 
| 27699 | 156  | 
or indeed the canonical embedding from Nat into the monoid.  | 
157  | 
||
| 27933 | 158  | 
lemma finsum_const:  | 
| 27699 | 159  | 
assumes fin [simp]: "finite A"  | 
160  | 
and a [simp]: "a : carrier G"  | 
|
161  | 
shows "finsum G (%x. a) A = a (^) card A"  | 
|
162  | 
using fin apply induct  | 
|
163  | 
apply force  | 
|
164  | 
apply (subst finsum_insert)  | 
|
165  | 
apply auto  | 
|
166  | 
apply (force simp add: Pi_def)  | 
|
167  | 
apply (subst m_comm)  | 
|
168  | 
apply auto  | 
|
169  | 
done  | 
|
170  | 
*)  | 
|
171  | 
||
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
172  | 
lemmas finsum_singleton = add.finprod_singleton  | 
| 27933 | 173  | 
|
174  | 
end  | 
|
175  | 
||
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
176  | 
sublocale abelian_group <  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
177  | 
add!: group "(| carrier = carrier G, mult = add G, one = zero G |)"  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
178  | 
where "carrier (| carrier = carrier G, mult = add G, one = zero G |) = carrier G"  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
179  | 
and "mult (| carrier = carrier G, mult = add G, one = zero G |) = add G"  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
180  | 
and "one (| carrier = carrier G, mult = add G, one = zero G |) = zero G"  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
181  | 
and "m_inv (| carrier = carrier G, mult = add G, one = zero G |) = a_inv G"  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
182  | 
by (rule a_group) (auto simp: m_inv_def a_inv_def)  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
183  | 
|
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
184  | 
context abelian_group begin  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
185  | 
|
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
186  | 
lemmas a_inv_closed = add.inv_closed  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
187  | 
|
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
188  | 
lemma minus_closed [intro, simp]:  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
189  | 
"[| x \<in> carrier G; y \<in> carrier G |] ==> x \<ominus> y \<in> carrier G"  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
190  | 
by (simp add: a_minus_def)  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
191  | 
|
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
192  | 
lemmas a_l_cancel = add.l_cancel  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
193  | 
lemmas a_r_cancel = add.r_cancel  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
194  | 
lemmas l_neg = add.l_inv [simp del]  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
195  | 
lemmas r_neg = add.r_inv [simp del]  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
196  | 
lemmas minus_zero = add.inv_one  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
197  | 
lemmas minus_minus = add.inv_inv  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
198  | 
lemmas a_inv_inj = add.inv_inj  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
199  | 
lemmas minus_equality = add.inv_equality  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
200  | 
|
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
201  | 
end  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
202  | 
|
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
203  | 
sublocale abelian_group <  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
204  | 
add!: comm_group "(| carrier = carrier G, mult = add G, one = zero G |)"  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
205  | 
where "carrier (| carrier = carrier G, mult = add G, one = zero G |) = carrier G"  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
206  | 
and "mult (| carrier = carrier G, mult = add G, one = zero G |) = add G"  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
207  | 
and "one (| carrier = carrier G, mult = add G, one = zero G |) = zero G"  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
208  | 
and "m_inv (| carrier = carrier G, mult = add G, one = zero G |) = a_inv G"  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
209  | 
and "finprod (| carrier = carrier G, mult = add G, one = zero G |) = finsum G"  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
210  | 
by (rule a_comm_group) (auto simp: m_inv_def a_inv_def finsum_def)  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
211  | 
|
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
212  | 
lemmas (in abelian_group) minus_add = add.inv_mult  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
213  | 
|
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
214  | 
text {* Derive an @{text "abelian_group"} from a @{text "comm_group"} *}
 | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
215  | 
|
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
216  | 
lemma comm_group_abelian_groupI:  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
217  | 
fixes G (structure)  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
218  | 
assumes cg: "comm_group \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
219  | 
shows "abelian_group G"  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
220  | 
proof -  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
221  | 
interpret comm_group "\<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
222  | 
by (rule cg)  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
223  | 
show "abelian_group G" ..  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
224  | 
qed  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
225  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
226  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
227  | 
subsection {* Rings: Basic Definitions *}
 | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
228  | 
|
| 29237 | 229  | 
locale ring = abelian_group R + monoid R for R (structure) +  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
230  | 
assumes l_distr: "[| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |]  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
231  | 
==> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
232  | 
and r_distr: "[| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |]  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
233  | 
==> z \<otimes> (x \<oplus> y) = z \<otimes> x \<oplus> z \<otimes> y"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
234  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
235  | 
locale cring = ring + comm_monoid R  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
236  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
237  | 
locale "domain" = cring +  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
238  | 
assumes one_not_zero [simp]: "\<one> ~= \<zero>"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
239  | 
and integral: "[| a \<otimes> b = \<zero>; a \<in> carrier R; b \<in> carrier R |] ==>  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
240  | 
a = \<zero> | b = \<zero>"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
241  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
242  | 
locale field = "domain" +  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
243  | 
  assumes field_Units: "Units R = carrier R - {\<zero>}"
 | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
244  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
245  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
246  | 
subsection {* Rings *}
 | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
247  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
248  | 
lemma ringI:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
249  | 
fixes R (structure)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
250  | 
assumes abelian_group: "abelian_group R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
251  | 
and monoid: "monoid R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
252  | 
and l_distr: "!!x y z. [| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |]  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
253  | 
==> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
254  | 
and r_distr: "!!x y z. [| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |]  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
255  | 
==> z \<otimes> (x \<oplus> y) = z \<otimes> x \<oplus> z \<otimes> y"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
256  | 
shows "ring R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
257  | 
by (auto intro: ring.intro  | 
| 
27714
 
27b4d7c01f8b
Tuned (for the sake of a meaningless log entry).
 
ballarin 
parents: 
27699 
diff
changeset
 | 
258  | 
abelian_group.axioms ring_axioms.intro assms)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
259  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
260  | 
context ring begin  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
261  | 
|
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
262  | 
lemma is_abelian_group:  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
263  | 
"abelian_group R"  | 
| 28823 | 264  | 
..  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
265  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
266  | 
lemma is_monoid:  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
267  | 
"monoid R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
268  | 
by (auto intro!: monoidI m_assoc)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
269  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
270  | 
lemma is_ring:  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
271  | 
"ring R"  | 
| 26202 | 272  | 
by (rule ring_axioms)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
273  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
274  | 
end  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
275  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
276  | 
lemmas ring_record_simps = monoid_record_simps ring.simps  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
277  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
278  | 
lemma cringI:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
279  | 
fixes R (structure)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
280  | 
assumes abelian_group: "abelian_group R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
281  | 
and comm_monoid: "comm_monoid R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
282  | 
and l_distr: "!!x y z. [| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |]  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
283  | 
==> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
284  | 
shows "cring R"  | 
| 23350 | 285  | 
proof (intro cring.intro ring.intro)  | 
286  | 
show "ring_axioms R"  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
287  | 
    -- {* Right-distributivity follows from left-distributivity and
 | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
288  | 
commutativity. *}  | 
| 23350 | 289  | 
proof (rule ring_axioms.intro)  | 
290  | 
fix x y z  | 
|
291  | 
assume R: "x \<in> carrier R" "y \<in> carrier R" "z \<in> carrier R"  | 
|
292  | 
note [simp] = comm_monoid.axioms [OF comm_monoid]  | 
|
293  | 
abelian_group.axioms [OF abelian_group]  | 
|
294  | 
abelian_monoid.a_closed  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
295  | 
|
| 23350 | 296  | 
from R have "z \<otimes> (x \<oplus> y) = (x \<oplus> y) \<otimes> z"  | 
297  | 
by (simp add: comm_monoid.m_comm [OF comm_monoid.intro])  | 
|
298  | 
also from R have "... = x \<otimes> z \<oplus> y \<otimes> z" by (simp add: l_distr)  | 
|
299  | 
also from R have "... = z \<otimes> x \<oplus> z \<otimes> y"  | 
|
300  | 
by (simp add: comm_monoid.m_comm [OF comm_monoid.intro])  | 
|
301  | 
finally show "z \<otimes> (x \<oplus> y) = z \<otimes> x \<oplus> z \<otimes> y" .  | 
|
302  | 
qed (rule l_distr)  | 
|
303  | 
qed (auto intro: cring.intro  | 
|
| 
27714
 
27b4d7c01f8b
Tuned (for the sake of a meaningless log entry).
 
ballarin 
parents: 
27699 
diff
changeset
 | 
304  | 
abelian_group.axioms comm_monoid.axioms ring_axioms.intro assms)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
305  | 
|
| 27699 | 306  | 
(*  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
307  | 
lemma (in cring) is_comm_monoid:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
308  | 
"comm_monoid R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
309  | 
by (auto intro!: comm_monoidI m_assoc m_comm)  | 
| 27699 | 310  | 
*)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
311  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
312  | 
lemma (in cring) is_cring:  | 
| 26202 | 313  | 
"cring R" by (rule cring_axioms)  | 
| 23350 | 314  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
315  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
316  | 
subsubsection {* Normaliser for Rings *}
 | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
317  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
318  | 
lemma (in abelian_group) r_neg2:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
319  | 
"[| x \<in> carrier G; y \<in> carrier G |] ==> x \<oplus> (\<ominus> x \<oplus> y) = y"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
320  | 
proof -  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
321  | 
assume G: "x \<in> carrier G" "y \<in> carrier G"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
322  | 
then have "(x \<oplus> \<ominus> x) \<oplus> y = y"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
323  | 
by (simp only: r_neg l_zero)  | 
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
324  | 
with G show ?thesis  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
325  | 
by (simp add: a_ac)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
326  | 
qed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
327  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
328  | 
lemma (in abelian_group) r_neg1:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
329  | 
"[| x \<in> carrier G; y \<in> carrier G |] ==> \<ominus> x \<oplus> (x \<oplus> y) = y"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
330  | 
proof -  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
331  | 
assume G: "x \<in> carrier G" "y \<in> carrier G"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
332  | 
then have "(\<ominus> x \<oplus> x) \<oplus> y = y"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
333  | 
by (simp only: l_neg l_zero)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
334  | 
with G show ?thesis by (simp add: a_ac)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
335  | 
qed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
336  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
337  | 
context ring begin  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
338  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
339  | 
text {* 
 | 
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
340  | 
The following proofs are from Jacobson, Basic Algebra I, pp.~88--89.  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
341  | 
*}  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
342  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
343  | 
lemma l_null [simp]:  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
344  | 
"x \<in> carrier R ==> \<zero> \<otimes> x = \<zero>"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
345  | 
proof -  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
346  | 
assume R: "x \<in> carrier R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
347  | 
then have "\<zero> \<otimes> x \<oplus> \<zero> \<otimes> x = (\<zero> \<oplus> \<zero>) \<otimes> x"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
348  | 
by (simp add: l_distr del: l_zero r_zero)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
349  | 
also from R have "... = \<zero> \<otimes> x \<oplus> \<zero>" by simp  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
350  | 
finally have "\<zero> \<otimes> x \<oplus> \<zero> \<otimes> x = \<zero> \<otimes> x \<oplus> \<zero>" .  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
351  | 
with R show ?thesis by (simp del: r_zero)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
352  | 
qed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
353  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
354  | 
lemma r_null [simp]:  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
355  | 
"x \<in> carrier R ==> x \<otimes> \<zero> = \<zero>"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
356  | 
proof -  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
357  | 
assume R: "x \<in> carrier R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
358  | 
then have "x \<otimes> \<zero> \<oplus> x \<otimes> \<zero> = x \<otimes> (\<zero> \<oplus> \<zero>)"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
359  | 
by (simp add: r_distr del: l_zero r_zero)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
360  | 
also from R have "... = x \<otimes> \<zero> \<oplus> \<zero>" by simp  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
361  | 
finally have "x \<otimes> \<zero> \<oplus> x \<otimes> \<zero> = x \<otimes> \<zero> \<oplus> \<zero>" .  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
362  | 
with R show ?thesis by (simp del: r_zero)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
363  | 
qed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
364  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
365  | 
lemma l_minus:  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
366  | 
"[| x \<in> carrier R; y \<in> carrier R |] ==> \<ominus> x \<otimes> y = \<ominus> (x \<otimes> y)"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
367  | 
proof -  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
368  | 
assume R: "x \<in> carrier R" "y \<in> carrier R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
369  | 
then have "(\<ominus> x) \<otimes> y \<oplus> x \<otimes> y = (\<ominus> x \<oplus> x) \<otimes> y" by (simp add: l_distr)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
370  | 
also from R have "... = \<zero>" by (simp add: l_neg l_null)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
371  | 
finally have "(\<ominus> x) \<otimes> y \<oplus> x \<otimes> y = \<zero>" .  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
372  | 
with R have "(\<ominus> x) \<otimes> y \<oplus> x \<otimes> y \<oplus> \<ominus> (x \<otimes> y) = \<zero> \<oplus> \<ominus> (x \<otimes> y)" by simp  | 
| 
21896
 
9a7949815a84
Experimenting with interpretations of "definition".
 
ballarin 
parents: 
20318 
diff
changeset
 | 
373  | 
with R show ?thesis by (simp add: a_assoc r_neg)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
374  | 
qed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
375  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
376  | 
lemma r_minus:  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
377  | 
"[| x \<in> carrier R; y \<in> carrier R |] ==> x \<otimes> \<ominus> y = \<ominus> (x \<otimes> y)"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
378  | 
proof -  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
379  | 
assume R: "x \<in> carrier R" "y \<in> carrier R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
380  | 
then have "x \<otimes> (\<ominus> y) \<oplus> x \<otimes> y = x \<otimes> (\<ominus> y \<oplus> y)" by (simp add: r_distr)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
381  | 
also from R have "... = \<zero>" by (simp add: l_neg r_null)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
382  | 
finally have "x \<otimes> (\<ominus> y) \<oplus> x \<otimes> y = \<zero>" .  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
383  | 
with R have "x \<otimes> (\<ominus> y) \<oplus> x \<otimes> y \<oplus> \<ominus> (x \<otimes> y) = \<zero> \<oplus> \<ominus> (x \<otimes> y)" by simp  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
384  | 
with R show ?thesis by (simp add: a_assoc r_neg )  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
385  | 
qed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
386  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
387  | 
end  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
388  | 
|
| 
23957
 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 
ballarin 
parents: 
23350 
diff
changeset
 | 
389  | 
lemma (in abelian_group) minus_eq:  | 
| 
 
54fab60ddc97
Interpretation of rings (as integers) maps defined operations to defined
 
ballarin 
parents: 
23350 
diff
changeset
 | 
390  | 
"[| x \<in> carrier G; y \<in> carrier G |] ==> x \<ominus> y = x \<oplus> \<ominus> y"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
391  | 
by (simp only: a_minus_def)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
392  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
393  | 
text {* Setup algebra method:
 | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
394  | 
compute distributive normal form in locale contexts *}  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
395  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
396  | 
use "ringsimp.ML"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
397  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
398  | 
setup Algebra.setup  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
399  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
400  | 
lemmas (in ring) ring_simprules  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
401  | 
[algebra ring "zero R" "add R" "a_inv R" "a_minus R" "one R" "mult R"] =  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
402  | 
a_closed zero_closed a_inv_closed minus_closed m_closed one_closed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
403  | 
a_assoc l_zero l_neg a_comm m_assoc l_one l_distr minus_eq  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
404  | 
r_zero r_neg r_neg2 r_neg1 minus_add minus_minus minus_zero  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
405  | 
a_lcomm r_distr l_null r_null l_minus r_minus  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
406  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
407  | 
lemmas (in cring)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
408  | 
[algebra del: ring "zero R" "add R" "a_inv R" "a_minus R" "one R" "mult R"] =  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
409  | 
_  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
410  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
411  | 
lemmas (in cring) cring_simprules  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
412  | 
[algebra add: cring "zero R" "add R" "a_inv R" "a_minus R" "one R" "mult R"] =  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
413  | 
a_closed zero_closed a_inv_closed minus_closed m_closed one_closed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
414  | 
a_assoc l_zero l_neg a_comm m_assoc l_one l_distr m_comm minus_eq  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
415  | 
r_zero r_neg r_neg2 r_neg1 minus_add minus_minus minus_zero  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
416  | 
a_lcomm m_lcomm r_distr l_null r_null l_minus r_minus  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
417  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
418  | 
lemma (in cring) nat_pow_zero:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
419  | 
"(n::nat) ~= 0 ==> \<zero> (^) n = \<zero>"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
420  | 
by (induct n) simp_all  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
421  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
422  | 
context ring begin  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
423  | 
|
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
424  | 
lemma one_zeroD:  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
425  | 
assumes onezero: "\<one> = \<zero>"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
426  | 
  shows "carrier R = {\<zero>}"
 | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
427  | 
proof (rule, rule)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
428  | 
fix x  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
429  | 
assume xcarr: "x \<in> carrier R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
430  | 
from xcarr  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
431  | 
have "x = x \<otimes> \<one>" by simp  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
432  | 
from this and onezero  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
433  | 
have "x = x \<otimes> \<zero>" by simp  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
434  | 
from this and xcarr  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
435  | 
have "x = \<zero>" by simp  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
436  | 
  thus "x \<in> {\<zero>}" by fast
 | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
437  | 
qed fast  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
438  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
439  | 
lemma one_zeroI:  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
440  | 
  assumes carrzero: "carrier R = {\<zero>}"
 | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
441  | 
shows "\<one> = \<zero>"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
442  | 
proof -  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
443  | 
from one_closed and carrzero  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
444  | 
show "\<one> = \<zero>" by simp  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
445  | 
qed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
446  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
447  | 
lemma carrier_one_zero:  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
448  | 
  shows "(carrier R = {\<zero>}) = (\<one> = \<zero>)"
 | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
449  | 
by (rule, erule one_zeroI, erule one_zeroD)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
450  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
451  | 
lemma carrier_one_not_zero:  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
452  | 
  shows "(carrier R \<noteq> {\<zero>}) = (\<one> \<noteq> \<zero>)"
 | 
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
453  | 
by (simp add: carrier_one_zero)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
454  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
455  | 
end  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
456  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
457  | 
text {* Two examples for use of method algebra *}
 | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
458  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
459  | 
lemma  | 
| 27611 | 460  | 
fixes R (structure) and S (structure)  | 
461  | 
assumes "ring R" "cring S"  | 
|
462  | 
assumes RS: "a \<in> carrier R" "b \<in> carrier R" "c \<in> carrier S" "d \<in> carrier S"  | 
|
463  | 
shows "a \<oplus> \<ominus> (a \<oplus> \<ominus> b) = b & c \<otimes>\<^bsub>S\<^esub> d = d \<otimes>\<^bsub>S\<^esub> c"  | 
|
464  | 
proof -  | 
|
| 29237 | 465  | 
interpret ring R by fact  | 
466  | 
interpret cring S by fact  | 
|
| 27611 | 467  | 
ML_val {* Algebra.print_structures @{context} *}
 | 
468  | 
from RS show ?thesis by algebra  | 
|
469  | 
qed  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
470  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
471  | 
lemma  | 
| 27611 | 472  | 
fixes R (structure)  | 
473  | 
assumes "ring R"  | 
|
474  | 
assumes R: "a \<in> carrier R" "b \<in> carrier R"  | 
|
475  | 
shows "a \<ominus> (a \<ominus> b) = b"  | 
|
476  | 
proof -  | 
|
| 29237 | 477  | 
interpret ring R by fact  | 
| 27611 | 478  | 
from R show ?thesis by algebra  | 
479  | 
qed  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
480  | 
|
| 35849 | 481  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
482  | 
subsubsection {* Sums over Finite Sets *}
 | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
483  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
484  | 
lemma (in ring) finsum_ldistr:  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
485  | 
"[| finite A; a \<in> carrier R; f \<in> A -> carrier R |] ==>  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
486  | 
finsum R f A \<otimes> a = finsum R (%i. f i \<otimes> a) A"  | 
| 22265 | 487  | 
proof (induct set: finite)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
488  | 
case empty then show ?case by simp  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
489  | 
next  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
490  | 
case (insert x F) then show ?case by (simp add: Pi_def l_distr)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
491  | 
qed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
492  | 
|
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
493  | 
lemma (in ring) finsum_rdistr:  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
494  | 
"[| finite A; a \<in> carrier R; f \<in> A -> carrier R |] ==>  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
495  | 
a \<otimes> finsum R f A = finsum R (%i. a \<otimes> f i) A"  | 
| 22265 | 496  | 
proof (induct set: finite)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
497  | 
case empty then show ?case by simp  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
498  | 
next  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
499  | 
case (insert x F) then show ?case by (simp add: Pi_def r_distr)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
500  | 
qed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
501  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
502  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
503  | 
subsection {* Integral Domains *}
 | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
504  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
505  | 
context "domain" begin  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
506  | 
|
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
507  | 
lemma zero_not_one [simp]:  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
508  | 
"\<zero> ~= \<one>"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
509  | 
by (rule not_sym) simp  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
510  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
511  | 
lemma integral_iff: (* not by default a simp rule! *)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
512  | 
"[| a \<in> carrier R; b \<in> carrier R |] ==> (a \<otimes> b = \<zero>) = (a = \<zero> | b = \<zero>)"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
513  | 
proof  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
514  | 
assume "a \<in> carrier R" "b \<in> carrier R" "a \<otimes> b = \<zero>"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
515  | 
then show "a = \<zero> | b = \<zero>" by (simp add: integral)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
516  | 
next  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
517  | 
assume "a \<in> carrier R" "b \<in> carrier R" "a = \<zero> | b = \<zero>"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
518  | 
then show "a \<otimes> b = \<zero>" by auto  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
519  | 
qed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
520  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
521  | 
lemma m_lcancel:  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
522  | 
assumes prem: "a ~= \<zero>"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
523  | 
and R: "a \<in> carrier R" "b \<in> carrier R" "c \<in> carrier R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
524  | 
shows "(a \<otimes> b = a \<otimes> c) = (b = c)"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
525  | 
proof  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
526  | 
assume eq: "a \<otimes> b = a \<otimes> c"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
527  | 
with R have "a \<otimes> (b \<ominus> c) = \<zero>" by algebra  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
528  | 
with R have "a = \<zero> | (b \<ominus> c) = \<zero>" by (simp add: integral_iff)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
529  | 
with prem and R have "b \<ominus> c = \<zero>" by auto  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
530  | 
with R have "b = b \<ominus> (b \<ominus> c)" by algebra  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
531  | 
also from R have "b \<ominus> (b \<ominus> c) = c" by algebra  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
532  | 
finally show "b = c" .  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
533  | 
next  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
534  | 
assume "b = c" then show "a \<otimes> b = a \<otimes> c" by simp  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
535  | 
qed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
536  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
537  | 
lemma m_rcancel:  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
538  | 
assumes prem: "a ~= \<zero>"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
539  | 
and R: "a \<in> carrier R" "b \<in> carrier R" "c \<in> carrier R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
540  | 
shows conc: "(b \<otimes> a = c \<otimes> a) = (b = c)"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
541  | 
proof -  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
542  | 
from prem and R have "(a \<otimes> b = a \<otimes> c) = (b = c)" by (rule m_lcancel)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
543  | 
with R show ?thesis by algebra  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
544  | 
qed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
545  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
546  | 
end  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
547  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
548  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
549  | 
subsection {* Fields *}
 | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
550  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
551  | 
text {* Field would not need to be derived from domain, the properties
 | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
552  | 
for domain follow from the assumptions of field *}  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
553  | 
lemma (in cring) cring_fieldI:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
554  | 
  assumes field_Units: "Units R = carrier R - {\<zero>}"
 | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
555  | 
shows "field R"  | 
| 28823 | 556  | 
proof  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
557  | 
from field_Units  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
558  | 
have a: "\<zero> \<notin> Units R" by fast  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
559  | 
have "\<one> \<in> Units R" by fast  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
560  | 
from this and a  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
561  | 
show "\<one> \<noteq> \<zero>" by force  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
562  | 
next  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
563  | 
fix a b  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
564  | 
assume acarr: "a \<in> carrier R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
565  | 
and bcarr: "b \<in> carrier R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
566  | 
and ab: "a \<otimes> b = \<zero>"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
567  | 
show "a = \<zero> \<or> b = \<zero>"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
568  | 
proof (cases "a = \<zero>", simp)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
569  | 
assume "a \<noteq> \<zero>"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
570  | 
from this and field_Units and acarr  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
571  | 
have aUnit: "a \<in> Units R" by fast  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
572  | 
from bcarr  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
573  | 
have "b = \<one> \<otimes> b" by algebra  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
574  | 
also from aUnit acarr  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
575  | 
have "... = (inv a \<otimes> a) \<otimes> b" by (simp add: Units_l_inv)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
576  | 
also from acarr bcarr aUnit[THEN Units_inv_closed]  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
577  | 
have "... = (inv a) \<otimes> (a \<otimes> b)" by algebra  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
578  | 
also from ab and acarr bcarr aUnit  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
579  | 
have "... = (inv a) \<otimes> \<zero>" by simp  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
580  | 
also from aUnit[THEN Units_inv_closed]  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
581  | 
have "... = \<zero>" by algebra  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
582  | 
finally  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
583  | 
have "b = \<zero>" .  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
584  | 
thus "a = \<zero> \<or> b = \<zero>" by simp  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
585  | 
qed  | 
| 23350 | 586  | 
qed (rule field_Units)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
587  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
588  | 
text {* Another variant to show that something is a field *}
 | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
589  | 
lemma (in cring) cring_fieldI2:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
590  | 
assumes notzero: "\<zero> \<noteq> \<one>"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
591  | 
and invex: "\<And>a. \<lbrakk>a \<in> carrier R; a \<noteq> \<zero>\<rbrakk> \<Longrightarrow> \<exists>b\<in>carrier R. a \<otimes> b = \<one>"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
592  | 
shows "field R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
593  | 
apply (rule cring_fieldI, simp add: Units_def)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
594  | 
apply (rule, clarsimp)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
595  | 
apply (simp add: notzero)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
596  | 
proof (clarsimp)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
597  | 
fix x  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
598  | 
assume xcarr: "x \<in> carrier R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
599  | 
and "x \<noteq> \<zero>"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
600  | 
from this  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
601  | 
have "\<exists>y\<in>carrier R. x \<otimes> y = \<one>" by (rule invex)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
602  | 
from this  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
603  | 
obtain y  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
604  | 
where ycarr: "y \<in> carrier R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
605  | 
and xy: "x \<otimes> y = \<one>"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
606  | 
by fast  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
607  | 
from xy xcarr ycarr have "y \<otimes> x = \<one>" by (simp add: m_comm)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
608  | 
from ycarr and this and xy  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
609  | 
show "\<exists>y\<in>carrier R. y \<otimes> x = \<one> \<and> x \<otimes> y = \<one>" by fast  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
610  | 
qed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
611  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
612  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
613  | 
subsection {* Morphisms *}
 | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
614  | 
|
| 35847 | 615  | 
definition  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
616  | 
  ring_hom :: "[('a, 'm) ring_scheme, ('b, 'n) ring_scheme] => ('a => 'b) set"
 | 
| 
35848
 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 
wenzelm 
parents: 
35847 
diff
changeset
 | 
617  | 
where "ring_hom R S =  | 
| 35847 | 618  | 
    {h. h \<in> carrier R -> carrier S &
 | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
619  | 
(ALL x y. x \<in> carrier R & y \<in> carrier R -->  | 
| 35847 | 620  | 
h (x \<otimes>\<^bsub>R\<^esub> y) = h x \<otimes>\<^bsub>S\<^esub> h y & h (x \<oplus>\<^bsub>R\<^esub> y) = h x \<oplus>\<^bsub>S\<^esub> h y) &  | 
621  | 
h \<one>\<^bsub>R\<^esub> = \<one>\<^bsub>S\<^esub>}"  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
622  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
623  | 
lemma ring_hom_memI:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
624  | 
fixes R (structure) and S (structure)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
625  | 
assumes hom_closed: "!!x. x \<in> carrier R ==> h x \<in> carrier S"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
626  | 
and hom_mult: "!!x y. [| x \<in> carrier R; y \<in> carrier R |] ==>  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
627  | 
h (x \<otimes> y) = h x \<otimes>\<^bsub>S\<^esub> h y"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
628  | 
and hom_add: "!!x y. [| x \<in> carrier R; y \<in> carrier R |] ==>  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
629  | 
h (x \<oplus> y) = h x \<oplus>\<^bsub>S\<^esub> h y"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
630  | 
and hom_one: "h \<one> = \<one>\<^bsub>S\<^esub>"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
631  | 
shows "h \<in> ring_hom R S"  | 
| 
27714
 
27b4d7c01f8b
Tuned (for the sake of a meaningless log entry).
 
ballarin 
parents: 
27699 
diff
changeset
 | 
632  | 
by (auto simp add: ring_hom_def assms Pi_def)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
633  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
634  | 
lemma ring_hom_closed:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
635  | 
"[| h \<in> ring_hom R S; x \<in> carrier R |] ==> h x \<in> carrier S"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
636  | 
by (auto simp add: ring_hom_def funcset_mem)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
637  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
638  | 
lemma ring_hom_mult:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
639  | 
fixes R (structure) and S (structure)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
640  | 
shows  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
641  | 
"[| h \<in> ring_hom R S; x \<in> carrier R; y \<in> carrier R |] ==>  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
642  | 
h (x \<otimes> y) = h x \<otimes>\<^bsub>S\<^esub> h y"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
643  | 
by (simp add: ring_hom_def)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
644  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
645  | 
lemma ring_hom_add:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
646  | 
fixes R (structure) and S (structure)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
647  | 
shows  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
648  | 
"[| h \<in> ring_hom R S; x \<in> carrier R; y \<in> carrier R |] ==>  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
649  | 
h (x \<oplus> y) = h x \<oplus>\<^bsub>S\<^esub> h y"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
650  | 
by (simp add: ring_hom_def)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
651  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
652  | 
lemma ring_hom_one:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
653  | 
fixes R (structure) and S (structure)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
654  | 
shows "h \<in> ring_hom R S ==> h \<one> = \<one>\<^bsub>S\<^esub>"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
655  | 
by (simp add: ring_hom_def)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
656  | 
|
| 29237 | 657  | 
locale ring_hom_cring = R: cring R + S: cring S  | 
658  | 
for R (structure) and S (structure) +  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
659  | 
fixes h  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
660  | 
assumes homh [simp, intro]: "h \<in> ring_hom R S"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
661  | 
notes hom_closed [simp, intro] = ring_hom_closed [OF homh]  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
662  | 
and hom_mult [simp] = ring_hom_mult [OF homh]  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
663  | 
and hom_add [simp] = ring_hom_add [OF homh]  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
664  | 
and hom_one [simp] = ring_hom_one [OF homh]  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
665  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
666  | 
lemma (in ring_hom_cring) hom_zero [simp]:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
667  | 
"h \<zero> = \<zero>\<^bsub>S\<^esub>"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
668  | 
proof -  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
669  | 
have "h \<zero> \<oplus>\<^bsub>S\<^esub> h \<zero> = h \<zero> \<oplus>\<^bsub>S\<^esub> \<zero>\<^bsub>S\<^esub>"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
670  | 
by (simp add: hom_add [symmetric] del: hom_add)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
671  | 
then show ?thesis by (simp del: S.r_zero)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
672  | 
qed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
673  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
674  | 
lemma (in ring_hom_cring) hom_a_inv [simp]:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
675  | 
"x \<in> carrier R ==> h (\<ominus> x) = \<ominus>\<^bsub>S\<^esub> h x"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
676  | 
proof -  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
677  | 
assume R: "x \<in> carrier R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
678  | 
then have "h x \<oplus>\<^bsub>S\<^esub> h (\<ominus> x) = h x \<oplus>\<^bsub>S\<^esub> (\<ominus>\<^bsub>S\<^esub> h x)"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
679  | 
by (simp add: hom_add [symmetric] R.r_neg S.r_neg del: hom_add)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
680  | 
with R show ?thesis by simp  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
681  | 
qed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
682  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
683  | 
lemma (in ring_hom_cring) hom_finsum [simp]:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
684  | 
"[| finite A; f \<in> A -> carrier R |] ==>  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
685  | 
h (finsum R f A) = finsum S (h o f) A"  | 
| 22265 | 686  | 
proof (induct set: finite)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
687  | 
case empty then show ?case by simp  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
688  | 
next  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
689  | 
case insert then show ?case by (simp add: Pi_def)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
690  | 
qed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
691  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
692  | 
lemma (in ring_hom_cring) hom_finprod:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
693  | 
"[| finite A; f \<in> A -> carrier R |] ==>  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
694  | 
h (finprod R f A) = finprod S (h o f) A"  | 
| 22265 | 695  | 
proof (induct set: finite)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
696  | 
case empty then show ?case by simp  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
697  | 
next  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
698  | 
case insert then show ?case by (simp add: Pi_def)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
699  | 
qed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
700  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
701  | 
declare ring_hom_cring.hom_finprod [simp]  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
702  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
703  | 
lemma id_ring_hom [simp]:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
704  | 
"id \<in> ring_hom R R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
705  | 
by (auto intro!: ring_hom_memI)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
706  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
707  | 
end  |