| author | wenzelm | 
| Fri, 30 Mar 2012 17:22:17 +0200 | |
| changeset 47230 | 6584098d5378 | 
| parent 35849 | b5522b51cb1e | 
| child 58860 | fee7cfa69c50 | 
| permissions | -rw-r--r-- | 
| 14706 | 1 | (* Title: HOL/Algebra/Module.thy | 
| 13936 | 2 | Author: Clemens Ballarin, started 15 April 2003 | 
| 3 | Copyright: Clemens Ballarin | |
| 4 | *) | |
| 5 | ||
| 35849 | 6 | theory Module | 
| 7 | imports Ring | |
| 8 | begin | |
| 13936 | 9 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
20168diff
changeset | 10 | section {* Modules over an Abelian Group *}
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
20168diff
changeset | 11 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
20168diff
changeset | 12 | subsection {* Definitions *}
 | 
| 13936 | 13 | |
| 14 | record ('a, 'b) module = "'b ring" +
 | |
| 15 | smult :: "['a, 'b] => 'b" (infixl "\<odot>\<index>" 70) | |
| 16 | ||
| 29237 | 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: 
14706diff
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: 
14706diff
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: 
14706diff
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: 
14706diff
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: 
14706diff
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: 
14706diff
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: 
14706diff
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: 
14706diff
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: 
14706diff
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: 
14706diff
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: 
14706diff
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: 
20318diff
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: 
14706diff
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: 
14706diff
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: 
14706diff
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: 
14706diff
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: 
14706diff
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: 
14706diff
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" | 
| 19984 
29bb4659f80a
Method intro_locales replaced by intro_locales and unfold_locales.
 ballarin parents: 
19931diff
changeset | 79 | apply intro_locales | 
| 27714 
27b4d7c01f8b
Tuned (for the sake of a meaningless log entry).
 ballarin parents: 
20318diff
changeset | 80 | apply (rule cring.axioms ring.axioms abelian_group.axioms comm_monoid.axioms assms)+ | 
| 19931 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 ballarin parents: 
19783diff
changeset | 81 | apply (rule module_axioms.intro) | 
| 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 ballarin parents: 
19783diff
changeset | 82 | apply (simp add: smult_closed) | 
| 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 ballarin parents: 
19783diff
changeset | 83 | apply (simp add: smult_l_distr) | 
| 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 ballarin parents: 
19783diff
changeset | 84 | apply (simp add: smult_r_distr) | 
| 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 ballarin parents: 
19783diff
changeset | 85 | apply (simp add: smult_assoc1) | 
| 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 ballarin parents: 
19783diff
changeset | 86 | apply (simp add: smult_one) | 
| 27714 
27b4d7c01f8b
Tuned (for the sake of a meaningless log entry).
 ballarin parents: 
20318diff
changeset | 87 | apply (rule cring.axioms ring.axioms abelian_group.axioms comm_monoid.axioms assms)+ | 
| 19931 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 ballarin parents: 
19783diff
changeset | 88 | apply (rule algebra_axioms.intro) | 
| 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 ballarin parents: 
19783diff
changeset | 89 | apply (simp add: smult_assoc2) | 
| 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 ballarin parents: 
19783diff
changeset | 90 | done | 
| 13936 | 91 | |
| 92 | lemma (in algebra) R_cring: | |
| 93 | "cring R" | |
| 28823 | 94 | .. | 
| 13936 | 95 | |
| 96 | lemma (in algebra) M_cring: | |
| 97 | "cring M" | |
| 28823 | 98 | .. | 
| 13936 | 99 | |
| 100 | lemma (in algebra) module: | |
| 101 | "module R M" | |
| 102 | by (auto intro: moduleI R_cring is_abelian_group | |
| 103 | smult_l_distr smult_r_distr smult_assoc1) | |
| 104 | ||
| 14651 | 105 | |
| 13936 | 106 | subsection {* Basic Properties of Algebras *}
 | 
| 107 | ||
| 108 | lemma (in algebra) smult_l_null [simp]: | |
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14706diff
changeset | 109 | "x \<in> carrier M ==> \<zero> \<odot>\<^bsub>M\<^esub> x = \<zero>\<^bsub>M\<^esub>" | 
| 13936 | 110 | proof - | 
| 111 | assume M: "x \<in> carrier M" | |
| 20168 
ed7bced29e1b
Reimplemented algebra method; now controlled by attribute.
 ballarin parents: 
