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