| author | paulson <lp15@cam.ac.uk> | 
| Sat, 30 Jun 2018 15:44:04 +0100 | |
| changeset 68551 | b680e74eb6f2 | 
| parent 68517 | 6b5f15387353 | 
| child 68552 | 391e89e03eef | 
| 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  | 
begin  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
9  | 
|
| 61382 | 10  | 
section \<open>The Algebraic Hierarchy of Rings\<close>  | 
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
11  | 
|
| 61382 | 12  | 
subsection \<open>Abelian Groups\<close>  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
13  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
14  | 
record 'a ring = "'a monoid" +  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
15  | 
  zero :: 'a ("\<zero>\<index>")
 | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
16  | 
add :: "['a, 'a] \<Rightarrow> 'a" (infixl "\<oplus>\<index>" 65)  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
17  | 
|
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
18  | 
abbreviation  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
19  | 
  add_monoid :: "('a, 'm) ring_scheme \<Rightarrow> ('a, 'm) monoid_scheme"
 | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
20  | 
where "add_monoid R \<equiv> \<lparr> carrier = carrier R, mult = add R, one = zero R, \<dots> = (undefined :: 'm) \<rparr>"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
21  | 
|
| 61382 | 22  | 
text \<open>Derived operations.\<close>  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
23  | 
|
| 35847 | 24  | 
definition  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
25  | 
  a_inv :: "[('a, 'm) ring_scheme, 'a ] \<Rightarrow> 'a" ("\<ominus>\<index> _" [81] 80)
 | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
26  | 
where "a_inv R = m_inv (add_monoid R)"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
27  | 
|
| 35847 | 28  | 
definition  | 
| 67398 | 29  | 
  a_minus :: "[('a, 'm) ring_scheme, 'a, 'a] => 'a" ("(_ \<ominus>\<index> _)" [65,66] 65)
 | 
| 
68445
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
30  | 
where "x \<ominus>\<^bsub>R\<^esub> y = x \<oplus>\<^bsub>R\<^esub> (\<ominus>\<^bsub>R\<^esub> y)"  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
31  | 
|
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
32  | 
definition  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
33  | 
  add_pow :: "[_, ('b :: semiring_1), 'a] \<Rightarrow> 'a" ("[_] \<cdot>\<index> _" [81, 81] 80)
 | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
34  | 
where "add_pow R k a = pow (add_monoid R) a k"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
35  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
36  | 
locale abelian_monoid =  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
37  | 
fixes G (structure)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
38  | 
assumes a_comm_monoid:  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
39  | 
"comm_monoid (add_monoid G)"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
40  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
41  | 
definition  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
42  | 
  finsum :: "[('b, 'm) ring_scheme, 'a \<Rightarrow> 'b, 'a set] \<Rightarrow> 'b" where
 | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
43  | 
"finsum G = finprod (add_monoid G)"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
44  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
45  | 
syntax  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
46  | 
"_finsum" :: "index \<Rightarrow> idt \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  | 
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
47  | 
      ("(3\<Oplus>__\<in>_. _)" [1000, 0, 51, 10] 10)
 | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
48  | 
translations  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
49  | 
"\<Oplus>\<^bsub>G\<^esub>i\<in>A. b" \<rightleftharpoons> "CONST finsum G (\<lambda>i. b) A"  | 
| 63167 | 50  | 
\<comment> \<open>Beware of argument permutation!\<close>  | 
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
51  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
52  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
53  | 
locale abelian_group = abelian_monoid +  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
54  | 
assumes a_comm_group:  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
55  | 
"comm_group (add_monoid G)"  | 
| 
20318
 
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  | 
|
| 61382 | 58  | 
subsection \<open>Basic Properties\<close>  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
59  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
60  | 
lemma abelian_monoidI:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
61  | 
fixes R (structure)  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
62  | 
assumes "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> x \<oplus> y \<in> carrier R"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
63  | 
and "\<zero> \<in> carrier R"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
64  | 
and "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> (x \<oplus> y) \<oplus> z = x \<oplus> (y \<oplus> z)"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
65  | 
and "\<And>x. x \<in> carrier R \<Longrightarrow> \<zero> \<oplus> x = x"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
66  | 
and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> x \<oplus> y = y \<oplus> x"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
67  | 
shows "abelian_monoid R"  | 
| 
27714
 
27b4d7c01f8b
Tuned (for the sake of a meaningless log entry).
 
ballarin 
parents: 
27699 
diff
changeset
 | 
68  | 
by (auto intro!: abelian_monoid.intro comm_monoidI intro: assms)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
69  | 
|
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
70  | 
lemma abelian_monoidE:  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
71  | 
fixes R (structure)  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
72  | 
assumes "abelian_monoid R"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
73  | 
shows "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> x \<oplus> y \<in> carrier R"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
74  | 
and "\<zero> \<in> carrier R"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
75  | 
and "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> (x \<oplus> y) \<oplus> z = x \<oplus> (y \<oplus> z)"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
76  | 
and "\<And>x. x \<in> carrier R \<Longrightarrow> \<zero> \<oplus> x = x"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
77  | 
and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> x \<oplus> y = y \<oplus> x"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
78  | 
using assms unfolding abelian_monoid_def comm_monoid_def comm_monoid_axioms_def monoid_def by auto  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
79  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
80  | 
lemma abelian_groupI:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
81  | 
fixes R (structure)  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
82  | 
assumes "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> x \<oplus> y \<in> carrier R"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
83  | 
and "\<zero> \<in> carrier R"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
84  | 
and "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> (x \<oplus> y) \<oplus> z = x \<oplus> (y \<oplus> z)"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
85  | 
and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> x \<oplus> y = y \<oplus> x"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
86  | 
and "\<And>x. x \<in> carrier R \<Longrightarrow> \<zero> \<oplus> x = x"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
87  | 
and "\<And>x. x \<in> carrier R \<Longrightarrow> \<exists>y \<in> carrier R. y \<oplus> x = \<zero>"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
88  | 
shows "abelian_group R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
89  | 
by (auto intro!: abelian_group.intro abelian_monoidI  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
90  | 
abelian_group_axioms.intro comm_monoidI comm_groupI  | 
| 
27714
 
27b4d7c01f8b
Tuned (for the sake of a meaningless log entry).
 
ballarin 
parents: 
27699 
diff
changeset
 | 
91  | 
intro: assms)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
92  | 
|
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
93  | 
lemma abelian_groupE:  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
94  | 
fixes R (structure)  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
95  | 
assumes "abelian_group R"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
96  | 
shows "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> x \<oplus> y \<in> carrier R"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
97  | 
and "\<zero> \<in> carrier R"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
98  | 
and "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> (x \<oplus> y) \<oplus> z = x \<oplus> (y \<oplus> z)"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
99  | 
and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> x \<oplus> y = y \<oplus> x"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
100  | 
and "\<And>x. x \<in> carrier R \<Longrightarrow> \<zero> \<oplus> x = x"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
101  | 
and "\<And>x. x \<in> carrier R \<Longrightarrow> \<exists>y \<in> carrier R. y \<oplus> x = \<zero>"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
102  | 
using abelian_group.a_comm_group assms comm_groupE by fastforce+  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
103  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
104  | 
lemma (in abelian_monoid) a_monoid:  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
105  | 
"monoid (add_monoid G)"  | 
| 68517 | 106  | 
by (rule comm_monoid.axioms, rule a_comm_monoid)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
107  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
108  | 
lemma (in abelian_group) a_group:  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
109  | 
"group (add_monoid G)"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
110  | 
by (simp add: group_def a_monoid)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
111  | 
(simp add: comm_group.axioms group.axioms a_comm_group)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
112  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
113  | 
lemmas monoid_record_simps = partial_object.simps monoid.simps  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
114  | 
|
| 61382 | 115  | 
text \<open>Transfer facts from multiplicative structures via interpretation.\<close>  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
116  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
117  | 
sublocale abelian_monoid <  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
118  | 
add: monoid "(add_monoid G)"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
119  | 
rewrites "carrier (add_monoid G) = carrier G"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
120  | 
and "mult (add_monoid G) = add G"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
121  | 
and "one (add_monoid G) = zero G"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
122  | 
and "(\<lambda>a k. pow (add_monoid G) a k) = (\<lambda>a k. add_pow G k a)"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
123  | 
by (rule a_monoid) (auto simp add: add_pow_def)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
124  | 
|
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
125  | 
context abelian_monoid  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
126  | 
begin  | 
| 27933 | 127  | 
|
| 68517 | 128  | 
lemmas a_closed = add.m_closed  | 
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
129  | 
lemmas zero_closed = add.one_closed  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
130  | 
lemmas a_assoc = add.m_assoc  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
131  | 
lemmas l_zero = add.l_one  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
132  | 
lemmas r_zero = add.r_one  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
133  | 
lemmas minus_unique = add.inv_unique  | 
| 
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  | 
end  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
136  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
137  | 
sublocale abelian_monoid <  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
138  | 
add: comm_monoid "(add_monoid G)"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
139  | 
rewrites "carrier (add_monoid G) = carrier G"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
140  | 
and "mult (add_monoid G) = add G"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
141  | 
and "one (add_monoid G) = zero G"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
142  | 
and "finprod (add_monoid G) = finsum G"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
143  | 
and "pow (add_monoid G) = (\<lambda>a k. add_pow G k a)"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
144  | 
by (rule a_comm_monoid) (auto simp: finsum_def add_pow_def)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
145  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
146  | 
context abelian_monoid begin  | 
| 
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 a_comm = add.m_comm  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
149  | 
lemmas a_lcomm = add.m_lcomm  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
150  | 
lemmas a_ac = a_assoc a_comm a_lcomm  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
151  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
152  | 
lemmas finsum_empty = add.finprod_empty  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
153  | 
lemmas finsum_insert = add.finprod_insert  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
154  | 
lemmas finsum_zero = add.finprod_one  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
155  | 
lemmas finsum_closed = add.finprod_closed  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
156  | 
lemmas finsum_Un_Int = add.finprod_Un_Int  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
157  | 
lemmas finsum_Un_disjoint = add.finprod_Un_disjoint  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
158  | 
lemmas finsum_addf = add.finprod_multf  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
159  | 
lemmas finsum_cong' = add.finprod_cong'  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
160  | 
lemmas finsum_0 = add.finprod_0  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
161  | 
lemmas finsum_Suc = add.finprod_Suc  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
162  | 
lemmas finsum_Suc2 = add.finprod_Suc2  | 
| 
60112
 
