author | wenzelm |
Sat, 29 Mar 2008 19:14:08 +0100 | |
changeset 26486 | b65a5272b360 |
parent 17274 | 746bb4c56800 |
child 26871 | 996add9defab |
permissions | -rw-r--r-- |
8903 | 1 |
|
9146 | 2 |
header {* Basic group theory *} |
8890 | 3 |
|
16417 | 4 |
theory Group imports Main begin |
8890 | 5 |
|
8903 | 6 |
text {* |
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
7 |
\medskip\noindent The meta-level type system of Isabelle supports |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
8 |
\emph{intersections} and \emph{inclusions} of type classes. These |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
9 |
directly correspond to intersections and inclusions of type |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
10 |
predicates in a purely set theoretic sense. This is sufficient as a |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
11 |
means to describe simple hierarchies of structures. As an |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
12 |
illustration, we use the well-known example of semigroups, monoids, |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
13 |
general groups and Abelian groups. |
9146 | 14 |
*} |
8903 | 15 |
|
9146 | 16 |
subsection {* Monoids and Groups *} |
8903 | 17 |
|
18 |
text {* |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
19 |
First we declare some polymorphic constants required later for the |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
20 |
signature parts of our structures. |
9146 | 21 |
*} |
8890 | 22 |
|
23 |
consts |
|
10140 | 24 |
times :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<odot>" 70) |
11071 | 25 |
invers :: "'a \<Rightarrow> 'a" ("(_\<inv>)" [1000] 999) |
12344 | 26 |
one :: 'a ("\<one>") |
8890 | 27 |
|
8903 | 28 |
text {* |
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
29 |
\noindent Next we define class @{text monoid} of monoids with |
12344 | 30 |
operations @{text \<odot>} and @{text \<one>}. Note that multiple class |
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
31 |
axioms are allowed for user convenience --- they simply represent |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
32 |
the conjunction of their respective universal closures. |
9146 | 33 |
*} |
8890 | 34 |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
35 |
axclass monoid \<subseteq> type |
10140 | 36 |
assoc: "(x \<odot> y) \<odot> z = x \<odot> (y \<odot> z)" |
12344 | 37 |
left_unit: "\<one> \<odot> x = x" |
38 |
right_unit: "x \<odot> \<one> = x" |
|
8890 | 39 |
|
8903 | 40 |
text {* |
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
41 |
\noindent So class @{text monoid} contains exactly those types |
12344 | 42 |
@{text \<tau>} where @{text "\<odot> \<Colon> \<tau> \<Rightarrow> \<tau> \<Rightarrow> \<tau>"} and @{text "\<one> \<Colon> \<tau>"} |
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
43 |
are specified appropriately, such that @{text \<odot>} is associative and |
12344 | 44 |
@{text \<one>} is a left and right unit element for the @{text \<odot>} |
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
45 |
operation. |
9146 | 46 |
*} |
8903 | 47 |
|
48 |
text {* |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
49 |
\medskip Independently of @{text monoid}, we now define a linear |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
50 |
hierarchy of semigroups, general groups and Abelian groups. Note |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
51 |
that the names of class axioms are automatically qualified with each |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
52 |
class name, so we may re-use common names such as @{text assoc}. |
9146 | 53 |
*} |
8890 | 54 |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
55 |
axclass semigroup \<subseteq> type |
10140 | 56 |
assoc: "(x \<odot> y) \<odot> z = x \<odot> (y \<odot> z)" |
8890 | 57 |
|
11099 | 58 |
axclass group \<subseteq> semigroup |
12344 | 59 |
left_unit: "\<one> \<odot> x = x" |
60 |
left_inverse: "x\<inv> \<odot> x = \<one>" |
|
8890 | 61 |
|
11099 | 62 |
axclass agroup \<subseteq> group |
10140 | 63 |
commute: "x \<odot> y = y \<odot> x" |
8903 | 64 |
|
65 |
text {* |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
66 |
\noindent Class @{text group} inherits associativity of @{text \<odot>} |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
67 |
from @{text semigroup} and adds two further group axioms. Similarly, |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
68 |
@{text agroup} is defined as the subset of @{text group} such that |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
69 |
for all of its elements @{text \<tau>}, the operation @{text "\<odot> \<Colon> \<tau> \<Rightarrow> \<tau> \<Rightarrow> |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
70 |
\<tau>"} is even commutative. |
9146 | 71 |
*} |
8903 | 72 |
|
73 |
||
9146 | 74 |
subsection {* Abstract reasoning *} |
8890 | 75 |
|
76 |
text {* |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
77 |
In a sense, axiomatic type classes may be viewed as \emph{abstract |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
78 |
theories}. Above class definitions gives rise to abstract axioms |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
79 |
@{text assoc}, @{text left_unit}, @{text left_inverse}, @{text |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
80 |
commute}, where any of these contain a type variable @{text "'a \<Colon> |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
81 |
c"} that is restricted to types of the corresponding class @{text |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
82 |
c}. \emph{Sort constraints} like this express a logical |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
83 |
precondition for the whole formula. For example, @{text assoc} |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
84 |
states that for all @{text \<tau>}, provided that @{text "\<tau> \<Colon> |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
85 |
semigroup"}, the operation @{text "\<odot> \<Colon> \<tau> \<Rightarrow> \<tau> \<Rightarrow> \<tau>"} is associative. |
8903 | 86 |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
87 |
\medskip From a technical point of view, abstract axioms are just |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
88 |
ordinary Isabelle theorems, which may be used in proofs without |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
89 |
special treatment. Such ``abstract proofs'' usually yield new |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
90 |
``abstract theorems''. For example, we may now derive the following |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
91 |
well-known laws of general groups. |
9146 | 92 |
*} |
8890 | 93 |
|
12344 | 94 |
theorem group_right_inverse: "x \<odot> x\<inv> = (\<one>\<Colon>'a\<Colon>group)" |
9146 | 95 |
proof - |
12344 | 96 |
have "x \<odot> x\<inv> = \<one> \<odot> (x \<odot> x\<inv>)" |
17274
746bb4c56800
axclass: name space prefix is now "c_class" instead of just "c";
wenzelm
parents:
16417
diff
changeset
|
97 |
by (simp only: group_class.left_unit) |
12344 | 98 |
also have "... = \<one> \<odot> x \<odot> x\<inv>" |
17274
746bb4c56800
axclass: name space prefix is now "c_class" instead of just "c";
wenzelm
parents:
16417
diff
changeset
|
99 |
by (simp only: semigroup_class.assoc) |
10140 | 100 |
also have "... = (x\<inv>)\<inv> \<odot> x\<inv> \<odot> x \<odot> x\<inv>" |
17274
746bb4c56800
axclass: name space prefix is now "c_class" instead of just "c";
wenzelm
parents:
16417
diff
changeset
|
101 |
by (simp only: group_class.left_inverse) |
10140 | 102 |
also have "... = (x\<inv>)\<inv> \<odot> (x\<inv> \<odot> x) \<odot> x\<inv>" |
17274
746bb4c56800
axclass: name space prefix is now "c_class" instead of just "c";
wenzelm
parents:
16417
diff
changeset
|
103 |
by (simp only: semigroup_class.assoc) |
12344 | 104 |
also have "... = (x\<inv>)\<inv> \<odot> \<one> \<odot> x\<inv>" |
17274
746bb4c56800
axclass: name space prefix is now "c_class" instead of just "c";
wenzelm
parents:
16417
diff
changeset
|
105 |
by (simp only: group_class.left_inverse) |
12344 | 106 |
also have "... = (x\<inv>)\<inv> \<odot> (\<one> \<odot> x\<inv>)" |
17274
746bb4c56800
axclass: name space prefix is now "c_class" instead of just "c";
wenzelm
parents:
16417
diff
changeset
|
107 |
by (simp only: semigroup_class.assoc) |
10140 | 108 |
also have "... = (x\<inv>)\<inv> \<odot> x\<inv>" |
17274
746bb4c56800
axclass: name space prefix is now "c_class" instead of just "c";
wenzelm
parents:
16417
diff
changeset
|
109 |
by (simp only: group_class.left_unit) |
12344 | 110 |
also have "... = \<one>" |
17274
746bb4c56800
axclass: name space prefix is now "c_class" instead of just "c";
wenzelm
parents:
16417
diff
changeset
|
111 |
by (simp only: group_class.left_inverse) |
9146 | 112 |
finally show ?thesis . |
113 |
qed |
|
8890 | 114 |
|
115 |
text {* |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
116 |
\noindent With @{text group_right_inverse} already available, @{text |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
117 |
group_right_unit}\label{thm:group-right-unit} is now established |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
118 |
much easier. |
9146 | 119 |
*} |
8890 | 120 |
|
12344 | 121 |
theorem group_right_unit: "x \<odot> \<one> = (x\<Colon>'a\<Colon>group)" |
9146 | 122 |
proof - |
12344 | 123 |
have "x \<odot> \<one> = x \<odot> (x\<inv> \<odot> x)" |
17274
746bb4c56800
axclass: name space prefix is now "c_class" instead of just "c";
wenzelm
parents:
16417
diff
changeset
|
124 |
by (simp only: group_class.left_inverse) |
10140 | 125 |
also have "... = x \<odot> x\<inv> \<odot> x" |
17274
746bb4c56800
axclass: name space prefix is now "c_class" instead of just "c";
wenzelm
parents:
16417
diff
changeset
|
126 |
by (simp only: semigroup_class.assoc) |
12344 | 127 |
also have "... = \<one> \<odot> x" |
9146 | 128 |
by (simp only: group_right_inverse) |
129 |
also have "... = x" |
|
17274
746bb4c56800
axclass: name space prefix is now "c_class" instead of just "c";
wenzelm
parents:
16417
diff
changeset
|
130 |
by (simp only: group_class.left_unit) |
9146 | 131 |
finally show ?thesis . |
132 |
qed |
|
8890 | 133 |
|
8903 | 134 |
text {* |
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
135 |
\medskip Abstract theorems may be instantiated to only those types |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
136 |
@{text \<tau>} where the appropriate class membership @{text "\<tau> \<Colon> c"} is |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
137 |
known at Isabelle's type signature level. Since we have @{text |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
138 |
"agroup \<subseteq> group \<subseteq> semigroup"} by definition, all theorems of @{text |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
139 |
semigroup} and @{text group} are automatically inherited by @{text |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
140 |
group} and @{text agroup}. |
9146 | 141 |
*} |
8890 | 142 |
|
143 |
||
9146 | 144 |
subsection {* Abstract instantiation *} |
8903 | 145 |
|
146 |
text {* |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
147 |
From the definition, the @{text monoid} and @{text group} classes |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
148 |
have been independent. Note that for monoids, @{text right_unit} |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
149 |
had to be included as an axiom, but for groups both @{text |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
150 |
right_unit} and @{text right_inverse} are derivable from the other |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
151 |
axioms. With @{text group_right_unit} derived as a theorem of group |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
152 |
theory (see page~\pageref{thm:group-right-unit}), we may now |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
153 |
instantiate @{text "monoid \<subseteq> semigroup"} and @{text "group \<subseteq> |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
154 |
monoid"} properly as follows (cf.\ \figref{fig:monoid-group}). |
8903 | 155 |
|
156 |
\begin{figure}[htbp] |
|
157 |
\begin{center} |
|
158 |
\small |
|
159 |
\unitlength 0.6mm |
|
160 |
\begin{picture}(65,90)(0,-10) |
|
161 |
\put(15,10){\line(0,1){10}} \put(15,30){\line(0,1){10}} |
|
162 |
\put(15,50){\line(1,1){10}} \put(35,60){\line(1,-1){10}} |
|
10140 | 163 |
\put(15,5){\makebox(0,0){@{text agroup}}} |
164 |
\put(15,25){\makebox(0,0){@{text group}}} |
|
165 |
\put(15,45){\makebox(0,0){@{text semigroup}}} |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
166 |
\put(30,65){\makebox(0,0){@{text type}}} \put(50,45){\makebox(0,0){@{text monoid}}} |
8903 | 167 |
\end{picture} |
168 |
\hspace{4em} |
|
169 |
\begin{picture}(30,90)(0,0) |
|
170 |
\put(15,10){\line(0,1){10}} \put(15,30){\line(0,1){10}} |
|
171 |
\put(15,50){\line(0,1){10}} \put(15,70){\line(0,1){10}} |
|
10140 | 172 |
\put(15,5){\makebox(0,0){@{text agroup}}} |
173 |
\put(15,25){\makebox(0,0){@{text group}}} |
|
174 |
\put(15,45){\makebox(0,0){@{text monoid}}} |
|
175 |
\put(15,65){\makebox(0,0){@{text semigroup}}} |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
176 |
\put(15,85){\makebox(0,0){@{text type}}} |
8903 | 177 |
\end{picture} |
178 |
\caption{Monoids and groups: according to definition, and by proof} |
|
179 |
\label{fig:monoid-group} |
|
180 |
\end{center} |
|
181 |
\end{figure} |
|
9146 | 182 |
*} |
8890 | 183 |
|
11099 | 184 |
instance monoid \<subseteq> semigroup |
10309 | 185 |
proof |
10140 | 186 |
fix x y z :: "'a\<Colon>monoid" |
187 |
show "x \<odot> y \<odot> z = x \<odot> (y \<odot> z)" |
|
17274
746bb4c56800
axclass: name space prefix is now "c_class" instead of just "c";
wenzelm
parents:
16417
diff
changeset
|
188 |
by (rule monoid_class.assoc) |
9146 | 189 |
qed |
8890 | 190 |
|
11099 | 191 |
instance group \<subseteq> monoid |
10309 | 192 |
proof |
10140 | 193 |
fix x y z :: "'a\<Colon>group" |
194 |
show "x \<odot> y \<odot> z = x \<odot> (y \<odot> z)" |
|
17274
746bb4c56800
axclass: name space prefix is now "c_class" instead of just "c";
wenzelm
parents:
16417
diff
changeset
|
195 |
by (rule semigroup_class.assoc) |
12344 | 196 |
show "\<one> \<odot> x = x" |
17274
746bb4c56800
axclass: name space prefix is now "c_class" instead of just "c";
wenzelm
parents:
16417
diff
changeset
|
197 |
by (rule group_class.left_unit) |
12344 | 198 |
show "x \<odot> \<one> = x" |
9146 | 199 |
by (rule group_right_unit) |
200 |
qed |
|
8890 | 201 |
|
8903 | 202 |
text {* |
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
203 |
\medskip The $\INSTANCE$ command sets up an appropriate goal that |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
204 |
represents the class inclusion (or type arity, see |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
205 |
\secref{sec:inst-arity}) to be proven (see also |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
206 |
\cite{isabelle-isar-ref}). The initial proof step causes |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
207 |
back-chaining of class membership statements wrt.\ the hierarchy of |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
208 |
any classes defined in the current theory; the effect is to reduce |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
209 |
to the initial statement to a number of goals that directly |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
210 |
correspond to any class axioms encountered on the path upwards |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
211 |
through the class hierarchy. |
9146 | 212 |
*} |
8890 | 213 |
|
214 |
||
9146 | 215 |
subsection {* Concrete instantiation \label{sec:inst-arity} *} |
8903 | 216 |
|
217 |
text {* |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
218 |
So far we have covered the case of the form $\INSTANCE$~@{text |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
219 |
"c\<^sub>1 \<subseteq> c\<^sub>2"}, namely \emph{abstract instantiation} --- |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
220 |
$c@1$ is more special than @{text "c\<^sub>1"} and thus an instance |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
221 |
of @{text "c\<^sub>2"}. Even more interesting for practical |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
222 |
applications are \emph{concrete instantiations} of axiomatic type |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
223 |
classes. That is, certain simple schemes @{text "(\<alpha>\<^sub>1, \<dots>, |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
224 |
\<alpha>\<^sub>n) t \<Colon> c"} of class membership may be established at the |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
225 |
logical level and then transferred to Isabelle's type signature |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
226 |
level. |
8903 | 227 |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
228 |
\medskip As a typical example, we show that type @{typ bool} with |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
229 |
exclusive-or as @{text \<odot>} operation, identity as @{text \<inv>}, and |
12344 | 230 |
@{term False} as @{text \<one>} forms an Abelian group. |
9146 | 231 |
*} |
8903 | 232 |
|
9306 | 233 |
defs (overloaded) |
10140 | 234 |
times_bool_def: "x \<odot> y \<equiv> x \<noteq> (y\<Colon>bool)" |
235 |
inverse_bool_def: "x\<inv> \<equiv> x\<Colon>bool" |
|
12344 | 236 |
unit_bool_def: "\<one> \<equiv> False" |
8890 | 237 |
|
8903 | 238 |
text {* |
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
239 |
\medskip It is important to note that above $\DEFS$ are just |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
240 |
overloaded meta-level constant definitions, where type classes are |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
241 |
not yet involved at all. This form of constant definition with |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
242 |
overloading (and optional recursion over the syntactic structure of |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
243 |
simple types) are admissible as definitional extensions of plain HOL |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
244 |
\cite{Wenzel:1997:TPHOL}. The Haskell-style type system is not |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
245 |
required for overloading. Nevertheless, overloaded definitions are |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
246 |
best applied in the context of type classes. |
8903 | 247 |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
248 |
\medskip Since we have chosen above $\DEFS$ of the generic group |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
249 |
operations on type @{typ bool} appropriately, the class membership |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
250 |
@{text "bool \<Colon> agroup"} may be now derived as follows. |
9146 | 251 |
*} |
8903 | 252 |
|
9146 | 253 |
instance bool :: agroup |
8890 | 254 |
proof (intro_classes, |
9146 | 255 |
unfold times_bool_def inverse_bool_def unit_bool_def) |
256 |
fix x y z |
|
10140 | 257 |
show "((x \<noteq> y) \<noteq> z) = (x \<noteq> (y \<noteq> z))" by blast |
258 |
show "(False \<noteq> x) = x" by blast |
|
259 |
show "(x \<noteq> x) = False" by blast |
|
260 |
show "(x \<noteq> y) = (y \<noteq> x)" by blast |
|
9146 | 261 |
qed |
8890 | 262 |
|
8903 | 263 |
text {* |
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
264 |
The result of an $\INSTANCE$ statement is both expressed as a |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
265 |
theorem of Isabelle's meta-logic, and as a type arity of the type |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
266 |
signature. The latter enables type-inference system to take care of |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
267 |
this new instance automatically. |
8903 | 268 |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
269 |
\medskip We could now also instantiate our group theory classes to |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
270 |
many other concrete types. For example, @{text "int \<Colon> agroup"} |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
271 |
(e.g.\ by defining @{text \<odot>} as addition, @{text \<inv>} as negation |
12344 | 272 |
and @{text \<one>} as zero) or @{text "list \<Colon> (type) semigroup"} |
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
273 |
(e.g.\ if @{text \<odot>} is defined as list append). Thus, the |
12344 | 274 |
characteristic constants @{text \<odot>}, @{text \<inv>}, @{text \<one>} |
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
275 |
really become overloaded, i.e.\ have different meanings on different |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
276 |
types. |
9146 | 277 |
*} |
8903 | 278 |
|
279 |
||
9146 | 280 |
subsection {* Lifting and Functors *} |
8903 | 281 |
|
282 |
text {* |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
283 |
As already mentioned above, overloading in the simply-typed HOL |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
284 |
systems may include recursion over the syntactic structure of types. |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
285 |
That is, definitional equations @{text "c\<^sup>\<tau> \<equiv> t"} may also |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
286 |
contain constants of name @{text c} on the right-hand side --- if |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
287 |
these have types that are structurally simpler than @{text \<tau>}. |
8903 | 288 |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
289 |
This feature enables us to \emph{lift operations}, say to Cartesian |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
290 |
products, direct sums or function spaces. Subsequently we lift |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
291 |
@{text \<odot>} component-wise to binary products @{typ "'a \<times> 'b"}. |
9146 | 292 |
*} |
8890 | 293 |
|
9306 | 294 |
defs (overloaded) |
10140 | 295 |
times_prod_def: "p \<odot> q \<equiv> (fst p \<odot> fst q, snd p \<odot> snd q)" |
8890 | 296 |
|
8903 | 297 |
text {* |
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
298 |
It is very easy to see that associativity of @{text \<odot>} on @{typ 'a} |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
299 |
and @{text \<odot>} on @{typ 'b} transfers to @{text \<odot>} on @{typ "'a \<times> |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
300 |
'b"}. Hence the binary type constructor @{text \<odot>} maps semigroups |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
301 |
to semigroups. This may be established formally as follows. |
9146 | 302 |
*} |
8903 | 303 |
|
9146 | 304 |
instance * :: (semigroup, semigroup) semigroup |
305 |
proof (intro_classes, unfold times_prod_def) |
|
10140 | 306 |
fix p q r :: "'a\<Colon>semigroup \<times> 'b\<Colon>semigroup" |
8890 | 307 |
show |
10140 | 308 |
"(fst (fst p \<odot> fst q, snd p \<odot> snd q) \<odot> fst r, |
309 |
snd (fst p \<odot> fst q, snd p \<odot> snd q) \<odot> snd r) = |
|
310 |
(fst p \<odot> fst (fst q \<odot> fst r, snd q \<odot> snd r), |
|
311 |
snd p \<odot> snd (fst q \<odot> fst r, snd q \<odot> snd r))" |
|
17274
746bb4c56800
axclass: name space prefix is now "c_class" instead of just "c";
wenzelm
parents:
16417
diff
changeset
|
312 |
by (simp add: semigroup_class.assoc) |
9146 | 313 |
qed |
8890 | 314 |
|
8903 | 315 |
text {* |
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
316 |
Thus, if we view class instances as ``structures'', then overloaded |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
317 |
constant definitions with recursion over types indirectly provide |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
318 |
some kind of ``functors'' --- i.e.\ mappings between abstract |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
319 |
theories. |
9146 | 320 |
*} |
8903 | 321 |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
11099
diff
changeset
|
322 |
end |