author | wenzelm |
Fri, 03 May 2019 20:03:45 +0200 | |
changeset 70244 | 2ca87b481077 |
parent 68596 | 81086e6f5429 |
child 80914 | d97fdabd9e2b |
permissions | -rw-r--r-- |
14706 | 1 |
(* Title: HOL/Algebra/Module.thy |
13936 | 2 |
Author: Clemens Ballarin, started 15 April 2003 |
68581 | 3 |
with contributions by Martin Baillon |
13936 | 4 |
*) |
5 |
||
35849 | 6 |
theory Module |
7 |
imports Ring |
|
8 |
begin |
|
13936 | 9 |
|
61382 | 10 |
section \<open>Modules over an Abelian Group\<close> |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
20168
diff
changeset
|
11 |
|
61382 | 12 |
subsection \<open>Definitions\<close> |
13936 | 13 |
|
14 |
record ('a, 'b) module = "'b ring" + |
|
15 |
smult :: "['a, 'b] => 'b" (infixl "\<odot>\<index>" 70) |
|
16 |
||
61565
352c73a689da
Qualifiers in locale expressions default to mandatory regardless of the command.
ballarin
parents:
61382
diff
changeset
|
17 |
locale module = R?: cring + M?: abelian_group M for M (structure) + |
13936 | 18 |
assumes smult_closed [simp, intro]: |
15095
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14706
diff
changeset
|
19 |
"[| a \<in> carrier R; x \<in> carrier M |] ==> a \<odot>\<^bsub>M\<^esub> x \<in> carrier M" |
13936 | 20 |
and smult_l_distr: |
21 |
"[| a \<in> carrier R; b \<in> carrier R; x \<in> carrier M |] ==> |
|
15095
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14706
diff
changeset
|
22 |
(a \<oplus> b) \<odot>\<^bsub>M\<^esub> x = a \<odot>\<^bsub>M\<^esub> x \<oplus>\<^bsub>M\<^esub> b \<odot>\<^bsub>M\<^esub> x" |
13936 | 23 |
and smult_r_distr: |
24 |
"[| a \<in> carrier R; x \<in> carrier M; y \<in> carrier M |] ==> |
|
15095
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14706
diff
changeset
|
25 |
a \<odot>\<^bsub>M\<^esub> (x \<oplus>\<^bsub>M\<^esub> y) = a \<odot>\<^bsub>M\<^esub> x \<oplus>\<^bsub>M\<^esub> a \<odot>\<^bsub>M\<^esub> y" |
13936 | 26 |
and smult_assoc1: |
27 |
"[| a \<in> carrier R; b \<in> carrier R; x \<in> carrier M |] ==> |
|
15095
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14706
diff
changeset
|
28 |
(a \<otimes> b) \<odot>\<^bsub>M\<^esub> x = a \<odot>\<^bsub>M\<^esub> (b \<odot>\<^bsub>M\<^esub> x)" |
13936 | 29 |
and smult_one [simp]: |
15095
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14706
diff
changeset
|
30 |
"x \<in> carrier M ==> \<one> \<odot>\<^bsub>M\<^esub> x = x" |
13936 | 31 |
|
29237 | 32 |
locale algebra = module + cring M + |
13936 | 33 |
assumes smult_assoc2: |
34 |
"[| a \<in> carrier R; x \<in> carrier M; y \<in> carrier M |] ==> |
|
15095
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14706
diff
changeset
|
35 |
(a \<odot>\<^bsub>M\<^esub> x) \<otimes>\<^bsub>M\<^esub> y = a \<odot>\<^bsub>M\<^esub> (x \<otimes>\<^bsub>M\<^esub> y)" |
13936 | 36 |
|
37 |
lemma moduleI: |
|
19783 | 38 |
fixes R (structure) and M (structure) |
13936 | 39 |
assumes cring: "cring R" |
40 |
and abelian_group: "abelian_group M" |
|
41 |
and smult_closed: |
|
15095
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14706
diff
changeset
|
42 |
"!!a x. [| a \<in> carrier R; x \<in> carrier M |] ==> a \<odot>\<^bsub>M\<^esub> x \<in> carrier M" |
13936 | 43 |
and smult_l_distr: |
44 |
"!!a b x. [| a \<in> carrier R; b \<in> carrier R; x \<in> carrier M |] ==> |
|
15095
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14706
diff
changeset
|
45 |
(a \<oplus> b) \<odot>\<^bsub>M\<^esub> x = (a \<odot>\<^bsub>M\<^esub> x) \<oplus>\<^bsub>M\<^esub> (b \<odot>\<^bsub>M\<^esub> x)" |
13936 | 46 |
and smult_r_distr: |
47 |
"!!a x y. [| a \<in> carrier R; x \<in> carrier M; y \<in> carrier M |] ==> |
|
15095
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14706
diff
changeset
|
48 |
a \<odot>\<^bsub>M\<^esub> (x \<oplus>\<^bsub>M\<^esub> y) = (a \<odot>\<^bsub>M\<^esub> x) \<oplus>\<^bsub>M\<^esub> (a \<odot>\<^bsub>M\<^esub> y)" |
13936 | 49 |
and smult_assoc1: |
50 |
"!!a b x. [| a \<in> carrier R; b \<in> carrier R; x \<in> carrier M |] ==> |
|
15095
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14706
diff
changeset
|
51 |
(a \<otimes> b) \<odot>\<^bsub>M\<^esub> x = a \<odot>\<^bsub>M\<^esub> (b \<odot>\<^bsub>M\<^esub> x)" |
13936 | 52 |
and smult_one: |
15095
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14706
diff
changeset
|
53 |
"!!x. x \<in> carrier M ==> \<one> \<odot>\<^bsub>M\<^esub> x = x" |
13936 | 54 |
shows "module R M" |
55 |
by (auto intro: module.intro cring.axioms abelian_group.axioms |
|
27714
27b4d7c01f8b
Tuned (for the sake of a meaningless log entry).
ballarin
parents:
20318
diff
changeset
|
56 |
module_axioms.intro assms) |
13936 | 57 |
|
58 |
lemma algebraI: |
|
19783 | 59 |
fixes R (structure) and M (structure) |
13936 | 60 |
assumes R_cring: "cring R" |
61 |
and M_cring: "cring M" |
|
62 |
and smult_closed: |
|
15095
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14706
diff
changeset
|
63 |
"!!a x. [| a \<in> carrier R; x \<in> carrier M |] ==> a \<odot>\<^bsub>M\<^esub> x \<in> carrier M" |
13936 | 64 |
and smult_l_distr: |
65 |
"!!a b x. [| a \<in> carrier R; b \<in> carrier R; x \<in> carrier M |] ==> |
|
15095
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14706
diff
changeset
|
66 |
(a \<oplus> b) \<odot>\<^bsub>M\<^esub> x = (a \<odot>\<^bsub>M\<^esub> x) \<oplus>\<^bsub>M\<^esub> (b \<odot>\<^bsub>M\<^esub> x)" |
13936 | 67 |
and smult_r_distr: |
68 |
"!!a x y. [| a \<in> carrier R; x \<in> carrier M; y \<in> carrier M |] ==> |
|
15095
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14706
diff
changeset
|
69 |
a \<odot>\<^bsub>M\<^esub> (x \<oplus>\<^bsub>M\<^esub> y) = (a \<odot>\<^bsub>M\<^esub> x) \<oplus>\<^bsub>M\<^esub> (a \<odot>\<^bsub>M\<^esub> y)" |
13936 | 70 |
and smult_assoc1: |
71 |
"!!a b x. [| a \<in> carrier R; b \<in> carrier R; x \<in> carrier M |] ==> |
|
15095
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14706
diff
changeset
|
72 |
(a \<otimes> b) \<odot>\<^bsub>M\<^esub> x = a \<odot>\<^bsub>M\<^esub> (b \<odot>\<^bsub>M\<^esub> x)" |
13936 | 73 |
and smult_one: |
15095
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14706
diff
changeset
|
74 |
"!!x. x \<in> carrier M ==> (one R) \<odot>\<^bsub>M\<^esub> x = x" |
13936 | 75 |
and smult_assoc2: |
76 |
"!!a x y. [| a \<in> carrier R; x \<in> carrier M; y \<in> carrier M |] ==> |
|
15095
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14706
diff
changeset
|
77 |
(a \<odot>\<^bsub>M\<^esub> x) \<otimes>\<^bsub>M\<^esub> y = a \<odot>\<^bsub>M\<^esub> (x \<otimes>\<^bsub>M\<^esub> y)" |
13936 | 78 |
shows "algebra R M" |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
79 |
apply intro_locales |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
80 |
apply (rule cring.axioms ring.axioms abelian_group.axioms comm_monoid.axioms assms)+ |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
81 |
apply (rule module_axioms.intro) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
82 |
apply (simp add: smult_closed) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
83 |
apply (simp add: smult_l_distr) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
84 |
apply (simp add: smult_r_distr) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
85 |
apply (simp add: smult_assoc1) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
86 |
apply (simp add: smult_one) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
87 |
apply (rule cring.axioms ring.axioms abelian_group.axioms comm_monoid.axioms assms)+ |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
88 |
apply (rule algebra_axioms.intro) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
89 |
apply (simp add: smult_assoc2) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
90 |
done |
13936 | 91 |
|
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
92 |
lemma (in algebra) R_cring: "cring R" .. |
13936 | 93 |
|
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
94 |
lemma (in algebra) M_cring: "cring M" .. |
13936 | 95 |
|
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
96 |
lemma (in algebra) module: "module R M" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
97 |
by (auto intro: moduleI R_cring is_abelian_group smult_l_distr smult_r_distr smult_assoc1) |
13936 | 98 |
|
14651 | 99 |
|
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
100 |
subsection \<open>Basic Properties of Modules\<close> |
13936 | 101 |
|
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
102 |
lemma (in module) smult_l_null [simp]: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
103 |
"x \<in> carrier M ==> \<zero> \<odot>\<^bsub>M\<^esub> x = \<zero>\<^bsub>M\<^esub>" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
104 |
proof- |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
105 |
assume M : "x \<in> carrier M" |
20168
ed7bced29e1b
Reimplemented algebra method; now controlled by attribute.
ballarin
parents:
19984
diff
changeset
|
106 |
note facts = M smult_closed [OF R.zero_closed] |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
107 |
from facts have "\<zero> \<odot>\<^bsub>M\<^esub> x = (\<zero> \<oplus> \<zero>) \<odot>\<^bsub>M\<^esub> x " |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
108 |
using smult_l_distr by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
109 |
also have "... = \<zero> \<odot>\<^bsub>M\<^esub> x \<oplus>\<^bsub>M\<^esub> \<zero> \<odot>\<^bsub>M\<^esub> x" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
110 |
using smult_l_distr[of \<zero> \<zero> x] facts by auto |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
111 |
finally show "\<zero> \<odot>\<^bsub>M\<^esub> x = \<zero>\<^bsub>M\<^esub>" using facts |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
112 |
by (metis M.add.r_cancel_one') |
13936 | 113 |
qed |
114 |
||
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
115 |
lemma (in module) smult_r_null [simp]: |
58860 | 116 |
"a \<in> carrier R ==> a \<odot>\<^bsub>M\<^esub> \<zero>\<^bsub>M\<^esub> = \<zero>\<^bsub>M\<^esub>" |
13936 | 117 |
proof - |
118 |
assume R: "a \<in> carrier R" |
|
119 |
note facts = R smult_closed |
|
15095
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14706
diff
changeset
|
120 |
from facts have "a \<odot>\<^bsub>M\<^esub> \<zero>\<^bsub>M\<^esub> = (a \<odot>\<^bsub>M\<^esub> \<zero>\<^bsub>M\<^esub> \<oplus>\<^bsub>M\<^esub> a \<odot>\<^bsub>M\<^esub> \<zero>\<^bsub>M\<^esub>) \<oplus>\<^bsub>M\<^esub> \<ominus>\<^bsub>M\<^esub> (a \<odot>\<^bsub>M\<^esub> \<zero>\<^bsub>M\<^esub>)" |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
121 |
by (simp add: M.add.inv_solve_right) |
15095
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14706
diff
changeset
|
122 |
also from R have "... = a \<odot>\<^bsub>M\<^esub> (\<zero>\<^bsub>M\<^esub> \<oplus>\<^bsub>M\<^esub> \<zero>\<^bsub>M\<^esub>) \<oplus>\<^bsub>M\<^esub> \<ominus>\<^bsub>M\<^esub> (a \<odot>\<^bsub>M\<^esub> \<zero>\<^bsub>M\<^esub>)" |
13936 | 123 |
by (simp add: smult_r_distr del: M.l_zero M.r_zero) |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
124 |
also from facts have "... = \<zero>\<^bsub>M\<^esub>" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
125 |
by (simp add: M.r_neg) |
13936 | 126 |
finally show ?thesis . |
127 |
qed |
|
128 |
||
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
129 |
lemma (in module) smult_l_minus: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
130 |
"\<lbrakk> a \<in> carrier R; x \<in> carrier M \<rbrakk> \<Longrightarrow> (\<ominus>a) \<odot>\<^bsub>M\<^esub> x = \<ominus>\<^bsub>M\<^esub> (a \<odot>\<^bsub>M\<^esub> x)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
131 |
proof- |
13936 | 132 |
assume RM: "a \<in> carrier R" "x \<in> carrier M" |
20168
ed7bced29e1b
Reimplemented algebra method; now controlled by attribute.
ballarin
parents:
19984
diff
changeset
|
133 |
from RM have a_smult: "a \<odot>\<^bsub>M\<^esub> x \<in> carrier M" by simp |
ed7bced29e1b
Reimplemented algebra method; now controlled by attribute.
ballarin
parents:
19984
diff
changeset
|
134 |
from RM have ma_smult: "\<ominus>a \<odot>\<^bsub>M\<^esub> x \<in> carrier M" by simp |
ed7bced29e1b
Reimplemented algebra method; now controlled by attribute.
ballarin
parents:
19984
diff
changeset
|
135 |
note facts = RM a_smult ma_smult |
15095
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14706
diff
changeset
|
136 |
from facts have "(\<ominus>a) \<odot>\<^bsub>M\<^esub> x = (\<ominus>a \<odot>\<^bsub>M\<^esub> x \<oplus>\<^bsub>M\<^esub> a \<odot>\<^bsub>M\<^esub> x) \<oplus>\<^bsub>M\<^esub> \<ominus>\<^bsub>M\<^esub>(a \<odot>\<^bsub>M\<^esub> x)" |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
137 |
by (simp add: M.add.inv_solve_right) |
15095
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14706
diff
changeset
|
138 |
also from RM have "... = (\<ominus>a \<oplus> a) \<odot>\<^bsub>M\<^esub> x \<oplus>\<^bsub>M\<^esub> \<ominus>\<^bsub>M\<^esub>(a \<odot>\<^bsub>M\<^esub> x)" |
13936 | 139 |
by (simp add: smult_l_distr) |
20168
ed7bced29e1b
Reimplemented algebra method; now controlled by attribute.
ballarin
parents:
19984
diff
changeset
|
140 |
also from facts smult_l_null have "... = \<ominus>\<^bsub>M\<^esub>(a \<odot>\<^bsub>M\<^esub> x)" |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
141 |
by (simp add: R.l_neg) |
13936 | 142 |
finally show ?thesis . |
143 |
qed |
|
144 |
||
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
145 |
lemma (in module) smult_r_minus: |
15095
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14706
diff
changeset
|
146 |
"[| a \<in> carrier R; x \<in> carrier M |] ==> a \<odot>\<^bsub>M\<^esub> (\<ominus>\<^bsub>M\<^esub>x) = \<ominus>\<^bsub>M\<^esub> (a \<odot>\<^bsub>M\<^esub> x)" |
13936 | 147 |
proof - |
148 |
assume RM: "a \<in> carrier R" "x \<in> carrier M" |
|
149 |
note facts = RM smult_closed |
|
15095
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14706
diff
changeset
|
150 |
from facts have "a \<odot>\<^bsub>M\<^esub> (\<ominus>\<^bsub>M\<^esub>x) = (a \<odot>\<^bsub>M\<^esub> \<ominus>\<^bsub>M\<^esub>x \<oplus>\<^bsub>M\<^esub> a \<odot>\<^bsub>M\<^esub> x) \<oplus>\<^bsub>M\<^esub> \<ominus>\<^bsub>M\<^esub>(a \<odot>\<^bsub>M\<^esub> x)" |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
151 |
by (simp add: M.add.inv_solve_right) |
15095
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14706
diff
changeset
|
152 |
also from RM have "... = a \<odot>\<^bsub>M\<^esub> (\<ominus>\<^bsub>M\<^esub>x \<oplus>\<^bsub>M\<^esub> x) \<oplus>\<^bsub>M\<^esub> \<ominus>\<^bsub>M\<^esub>(a \<odot>\<^bsub>M\<^esub> x)" |
13936 | 153 |
by (simp add: smult_r_distr) |
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
154 |
also from facts smult_l_null have "... = \<ominus>\<^bsub>M\<^esub>(a \<odot>\<^bsub>M\<^esub> x)" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
155 |
by (metis M.add.inv_closed M.add.inv_solve_right M.l_neg R.zero_closed r_null smult_assoc1) |
13936 | 156 |
finally show ?thesis . |
157 |
qed |
|
158 |
||
68551
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
159 |
lemma (in module) finsum_smult_ldistr: |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
160 |
"\<lbrakk> finite A; a \<in> carrier R; f: A \<rightarrow> carrier M \<rbrakk> \<Longrightarrow> |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
161 |
a \<odot>\<^bsub>M\<^esub> (\<Oplus>\<^bsub>M\<^esub> i \<in> A. (f i)) = (\<Oplus>\<^bsub>M\<^esub> i \<in> A. ( a \<odot>\<^bsub>M\<^esub> (f i)))" |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
162 |
proof (induct set: finite) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
163 |
case empty then show ?case |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
164 |
by (metis M.finsum_empty M.zero_closed R.zero_closed r_null smult_assoc1 smult_l_null) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
165 |
next |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
166 |
case (insert x F) then show ?case |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
167 |
by (simp add: Pi_def smult_r_distr) |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
168 |
qed |
b680e74eb6f2
More on Algebra by Paulo and Martin
paulson <lp15@cam.ac.uk>
parents:
61565
diff
changeset
|
169 |
|
68592 | 170 |
|
171 |
||
172 |
subsection \<open>Submodules\<close> |
|
173 |
||
174 |
locale submodule = subgroup H "add_monoid M" for H and R :: "('a, 'b) ring_scheme" and M (structure) |
|
175 |
+ assumes smult_closed [simp, intro]: |
|
176 |
"\<lbrakk>a \<in> carrier R; x \<in> H\<rbrakk> \<Longrightarrow> a \<odot>\<^bsub>M\<^esub> x \<in> H" |
|
177 |
||
178 |
||
179 |
lemma (in module) submoduleI : |
|
180 |
assumes subset: "H \<subseteq> carrier M" |
|
181 |
and zero: "\<zero>\<^bsub>M\<^esub> \<in> H" |
|
182 |
and a_inv: "!!a. a \<in> H \<Longrightarrow> \<ominus>\<^bsub>M\<^esub> a \<in> H" |
|
183 |
and add : "\<And> a b. \<lbrakk>a \<in> H ; b \<in> H\<rbrakk> \<Longrightarrow> a \<oplus>\<^bsub>M\<^esub> b \<in> H" |
|
184 |
and smult_closed : "\<And> a x. \<lbrakk>a \<in> carrier R; x \<in> H\<rbrakk> \<Longrightarrow> a \<odot>\<^bsub>M\<^esub> x \<in> H" |
|
185 |
shows "submodule H R M" |
|
186 |
apply (intro submodule.intro subgroup.intro) |
|
187 |
using assms unfolding submodule_axioms_def |
|
188 |
by (simp_all add : a_inv_def) |
|
189 |
||
190 |
||
191 |
lemma (in module) submoduleE : |
|
192 |
assumes "submodule H R M" |
|
193 |
shows "H \<subseteq> carrier M" |
|
194 |
and "H \<noteq> {}" |
|
195 |
and "\<And>a. a \<in> H \<Longrightarrow> \<ominus>\<^bsub>M\<^esub> a \<in> H" |
|
196 |
and "\<And>a b. \<lbrakk>a \<in> carrier R; b \<in> H\<rbrakk> \<Longrightarrow> a \<odot>\<^bsub>M\<^esub> b \<in> H" |
|
197 |
and "\<And> a b. \<lbrakk>a \<in> H ; b \<in> H\<rbrakk> \<Longrightarrow> a \<oplus>\<^bsub>M\<^esub> b \<in> H" |
|
198 |
and "\<And> x. x \<in> H \<Longrightarrow> (a_inv M x) \<in> H" |
|
199 |
using group.subgroupE[of "add_monoid M" H, OF _ submodule.axioms(1)[OF assms]] a_comm_group |
|
200 |
submodule.smult_closed[OF assms] |
|
201 |
unfolding comm_group_def a_inv_def |
|
202 |
by auto |
|
203 |
||
204 |
||
205 |
lemma (in module) carrier_is_submodule : |
|
206 |
"submodule (carrier M) R M" |
|
207 |
apply (intro submoduleI) |
|
208 |
using a_comm_group group.inv_closed unfolding comm_group_def a_inv_def group_def monoid_def |
|
209 |
by auto |
|
210 |
||
211 |
lemma (in submodule) submodule_is_module : |
|
212 |
assumes "module R M" |
|
213 |
shows "module R (M\<lparr>carrier := H\<rparr>)" |
|
214 |
proof (auto intro! : moduleI abelian_group.intro) |
|
215 |
show "cring (R)" using assms unfolding module_def by auto |
|
216 |
show "abelian_monoid (M\<lparr>carrier := H\<rparr>)" |
|
217 |
using comm_monoid.submonoid_is_comm_monoid[OF _ subgroup_is_submonoid] assms |
|
218 |
unfolding abelian_monoid_def module_def abelian_group_def |
|
219 |
by auto |
|
220 |
thus "abelian_group_axioms (M\<lparr>carrier := H\<rparr>)" |
|
221 |
using subgroup_is_group assms |
|
222 |
unfolding abelian_group_axioms_def comm_group_def abelian_monoid_def module_def abelian_group_def |
|
223 |
by auto |
|
224 |
show "\<And>z. z \<in> H \<Longrightarrow> \<one>\<^bsub>R\<^esub> \<odot> z = z" |
|
68596
81086e6f5429
replaced subgroup_imp_subset in Modules
paulson <lp15@cam.ac.uk>
parents:
68592
diff
changeset
|
225 |
using subgroup.subset[OF subgroup_axioms] module.smult_one[OF assms] |
68592 | 226 |
by auto |
227 |
fix a b x y assume a : "a \<in> carrier R" and b : "b \<in> carrier R" and x : "x \<in> H" and y : "y \<in> H" |
|
228 |
show "(a \<oplus>\<^bsub>R\<^esub> b) \<odot> x = a \<odot> x \<oplus> b \<odot> x" |
|
68596
81086e6f5429
replaced subgroup_imp_subset in Modules
paulson <lp15@cam.ac.uk>
parents:
68592
diff
changeset
|
229 |
using a b x module.smult_l_distr[OF assms] subgroup.subset[OF subgroup_axioms] |
68592 | 230 |
by auto |
231 |
show "a \<odot> (x \<oplus> y) = a \<odot> x \<oplus> a \<odot> y" |
|
68596
81086e6f5429
replaced subgroup_imp_subset in Modules
paulson <lp15@cam.ac.uk>
parents:
68592
diff
changeset
|
232 |
using a x y module.smult_r_distr[OF assms] subgroup.subset[OF subgroup_axioms] |
68592 | 233 |
by auto |
234 |
show "a \<otimes>\<^bsub>R\<^esub> b \<odot> x = a \<odot> (b \<odot> x)" |
|
68596
81086e6f5429
replaced subgroup_imp_subset in Modules
paulson <lp15@cam.ac.uk>
parents:
68592
diff
changeset
|
235 |
using a b x module.smult_assoc1[OF assms] subgroup.subset[OF subgroup_axioms] |
68592 | 236 |
by auto |
237 |
qed |
|
238 |
||
239 |
||
240 |
lemma (in module) module_incl_imp_submodule : |
|
241 |
assumes "H \<subseteq> carrier M" |
|
242 |
and "module R (M\<lparr>carrier := H\<rparr>)" |
|
243 |
shows "submodule H R M" |
|
244 |
apply (intro submodule.intro) |
|
245 |
using add.group_incl_imp_subgroup[OF assms(1)] assms module.axioms(2)[OF assms(2)] |
|
246 |
module.smult_closed[OF assms(2)] |
|
247 |
unfolding abelian_group_def abelian_group_axioms_def comm_group_def submodule_axioms_def |
|
248 |
by simp_all |
|
249 |
||
250 |
||
13936 | 251 |
end |