3eab4acaa035
finprod takes 1 in case of infinite sets => remove several "finite A" assumptions
 
Rene Thiemann <rene.thiemann@uibk.ac.at> 
parents: 
59851 
diff
changeset
 | 
163  | 
lemmas finsum_infinite = add.finprod_infinite  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
164  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
165  | 
lemmas finsum_cong = add.finprod_cong  | 
| 61382 | 166  | 
text \<open>Usually, if this rule causes a failed congruence proof error,  | 
| 63167 | 167  | 
the reason is that the premise \<open>g \<in> B \<rightarrow> carrier G\<close> cannot be shown.  | 
| 61382 | 168  | 
   Adding @{thm [source] Pi_def} to the simpset is often useful.\<close>
 | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
169  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
170  | 
lemmas finsum_reindex = add.finprod_reindex  | 
| 27699 | 171  | 
|
| 
67341
 
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
 
nipkow 
parents: 
67091 
diff
changeset
 | 
172  | 
(* The following would be wrong. Needed is the equivalent of [^] for addition,  | 
| 27699 | 173  | 
or indeed the canonical embedding from Nat into the monoid.  | 
174  | 
||
| 27933 | 175  | 
lemma finsum_const:  | 
| 27699 | 176  | 
assumes fin [simp]: "finite A"  | 
177  | 
and a [simp]: "a : carrier G"  | 
|
| 
67341
 
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
 
nipkow 
parents: 
67091 
diff
changeset
 | 
178  | 
shows "finsum G (%x. a) A = a [^] card A"  | 
| 27699 | 179  | 
using fin apply induct  | 
180  | 
apply force  | 
|
181  | 
apply (subst finsum_insert)  | 
|
182  | 
apply auto  | 
|
183  | 
apply (force simp add: Pi_def)  | 
|
184  | 
apply (subst m_comm)  | 
|
185  | 
apply auto  | 
|
186  | 
done  | 
|
187  | 
*)  | 
|
188  | 
||
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
189  | 
lemmas finsum_singleton = add.finprod_singleton  | 
| 27933 | 190  | 
|
191  | 
end  | 
|
192  | 
||
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
193  | 
sublocale abelian_group <  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
194  | 
add: group "(add_monoid G)"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
195  | 
rewrites "carrier (add_monoid G) = carrier G"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
196  | 
and "mult (add_monoid G) = add G"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
197  | 
and "one (add_monoid G) = zero G"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
198  | 
and "m_inv (add_monoid G) = a_inv G"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
199  | 
and "pow (add_monoid G) = (\<lambda>a k. add_pow G k a)"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
200  | 
by (rule a_group) (auto simp: m_inv_def a_inv_def add_pow_def)  | 
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
201  | 
|
| 55926 | 202  | 
context abelian_group  | 
203  | 
begin  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
204  | 
|
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
205  | 
lemmas a_inv_closed = add.inv_closed  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
206  | 
|
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
207  | 
lemma minus_closed [intro, simp]:  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
208  | 
"[| 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
 | 
209  | 
by (simp add: a_minus_def)  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
210  | 
|
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
211  | 
lemmas l_neg = add.l_inv [simp del]  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
212  | 
lemmas r_neg = add.r_inv [simp del]  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
213  | 
lemmas minus_minus = add.inv_inv  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
214  | 
lemmas a_inv_inj = add.inv_inj  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
215  | 
lemmas minus_equality = add.inv_equality  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
216  | 
|
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
217  | 
end  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
218  | 
|
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
219  | 
sublocale abelian_group <  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
220  | 
add: comm_group "(add_monoid G)"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
221  | 
rewrites "carrier (add_monoid G) = carrier G"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
222  | 
and "mult (add_monoid G) = add G"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
223  | 
and "one (add_monoid G) = zero G"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
224  | 
and "m_inv (add_monoid G) = a_inv G"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
225  | 
and "finprod (add_monoid G) = finsum G"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
226  | 
and "pow (add_monoid G) = (\<lambda>a k. add_pow G k a)"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
227  | 
by (rule a_comm_group) (auto simp: m_inv_def a_inv_def finsum_def add_pow_def)  | 
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
228  | 
|
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
229  | 
lemmas (in abelian_group) minus_add = add.inv_mult  | 
| 68517 | 230  | 
|
| 63167 | 231  | 
text \<open>Derive an \<open>abelian_group\<close> from a \<open>comm_group\<close>\<close>  | 
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
232  | 
|
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
233  | 
lemma comm_group_abelian_groupI:  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
234  | 
fixes G (structure)  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
235  | 
assumes cg: "comm_group (add_monoid G)"  | 
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
236  | 
shows "abelian_group G"  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
237  | 
proof -  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
238  | 
interpret comm_group "(add_monoid G)"  | 
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
239  | 
by (rule cg)  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
240  | 
show "abelian_group G" ..  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
241  | 
qed  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
242  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
243  | 
|
| 61382 | 244  | 
subsection \<open>Rings: Basic Definitions\<close>  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
245  | 
|
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
246  | 
locale semiring = abelian_monoid (* for add *) R + monoid (* for mult *) R for R (structure) +  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
247  | 
assumes l_distr: "\<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
248  | 
and r_distr: "\<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> z \<otimes> (x \<oplus> y) = z \<otimes> x \<oplus> z \<otimes> y"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
249  | 
and l_null[simp]: "x \<in> carrier R \<Longrightarrow> \<zero> \<otimes> x = \<zero>"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
250  | 
and r_null[simp]: "x \<in> carrier R \<Longrightarrow> x \<otimes> \<zero> = \<zero>"  | 
| 
59851
 
43b1870b3e61
added locale for semirings
 
Rene Thiemann <rene.thiemann@uibk.ac.at> 
parents: 
58811 
diff
changeset
 | 
251  | 
|
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
252  | 
locale ring = abelian_group (* for add *) R + monoid (* for mult *) R for R (structure) +  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
253  | 
assumes "\<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
254  | 
and "\<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> z \<otimes> (x \<oplus> y) = z \<otimes> x \<oplus> z \<otimes> y"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
255  | 
|
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
256  | 
locale cring = ring + comm_monoid (* for mult *) R  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
257  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
258  | 
locale "domain" = cring +  | 
| 67091 | 259  | 
assumes one_not_zero [simp]: "\<one> \<noteq> \<zero>"  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
260  | 
and integral: "\<lbrakk> a \<otimes> b = \<zero>; a \<in> carrier R; b \<in> carrier R \<rbrakk> \<Longrightarrow> a = \<zero> \<or> b = \<zero>"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
261  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
262  | 
locale field = "domain" +  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
263  | 
  assumes field_Units: "Units R = carrier R - {\<zero>}"
 | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
