author | wenzelm |
Tue, 01 Jun 2004 12:33:50 +0200 | |
changeset 14854 | 61bdf2ae4dc5 |
parent 14254 | 342634f38451 |
child 14981 | e73f8140af78 |
permissions | -rw-r--r-- |
12075 | 1 |
(* Title: HOL/ex/Locales.thy |
2 |
ID: $Id$ |
|
12574 | 3 |
Author: Markus Wenzel, LMU München |
12075 | 4 |
License: GPL (GNU GENERAL PUBLIC LICENSE) |
5 |
*) |
|
6 |
||
12507 | 7 |
header {* Using locales in Isabelle/Isar *} |
12075 | 8 |
|
9 |
theory Locales = Main: |
|
10 |
||
12507 | 11 |
text_raw {* |
12 |
\newcommand{\isasyminv}{\isasyminverse} |
|
13 |
\newcommand{\isasymIN}{\isatext{\isakeyword{in}}} |
|
12965 | 14 |
\newcommand{\isasymINCLUDES}{\isatext{\isakeyword{includes}}} |
12507 | 15 |
*} |
16 |
||
17 |
subsection {* Overview *} |
|
18 |
||
12091 | 19 |
text {* |
12105 | 20 |
Locales provide a mechanism for encapsulating local contexts. The |
12574 | 21 |
original version due to Florian Kammüller \cite{Kamm-et-al:1999} |
13538 | 22 |
refers directly to Isabelle's meta-logic \cite{Paulson:1989}, which |
23 |
is minimal higher-order logic with connectives @{text "\<And>"} |
|
24 |
(universal quantification), @{text "\<Longrightarrow>"} (implication), and @{text |
|
25 |
"\<equiv>"} (equality). |
|
26 |
||
27 |
From this perspective, a locale is essentially a meta-level |
|
28 |
predicate, together with some infrastructure to manage the |
|
29 |
abstracted parameters (@{text "\<And>"}), assumptions (@{text "\<Longrightarrow>"}), and |
|
30 |
definitions for (@{text "\<equiv>"}) in a reasonable way during the proof |
|
31 |
process. This simple predicate view also provides a solid |
|
32 |
semantical basis for our specification concepts to be developed |
|
33 |
later. |
|
12507 | 34 |
|
13538 | 35 |
\medskip The present version of locales for Isabelle/Isar builds on |
36 |
top of the rich infrastructure of proof contexts |
|
37 |
\cite{Wenzel:1999,Wenzel:2001:isar-ref,Wenzel:2001:Thesis}, which in |
|
38 |
turn is based on the same meta-logic. Thus we achieve a tight |
|
39 |
integration with Isar proof texts, and a slightly more abstract view |
|
40 |
of the underlying logical concepts. An Isar proof context |
|
41 |
encapsulates certain language elements that correspond to @{text |
|
42 |
"\<And>/\<Longrightarrow>/\<equiv>"} at the level of structure proof texts. Moreover, there are |
|
43 |
extra-logical concepts like term abbreviations or local theorem |
|
44 |
attributes (declarations of simplification rules etc.) that are |
|
45 |
useful in applications (e.g.\ consider standard simplification rules |
|
46 |
declared in a group context). |
|
12507 | 47 |
|
13538 | 48 |
Locales also support concrete syntax, i.e.\ a localized version of |
12507 | 49 |
the existing concept of mixfix annotations of Isabelle |
50 |
\cite{Paulson:1994:Isabelle}. Furthermore, there is a separate |
|
51 |
concept of ``implicit structures'' that admits to refer to |
|
13538 | 52 |
particular locale parameters in a casual manner (basically a |
53 |
simplified version of the idea of ``anti-quotations'', or |
|
54 |
generalized de-Bruijn indexes as demonstrated elsewhere |
|
55 |
\cite[\S13--14]{Wenzel:2001:Isar-examples}). |
|
12105 | 56 |
|
12507 | 57 |
Implicit structures work particular well together with extensible |
13538 | 58 |
records in HOL \cite{Naraschewski-Wenzel:1998} (without the |
59 |
``object-oriented'' features discussed there as well). Thus we |
|
60 |
achieve a specification technique where record type schemes |
|
61 |
represent polymorphic signatures of mathematical structures, and |
|
62 |
actual locales describe the corresponding logical properties. |
|
63 |
Semantically speaking, such abstract mathematical structures are |
|
64 |
just predicates over record types. Due to type inference of |
|
65 |
simply-typed records (which subsumes structural subtyping) we arrive |
|
66 |
at succinct specification texts --- ``signature morphisms'' |
|
67 |
degenerate to implicit type-instantiations. Additional eye-candy is |
|
68 |
provided by the separate concept of ``indexed concrete syntax'' used |
|
69 |
for record selectors, so we get close to informal mathematical |
|
70 |
notation. |
|
12507 | 71 |
|
13538 | 72 |
\medskip Operations for building up locale contexts from existing |
73 |
ones include \emph{merge} (disjoint union) and \emph{rename} (of |
|
74 |
term parameters only, types are inferred automatically). Here we |
|
75 |
draw from existing traditions of algebraic specification languages. |
|
76 |
A structured specification corresponds to a directed acyclic graph |
|
77 |
of potentially renamed nodes (due to distributivity renames may |
|
78 |
pushed inside of merges). The result is a ``flattened'' list of |
|
79 |
primitive context elements in canonical order (corresponding to |
|
80 |
left-to-right reading of merges, while suppressing duplicates). |
|
12507 | 81 |
|
13538 | 82 |
\medskip The present version of Isabelle/Isar locales still lacks |
83 |
some important specification concepts. |
|
84 |
||
12507 | 85 |
\begin{itemize} |
12099 | 86 |
|
13538 | 87 |
\item Separate language elements for \emph{instantiation} of |
12507 | 88 |
locales. |
89 |
||
90 |
Currently users may simulate this to some extend by having primitive |
|
91 |
Isabelle/Isar operations (@{text of} for substitution and @{text OF} |
|
92 |
for composition, \cite{Wenzel:2001:isar-ref}) act on the |
|
93 |
automatically exported results stemming from different contexts. |
|
94 |
||
13538 | 95 |
\item Interpretation of locales (think of ``views'', ``functors'' |
96 |
etc.). |
|
12507 | 97 |
|
98 |
In principle one could directly work with functions over structures |
|
99 |
(extensible records), and predicates being derived from locale |
|
100 |
definitions. |
|
101 |
||
102 |
\end{itemize} |
|
103 |
||
104 |
\medskip Subsequently, we demonstrate some readily available |
|
105 |
concepts of Isabelle/Isar locales by some simple examples of |
|
106 |
abstract algebraic reasoning. |
|
12091 | 107 |
*} |
108 |
||
13538 | 109 |
|
12091 | 110 |
subsection {* Local contexts as mathematical structures *} |
111 |
||
12099 | 112 |
text {* |
13538 | 113 |
The following definitions of @{text group_context} and @{text |
114 |
abelian_group_context} merely encapsulate local parameters (with |
|
115 |
private syntax) and assumptions; local definitions of derived |
|
116 |
concepts could be given, too, but are unused below. |
|
12075 | 117 |
*} |
118 |
||
12574 | 119 |
locale group_context = |
12075 | 120 |
fixes prod :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<cdot>" 70) |
121 |
and inv :: "'a \<Rightarrow> 'a" ("(_\<inv>)" [1000] 999) |
|
122 |
and one :: 'a ("\<one>") |
|
123 |
assumes assoc: "(x \<cdot> y) \<cdot> z = x \<cdot> (y \<cdot> z)" |
|
124 |
and left_inv: "x\<inv> \<cdot> x = \<one>" |
|
125 |
and left_one: "\<one> \<cdot> x = x" |
|
126 |
||
12574 | 127 |
locale abelian_group_context = group_context + |
12075 | 128 |
assumes commute: "x \<cdot> y = y \<cdot> x" |
129 |
||
12099 | 130 |
text {* |
131 |
\medskip We may now prove theorems within a local context, just by |
|
132 |
including a directive ``@{text "(\<IN> name)"}'' in the goal |
|
133 |
specification. The final result will be stored within the named |
|
13538 | 134 |
locale, still holding the context; a second copy is exported to the |
135 |
enclosing theory context (with qualified name). |
|
12099 | 136 |
*} |
137 |
||
12574 | 138 |
theorem (in group_context) |
12075 | 139 |
right_inv: "x \<cdot> x\<inv> = \<one>" |
140 |
proof - |
|
141 |
have "x \<cdot> x\<inv> = \<one> \<cdot> (x \<cdot> x\<inv>)" by (simp only: left_one) |
|
142 |
also have "\<dots> = \<one> \<cdot> x \<cdot> x\<inv>" by (simp only: assoc) |
|
143 |
also have "\<dots> = (x\<inv>)\<inv> \<cdot> x\<inv> \<cdot> x \<cdot> x\<inv>" by (simp only: left_inv) |
|
144 |
also have "\<dots> = (x\<inv>)\<inv> \<cdot> (x\<inv> \<cdot> x) \<cdot> x\<inv>" by (simp only: assoc) |
|
145 |
also have "\<dots> = (x\<inv>)\<inv> \<cdot> \<one> \<cdot> x\<inv>" by (simp only: left_inv) |
|
146 |
also have "\<dots> = (x\<inv>)\<inv> \<cdot> (\<one> \<cdot> x\<inv>)" by (simp only: assoc) |
|
147 |
also have "\<dots> = (x\<inv>)\<inv> \<cdot> x\<inv>" by (simp only: left_one) |
|
148 |
also have "\<dots> = \<one>" by (simp only: left_inv) |
|
149 |
finally show ?thesis . |
|
150 |
qed |
|
151 |
||
12574 | 152 |
theorem (in group_context) |
12075 | 153 |
right_one: "x \<cdot> \<one> = x" |
154 |
proof - |
|
155 |
have "x \<cdot> \<one> = x \<cdot> (x\<inv> \<cdot> x)" by (simp only: left_inv) |
|
156 |
also have "\<dots> = x \<cdot> x\<inv> \<cdot> x" by (simp only: assoc) |
|
157 |
also have "\<dots> = \<one> \<cdot> x" by (simp only: right_inv) |
|
158 |
also have "\<dots> = x" by (simp only: left_one) |
|
159 |
finally show ?thesis . |
|
160 |
qed |
|
161 |
||
12099 | 162 |
text {* |
13538 | 163 |
Facts like @{text right_one} are available @{text group_context} as |
164 |
stated above. The exported version looses the additional |
|
165 |
infrastructure of Isar proof contexts (syntax etc.) retaining only |
|
166 |
the pure logical content: @{thm [source] group_context.right_one} |
|
167 |
becomes @{thm group_context.right_one} (in Isabelle outermost @{text |
|
168 |
\<And>} quantification is replaced by schematic variables). |
|
169 |
||
170 |
\medskip Apart from a named locale we may also refer to further |
|
171 |
context elements (parameters, assumptions, etc.) in an ad-hoc |
|
172 |
fashion, just for this particular statement. In the result (local |
|
173 |
or global), any additional elements are discharged as usual. |
|
12099 | 174 |
*} |
175 |
||
12574 | 176 |
theorem (in group_context) |
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12574
diff
changeset
|
177 |
assumes eq: "e \<cdot> x = x" |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12574
diff
changeset
|
178 |
shows one_equality: "\<one> = e" |
12075 | 179 |
proof - |
180 |
have "\<one> = x \<cdot> x\<inv>" by (simp only: right_inv) |
|
181 |
also have "\<dots> = (e \<cdot> x) \<cdot> x\<inv>" by (simp only: eq) |
|
182 |
also have "\<dots> = e \<cdot> (x \<cdot> x\<inv>)" by (simp only: assoc) |
|
183 |
also have "\<dots> = e \<cdot> \<one>" by (simp only: right_inv) |
|
184 |
also have "\<dots> = e" by (simp only: right_one) |
|
185 |
finally show ?thesis . |
|
186 |
qed |
|
187 |
||
12574 | 188 |
theorem (in group_context) |
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12574
diff
changeset
|
189 |
assumes eq: "x' \<cdot> x = \<one>" |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12574
diff
changeset
|
190 |
shows inv_equality: "x\<inv> = x'" |
12075 | 191 |
proof - |
192 |
have "x\<inv> = \<one> \<cdot> x\<inv>" by (simp only: left_one) |
|
193 |
also have "\<dots> = (x' \<cdot> x) \<cdot> x\<inv>" by (simp only: eq) |
|
194 |
also have "\<dots> = x' \<cdot> (x \<cdot> x\<inv>)" by (simp only: assoc) |
|
195 |
also have "\<dots> = x' \<cdot> \<one>" by (simp only: right_inv) |
|
196 |
also have "\<dots> = x'" by (simp only: right_one) |
|
197 |
finally show ?thesis . |
|
198 |
qed |
|
199 |
||
12574 | 200 |
theorem (in group_context) |
12075 | 201 |
inv_prod: "(x \<cdot> y)\<inv> = y\<inv> \<cdot> x\<inv>" |
202 |
proof (rule inv_equality) |
|
203 |
show "(y\<inv> \<cdot> x\<inv>) \<cdot> (x \<cdot> y) = \<one>" |
|
204 |
proof - |
|
205 |
have "(y\<inv> \<cdot> x\<inv>) \<cdot> (x \<cdot> y) = (y\<inv> \<cdot> (x\<inv> \<cdot> x)) \<cdot> y" by (simp only: assoc) |
|
206 |
also have "\<dots> = (y\<inv> \<cdot> \<one>) \<cdot> y" by (simp only: left_inv) |
|
207 |
also have "\<dots> = y\<inv> \<cdot> y" by (simp only: right_one) |
|
208 |
also have "\<dots> = \<one>" by (simp only: left_inv) |
|
209 |
finally show ?thesis . |
|
210 |
qed |
|
211 |
qed |
|
212 |
||
12099 | 213 |
text {* |
12105 | 214 |
\medskip Established results are automatically propagated through |
215 |
the hierarchy of locales. Below we establish a trivial fact in |
|
216 |
commutative groups, while referring both to theorems of @{text |
|
217 |
group} and the additional assumption of @{text abelian_group}. |
|
12099 | 218 |
*} |
219 |
||
12574 | 220 |
theorem (in abelian_group_context) |
12075 | 221 |
inv_prod': "(x \<cdot> y)\<inv> = x\<inv> \<cdot> y\<inv>" |
222 |
proof - |
|
223 |
have "(x \<cdot> y)\<inv> = y\<inv> \<cdot> x\<inv>" by (rule inv_prod) |
|
224 |
also have "\<dots> = x\<inv> \<cdot> y\<inv>" by (rule commute) |
|
225 |
finally show ?thesis . |
|
226 |
qed |
|
227 |
||
12099 | 228 |
text {* |
12507 | 229 |
We see that the initial import of @{text group} within the |
230 |
definition of @{text abelian_group} is actually evaluated |
|
231 |
dynamically. Thus any results in @{text group} are made available |
|
232 |
to the derived context of @{text abelian_group} as well. Note that |
|
12965 | 233 |
the alternative context element @{text \<INCLUDES>} would import |
12507 | 234 |
existing locales in a static fashion, without participating in |
235 |
further facts emerging later on. |
|
236 |
||
237 |
\medskip Some more properties of inversion in general group theory |
|
238 |
follow. |
|
12099 | 239 |
*} |
240 |
||
12574 | 241 |
theorem (in group_context) |
12075 | 242 |
inv_inv: "(x\<inv>)\<inv> = x" |
243 |
proof (rule inv_equality) |
|
244 |
show "x \<cdot> x\<inv> = \<one>" by (simp only: right_inv) |
|
245 |
qed |
|
246 |
||
12574 | 247 |
theorem (in group_context) |
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12574
diff
changeset
|
248 |
assumes eq: "x\<inv> = y\<inv>" |
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12574
diff
changeset
|
249 |
shows inv_inject: "x = y" |
12075 | 250 |
proof - |
251 |
have "x = x \<cdot> \<one>" by (simp only: right_one) |
|
252 |
also have "\<dots> = x \<cdot> (y\<inv> \<cdot> y)" by (simp only: left_inv) |
|
253 |
also have "\<dots> = x \<cdot> (x\<inv> \<cdot> y)" by (simp only: eq) |
|
254 |
also have "\<dots> = (x \<cdot> x\<inv>) \<cdot> y" by (simp only: assoc) |
|
255 |
also have "\<dots> = \<one> \<cdot> y" by (simp only: right_inv) |
|
256 |
also have "\<dots> = y" by (simp only: left_one) |
|
257 |
finally show ?thesis . |
|
258 |
qed |
|
259 |
||
12099 | 260 |
text {* |
261 |
\bigskip We see that this representation of structures as local |
|
262 |
contexts is rather light-weight and convenient to use for abstract |
|
263 |
reasoning. Here the ``components'' (the group operations) have been |
|
264 |
exhibited directly as context parameters; logically this corresponds |
|
13383 | 265 |
to a curried predicate definition: |
266 |
||
13419 | 267 |
@{thm [display] group_context_def [no_vars]} |
268 |
||
269 |
The corresponding introduction rule is as follows: |
|
270 |
||
13457 | 271 |
@{thm [display, mode = no_brackets] group_context.intro [no_vars]} |
13383 | 272 |
|
273 |
Occasionally, this ``externalized'' version of the informal idea of |
|
274 |
classes of tuple structures may cause some inconveniences, |
|
275 |
especially in meta-theoretical studies (involving functors from |
|
276 |
groups to groups, for example). |
|
12091 | 277 |
|
12507 | 278 |
Another minor drawback of the naive approach above is that concrete |
279 |
syntax will get lost on any kind of operation on the locale itself |
|
280 |
(such as renaming, copying, or instantiation). Whenever the |
|
281 |
particular terminology of local parameters is affected the |
|
282 |
associated syntax would have to be changed as well, which is hard to |
|
283 |
achieve formally. |
|
12099 | 284 |
*} |
285 |
||
286 |
||
12105 | 287 |
subsection {* Explicit structures referenced implicitly *} |
12091 | 288 |
|
12099 | 289 |
text {* |
12507 | 290 |
We introduce the same hierarchy of basic group structures as above, |
291 |
this time using extensible record types for the signature part, |
|
292 |
together with concrete syntax for selector functions. |
|
12099 | 293 |
*} |
294 |
||
12507 | 295 |
record 'a semigroup = |
296 |
prod :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<cdot>\<index>" 70) |
|
297 |
||
298 |
record 'a group = "'a semigroup" + |
|
299 |
inv :: "'a \<Rightarrow> 'a" ("(_\<inv>\<index>)" [1000] 999) |
|
300 |
one :: 'a ("\<one>\<index>") |
|
12099 | 301 |
|
302 |
text {* |
|
12507 | 303 |
The mixfix annotations above include a special ``structure index |
13457 | 304 |
indicator'' @{text \<index>} that makes grammar productions dependent on |
12507 | 305 |
certain parameters that have been declared as ``structure'' in a |
306 |
locale context later on. Thus we achieve casual notation as |
|
307 |
encountered in informal mathematics, e.g.\ @{text "x \<cdot> y"} for |
|
308 |
@{text "prod G x y"}. |
|
12099 | 309 |
|
12507 | 310 |
\medskip The following locale definitions introduce operate on a |
311 |
single parameter declared as ``\isakeyword{structure}''. Type |
|
312 |
inference takes care to fill in the appropriate record type schemes |
|
313 |
internally. |
|
12099 | 314 |
*} |
12091 | 315 |
|
12574 | 316 |
locale semigroup = |
13457 | 317 |
fixes S (structure) |
12507 | 318 |
assumes assoc: "(x \<cdot> y) \<cdot> z = x \<cdot> (y \<cdot> z)" |
12099 | 319 |
|
12574 | 320 |
locale group = semigroup G + |
12507 | 321 |
assumes left_inv: "x\<inv> \<cdot> x = \<one>" |
12099 | 322 |
and left_one: "\<one> \<cdot> x = x" |
323 |
||
14254
342634f38451
Isar/Locales: <loc>.intro and <loc>.axioms no longer intro? and elim? by
ballarin
parents:
13538
diff
changeset
|
324 |
declare semigroup.intro [intro?] |
342634f38451
Isar/Locales: <loc>.intro and <loc>.axioms no longer intro? and elim? by
ballarin
parents:
13538
diff
changeset
|
325 |
group.intro [intro?] group_axioms.intro [intro?] |
342634f38451
Isar/Locales: <loc>.intro and <loc>.axioms no longer intro? and elim? by
ballarin
parents:
13538
diff
changeset
|
326 |
|
12099 | 327 |
text {* |
12507 | 328 |
Note that we prefer to call the @{text group} record structure |
329 |
@{text G} rather than @{text S} inherited from @{text semigroup}. |
|
330 |
This does not affect our concrete syntax, which is only dependent on |
|
331 |
the \emph{positional} arrangements of currently active structures |
|
332 |
(actually only one above), rather than names. In fact, these |
|
333 |
parameter names rarely occur in the term language at all (due to the |
|
334 |
``indexed syntax'' facility of Isabelle). On the other hand, names |
|
335 |
of locale facts will get qualified accordingly, e.g. @{text S.assoc} |
|
336 |
versus @{text G.assoc}. |
|
337 |
||
12574 | 338 |
\medskip We may now proceed to prove results within @{text group} |
339 |
just as before for @{text group}. The subsequent proof texts are |
|
340 |
exactly the same as despite the more advanced internal arrangement. |
|
12099 | 341 |
*} |
342 |
||
12574 | 343 |
theorem (in group) |
12099 | 344 |
right_inv: "x \<cdot> x\<inv> = \<one>" |
345 |
proof - |
|
346 |
have "x \<cdot> x\<inv> = \<one> \<cdot> (x \<cdot> x\<inv>)" by (simp only: left_one) |
|
347 |
also have "\<dots> = \<one> \<cdot> x \<cdot> x\<inv>" by (simp only: assoc) |
|
348 |
also have "\<dots> = (x\<inv>)\<inv> \<cdot> x\<inv> \<cdot> x \<cdot> x\<inv>" by (simp only: left_inv) |
|
349 |
also have "\<dots> = (x\<inv>)\<inv> \<cdot> (x\<inv> \<cdot> x) \<cdot> x\<inv>" by (simp only: assoc) |
|
350 |
also have "\<dots> = (x\<inv>)\<inv> \<cdot> \<one> \<cdot> x\<inv>" by (simp only: left_inv) |
|
351 |
also have "\<dots> = (x\<inv>)\<inv> \<cdot> (\<one> \<cdot> x\<inv>)" by (simp only: assoc) |
|
352 |
also have "\<dots> = (x\<inv>)\<inv> \<cdot> x\<inv>" by (simp only: left_one) |
|
353 |
also have "\<dots> = \<one>" by (simp only: left_inv) |
|
354 |
finally show ?thesis . |
|
355 |
qed |
|
356 |
||
12574 | 357 |
theorem (in group) |
12099 | 358 |
right_one: "x \<cdot> \<one> = x" |
359 |
proof - |
|
360 |
have "x \<cdot> \<one> = x \<cdot> (x\<inv> \<cdot> x)" by (simp only: left_inv) |
|
361 |
also have "\<dots> = x \<cdot> x\<inv> \<cdot> x" by (simp only: assoc) |
|
362 |
also have "\<dots> = \<one> \<cdot> x" by (simp only: right_inv) |
|
363 |
also have "\<dots> = x" by (simp only: left_one) |
|
364 |
finally show ?thesis . |
|
365 |
qed |
|
366 |
||
367 |
text {* |
|
12574 | 368 |
\medskip Several implicit structures may be active at the same time. |
369 |
The concrete syntax facility for locales actually maintains indexed |
|
370 |
structures that may be references implicitly --- via mixfix |
|
371 |
annotations that have been decorated by an ``index argument'' |
|
372 |
(@{text \<index>}). |
|
12099 | 373 |
|
12574 | 374 |
The following synthetic example demonstrates how to refer to several |
375 |
structures of type @{text group} succinctly. We work with two |
|
376 |
versions of the @{text group} locale above. |
|
12099 | 377 |
*} |
12091 | 378 |
|
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12574
diff
changeset
|
379 |
lemma |
12955 | 380 |
includes group G |
381 |
includes group H |
|
13107 | 382 |
shows "x \<cdot> y \<cdot> \<one> = prod G (prod G x y) (one G)" |
383 |
and "x \<cdot>\<^sub>2 y \<cdot>\<^sub>2 \<one>\<^sub>2 = prod H (prod H x y) (one H)" |
|
384 |
and "x \<cdot> \<one>\<^sub>2 = prod G x (one H)" |
|
385 |
by (rule refl)+ |
|
12574 | 386 |
|
387 |
text {* |
|
388 |
Note that the trivial statements above need to be given as a |
|
389 |
simultaneous goal in order to have type-inference make the implicit |
|
390 |
typing of structures @{text G} and @{text H} agree. |
|
391 |
*} |
|
12091 | 392 |
|
13457 | 393 |
|
394 |
subsection {* Simple meta-theory of structures *} |
|
395 |
||
396 |
text {* |
|
397 |
The packaging of the logical specification as a predicate and the |
|
398 |
syntactic structure as a record type provides a reasonable starting |
|
399 |
point for simple meta-theoretic studies of mathematical structures. |
|
400 |
This includes operations on structures (also known as ``functors''), |
|
401 |
and statements about such constructions. |
|
402 |
||
403 |
For example, the direct product of semigroups works as follows. |
|
404 |
*} |
|
405 |
||
406 |
constdefs |
|
407 |
semigroup_product :: "'a semigroup \<Rightarrow> 'b semigroup \<Rightarrow> ('a \<times> 'b) semigroup" |
|
13538 | 408 |
"semigroup_product S T \<equiv> |
409 |
\<lparr>prod = \<lambda>p q. (prod S (fst p) (fst q), prod T (snd p) (snd q))\<rparr>" |
|
13457 | 410 |
|
411 |
lemma semigroup_product [intro]: |
|
412 |
assumes S: "semigroup S" |
|
413 |
and T: "semigroup T" |
|
414 |
shows "semigroup (semigroup_product S T)" |
|
415 |
proof |
|
416 |
fix p q r :: "'a \<times> 'b" |
|
417 |
have "prod S (prod S (fst p) (fst q)) (fst r) = |
|
418 |
prod S (fst p) (prod S (fst q) (fst r))" |
|
419 |
by (rule semigroup.assoc [OF S]) |
|
420 |
moreover have "prod T (prod T (snd p) (snd q)) (snd r) = |
|
421 |
prod T (snd p) (prod T (snd q) (snd r))" |
|
422 |
by (rule semigroup.assoc [OF T]) |
|
13538 | 423 |
ultimately |
424 |
show "prod (semigroup_product S T) (prod (semigroup_product S T) p q) r = |
|
13457 | 425 |
prod (semigroup_product S T) p (prod (semigroup_product S T) q r)" |
426 |
by (simp add: semigroup_product_def) |
|
427 |
qed |
|
428 |
||
429 |
text {* |
|
430 |
The above proof is fairly easy, but obscured by the lack of concrete |
|
431 |
syntax. In fact, we didn't make use of the infrastructure of |
|
432 |
locales, apart from the raw predicate definition of @{text |
|
433 |
semigroup}. |
|
434 |
||
435 |
The alternative version below uses local context expressions to |
|
436 |
achieve a succinct proof body. The resulting statement is exactly |
|
437 |
the same as before, even though its specification is a bit more |
|
438 |
complex. |
|
439 |
*} |
|
440 |
||
441 |
lemma |
|
442 |
includes semigroup S + semigroup T |
|
443 |
fixes U (structure) |
|
444 |
defines "U \<equiv> semigroup_product S T" |
|
445 |
shows "semigroup U" |
|
446 |
proof |
|
447 |
fix p q r :: "'a \<times> 'b" |
|
448 |
have "(fst p \<cdot>\<^sub>1 fst q) \<cdot>\<^sub>1 fst r = fst p \<cdot>\<^sub>1 (fst q \<cdot>\<^sub>1 fst r)" |
|
449 |
by (rule S.assoc) |
|
450 |
moreover have "(snd p \<cdot>\<^sub>2 snd q) \<cdot>\<^sub>2 snd r = snd p \<cdot>\<^sub>2 (snd q \<cdot>\<^sub>2 snd r)" |
|
451 |
by (rule T.assoc) |
|
452 |
ultimately show "(p \<cdot>\<^sub>3 q) \<cdot>\<^sub>3 r = p \<cdot>\<^sub>3 (q \<cdot>\<^sub>3 r)" |
|
453 |
by (simp add: U_def semigroup_product_def semigroup.defs) |
|
454 |
qed |
|
455 |
||
456 |
text {* |
|
457 |
Direct products of group structures may be defined in a similar |
|
458 |
manner, taking two further operations into account. Subsequently, |
|
459 |
we use high-level record operations to convert between different |
|
460 |
signature types explicitly; see also |
|
461 |
\cite[\S8.3]{isabelle-hol-book}. |
|
462 |
*} |
|
463 |
||
464 |
constdefs |
|
465 |
group_product :: "'a group \<Rightarrow> 'b group \<Rightarrow> ('a \<times> 'b) group" |
|
466 |
"group_product G H \<equiv> |
|
467 |
semigroup.extend |
|
468 |
(semigroup_product (semigroup.truncate G) (semigroup.truncate H)) |
|
469 |
(group.fields (\<lambda>p. (inv G (fst p), inv H (snd p))) (one G, one H))" |
|
470 |
||
471 |
lemma group_product_aux: |
|
472 |
includes group G + group H |
|
473 |
fixes I (structure) |
|
474 |
defines "I \<equiv> group_product G H" |
|
475 |
shows "group I" |
|
476 |
proof |
|
477 |
show "semigroup I" |
|
478 |
proof - |
|
479 |
let ?G' = "semigroup.truncate G" and ?H' = "semigroup.truncate H" |
|
480 |
have "prod (semigroup_product ?G' ?H') = prod I" |
|
13538 | 481 |
by (simp add: I_def group_product_def group.defs |
482 |
semigroup_product_def semigroup.defs) |
|
13457 | 483 |
moreover |
484 |
have "semigroup ?G'" and "semigroup ?H'" |
|
485 |
using prems by (simp_all add: semigroup_def semigroup.defs) |
|
486 |
then have "semigroup (semigroup_product ?G' ?H')" .. |
|
487 |
ultimately show ?thesis by (simp add: I_def semigroup_def) |
|
488 |
qed |
|
489 |
show "group_axioms I" |
|
490 |
proof |
|
491 |
fix p :: "'a \<times> 'b" |
|
492 |
have "(fst p)\<inv>\<^sub>1 \<cdot>\<^sub>1 fst p = \<one>\<^sub>1" |
|
493 |
by (rule G.left_inv) |
|
494 |
moreover have "(snd p)\<inv>\<^sub>2 \<cdot>\<^sub>2 snd p = \<one>\<^sub>2" |
|
495 |
by (rule H.left_inv) |
|
496 |
ultimately show "p\<inv>\<^sub>3 \<cdot>\<^sub>3 p = \<one>\<^sub>3" |
|
13538 | 497 |
by (simp add: I_def group_product_def group.defs |
498 |
semigroup_product_def semigroup.defs) |
|
13457 | 499 |
have "\<one>\<^sub>1 \<cdot>\<^sub>1 fst p = fst p" by (rule G.left_one) |
500 |
moreover have "\<one>\<^sub>2 \<cdot>\<^sub>2 snd p = snd p" by (rule H.left_one) |
|
501 |
ultimately show "\<one>\<^sub>3 \<cdot>\<^sub>3 p = p" |
|
13538 | 502 |
by (simp add: I_def group_product_def group.defs |
503 |
semigroup_product_def semigroup.defs) |
|
13457 | 504 |
qed |
505 |
qed |
|
506 |
||
507 |
theorem group_product: "group G \<Longrightarrow> group H \<Longrightarrow> group (group_product G H)" |
|
508 |
by (rule group_product_aux) (assumption | rule group.axioms)+ |
|
509 |
||
12099 | 510 |
end |