19984diff
changeset | 112 | note facts = M smult_closed [OF R.zero_closed] | 
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14706diff
changeset | 113 | from facts have "\<zero> \<odot>\<^bsub>M\<^esub> x = (\<zero> \<odot>\<^bsub>M\<^esub> x \<oplus>\<^bsub>M\<^esub> \<zero> \<odot>\<^bsub>M\<^esub> x) \<oplus>\<^bsub>M\<^esub> \<ominus>\<^bsub>M\<^esub> (\<zero> \<odot>\<^bsub>M\<^esub> x)" by algebra | 
| 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14706diff
changeset | 114 | also from M have "... = (\<zero> \<oplus> \<zero>) \<odot>\<^bsub>M\<^esub> x \<oplus>\<^bsub>M\<^esub> \<ominus>\<^bsub>M\<^esub> (\<zero> \<odot>\<^bsub>M\<^esub> x)" | 
| 13936 | 115 | by (simp add: smult_l_distr del: R.l_zero R.r_zero) | 
| 20168 
ed7bced29e1b
Reimplemented algebra method; now controlled by attribute.
 ballarin parents: 
19984diff
changeset | 116 | also from facts have "... = \<zero>\<^bsub>M\<^esub>" apply algebra apply algebra done | 
| 13936 | 117 | finally show ?thesis . | 
| 118 | qed | |
| 119 | ||
| 120 | lemma (in algebra) smult_r_null [simp]: | |
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14706diff
changeset | 121 | "a \<in> carrier R ==> a \<odot>\<^bsub>M\<^esub> \<zero>\<^bsub>M\<^esub> = \<zero>\<^bsub>M\<^esub>"; | 
| 13936 | 122 | proof - | 
| 123 | assume R: "a \<in> carrier R" | |
| 124 | note facts = R smult_closed | |
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14706diff
changeset | 125 | 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>)" | 
| 13936 | 126 | by algebra | 
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14706diff
changeset | 127 | 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 | 128 | by (simp add: smult_r_distr del: M.l_zero M.r_zero) | 
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14706diff
changeset | 129 | also from facts have "... = \<zero>\<^bsub>M\<^esub>" by algebra | 
| 13936 | 130 | finally show ?thesis . | 
| 131 | qed | |
| 132 | ||
| 133 | lemma (in algebra) smult_l_minus: | |
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14706diff
changeset | 134 | "[| a \<in> carrier R; x \<in> carrier M |] ==> (\<ominus>a) \<odot>\<^bsub>M\<^esub> x = \<ominus>\<^bsub>M\<^esub> (a \<odot>\<^bsub>M\<^esub> x)" | 
| 13936 | 135 | proof - | 
| 136 | assume RM: "a \<in> carrier R" "x \<in> carrier M" | |
| 20168 
ed7bced29e1b
Reimplemented algebra method; now controlled by attribute.
 ballarin parents: 
19984diff
changeset | 137 | 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: 
19984diff
changeset | 138 | 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: 
19984diff
changeset | 139 | note facts = RM a_smult ma_smult | 
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14706diff
changeset | 140 | 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)" | 
| 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14706diff
changeset | 141 | by algebra | 
| 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14706diff
changeset | 142 | 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 | 143 | by (simp add: smult_l_distr) | 
| 20168 
ed7bced29e1b
Reimplemented algebra method; now controlled by attribute.
 ballarin parents: 
19984diff
changeset | 144 | also from facts smult_l_null have "... = \<ominus>\<^bsub>M\<^esub>(a \<odot>\<^bsub>M\<^esub> x)" | 
| 
ed7bced29e1b
Reimplemented algebra method; now controlled by attribute.
 ballarin parents: 
19984diff
changeset | 145 | apply algebra apply algebra done | 
| 13936 | 146 | finally show ?thesis . | 
| 147 | qed | |
| 148 | ||
| 149 | lemma (in algebra) smult_r_minus: | |
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14706diff
changeset | 150 | "[| 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 | 151 | proof - | 
| 152 | assume RM: "a \<in> carrier R" "x \<in> carrier M" | |
| 153 | note facts = RM smult_closed | |
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14706diff
changeset | 154 | 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)" | 
| 13936 | 155 | by algebra | 
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14706diff
changeset | 156 | 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 | 157 | by (simp add: smult_r_distr) | 
| 15095 
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
 ballarin parents: 
14706diff
changeset | 158 | also from facts smult_r_null have "... = \<ominus>\<^bsub>M\<^esub>(a \<odot>\<^bsub>M\<^esub> x)" by algebra | 
| 13936 | 159 | finally show ?thesis . | 
| 160 | qed | |
| 161 | ||
| 162 | end |