264  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
265  | 
|
| 61382 | 266  | 
subsection \<open>Rings\<close>  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
267  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
268  | 
lemma ringI:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
269  | 
fixes R (structure)  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
270  | 
assumes "abelian_group R"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
271  | 
and "monoid R"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
272  | 
and "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
273  | 
and "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> z \<otimes> (x \<oplus> y) = z \<otimes> x \<oplus> z \<otimes> y"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
274  | 
shows "ring R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
275  | 
by (auto intro: ring.intro  | 
| 
27714
 
27b4d7c01f8b
Tuned (for the sake of a meaningless log entry).
 
ballarin 
parents: 
27699 
diff
changeset
 | 
276  | 
abelian_group.axioms ring_axioms.intro assms)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
277  | 
|
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
278  | 
lemma ringE:  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
279  | 
fixes R (structure)  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
280  | 
assumes "ring R"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
281  | 
shows "abelian_group R"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
282  | 
and "monoid R"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
283  | 
and "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
284  | 
and "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> z \<otimes> (x \<oplus> y) = z \<otimes> x \<oplus> z \<otimes> y"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
285  | 
using assms unfolding ring_def ring_axioms_def by auto  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
286  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
287  | 
context ring begin  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
288  | 
|
| 46721 | 289  | 
lemma is_abelian_group: "abelian_group R" ..  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
290  | 
|
| 46721 | 291  | 
lemma is_monoid: "monoid R"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
292  | 
by (auto intro!: monoidI m_assoc)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
293  | 
|
| 46721 | 294  | 
lemma is_ring: "ring R"  | 
| 26202 | 295  | 
by (rule ring_axioms)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
296  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
297  | 
end  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
298  | 
thm monoid_record_simps  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
299  | 
lemmas ring_record_simps = monoid_record_simps ring.simps  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
300  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
301  | 
lemma cringI:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
302  | 
fixes R (structure)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
303  | 
assumes abelian_group: "abelian_group R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
304  | 
and comm_monoid: "comm_monoid R"  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
305  | 
and l_distr: "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow>  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
306  | 
(x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
307  | 
shows "cring R"  | 
| 23350 | 308  | 
proof (intro cring.intro ring.intro)  | 
309  | 
show "ring_axioms R"  | 
|
| 63167 | 310  | 
\<comment> \<open>Right-distributivity follows from left-distributivity and  | 
| 61382 | 311  | 
commutativity.\<close>  | 
| 23350 | 312  | 
proof (rule ring_axioms.intro)  | 
313  | 
fix x y z  | 
|
314  | 
assume R: "x \<in> carrier R" "y \<in> carrier R" "z \<in> carrier R"  | 
|
315  | 
note [simp] = comm_monoid.axioms [OF comm_monoid]  | 
|
316  | 
abelian_group.axioms [OF abelian_group]  | 
|
317  | 
abelian_monoid.a_closed  | 
|
| 68517 | 318  | 
|
| 23350 | 319  | 
from R have "z \<otimes> (x \<oplus> y) = (x \<oplus> y) \<otimes> z"  | 
320  | 
by (simp add: comm_monoid.m_comm [OF comm_monoid.intro])  | 
|
321  | 
also from R have "... = x \<otimes> z \<oplus> y \<otimes> z" by (simp add: l_distr)  | 
|
322  | 
also from R have "... = z \<otimes> x \<oplus> z \<otimes> y"  | 
|
323  | 
by (simp add: comm_monoid.m_comm [OF comm_monoid.intro])  | 
|
324  | 
finally show "z \<otimes> (x \<oplus> y) = z \<otimes> x \<oplus> z \<otimes> y" .  | 
|
325  | 
qed (rule l_distr)  | 
|
326  | 
qed (auto intro: cring.intro  | 
|
| 
27714
 
27b4d7c01f8b
Tuned (for the sake of a meaningless log entry).
 
ballarin 
parents: 
27699 
diff
changeset
 | 
327  | 
abelian_group.axioms comm_monoid.axioms ring_axioms.intro assms)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
328  | 
|
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
329  | 
lemma cringE:  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
330  | 
fixes R (structure)  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
331  | 
assumes "cring R"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
332  | 
shows "comm_monoid R"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
333  | 
and "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
334  | 
using assms cring_def apply auto by (simp add: assms cring.axioms(1) ringE(3))  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
335  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
336  | 
lemma (in cring) is_cring:  | 
| 26202 | 337  | 
"cring R" by (rule cring_axioms)  | 
| 23350 | 338  | 
|
| 
68445
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
339  | 
lemma (in ring) minus_zero [simp]: "\<ominus> \<zero> = \<zero>"  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
340  | 
by (simp add: a_inv_def)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
341  | 
|
| 61382 | 342  | 
subsubsection \<open>Normaliser for Rings\<close>  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
343  | 
|
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
344  | 
lemma (in abelian_group) r_neg1:  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
345  | 
"\<lbrakk> x \<in> carrier G; y \<in> carrier G \<rbrakk> \<Longrightarrow> (\<ominus> x) \<oplus> (x \<oplus> y) = y"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
346  | 
proof -  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
347  | 
assume G: "x \<in> carrier G" "y \<in> carrier G"  | 
| 68517 | 348  | 
then have "(\<ominus> x \<oplus> x) \<oplus> y = y"  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
349  | 
by (simp only: l_neg l_zero)  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
350  | 
with G show ?thesis by (simp add: a_ac)  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
351  | 
qed  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
352  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
353  | 
lemma (in abelian_group) r_neg2:  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
354  | 
"\<lbrakk> x \<in> carrier G; y \<in> carrier G \<rbrakk> \<Longrightarrow> x \<oplus> ((\<ominus> x) \<oplus> y) = y"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
355  | 
proof -  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
356  | 
assume G: "x \<in> carrier G" "y \<in> carrier G"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
357  | 
then have "(x \<oplus> \<ominus> x) \<oplus> y = y"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
358  | 
by (simp only: r_neg l_zero)  | 
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
359  | 
with G show ?thesis  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
360  | 
by (simp add: a_ac)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
361  | 
qed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
362  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
363  | 
context ring begin  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
364  | 
|
| 61382 | 365  | 
text \<open>  | 
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
366  | 
The following proofs are from Jacobson, Basic Algebra I, pp.~88--89.  | 
| 61382 | 367  | 
\<close>  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
368  | 
|
| 
59851
 
43b1870b3e61
added locale for semirings
 
Rene Thiemann <rene.thiemann@uibk.ac.at> 
parents: 
58811 
diff
changeset
 | 
369  | 
sublocale semiring  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
370  | 
proof -  | 
| 
59851
 
43b1870b3e61
added locale for semirings
 
Rene Thiemann <rene.thiemann@uibk.ac.at> 
parents: 
58811 
diff
changeset
 | 
371  | 
note [simp] = ring_axioms[unfolded ring_def ring_axioms_def]  | 
| 
 
43b1870b3e61
added locale for semirings
 
Rene Thiemann <rene.thiemann@uibk.ac.at> 
parents: 
58811 
diff
changeset
 | 
372  | 
show "semiring R"  | 
| 
 
43b1870b3e61
added locale for semirings
 
Rene Thiemann <rene.thiemann@uibk.ac.at> 
parents: 
58811 
diff
changeset
 | 
373  | 
proof (unfold_locales)  | 
| 
 
43b1870b3e61
added locale for semirings
 
Rene Thiemann <rene.thiemann@uibk.ac.at> 
parents: 
58811 
diff
changeset
 | 
374  | 
fix x  | 
| 
 
43b1870b3e61
added locale for semirings
 
Rene Thiemann <rene.thiemann@uibk.ac.at> 
parents: 
58811 
diff
changeset
 | 
375  | 
assume R: "x \<in> carrier R"  | 
| 
 
43b1870b3e61
added locale for semirings
 
Rene Thiemann <rene.thiemann@uibk.ac.at> 
parents: 
58811 
diff
changeset
 | 
376  | 
then have "\<zero> \<otimes> x \<oplus> \<zero> \<otimes> x = (\<zero> \<oplus> \<zero>) \<otimes> x"  | 
| 
 
43b1870b3e61
added locale for semirings
 
Rene Thiemann <rene.thiemann@uibk.ac.at> 
parents: 
58811 
diff
changeset
 | 
377  | 
by (simp del: l_zero r_zero)  | 
| 
 
43b1870b3e61
added locale for semirings
 
Rene Thiemann <rene.thiemann@uibk.ac.at> 
parents: 
58811 
diff
changeset
 | 
378  | 
also from R have "... = \<zero> \<otimes> x \<oplus> \<zero>" by simp  | 
| 
 
43b1870b3e61
added locale for semirings
 
Rene Thiemann <rene.thiemann@uibk.ac.at> 
parents: 
58811 
diff
changeset
 | 
379  | 
finally have "\<zero> \<otimes> x \<oplus> \<zero> \<otimes> x = \<zero> \<otimes> x \<oplus> \<zero>" .  | 
| 
 
43b1870b3e61
added locale for semirings
 
Rene Thiemann <rene.thiemann@uibk.ac.at> 
parents: 
58811 
diff
changeset
 | 
380  | 
with R show "\<zero> \<otimes> x = \<zero>" by (simp del: r_zero)  | 
| 
 
43b1870b3e61
added locale for semirings
 
Rene Thiemann <rene.thiemann@uibk.ac.at> 
parents: 
58811 
diff
changeset
 | 
381  | 
from R have "x \<otimes> \<zero> \<oplus> x \<otimes> \<zero> = x \<otimes> (\<zero> \<oplus> \<zero>)"  | 
| 
 
43b1870b3e61
added locale for semirings
 
Rene Thiemann <rene.thiemann@uibk.ac.at> 
parents: 
58811 
diff
changeset
 | 
382  | 
by (simp del: l_zero r_zero)  | 
| 
 
43b1870b3e61
added locale for semirings
 
Rene Thiemann <rene.thiemann@uibk.ac.at> 
parents: 
58811 
diff
changeset
 | 
383  | 
also from R have "... = x \<otimes> \<zero> \<oplus> \<zero>" by simp  | 
| 
 
43b1870b3e61
added locale for semirings
 
Rene Thiemann <rene.thiemann@uibk.ac.at> 
parents: 
58811 
diff
changeset
 | 
384  | 
finally have "x \<otimes> \<zero> \<oplus> x \<otimes> \<zero> = x \<otimes> \<zero> \<oplus> \<zero>" .  | 
| 
 
43b1870b3e61
added locale for semirings
 
Rene Thiemann <rene.thiemann@uibk.ac.at> 
parents: 
58811 
diff
changeset
 | 
385  | 
with R show "x \<otimes> \<zero> = \<zero>" by (simp del: r_zero)  | 
| 
 
43b1870b3e61
added locale for semirings
 
Rene Thiemann <rene.thiemann@uibk.ac.at> 
parents: 
58811 
diff
changeset
 | 
386  | 
qed auto  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
387  | 
qed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
388  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
389  | 
lemma l_minus:  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
390  | 
"\<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> (\<ominus> x) \<otimes> y = \<ominus> (x \<otimes> y)"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
391  | 
proof -  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
392  | 
assume R: "x \<in> carrier R" "y \<in> carrier R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
393  | 
then have "(\<ominus> x) \<otimes> y \<oplus> x \<otimes> y = (\<ominus> x \<oplus> x) \<otimes> y" by (simp add: l_distr)  | 
| 44677 | 394  | 
also from R have "... = \<zero>" by (simp add: l_neg)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
395  | 
finally have "(\<ominus> x) \<otimes> y \<oplus> x \<otimes> y = \<zero>" .  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
396  | 
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
 | 
397  | 
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
 | 
398  | 
qed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
399  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
400  | 
lemma r_minus:  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
401  | 
"\<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> x \<otimes> (\<ominus> y) = \<ominus> (x \<otimes> y)"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
402  | 
proof -  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
403  | 
assume R: "x \<in> carrier R" "y \<in> carrier R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
404  | 
then have "x \<otimes> (\<ominus> y) \<oplus> x \<otimes> y = x \<otimes> (\<ominus> y \<oplus> y)" by (simp add: r_distr)  | 
| 44677 | 405  | 
also from R have "... = \<zero>" by (simp add: l_neg)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
406  | 
finally have "x \<otimes> (\<ominus> y) \<oplus> x \<otimes> y = \<zero>" .  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
407  | 
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
 | 
408  | 
with R show ?thesis by (simp add: a_assoc r_neg )  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
409  | 
qed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
410  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
411  | 
end  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
412  | 
|
| 
68445
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
413  | 
lemma (in abelian_group) minus_eq: "x \<ominus> y = x \<oplus> (\<ominus> y)"  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
414  | 
by (rule a_minus_def)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
415  | 
|
| 61382 | 416  | 
text \<open>Setup algebra method:  | 
417  | 
compute distributive normal form in locale contexts\<close>  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
418  | 
|
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
419  | 
|
| 48891 | 420  | 
ML_file "ringsimp.ML"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
421  | 
|
| 61382 | 422  | 
attribute_setup algebra = \<open>  | 
| 58811 | 423  | 
Scan.lift ((Args.add >> K true || Args.del >> K false) --| Args.colon || Scan.succeed true)  | 
424  | 
-- Scan.lift Args.name -- Scan.repeat Args.term  | 
|
425  | 
>> (fn ((b, n), ts) => if b then Ringsimp.add_struct (n, ts) else Ringsimp.del_struct (n, ts))  | 
|
| 61382 | 426  | 
\<close> "theorems controlling algebra method"  | 
| 47701 | 427  | 
|
| 61382 | 428  | 
method_setup algebra = \<open>  | 
| 58811 | 429  | 
Scan.succeed (SIMPLE_METHOD' o Ringsimp.algebra_tac)  | 
| 61382 | 430  | 
\<close> "normalisation of algebraic structure"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
431  | 
|
| 
59851
 
43b1870b3e61
added locale for semirings
 
Rene Thiemann <rene.thiemann@uibk.ac.at> 
parents: 
58811 
diff
changeset
 | 
432  | 
lemmas (in semiring) semiring_simprules  | 
| 
 
43b1870b3e61
added locale for semirings
 
Rene Thiemann <rene.thiemann@uibk.ac.at> 
parents: 
58811 
diff
changeset
 | 
433  | 
[algebra ring "zero R" "add R" "a_inv R" "a_minus R" "one R" "mult R"] =  | 
| 
 
43b1870b3e61
added locale for semirings
 
Rene Thiemann <rene.thiemann@uibk.ac.at> 
parents: 
58811 
diff
changeset
 | 
434  | 
a_closed zero_closed m_closed one_closed  | 
| 
 
43b1870b3e61
added locale for semirings
 
Rene Thiemann <rene.thiemann@uibk.ac.at> 
parents: 
58811 
diff
changeset
 | 
435  | 
a_assoc l_zero a_comm m_assoc l_one l_distr r_zero  | 
| 68517 | 436  | 
a_lcomm r_distr l_null r_null  | 
| 
59851
 
43b1870b3e61
added locale for semirings
 
Rene Thiemann <rene.thiemann@uibk.ac.at> 
parents: 
58811 
diff
changeset
 | 
437  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
438  | 
lemmas (in ring) ring_simprules  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
439  | 
[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
 | 
440  | 
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
 | 
441  | 
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
 | 
442  | 
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
 | 
443  | 
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
 | 
444  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
445  | 
lemmas (in cring)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
446  | 
[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
 | 
447  | 
_  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
448  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
449  | 
lemmas (in cring) cring_simprules  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
450  | 
[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
 | 
451  | 
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
 | 
452  | 
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
 | 
453  | 
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
 | 
454  | 
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
 | 
455  | 
|
| 
59851
 
43b1870b3e61
added locale for semirings
 
Rene Thiemann <rene.thiemann@uibk.ac.at> 
parents: 
58811 
diff
changeset
 | 
456  | 
lemma (in semiring) nat_pow_zero:  | 
| 
67341
 
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
 
nipkow 
parents: 
67091 
diff
changeset
 | 
457  | 
"(n::nat) \<noteq> 0 \<Longrightarrow> \<zero> [^] n = \<zero>"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
458  | 
by (induct n) simp_all  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
459  | 
|
| 
59851
 
43b1870b3e61
added locale for semirings
 
Rene Thiemann <rene.thiemann@uibk.ac.at> 
parents: 
58811 
diff
changeset
 | 
460  | 
context semiring begin  | 
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
461  | 
|
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
462  | 
lemma one_zeroD:  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
463  | 
assumes onezero: "\<one> = \<zero>"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
464  | 
  shows "carrier R = {\<zero>}"
 | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
465  | 
proof (rule, rule)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
466  | 
fix x  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
467  | 
assume xcarr: "x \<in> carrier R"  | 
| 47409 | 468  | 
from xcarr have "x = x \<otimes> \<one>" by simp  | 
469  | 
with onezero have "x = x \<otimes> \<zero>" by simp  | 
|
470  | 
with xcarr have "x = \<zero>" by simp  | 
|
471  | 
  then show "x \<in> {\<zero>}" by fast
 | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
472  | 
qed fast  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
473  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
474  | 
lemma one_zeroI:  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
475  | 
  assumes carrzero: "carrier R = {\<zero>}"
 | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
476  | 
shows "\<one> = \<zero>"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
477  | 
proof -  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
478  | 
from one_closed and carrzero  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
479  | 
show "\<one> = \<zero>" by simp  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
480  | 
qed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
481  | 
|
| 46721 | 482  | 
lemma carrier_one_zero: "(carrier R = {\<zero>}) = (\<one> = \<zero>)"
 | 
483  | 
apply rule  | 
|
484  | 
apply (erule one_zeroI)  | 
|
485  | 
apply (erule one_zeroD)  | 
|
486  | 
done  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
487  | 
|
| 46721 | 488  | 
lemma carrier_one_not_zero: "(carrier R \<noteq> {\<zero>}) = (\<one> \<noteq> \<zero>)"
 | 
| 
27717
 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 
ballarin 
parents: 
27714 
diff
changeset
 | 
489  | 
by (simp add: carrier_one_zero)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
490  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
491  | 
end  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
492  | 
|
| 61382 | 493  | 
text \<open>Two examples for use of method algebra\<close>  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
494  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
495  | 
lemma  | 
| 27611 | 496  | 
fixes R (structure) and S (structure)  | 
497  | 
assumes "ring R" "cring S"  | 
|
498  | 
assumes RS: "a \<in> carrier R" "b \<in> carrier R" "c \<in> carrier S" "d \<in> carrier S"  | 
|
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
499  | 
shows "a \<oplus> (\<ominus> (a \<oplus> (\<ominus> b))) = b \<and> c \<otimes>\<^bsub>S\<^esub> d = d \<otimes>\<^bsub>S\<^esub> c"  | 
| 27611 | 500  | 
proof -  | 
| 29237 | 501  | 
interpret ring R by fact  | 
502  | 
interpret cring S by fact  | 
|
| 27611 | 503  | 
from RS show ?thesis by algebra  | 
504  | 
qed  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
505  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
506  | 
lemma  | 
| 27611 | 507  | 
fixes R (structure)  | 
508  | 
assumes "ring R"  | 
|
509  | 
assumes R: "a \<in> carrier R" "b \<in> carrier R"  | 
|
510  | 
shows "a \<ominus> (a \<ominus> b) = b"  | 
|
511  | 
proof -  | 
|
| 29237 | 512  | 
interpret ring R by fact  | 
| 27611 | 513  | 
from R show ?thesis by algebra  | 
514  | 
qed  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
515  | 
|
| 35849 | 516  | 
|
| 61382 | 517  | 
subsubsection \<open>Sums over Finite Sets\<close>  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
518  | 
|
| 
59851
 
43b1870b3e61
added locale for semirings
 
Rene Thiemann <rene.thiemann@uibk.ac.at> 
parents: 
58811 
diff
changeset
 | 
519  | 
lemma (in semiring) finsum_ldistr:  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
520  | 
"\<lbrakk> finite A; a \<in> carrier R; f: A \<rightarrow> carrier R \<rbrakk> \<Longrightarrow>  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
521  | 
(\<Oplus> i \<in> A. (f i)) \<otimes> a = (\<Oplus> i \<in> A. ((f i) \<otimes> a))"  | 
| 22265 | 522  | 
proof (induct set: finite)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
523  | 
case empty then show ?case by simp  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
524  | 
next  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
525  | 
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
 | 
526  | 
qed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
527  | 
|
| 
59851
 
43b1870b3e61
added locale for semirings
 
Rene Thiemann <rene.thiemann@uibk.ac.at> 
parents: 
58811 
diff
changeset
 | 
528  | 
lemma (in semiring) finsum_rdistr:  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
529  | 
"\<lbrakk> finite A; a \<in> carrier R; f: A \<rightarrow> carrier R \<rbrakk> \<Longrightarrow>  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
530  | 
a \<otimes> (\<Oplus> i \<in> A. (f i)) = (\<Oplus> i \<in> A. (a \<otimes> (f i)))"  | 
| 22265 | 531  | 
proof (induct set: finite)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
532  | 
case empty then show ?case by simp  | 
| 
 
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  | 
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
 | 
535  | 
qed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
536  | 
|
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
537  | 
(* ************************************************************************** *)  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
538  | 
(* Contributed by Paulo E. de Vilhena. *)  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
539  | 
|
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
540  | 
text \<open>A quick detour\<close>  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
541  | 
|
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
542  | 
lemma add_pow_int_ge: "(k :: int) \<ge> 0 \<Longrightarrow> [ k ] \<cdot>\<^bsub>R\<^esub> a = [ nat k ] \<cdot>\<^bsub>R\<^esub> a"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
543  | 
by (simp add: add_pow_def int_pow_def nat_pow_def)  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
544  | 
|
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
545  | 
lemma add_pow_int_lt: "(k :: int) < 0 \<Longrightarrow> [ k ] \<cdot>\<^bsub>R\<^esub> a = \<ominus>\<^bsub>R\<^esub> ([ nat (- k) ] \<cdot>\<^bsub>R\<^esub> a)"  | 
| 68517 | 546  | 
by (simp add: int_pow_def nat_pow_def a_inv_def add_pow_def)  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
547  | 
|
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
548  | 
corollary (in semiring) add_pow_ldistr:  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
549  | 
assumes "a \<in> carrier R" "b \<in> carrier R"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
550  | 
shows "([(k :: nat)] \<cdot> a) \<otimes> b = [k] \<cdot> (a \<otimes> b)"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
551  | 
proof -  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
552  | 
  have "([k] \<cdot> a) \<otimes> b = (\<Oplus> i \<in> {..< k}. a) \<otimes> b"
 | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
553  | 
    using add.finprod_const[OF assms(1), of "{..<k}"] by simp
 | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
554  | 
  also have " ... = (\<Oplus> i \<in> {..< k}. (a \<otimes> b))"
 | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
555  | 
    using finsum_ldistr[of "{..<k}" b "\<lambda>x. a"] assms by simp
 | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
556  | 
also have " ... = [k] \<cdot> (a \<otimes> b)"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
557  | 
    using add.finprod_const[of "a \<otimes> b" "{..<k}"] assms by simp
 | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
558  | 
finally show ?thesis .  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
559  | 
qed  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
560  | 
|
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
561  | 
corollary (in semiring) add_pow_rdistr:  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
562  | 
assumes "a \<in> carrier R" "b \<in> carrier R"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
563  | 
shows "a \<otimes> ([(k :: nat)] \<cdot> b) = [k] \<cdot> (a \<otimes> b)"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
564  | 
proof -  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
565  | 
  have "a \<otimes> ([k] \<cdot> b) = a \<otimes> (\<Oplus> i \<in> {..< k}. b)"
 | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
566  | 
    using add.finprod_const[OF assms(2), of "{..<k}"] by simp
 | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
567  | 
  also have " ... = (\<Oplus> i \<in> {..< k}. (a \<otimes> b))"
 | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
568  | 
    using finsum_rdistr[of "{..<k}" a "\<lambda>x. b"] assms by simp
 | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
569  | 
also have " ... = [k] \<cdot> (a \<otimes> b)"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
570  | 
    using add.finprod_const[of "a \<otimes> b" "{..<k}"] assms by simp
 | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
571  | 
finally show ?thesis .  | 
| 68517 | 572  | 
qed  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
573  | 
|
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
574  | 
(* For integers, we need the uniqueness of the additive inverse *)  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
575  | 
lemma (in ring) add_pow_ldistr_int:  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
576  | 
assumes "a \<in> carrier R" "b \<in> carrier R"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
577  | 
shows "([(k :: int)] \<cdot> a) \<otimes> b = [k] \<cdot> (a \<otimes> b)"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
578  | 
proof (cases "k \<ge> 0")  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
579  | 
case True thus ?thesis  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
580  | 
using add_pow_int_ge[of k R] add_pow_ldistr[OF assms] by auto  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
581  | 
next  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
582  | 
case False thus ?thesis  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
583  | 
using add_pow_int_lt[of k R a] add_pow_int_lt[of k R "a \<otimes> b"]  | 
| 68517 | 584  | 
add_pow_ldistr[OF assms, of "nat (- k)"] assms l_minus by auto  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
585  | 
qed  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
586  | 
|
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
587  | 
lemma (in ring) add_pow_rdistr_int:  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
588  | 
assumes "a \<in> carrier R" "b \<in> carrier R"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
589  | 
shows "a \<otimes> ([(k :: int)] \<cdot> b) = [k] \<cdot> (a \<otimes> b)"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
590  | 
proof (cases "k \<ge> 0")  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
591  | 
case True thus ?thesis  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
592  | 
using add_pow_int_ge[of k R] add_pow_rdistr[OF assms] by auto  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
593  | 
next  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
594  | 
case False thus ?thesis  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
595  | 
using add_pow_int_lt[of k R b] add_pow_int_lt[of k R "a \<otimes> b"]  | 
| 68517 | 596  | 
add_pow_rdistr[OF assms, of "nat (- k)"] assms r_minus by auto  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
597  | 
qed  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
598  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
599  | 
|
| 61382 | 600  | 
subsection \<open>Integral Domains\<close>  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
601  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
602  | 
context "domain" begin  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
603  | 
|
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
604  | 
lemma zero_not_one [simp]: "\<zero> \<noteq> \<one>"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
605  | 
by (rule not_sym) simp  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
606  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
607  | 
lemma integral_iff: (* not by default a simp rule! *)  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
608  | 
"\<lbrakk> a \<in> carrier R; b \<in> carrier R \<rbrakk> \<Longrightarrow> (a \<otimes> b = \<zero>) = (a = \<zero> \<or> b = \<zero>)"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
609  | 
proof  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
610  | 
assume "a \<in> carrier R" "b \<in> carrier R" "a \<otimes> b = \<zero>"  | 
| 67091 | 611  | 
then show "a = \<zero> \<or> b = \<zero>" by (simp add: integral)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
612  | 
next  | 
| 67091 | 613  | 
assume "a \<in> carrier R" "b \<in> carrier R" "a = \<zero> \<or> b = \<zero>"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
614  | 
then show "a \<otimes> b = \<zero>" by auto  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
615  | 
qed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
616  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
617  | 
lemma m_lcancel:  | 
| 67091 | 618  | 
assumes prem: "a \<noteq> \<zero>"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
619  | 
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
 | 
620  | 
shows "(a \<otimes> b = a \<otimes> c) = (b = c)"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
621  | 
proof  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
622  | 
assume eq: "a \<otimes> b = a \<otimes> c"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
623  | 
with R have "a \<otimes> (b \<ominus> c) = \<zero>" by algebra  | 
| 67091 | 624  | 
with R have "a = \<zero> \<or> (b \<ominus> c) = \<zero>" by (simp add: integral_iff)  | 
| 68517 | 625  | 
with prem and R have "b \<ominus> c = \<zero>" by auto  | 
| 
68445
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
626  | 
with R have "b = b \<ominus> (b \<ominus> c)" by algebra  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
627  | 
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
 | 
628  | 
finally show "b = c" .  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
629  | 
next  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
630  | 
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
 | 
631  | 
qed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
632  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
633  | 
lemma m_rcancel:  | 
| 67091 | 634  | 
assumes prem: "a \<noteq> \<zero>"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
635  | 
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
 | 
636  | 
shows conc: "(b \<otimes> a = c \<otimes> a) = (b = c)"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
637  | 
proof -  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
638  | 
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
 | 
639  | 
with R show ?thesis by algebra  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
640  | 
qed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
641  | 
|
| 
41433
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
642  | 
end  | 
| 
 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 
ballarin 
parents: 
35849 
diff
changeset
 | 
643  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
644  | 
|
| 61382 | 645  | 
subsection \<open>Fields\<close>  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
646  | 
|
| 61382 | 647  | 
text \<open>Field would not need to be derived from domain, the properties  | 
648  | 
for domain follow from the assumptions of field\<close>  | 
|
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
649  | 
|
| 
68551
 
b680e74eb6f2
More on Algebra by Paulo and Martin
 
paulson <lp15@cam.ac.uk> 
parents: 
68517 
diff
changeset
 | 
650  | 
lemma fieldE :  | 
| 
 
b680e74eb6f2
More on Algebra by Paulo and Martin
 
paulson <lp15@cam.ac.uk> 
parents: 
68517 
diff
changeset
 | 
651  | 
fixes R (structure)  | 
| 
 
b680e74eb6f2
More on Algebra by Paulo and Martin
 
paulson <lp15@cam.ac.uk> 
parents: 
68517 
diff
changeset
 | 
652  | 
assumes "field R"  | 
| 
 
b680e74eb6f2
More on Algebra by Paulo and Martin
 
paulson <lp15@cam.ac.uk> 
parents: 
68517 
diff
changeset
 | 
653  | 
shows "cring R"  | 
| 
 
b680e74eb6f2
More on Algebra by Paulo and Martin
 
paulson <lp15@cam.ac.uk> 
parents: 
68517 
diff
changeset
 | 
654  | 
and one_not_zero : "\<one> \<noteq> \<zero>"  | 
| 
 
b680e74eb6f2
More on Algebra by Paulo and Martin
 
paulson <lp15@cam.ac.uk> 
parents: 
68517 
diff
changeset
 | 
655  | 
and integral: "\<And>a b. \<lbrakk> a \<otimes> b = \<zero>; a \<in> carrier R; b \<in> carrier R \<rbrakk> \<Longrightarrow> a = \<zero> \<or> b = \<zero>"  | 
| 
 
b680e74eb6f2
More on Algebra by Paulo and Martin
 
paulson <lp15@cam.ac.uk> 
parents: 
68517 
diff
changeset
 | 
656  | 
  and field_Units: "Units R = carrier R - {\<zero>}"
 | 
| 
 
b680e74eb6f2
More on Algebra by Paulo and Martin
 
paulson <lp15@cam.ac.uk> 
parents: 
68517 
diff
changeset
 | 
657  | 
using assms unfolding field_def field_axioms_def domain_def domain_axioms_def by simp_all  | 
| 
 
b680e74eb6f2
More on Algebra by Paulo and Martin
 
paulson <lp15@cam.ac.uk> 
parents: 
68517 
diff
changeset
 | 
658  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
659  | 
lemma (in cring) cring_fieldI:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
660  | 
  assumes field_Units: "Units R = carrier R - {\<zero>}"
 | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
661  | 
shows "field R"  | 
| 28823 | 662  | 
proof  | 
| 47409 | 663  | 
from field_Units have "\<zero> \<notin> Units R" by fast  | 
664  | 
moreover have "\<one> \<in> Units R" by fast  | 
|
665  | 
ultimately show "\<one> \<noteq> \<zero>" by force  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
666  | 
next  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
667  | 
fix a b  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
668  | 
assume acarr: "a \<in> carrier R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
669  | 
and bcarr: "b \<in> carrier R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
670  | 
and ab: "a \<otimes> b = \<zero>"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
671  | 
show "a = \<zero> \<or> b = \<zero>"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
672  | 
proof (cases "a = \<zero>", simp)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
673  | 
assume "a \<noteq> \<zero>"  | 
| 47409 | 674  | 
with field_Units and acarr have aUnit: "a \<in> Units R" by fast  | 
675  | 
from bcarr have "b = \<one> \<otimes> b" by algebra  | 
|
676  | 
also from aUnit acarr have "... = (inv a \<otimes> a) \<otimes> b" by simp  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
677  | 
also from acarr bcarr aUnit[THEN Units_inv_closed]  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
678  | 
have "... = (inv a) \<otimes> (a \<otimes> b)" by algebra  | 
| 47409 | 679  | 
also from ab and acarr bcarr aUnit have "... = (inv a) \<otimes> \<zero>" by simp  | 
680  | 
also from aUnit[THEN Units_inv_closed] have "... = \<zero>" by algebra  | 
|
681  | 
finally have "b = \<zero>" .  | 
|
682  | 
then show "a = \<zero> \<or> b = \<zero>" by simp  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
683  | 
qed  | 
| 23350 | 684  | 
qed (rule field_Units)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
685  | 
|
| 61382 | 686  | 
text \<open>Another variant to show that something is a field\<close>  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
687  | 
lemma (in cring) cring_fieldI2:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
688  | 
assumes notzero: "\<zero> \<noteq> \<one>"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
689  | 
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
 | 
690  | 
shows "field R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
691  | 
apply (rule cring_fieldI, simp add: Units_def)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
692  | 
apply (rule, clarsimp)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
693  | 
apply (simp add: notzero)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
694  | 
proof (clarsimp)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
695  | 
fix x  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
696  | 
assume xcarr: "x \<in> carrier R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
697  | 
and "x \<noteq> \<zero>"  | 
| 47409 | 698  | 
then have "\<exists>y\<in>carrier R. x \<otimes> y = \<one>" by (rule invex)  | 
699  | 
then obtain y where ycarr: "y \<in> carrier R" and xy: "x \<otimes> y = \<one>" by fast  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
700  | 
from xy xcarr ycarr have "y \<otimes> x = \<one>" by (simp add: m_comm)  | 
| 47409 | 701  | 
with ycarr and xy show "\<exists>y\<in>carrier R. y \<otimes> x = \<one> \<and> x \<otimes> y = \<one>" by fast  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
702  | 
qed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
703  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
704  | 
|
| 61382 | 705  | 
subsection \<open>Morphisms\<close>  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
706  | 
|
| 35847 | 707  | 
definition  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
708  | 
  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
 | 
709  | 
where "ring_hom R S =  | 
| 67091 | 710  | 
    {h. h \<in> carrier R \<rightarrow> carrier S \<and>
 | 
711  | 
(\<forall>x y. x \<in> carrier R \<and> y \<in> carrier R \<longrightarrow>  | 
|
712  | 
h (x \<otimes>\<^bsub>R\<^esub> y) = h x \<otimes>\<^bsub>S\<^esub> h y \<and> h (x \<oplus>\<^bsub>R\<^esub> y) = h x \<oplus>\<^bsub>S\<^esub> h y) \<and>  | 
|
| 35847 | 713  | 
h \<one>\<^bsub>R\<^esub> = \<one>\<^bsub>S\<^esub>}"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
714  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
715  | 
lemma ring_hom_memI:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
716  | 
fixes R (structure) and S (structure)  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
717  | 
assumes "\<And>x. x \<in> carrier R \<Longrightarrow> h x \<in> carrier S"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
718  | 
and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> h (x \<otimes> y) = h x \<otimes>\<^bsub>S\<^esub> h y"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
719  | 
and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> h (x \<oplus> y) = h x \<oplus>\<^bsub>S\<^esub> h y"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
720  | 
and "h \<one> = \<one>\<^bsub>S\<^esub>"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
721  | 
shows "h \<in> ring_hom R S"  | 
| 
27714
 
27b4d7c01f8b
Tuned (for the sake of a meaningless log entry).
 
ballarin 
parents: 
27699 
diff
changeset
 | 
722  | 
by (auto simp add: ring_hom_def assms Pi_def)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
723  | 
|
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
724  | 
lemma ring_hom_memE:  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
725  | 
fixes R (structure) and S (structure)  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
726  | 
assumes "h \<in> ring_hom R S"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
727  | 
shows "\<And>x. x \<in> carrier R \<Longrightarrow> h x \<in> carrier S"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
728  | 
and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> h (x \<otimes> y) = h x \<otimes>\<^bsub>S\<^esub> h y"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
729  | 
and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> h (x \<oplus> y) = h x \<oplus>\<^bsub>S\<^esub> h y"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
730  | 
and "h \<one> = \<one>\<^bsub>S\<^esub>"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
731  | 
using assms unfolding ring_hom_def by auto  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
732  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
733  | 
lemma ring_hom_closed:  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
734  | 
"\<lbrakk> h \<in> ring_hom R S; x \<in> carrier R \<rbrakk> \<Longrightarrow> h x \<in> carrier S"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
735  | 
by (auto simp add: ring_hom_def funcset_mem)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
736  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
737  | 
lemma ring_hom_mult:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
738  | 
fixes R (structure) and S (structure)  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
739  | 
shows "\<lbrakk> h \<in> ring_hom R S; x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> h (x \<otimes> y) = h x \<otimes>\<^bsub>S\<^esub> h y"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
740  | 
by (simp add: ring_hom_def)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
741  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
742  | 
lemma ring_hom_add:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
743  | 
fixes R (structure) and S (structure)  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
744  | 
shows "\<lbrakk> h \<in> ring_hom R S; x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> h (x \<oplus> y) = h x \<oplus>\<^bsub>S\<^esub> h y"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
745  | 
by (simp add: ring_hom_def)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
746  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
747  | 
lemma ring_hom_one:  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
748  | 
fixes R (structure) and S (structure)  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
749  | 
shows "h \<in> ring_hom R S \<Longrightarrow> h \<one> = \<one>\<^bsub>S\<^esub>"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
750  | 
by (simp add: ring_hom_def)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
751  | 
|
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
752  | 
lemma ring_hom_zero:  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
753  | 
fixes R (structure) and S (structure)  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
754  | 
assumes "h \<in> ring_hom R S" "ring R" "ring S"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
755  | 
shows "h \<zero> = \<zero>\<^bsub>S\<^esub>"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
756  | 
proof -  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
757  | 
have "h \<zero> = h \<zero> \<oplus>\<^bsub>S\<^esub> h \<zero>"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
758  | 
using ring_hom_add[OF assms(1), of \<zero> \<zero>] assms(2)  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
759  | 
by (simp add: ring.ring_simprules(2) ring.ring_simprules(15))  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
760  | 
thus ?thesis  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
761  | 
by (metis abelian_group.l_neg assms ring.is_abelian_group ring.ring_simprules(18) ring.ring_simprules(2) ring_hom_closed)  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
762  | 
qed  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
763  | 
|
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
764  | 
locale ring_hom_cring =  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
765  | 
R?: cring R + S?: cring S for R (structure) and S (structure) + fixes h  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
766  | 
assumes homh [simp, intro]: "h \<in> ring_hom R S"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
767  | 
notes hom_closed [simp, intro] = ring_hom_closed [OF homh]  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
768  | 
and hom_mult [simp] = ring_hom_mult [OF homh]  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
769  | 
and hom_add [simp] = ring_hom_add [OF homh]  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
770  | 
and hom_one [simp] = ring_hom_one [OF homh]  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
771  | 
|
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
772  | 
lemma (in ring_hom_cring) hom_zero [simp]: "h \<zero> = \<zero>\<^bsub>S\<^esub>"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
773  | 
proof -  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
774  | 
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
 | 
775  | 
by (simp add: hom_add [symmetric] del: hom_add)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
776  | 
then show ?thesis by (simp del: S.r_zero)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
777  | 
qed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
778  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
779  | 
lemma (in ring_hom_cring) hom_a_inv [simp]:  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
780  | 
"x \<in> carrier R \<Longrightarrow> h (\<ominus> x) = \<ominus>\<^bsub>S\<^esub> h x"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
781  | 
proof -  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
782  | 
assume R: "x \<in> carrier R"  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
783  | 
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
 | 
784  | 
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
 | 
785  | 
with R show ?thesis by simp  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
786  | 
qed  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
787  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
788  | 
lemma (in ring_hom_cring) hom_finsum [simp]:  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
789  | 
assumes "f: A \<rightarrow> carrier R"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
790  | 
shows "h (\<Oplus> i \<in> A. f i) = (\<Oplus>\<^bsub>S\<^esub> i \<in> A. (h o f) i)"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
791  | 
using assms by (induct A rule: infinite_finite_induct, auto simp: Pi_def)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
792  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
793  | 
lemma (in ring_hom_cring) hom_finprod:  | 
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
794  | 
assumes "f: A \<rightarrow> carrier R"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
795  | 
shows "h (\<Otimes> i \<in> A. f i) = (\<Otimes>\<^bsub>S\<^esub> i \<in> A. (h o f) i)"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
796  | 
using assms by (induct A rule: infinite_finite_induct, auto simp: Pi_def)  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
797  | 
|
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
798  | 
declare ring_hom_cring.hom_finprod [simp]  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
799  | 
|
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
800  | 
lemma id_ring_hom [simp]: "id \<in> ring_hom R R"  | 
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
801  | 
by (auto intro!: ring_hom_memI)  | 
| 
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
802  | 
|
| 
68443
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
803  | 
(* Next lemma contributed by Paulo Emílio de Vilhena. *)  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
804  | 
|
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
805  | 
lemma ring_hom_trans:  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
806  | 
"\<lbrakk> f \<in> ring_hom R S; g \<in> ring_hom S T \<rbrakk> \<Longrightarrow> g \<circ> f \<in> ring_hom R T"  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
807  | 
by (rule ring_hom_memI) (auto simp add: ring_hom_closed ring_hom_mult ring_hom_add ring_hom_one)  | 
| 
 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 
paulson <lp15@cam.ac.uk> 
parents: 
68399 
diff
changeset
 | 
808  | 
|
| 
68445
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
809  | 
subsection\<open>Jeremy Avigad's @{text"More_Finite_Product"} material\<close>
 | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
810  | 
|
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
811  | 
(* need better simplification rules for rings *)  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
812  | 
(* the next one holds more generally for abelian groups *)  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
813  | 
|
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
814  | 
lemma (in cring) sum_zero_eq_neg: "x \<in> carrier R \<Longrightarrow> y \<in> carrier R \<Longrightarrow> x \<oplus> y = \<zero> \<Longrightarrow> x = \<ominus> y"  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
815  | 
by (metis minus_equality)  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
816  | 
|
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
817  | 
lemma (in domain) square_eq_one:  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
818  | 
fixes x  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
819  | 
assumes [simp]: "x \<in> carrier R"  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
820  | 
and "x \<otimes> x = \<one>"  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
821  | 
shows "x = \<one> \<or> x = \<ominus>\<one>"  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
822  | 
proof -  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
823  | 
have "(x \<oplus> \<one>) \<otimes> (x \<oplus> \<ominus> \<one>) = x \<otimes> x \<oplus> \<ominus> \<one>"  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
824  | 
by (simp add: ring_simprules)  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
825  | 
also from \<open>x \<otimes> x = \<one>\<close> have "\<dots> = \<zero>"  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
826  | 
by (simp add: ring_simprules)  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
827  | 
finally have "(x \<oplus> \<one>) \<otimes> (x \<oplus> \<ominus> \<one>) = \<zero>" .  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
828  | 
then have "(x \<oplus> \<one>) = \<zero> \<or> (x \<oplus> \<ominus> \<one>) = \<zero>"  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
829  | 
by (intro integral) auto  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
830  | 
then show ?thesis  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
831  | 
by (metis add.inv_closed add.inv_solve_right assms(1) l_zero one_closed zero_closed)  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
832  | 
qed  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
833  | 
|
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
834  | 
lemma (in domain) inv_eq_self: "x \<in> Units R \<Longrightarrow> x = inv x \<Longrightarrow> x = \<one> \<or> x = \<ominus>\<one>"  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
835  | 
by (metis Units_closed Units_l_inv square_eq_one)  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
836  | 
|
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
837  | 
|
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
838  | 
text \<open>  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
839  | 
The following translates theorems about groups to the facts about  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
840  | 
the units of a ring. (The list should be expanded as more things are  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
841  | 
needed.)  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
842  | 
\<close>  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
843  | 
|
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
844  | 
lemma (in ring) finite_ring_finite_units [intro]: "finite (carrier R) \<Longrightarrow> finite (Units R)"  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
845  | 
by (rule finite_subset) auto  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
846  | 
|
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
847  | 
lemma (in monoid) units_of_pow:  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
848  | 
fixes n :: nat  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
849  | 
shows "x \<in> Units G \<Longrightarrow> x [^]\<^bsub>units_of G\<^esub> n = x [^]\<^bsub>G\<^esub> n"  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
850  | 
apply (induct n)  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
851  | 
apply (auto simp add: units_group group.is_monoid  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
852  | 
monoid.nat_pow_0 monoid.nat_pow_Suc units_of_one units_of_mult)  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
853  | 
done  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
854  | 
|
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
855  | 
lemma (in cring) units_power_order_eq_one:  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
856  | 
"finite (Units R) \<Longrightarrow> a \<in> Units R \<Longrightarrow> a [^] card(Units R) = \<one>"  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
857  | 
by (metis comm_group.power_order_eq_one units_comm_group units_of_carrier units_of_one units_of_pow)  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
858  | 
|
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
859  | 
subsection\<open>Jeremy Avigad's @{text"More_Ring"} material\<close>
 | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
860  | 
|
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
861  | 
lemma (in cring) field_intro2: "\<zero>\<^bsub>R\<^esub> \<noteq> \<one>\<^bsub>R\<^esub> \<Longrightarrow> \<forall>x \<in> carrier R - {\<zero>\<^bsub>R\<^esub>}. x \<in> Units R \<Longrightarrow> field R"
 | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
862  | 
apply (unfold_locales)  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
863  | 
apply (use cring_axioms in auto)  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
864  | 
apply (rule trans)  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
865  | 
apply (subgoal_tac "a = (a \<otimes> b) \<otimes> inv b")  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
866  | 
apply assumption  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
867  | 
apply (subst m_assoc)  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
868  | 
apply auto  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
869  | 
apply (unfold Units_def)  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
870  | 
apply auto  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
871  | 
done  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
872  | 
|
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
873  | 
lemma (in monoid) inv_char:  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
874  | 
"x \<in> carrier G \<Longrightarrow> y \<in> carrier G \<Longrightarrow> x \<otimes> y = \<one> \<Longrightarrow> y \<otimes> x = \<one> \<Longrightarrow> inv x = y"  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
875  | 
apply (subgoal_tac "x \<in> Units G")  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
876  | 
apply (subgoal_tac "y = inv x \<otimes> \<one>")  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
877  | 
apply simp  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
878  | 
apply (erule subst)  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
879  | 
apply (subst m_assoc [symmetric])  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
880  | 
apply auto  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
881  | 
apply (unfold Units_def)  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
882  | 
apply auto  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
883  | 
done  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
884  | 
|
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
885  | 
lemma (in comm_monoid) comm_inv_char: "x \<in> carrier G \<Longrightarrow> y \<in> carrier G \<Longrightarrow> x \<otimes> y = \<one> \<Longrightarrow> inv x = y"  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
886  | 
by (simp add: inv_char m_comm)  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
887  | 
|
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
888  | 
lemma (in ring) inv_neg_one [simp]: "inv (\<ominus> \<one>) = \<ominus> \<one>"  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
889  | 
apply (rule inv_char)  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
890  | 
apply (auto simp add: l_minus r_minus)  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
891  | 
done  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
892  | 
|
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
893  | 
lemma (in monoid) inv_eq_imp_eq: "x \<in> Units G \<Longrightarrow> y \<in> Units G \<Longrightarrow> inv x = inv y \<Longrightarrow> x = y"  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
894  | 
apply (subgoal_tac "inv (inv x) = inv (inv y)")  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
895  | 
apply (subst (asm) Units_inv_inv)+  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
896  | 
apply auto  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
897  | 
done  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
898  | 
|
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
899  | 
lemma (in ring) Units_minus_one_closed [intro]: "\<ominus> \<one> \<in> Units R"  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
900  | 
apply (unfold Units_def)  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
901  | 
apply auto  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
902  | 
apply (rule_tac x = "\<ominus> \<one>" in bexI)  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
903  | 
apply auto  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
904  | 
apply (simp add: l_minus r_minus)  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
905  | 
done  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
906  | 
|
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
907  | 
lemma (in ring) inv_eq_neg_one_eq: "x \<in> Units R \<Longrightarrow> inv x = \<ominus> \<one> \<longleftrightarrow> x = \<ominus> \<one>"  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
908  | 
apply auto  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
909  | 
apply (subst Units_inv_inv [symmetric])  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
910  | 
apply auto  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
911  | 
done  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
912  | 
|
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
913  | 
lemma (in monoid) inv_eq_one_eq: "x \<in> Units G \<Longrightarrow> inv x = \<one> \<longleftrightarrow> x = \<one>"  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
914  | 
by (metis Units_inv_inv inv_one)  | 
| 
 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 
paulson <lp15@cam.ac.uk> 
parents: 
68443 
diff
changeset
 | 
915  | 
|
| 
20318
 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 
ballarin 
parents:  
diff
changeset
 | 
916  | 
end  |