author | haftmann |
Mon, 02 Oct 2006 23:00:49 +0200 | |
changeset 20833 | 4fcf8ddb54f5 |
parent 17567 | 20c0b69dd192 |
child 26480 | 544cef16045b |
permissions | -rw-r--r-- |
14586 | 1 |
(* |
2 |
Title: Locales in Isabelle/Isar |
|
3 |
Id: $Id$ |
|
4 |
Author: Clemens Ballarin, started 31 January 2003 |
|
5 |
Copyright (c) 2003 by Clemens Ballarin |
|
6 |
*) |
|
7 |
||
8 |
(*<*) |
|
16417 | 9 |
theory Locales imports Main begin |
14586 | 10 |
|
17567 | 11 |
ML {* print_mode := "type_brackets" :: !print_mode; *} |
14586 | 12 |
(*>*) |
13 |
||
14 |
section {* Overview *} |
|
15 |
||
16 |
text {* |
|
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
17 |
The present text is based on~\cite{Ballarin2004a}. It was updated |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
18 |
for for Isabelle2005, but does not cover locale interpretation. |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
19 |
|
14586 | 20 |
Locales are an extension of the Isabelle proof assistant. They |
21 |
provide support for modular reasoning. Locales were initially |
|
22 |
developed by Kamm\"uller~\cite{Kammuller2000} to support reasoning |
|
23 |
in abstract algebra, but are applied also in other domains --- for |
|
24 |
example, bytecode verification~\cite{Klein2003}. |
|
25 |
||
26 |
Kamm\"uller's original design, implemented in Isabelle99, provides, in |
|
27 |
addition to |
|
28 |
means for declaring locales, a set of ML functions that were used |
|
29 |
along with ML tactics in a proof. In the meantime, the input format |
|
30 |
for proof in Isabelle has changed and users write proof |
|
31 |
scripts in ML only rarely if at all. Two new proof styles are |
|
32 |
available, and can |
|
33 |
be used interchangeably: linear proof scripts that closely resemble ML |
|
34 |
tactics, and the structured Isar proof language by |
|
35 |
Wenzel~\cite{Wenzel2002a}. Subsequently, Wenzel re-implemented |
|
36 |
locales for |
|
37 |
the new proof format. The implementation, available with |
|
17567 | 38 |
Isabelle2002 or later, constitutes a complete re-design and exploits that |
14586 | 39 |
both Isar and locales are based on the notion of context, |
40 |
and thus locales are seen as a natural extension of Isar. |
|
41 |
Nevertheless, locales can also be used with proof scripts: |
|
42 |
their use does not require a deep understanding of the structured |
|
43 |
Isar proof style. |
|
44 |
||
45 |
At the same time, Wenzel considerably extended locales. The most |
|
46 |
important addition are locale expressions, which allow to combine |
|
47 |
locales more freely. Previously only |
|
48 |
linear inheritance was possible. Now locales support multiple |
|
49 |
inheritance through a normalisation algorithm. New are also |
|
50 |
structures, which provide special syntax for locale parameters that |
|
51 |
represent algebraic structures. |
|
52 |
||
53 |
Unfortunately, Wenzel provided only an implementation but hardly any |
|
54 |
documentation. Besides providing documentation, the present paper |
|
55 |
is a high-level description of locales, and in particular locale |
|
56 |
expressions. It is meant as a first step towards the semantics of |
|
57 |
locales, and also as a base for comparing locales with module concepts |
|
58 |
in other provers. It also constitutes the base for future |
|
59 |
extensions of locales in Isabelle. |
|
60 |
The description was derived mainly by experimenting |
|
61 |
with locales and partially also by inspecting the code. |
|
62 |
||
63 |
The main contribution of the author of the present paper is the |
|
64 |
abstract description of Wenzel's version of locales, and in |
|
65 |
particular of the normalisation algorithm for locale expressions (see |
|
66 |
Section~\ref{sec-normal-forms}). Contributions to the |
|
67 |
implementation are confined to bug fixes and to provisions that |
|
68 |
enable the use of locales with linear proof scripts. |
|
69 |
||
70 |
Concepts are introduced along with examples, so that the text can be |
|
71 |
used as tutorial. It is assumed that the reader is somewhat |
|
72 |
familiar with Isabelle proof scripts. Examples have been phrased as |
|
73 |
structured |
|
74 |
Isar proofs. However, in order to understand the key concepts, |
|
75 |
including locales expressions and their normalisation, detailed |
|
76 |
knowledge of Isabelle is not necessary. |
|
77 |
||
78 |
\nocite{Nipkow2003,Wenzel2002b,Wenzel2003} |
|
79 |
*} |
|
80 |
||
81 |
section {* Locales: Beyond Proof Contexts *} |
|
82 |
||
83 |
text {* |
|
84 |
In tactic-based provers the application of a sequence of proof |
|
85 |
tactics leads to a proof state. This state is usually hard to |
|
86 |
predict from looking at the tactic script, unless one replays the |
|
87 |
proof step-by-step. The structured proof language Isar is |
|
88 |
different. It is additionally based on \emph{proof contexts}, |
|
89 |
which are directly visible in Isar scripts, and since tactic |
|
90 |
sequences tend to be short, this commonly leads to clearer proof |
|
91 |
scripts. |
|
92 |
||
93 |
Goals are stated with the \textbf{theorem} |
|
94 |
command. This is followed by a proof. When discharging a goal |
|
95 |
requires an elaborate argument |
|
96 |
(rather than the application of a single tactic) a new context |
|
97 |
may be entered (\textbf{proof}). Inside the context, variables may |
|
98 |
be fixed (\textbf{fix}), assumptions made (\textbf{assume}) and |
|
99 |
intermediate goals stated (\textbf{have}) and proved. The |
|
100 |
assumptions must be dischargeable by premises of the surrounding |
|
101 |
goal, and once this goal has been proved (\textbf{show}) the proof context |
|
102 |
can be closed (\textbf{qed}). Contexts inherit from surrounding |
|
103 |
contexts, but it is not possible |
|
104 |
to export from them (with exception of the proved goal); |
|
105 |
they ``disappear'' after the closing \textbf{qed}. |
|
106 |
Facts may have attributes --- for example, identifying them as |
|
107 |
default to the simplifier or classical reasoner. |
|
108 |
||
109 |
Locales extend proof contexts in various ways: |
|
110 |
\begin{itemize} |
|
111 |
\item |
|
112 |
Locales are usually \emph{named}. This makes them persistent. |
|
113 |
\item |
|
114 |
Fixed variables may have \emph{syntax}. |
|
115 |
\item |
|
116 |
It is possible to \emph{add} and \emph{export} facts. |
|
117 |
\item |
|
118 |
Locales can be combined and modified with \emph{locale |
|
119 |
expressions}. |
|
120 |
\end{itemize} |
|
121 |
The Locales facility extends the Isar language: it provides new ways |
|
122 |
of stating and managing facts, but it does not modify the language |
|
123 |
for proofs. Its purpose is to support writing modular proofs. |
|
124 |
*} |
|
125 |
||
126 |
section {* Simple Locales *} |
|
127 |
||
128 |
subsection {* Syntax and Terminology *} |
|
129 |
||
130 |
text {* |
|
131 |
The grammar of Isar is extended by commands for locales as shown in |
|
132 |
Figure~\ref{fig-grammar}. |
|
133 |
A key concept, introduced by Wenzel, is that |
|
134 |
locales are (internally) lists |
|
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
135 |
of \emph{context elements}. There are five kinds, identified |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
136 |
by the keywords \textbf{fixes}, \textbf{constrains}, |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
137 |
\textbf{assumes}, \textbf{defines} and \textbf{notes}. |
14586 | 138 |
|
139 |
\begin{figure} |
|
140 |
\hrule |
|
141 |
\vspace{2ex} |
|
142 |
\begin{small} |
|
143 |
\begin{tabular}{l>$c<$l} |
|
144 |
\textit{attr-name} & ::= |
|
145 |
& \textit{name} $|$ \textit{attribute} $|$ |
|
146 |
\textit{name} \textit{attribute} \\ |
|
147 |
||
148 |
\textit{locale-expr} & ::= |
|
149 |
& \textit{locale-expr1} ( ``\textbf{+}'' \textit{locale-expr1} )$^*$ \\ |
|
150 |
\textit{locale-expr1} & ::= |
|
151 |
& ( \textit{qualified-name} $|$ |
|
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
152 |
``\textbf{(}'' \textit{locale-expr} ``\textbf{)}'' ) \\ |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
153 |
& & ( \textit{name} [ \textit{mixfix} ] $|$ ``\textbf{\_}'' )$^*$ \\ |
14586 | 154 |
|
155 |
\textit{fixes} & ::= |
|
156 |
& \textit{name} [ ``\textbf{::}'' \textit{type} ] |
|
157 |
[ ``\textbf{(}'' \textbf{structure} ``\textbf{)}'' $|$ |
|
158 |
\textit{mixfix} ] \\ |
|
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
159 |
\textit{constrains} & ::= |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
160 |
& \textit{name} ``\textbf{::}'' \textit{type} \\ |
14586 | 161 |
\textit{assumes} & ::= |
162 |
& [ \textit{attr-name} ``\textbf{:}'' ] \textit{proposition} \\ |
|
163 |
\textit{defines} & ::= |
|
164 |
& [ \textit{attr-name} ``\textbf{:}'' ] \textit{proposition} \\ |
|
165 |
\textit{notes} & ::= |
|
166 |
& [ \textit{attr-name} ``\textbf{=}'' ] |
|
167 |
( \textit{qualified-name} [ \textit{attribute} ] )$^+$ \\ |
|
168 |
||
169 |
\textit{element} & ::= |
|
170 |
& \textbf{fixes} \textit{fixes} ( \textbf{and} \textit{fixes} )$^*$ \\ |
|
171 |
& | |
|
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
172 |
& \textbf{constrains} \textit{constrains} |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
173 |
( \textbf{and} \textit{constrains} )$^*$ \\ |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
174 |
& | |
14586 | 175 |
& \textbf{assumes} \textit{assumes} ( \textbf{and} \textit{assumes} )$^*$ \\ |
176 |
& | |
|
177 |
& \textbf{defines} \textit{defines} ( \textbf{and} \textit{defines} )$^*$ \\ |
|
178 |
& | |
|
179 |
& \textbf{notes} \textit{notes} ( \textbf{and} \textit{notes} )$^*$ \\ |
|
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
180 |
\textit{element1} & ::= |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
181 |
& \textit{element} \\ |
14586 | 182 |
& | & \textbf{includes} \textit{locale-expr} \\ |
183 |
||
184 |
\textit{locale} & ::= |
|
185 |
& \textit{element}$^+$ \\ |
|
186 |
& | & \textit{locale-expr} [ ``\textbf{+}'' \textit{element}$^+$ ] \\ |
|
187 |
||
188 |
\textit{in-target} & ::= |
|
189 |
& ``\textbf{(}'' \textbf{in} \textit{qualified-name} ``\textbf{)}'' \\ |
|
190 |
||
191 |
\textit{theorem} & ::= & ( \textbf{theorem} $|$ \textbf{lemma} $|$ |
|
192 |
\textbf{corollary} ) [ \textit{in-target} ] [ \textit{attr-name} ] \\ |
|
193 |
||
194 |
\textit{theory-level} & ::= & \ldots \\ |
|
195 |
& | & \textbf{locale} \textit{name} [ ``\textbf{=}'' |
|
196 |
\textit{locale} ] \\ |
|
197 |
% note: legacy "locale (open)" omitted. |
|
198 |
& | & ( \textbf{theorems} $|$ \textbf{lemmas} ) \\ |
|
199 |
& & [ \textit{in-target} ] [ \textit{attr-name} ``\textbf{=}'' ] |
|
200 |
( \textit{qualified-name} [ \textit{attribute} ] )$^+$ \\ |
|
201 |
& | & \textbf{declare} [ \textit{in-target} ] ( \textit{qualified-name} |
|
202 |
[ \textit{attribute} ] )$^+$ \\ |
|
203 |
& | & \textit{theorem} \textit{proposition} \textit{proof} \\ |
|
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
204 |
& | & \textit{theorem} \textit{element1}$^*$ |
14586 | 205 |
\textbf{shows} \textit{proposition} \textit{proof} \\ |
206 |
& | & \textbf{print\_locale} \textit{locale} \\ |
|
207 |
& | & \textbf{print\_locales} |
|
208 |
\end{tabular} |
|
209 |
\end{small} |
|
210 |
\vspace{2ex} |
|
211 |
\hrule |
|
212 |
\caption{Locales extend the grammar of Isar.} |
|
213 |
\label{fig-grammar} |
|
214 |
\end{figure} |
|
215 |
||
216 |
At the theory level --- that is, at the outer syntactic level of an |
|
217 |
Isabelle input file --- \textbf{locale} declares a named |
|
218 |
locale. Other kinds of locales, |
|
219 |
locale expressions and unnamed locales, will be introduced later. When |
|
220 |
declaring a named locale, it is possible to \emph{import} another |
|
221 |
named locale, or indeed several ones by importing a locale |
|
222 |
expression. The second part of the declaration, also optional, |
|
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
223 |
consists of a number of context element declarations. |
14586 | 224 |
|
225 |
A number of Isar commands have an additional, optional \emph{target} |
|
226 |
argument, which always refers to a named locale. These commands |
|
227 |
are \textbf{theorem} (together with \textbf{lemma} and |
|
228 |
\textbf{corollary}), \textbf{theorems} (and |
|
229 |
\textbf{lemmas}), and \textbf{declare}. The effect of specifying a target is |
|
230 |
that these commands focus on the specified locale, not the |
|
231 |
surrounding theory. Commands that are used to |
|
232 |
prove new theorems will add them not to the theory, but to the |
|
233 |
locale. Similarly, \textbf{declare} modifies attributes of theorems |
|
234 |
that belong to the specified target. Additionally, for |
|
235 |
\textbf{theorem} (and related commands), theorems stored in the target |
|
236 |
can be used in the associated proof scripts. |
|
237 |
||
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
238 |
The Locales package provides a \emph{long goals format} for |
14586 | 239 |
propositions stated with \textbf{theorem} (and friends). While |
240 |
normally a goal is just a formula, a long goal is a list of context |
|
241 |
elements, followed by the keyword \textbf{shows}, followed by the |
|
242 |
formula. Roughly speaking, the context elements are |
|
243 |
(additional) premises. For an example, see |
|
244 |
Section~\ref{sec-includes}. The list of context elements in a long goal |
|
245 |
is also called \emph{unnamed locale}. |
|
246 |
||
247 |
Finally, there are two commands to inspect locales when working in |
|
248 |
interactive mode: \textbf{print\_locales} prints the names of all |
|
249 |
targets |
|
250 |
visible in the current theory, \textbf{print\_locale} outputs the |
|
251 |
elements of a named locale or locale expression. |
|
252 |
||
253 |
The following presentation will use notation of |
|
254 |
Isabelle's meta logic, hence a few sentences to explain this. |
|
255 |
The logical |
|
256 |
primitives are universal quantification (@{text "\<And>"}), entailment |
|
257 |
(@{text "\<Longrightarrow>"}) and equality (@{text "\<equiv>"}). Variables (not bound |
|
258 |
variables) are sometimes preceded by a question mark. The logic is |
|
259 |
typed. Type variables are denoted by @{text "'a"}, @{text "'b"} |
|
260 |
etc., and @{text "\<Rightarrow>"} is the function type. Double brackets @{text |
|
261 |
"\<lbrakk>"} and @{text "\<rbrakk>"} are used to abbreviate nested entailment. |
|
262 |
*} |
|
263 |
||
264 |
subsection {* Parameters, Assumptions and Facts *} |
|
265 |
||
266 |
text {* |
|
267 |
From a logical point of view a \emph{context} is a formula schema of |
|
268 |
the form |
|
269 |
\[ |
|
270 |
@{text "\<And>x\<^sub>1\<dots>x\<^sub>n. \<lbrakk> C\<^sub>1; \<dots> ;C\<^sub>m \<rbrakk> \<Longrightarrow> \<dots>"} |
|
271 |
\] |
|
272 |
The variables $@{text "x\<^sub>1"}, \ldots, @{text "x\<^sub>n"}$ are |
|
273 |
called \emph{parameters}, the premises $@{text "C\<^sub>1"}, \ldots, |
|
274 |
@{text "C\<^sub>n"}$ \emph{assumptions}. A formula @{text "F"} |
|
275 |
holds in this context if |
|
276 |
\begin{equation} |
|
277 |
\label{eq-fact-in-context} |
|
278 |
@{text "\<And>x\<^sub>1\<dots>x\<^sub>n. \<lbrakk> C\<^sub>1; \<dots> ;C\<^sub>m \<rbrakk> \<Longrightarrow> F"} |
|
279 |
\end{equation} |
|
280 |
is valid. The formula is called a \emph{fact} of the context. |
|
281 |
||
282 |
A locale allows fixing the parameters @{text |
|
283 |
"x\<^sub>1, \<dots>, x\<^sub>n"} and making the assumptions @{text |
|
284 |
"C\<^sub>1, \<dots>, C\<^sub>m"}. This implicitly builds the context in |
|
285 |
which the formula @{text "F"} can be established. |
|
286 |
Parameters of a locale correspond to the context element |
|
287 |
\textbf{fixes}, and assumptions may be declared with |
|
288 |
\textbf{assumes}. Using these context elements one can define |
|
289 |
the specification of semigroups. |
|
290 |
*} |
|
291 |
||
292 |
locale semi = |
|
293 |
fixes prod :: "['a, 'a] \<Rightarrow> 'a" (infixl "\<cdot>" 70) |
|
294 |
assumes assoc: "(x \<cdot> y) \<cdot> z = x \<cdot> (y \<cdot> z)" |
|
295 |
||
296 |
text {* |
|
297 |
The parameter @{term prod} has a |
|
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
298 |
syntax annotation enabling the infix ``@{text "\<cdot>"}'' in the |
14586 | 299 |
assumption of associativity. Parameters may have arbitrary mixfix |
300 |
syntax, like constants. In the example, the type of @{term prod} is |
|
301 |
specified explicitly. This is not necessary. If no type is |
|
302 |
specified, a most general type is inferred simultaneously for all |
|
303 |
parameters, taking into account all assumptions (and type |
|
304 |
specifications of parameters, if present).% |
|
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
305 |
\footnote{Type inference also takes into account type constraints, |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
306 |
definitions and import, as introduced later.} |
14586 | 307 |
|
308 |
Free variables in assumptions are implicitly universally quantified, |
|
309 |
unless they are parameters. Hence the context defined by the locale |
|
310 |
@{term "semi"} is |
|
311 |
\[ |
|
312 |
@{text "\<And>prod. \<lbrakk> \<And>x y z. prod (prod x y) z = prod x (prod y z) |
|
313 |
\<rbrakk> \<Longrightarrow> \<dots>"} |
|
314 |
\] |
|
315 |
The locale can be extended to commutative semigroups. |
|
316 |
*} |
|
317 |
||
318 |
locale comm_semi = semi + |
|
319 |
assumes comm: "x \<cdot> y = y \<cdot> x" |
|
320 |
||
321 |
text {* |
|
322 |
This locale \emph{imports} all elements of @{term "semi"}. The |
|
323 |
latter locale is called the import of @{term "comm_semi"}. The |
|
324 |
definition adds commutativity, hence its context is |
|
325 |
\begin{align*% |
|
326 |
} |
|
327 |
@{text "\<And>prod. \<lbrakk> "} & |
|
328 |
@{text "\<And>x y z. prod (prod x y) z = prod x (prod y z);"} \\ |
|
329 |
& @{text "\<And>x y. prod x y = prod y x \<rbrakk> \<Longrightarrow> \<dots>"} |
|
330 |
\end{align*% |
|
331 |
} |
|
332 |
One may now derive facts --- for example, left-commutativity --- in |
|
333 |
the context of @{term "comm_semi"} by specifying this locale as |
|
334 |
target, and by referring to the names of the assumptions @{text |
|
335 |
assoc} and @{text comm} in the proof. |
|
336 |
*} |
|
337 |
||
338 |
theorem (in comm_semi) lcomm: |
|
339 |
"x \<cdot> (y \<cdot> z) = y \<cdot> (x \<cdot> z)" |
|
340 |
proof - |
|
341 |
have "x \<cdot> (y \<cdot> z) = (x \<cdot> y) \<cdot> z" by (simp add: assoc) |
|
342 |
also have "\<dots> = (y \<cdot> x) \<cdot> z" by (simp add: comm) |
|
343 |
also have "\<dots> = y \<cdot> (x \<cdot> z)" by (simp add: assoc) |
|
344 |
finally show ?thesis . |
|
345 |
qed |
|
346 |
||
347 |
text {* |
|
348 |
In this equational Isar proof, ``@{text "\<dots>"}'' refers to the |
|
349 |
right hand side of the preceding equation. |
|
350 |
After the proof is finished, the fact @{text "lcomm"} is added to |
|
351 |
the locale @{term "comm_semi"}. This is done by adding a |
|
352 |
\textbf{notes} element to the internal representation of the locale, |
|
353 |
as explained the next section. |
|
354 |
*} |
|
355 |
||
356 |
subsection {* Locale Predicates and the Internal Representation of |
|
357 |
Locales *} |
|
358 |
||
359 |
text {* |
|
360 |
\label{sec-locale-predicates} |
|
361 |
In mathematical texts, often arbitrary but fixed objects with |
|
362 |
certain properties are considered --- for instance, an arbitrary but |
|
363 |
fixed group $G$ --- with the purpose of establishing facts valid for |
|
364 |
any group. These facts are subsequently used on other objects that |
|
365 |
also have these properties. |
|
366 |
||
367 |
Locales permit the same style of reasoning. Exporting a fact $F$ |
|
368 |
generalises the fixed parameters and leads to a (valid) formula of the |
|
369 |
form of equation~(\ref{eq-fact-in-context}). If a locale has many |
|
370 |
assumptions |
|
371 |
(possibly accumulated through a number of imports) this formula can |
|
372 |
become large and cumbersome. Therefore, Wenzel introduced |
|
373 |
predicates that abbreviate the assumptions of locales. These |
|
374 |
predicates are not confined to the locale but are visible in the |
|
375 |
surrounding theory. |
|
376 |
||
377 |
The definition of the locale @{term semi} generates the \emph{locale |
|
378 |
predicate} @{term semi} over the type of the parameter @{term prod}, |
|
379 |
hence the predicate's type is @{typ "(['a, 'a] \<Rightarrow> 'a) |
|
380 |
\<Rightarrow> bool"}. Its |
|
381 |
definition is |
|
382 |
\begin{equation} |
|
383 |
\tag*{@{thm [source] semi_def}:} @{thm semi_def}. |
|
384 |
\end{equation} |
|
385 |
In the case where the locale has no import, the generated |
|
386 |
predicate abbreviates all assumptions and is over the parameters |
|
387 |
that occur in these assumptions. |
|
388 |
||
389 |
The situation is more complicated when a locale extends |
|
390 |
another locale, as is the case for @{term comm_semi}. Two |
|
391 |
predicates are defined. The predicate |
|
392 |
@{term comm_semi_axioms} corresponds to the new assumptions and is |
|
393 |
called \emph{delta predicate}, the locale |
|
394 |
predicate @{term comm_semi} captures the content of all the locale, |
|
395 |
including the import. |
|
396 |
If a locale has neither assumptions nor import, no predicate is |
|
397 |
defined. If a locale has import but no assumptions, only the locale |
|
398 |
predicate is defined. |
|
399 |
*} |
|
400 |
(*<*) |
|
401 |
ML_setup {* |
|
402 |
val [comm_semi_ax1, comm_semi_ax2] = thms "comm_semi.axioms"; |
|
403 |
bind_thm ("comm_semi_ax1", comm_semi_ax1); |
|
404 |
bind_thm ("comm_semi_ax2", comm_semi_ax2); |
|
405 |
*} |
|
406 |
(*>*) |
|
407 |
text {* |
|
408 |
The Locales package generates a number of theorems for locale and |
|
409 |
delta predicates. All predicates have a definition and an |
|
410 |
introduction rule. Locale predicates that are defined in terms of |
|
411 |
other predicates (which is the case if and only if the locale has |
|
412 |
import) also have a number of elimination rules (called |
|
413 |
\emph{axioms}). All generated theorems for the predicates of the |
|
414 |
locales @{text semi} and @{text comm_semi} are shown in |
|
415 |
Figures~\ref{fig-theorems-semi} and~\ref{fig-theorems-comm-semi}, |
|
416 |
respectively. |
|
417 |
\begin{figure} |
|
418 |
\hrule |
|
419 |
\vspace{2ex} |
|
420 |
Theorems generated for the predicate @{term semi}. |
|
421 |
\begin{gather} |
|
422 |
\tag*{@{thm [source] semi_def}:} @{thm semi_def} \\ |
|
423 |
\tag*{@{thm [source] semi.intro}:} @{thm semi.intro} |
|
424 |
\end{gather} |
|
425 |
\hrule |
|
426 |
\caption{Theorems for the locale predicate @{term "semi"}.} |
|
427 |
\label{fig-theorems-semi} |
|
428 |
\end{figure} |
|
429 |
||
430 |
\begin{figure}[h] |
|
431 |
\hrule |
|
432 |
\vspace{2ex} |
|
433 |
Theorems generated for the predicate @{term "comm_semi_axioms"}. |
|
434 |
\begin{gather} |
|
435 |
\tag*{@{thm [source] comm_semi_axioms_def}:} @{thm |
|
436 |
comm_semi_axioms_def} \\ |
|
437 |
\tag*{@{thm [source] comm_semi_axioms.intro}:} @{thm |
|
438 |
comm_semi_axioms.intro} |
|
439 |
\end{gather} |
|
440 |
Theorems generated for the predicate @{term "comm_semi"}. |
|
441 |
\begin{gather} |
|
442 |
\tag*{@{thm [source] comm_semi_def}:} @{thm |
|
443 |
comm_semi_def} \\ |
|
444 |
\tag*{@{thm [source] comm_semi.intro}:} @{thm |
|
445 |
comm_semi.intro} \\ |
|
446 |
\tag*{@{thm [source] comm_semi.axioms}:} \mbox{} \\ |
|
447 |
\notag @{thm comm_semi_ax1} \\ |
|
448 |
\notag @{thm comm_semi_ax2} |
|
449 |
\end{gather} |
|
450 |
\hrule |
|
451 |
\caption{Theorems for the predicates @{term "comm_semi_axioms"} and |
|
452 |
@{term "comm_semi"}.} |
|
453 |
\label{fig-theorems-comm-semi} |
|
454 |
\end{figure} |
|
455 |
Note that the theorems generated by a locale |
|
456 |
definition may be inspected immediately after the definition in the |
|
457 |
Proof General interface \cite{Aspinall2000} of Isabelle through |
|
458 |
the menu item ``Isabelle/Isar$>$Show me $\ldots>$Theorems''. |
|
459 |
||
460 |
Locale and delta predicates are used also in the internal |
|
461 |
representation of locales as lists of context elements. While all |
|
462 |
\textbf{fixes} in a |
|
463 |
declaration generate internal \textbf{fixes}, all assumptions |
|
464 |
of one locale declaration contribute to one internal \textbf{assumes} |
|
465 |
element. The internal representation of @{term semi} is |
|
466 |
\[ |
|
467 |
\begin{array}{ll} |
|
468 |
\textbf{fixes} & @{text prod} :: @{typ [quotes] "['a, 'a] \<Rightarrow> 'a"} |
|
469 |
(\textbf{infixl} @{text [quotes] "\<cdot>"} 70) \\ |
|
470 |
\textbf{assumes} & @{text [quotes] "semi prod"} \\ |
|
471 |
\textbf{notes} & @{text assoc}: @{text [quotes] "?x \<cdot> ?y \<cdot> ?z = ?x \<cdot> (?y \<cdot> |
|
472 |
?z)"} |
|
473 |
\end{array} |
|
474 |
\] |
|
475 |
and the internal representation of @{text [quotes] comm_semi} is |
|
476 |
\begin{equation} |
|
477 |
\label{eq-comm-semi} |
|
478 |
\begin{array}{ll} |
|
479 |
\textbf{fixes} & @{text prod} :: @{typ [quotes] "['a, 'a] \<Rightarrow> 'a"} |
|
480 |
~(\textbf{infixl}~@{text [quotes] "\<cdot>"}~70) \\ |
|
481 |
\textbf{assumes} & @{text [quotes] "semi prod"} \\ |
|
482 |
\textbf{notes} & @{text assoc}: @{text [quotes] "?x \<cdot> ?y \<cdot> ?z = ?x \<cdot> (?y \<cdot> |
|
483 |
?z)"} \\ |
|
484 |
\textbf{assumes} & @{text [quotes] "comm_semi_axioms prod"} \\ |
|
485 |
\textbf{notes} & @{text comm}: @{text [quotes] "?x \<cdot> ?y = ?y \<cdot> ?x"} \\ |
|
486 |
\textbf{notes} & @{text lcomm}: @{text [quotes] "?x \<cdot> (?y \<cdot> ?z) = ?y \<cdot> (?x \<cdot> |
|
487 |
?z)"} |
|
488 |
\end{array} |
|
489 |
\end{equation} |
|
490 |
The \textbf{notes} elements store facts of the |
|
491 |
locales. The facts @{text assoc} and @{text comm} were added |
|
492 |
during the declaration of the locales. They stem from assumptions, |
|
493 |
which are trivially facts. The fact @{text lcomm} was |
|
494 |
added later, after finishing the proof in the respective |
|
495 |
\textbf{theorem} command above. |
|
496 |
||
497 |
By using \textbf{notes} in a declaration, facts can be added |
|
498 |
to a locale directly. Of course, these must be theorems. |
|
499 |
Typical use of this feature includes adding theorems that are not |
|
500 |
usually used as a default rewrite rules by the simplifier to the |
|
501 |
simpset of the locale by a \textbf{notes} element with the attribute |
|
502 |
@{text "[simp]"}. This way it is also possible to add specialised |
|
503 |
versions of |
|
504 |
theorems to a locale by instantiating locale parameters for unknowns |
|
505 |
or locale assumptions for premises. |
|
506 |
*} |
|
507 |
||
508 |
subsection {* Definitions *} |
|
509 |
||
510 |
text {* |
|
511 |
Definitions were available in Kamm\"uller's version of Locales, and |
|
512 |
they are in Wenzel's. |
|
513 |
The context element \textbf{defines} adds a definition of the form |
|
514 |
@{text "p x\<^sub>1 \<dots> x\<^sub>n \<equiv> t"} as an assumption, where @{term p} is a |
|
515 |
parameter of the locale (possibly an imported parameter), and @{text |
|
516 |
t} a term that may contain the @{text "x\<^sub>i"}. The parameter may |
|
517 |
neither occur in a previous \textbf{assumes} or \textbf{defines} |
|
518 |
element, nor on the right hand side of the definition. Hence |
|
519 |
recursion is not allowed. |
|
520 |
The parameter may, however, occur in subsequent \textbf{assumes} and |
|
521 |
on the right hand side of subsequent \textbf{defines}. We call |
|
522 |
@{term p} \emph{defined parameter}. |
|
523 |
*} |
|
524 |
||
525 |
locale semi2 = semi + |
|
526 |
fixes rprod (infixl "\<odot>" 70) |
|
527 |
defines rprod_def: "rprod x y \<equiv> y \<cdot> x " |
|
528 |
||
529 |
text {* |
|
530 |
This locale extends @{term semi} by a second binary operation @{text |
|
531 |
[quotes] \<odot>} that is like @{text [quotes] \<cdot>} but with |
|
532 |
reversed arguments. The |
|
533 |
definition of the locale generates the predicate @{term semi2}, |
|
534 |
which is equivalent to @{text semi}, but no @{term "semi2_axioms"}. |
|
535 |
The difference between \textbf{assumes} and \textbf{defines} lies in |
|
536 |
the way parameters are treated on export. |
|
537 |
*} |
|
538 |
||
539 |
subsection {* Export *} |
|
540 |
||
541 |
text {* |
|
542 |
A fact is exported out |
|
543 |
of a locale by generalising over the parameters and adding |
|
544 |
assumptions as premises. For brevity of the exported theorems, |
|
545 |
locale predicates are used. Exported facts are referenced by |
|
546 |
writing qualified names consisting of the locale name and the fact name --- |
|
547 |
for example, |
|
548 |
\begin{equation} |
|
549 |
\tag*{@{thm [source] semi.assoc}:} @{thm semi.assoc}. |
|
550 |
\end{equation} |
|
551 |
Defined parameters receive special treatment. Instead of adding a |
|
552 |
premise for the definition, the definition is unfolded in the |
|
553 |
exported theorem. In order to illustrate this we prove that the |
|
554 |
reverse operation @{text [quotes] \<odot>} defined in the locale |
|
555 |
@{text semi2} is also associative. |
|
556 |
*} |
|
557 |
||
558 |
theorem (in semi2) r_assoc: "(x \<odot> y) \<odot> z = x \<odot> (y \<odot> z)" |
|
559 |
by (simp only: rprod_def assoc) |
|
560 |
||
561 |
text {* |
|
562 |
The exported fact is |
|
563 |
\begin{equation} |
|
564 |
\tag*{@{thm [source] semi2.r_assoc}:} @{thm semi2.r_assoc}. |
|
565 |
\end{equation} |
|
566 |
The defined parameter is not present but is replaced by its |
|
567 |
definition. Note that the definition itself is not exported, hence |
|
568 |
there is no @{text "semi2.rprod_def"}.% |
|
569 |
\footnote{The definition could alternatively be exported using a |
|
570 |
let-construct if there was one in Isabelle's meta-logic. Let is |
|
571 |
usually defined in object-logics.} |
|
572 |
*} |
|
573 |
||
574 |
section {* Locale Expressions *} |
|
575 |
||
576 |
text {* |
|
577 |
Locale expressions provide a simple language for combining |
|
578 |
locales. They are an effective means of building complex |
|
579 |
specifications from simple ones. Locale expressions are the main |
|
580 |
innovation of the version of Locales discussed here. Locale |
|
581 |
expressions are also reason for introducing locale predicates. |
|
582 |
*} |
|
583 |
||
584 |
subsection {* Rename and Merge *} |
|
585 |
||
586 |
text {* |
|
587 |
The grammar of locale expressions is part of the grammar in |
|
588 |
Figure~\ref{fig-grammar}. Locale names are locale |
|
589 |
expressions, and |
|
590 |
further expressions are obtained by \emph{rename} and \emph{merge}. |
|
591 |
\begin{description} |
|
592 |
\item[Rename.] |
|
593 |
The locale expression $e\: q_1 \ldots q_n$ denotes |
|
594 |
the locale of $e$ where pa\-ra\-me\-ters, in the order in |
|
595 |
which they are fixed, are renamed to $q_1$ to $q_n$. |
|
596 |
The expression is only well-formed if $n$ does not |
|
597 |
exceed the number of parameters of $e$. Underscores denote |
|
598 |
parameters that are not renamed. |
|
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
599 |
Renaming by default removes mixfix syntax, but new syntax may be |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
600 |
specified. |
14586 | 601 |
\item[Merge.] |
602 |
The locale expression $e_1 + e_2$ denotes |
|
603 |
the locale obtained by merging the locales of $e_1$ |
|
604 |
and $e_2$. This locale contains the context elements |
|
605 |
of $e_1$, followed by the context elements of $e_2$. |
|
606 |
||
607 |
In actual fact, the semantics of the merge operation |
|
608 |
is more complicated. If $e_1$ and $e_2$ are expressions containing |
|
609 |
the same name, followed by |
|
610 |
identical parameter lists, then the merge of both will contain |
|
611 |
the elements of those locales only once. Details are explained in |
|
612 |
Section~\ref{sec-normal-forms} below. |
|
613 |
||
614 |
The merge operation is associative but not commutative. The latter |
|
615 |
is because parameters of $e_1$ appear |
|
616 |
before parameters of $e_2$ in the composite expression. |
|
617 |
\end{description} |
|
618 |
||
619 |
Rename can be used if a different parameter name seems more |
|
620 |
appropriate --- for example, when moving from groups to rings, a |
|
621 |
parameter @{text G} representing the group could be changed to |
|
622 |
@{text R}. Besides of this stylistic use, renaming is important in |
|
623 |
combination with merge. Both operations are used in the following |
|
624 |
specification of semigroup homomorphisms. |
|
625 |
*} |
|
626 |
||
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
627 |
locale semi_hom = comm_semi sum (infixl "\<oplus>" 65) + comm_semi + |
14586 | 628 |
fixes hom |
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
629 |
assumes hom: "hom (x \<oplus> y) = hom x \<cdot> hom y" |
14586 | 630 |
|
631 |
text {* |
|
632 |
This locale defines a context with three parameters @{text "sum"}, |
|
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
633 |
@{text "prod"} and @{text "hom"}. The first two parameters have |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
634 |
mixfix syntax. They are associative operations, |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
635 |
the first of type @{typeof [locale=semi_hom] prod}, the second of |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
636 |
type @{typeof [locale=semi_hom] sum}. |
14586 | 637 |
|
638 |
How are facts that are imported via a locale expression identified? |
|
639 |
Facts are always introduced in a named locale (either in the |
|
640 |
locale's declaration, or by using the locale as target in |
|
641 |
\textbf{theorem}), and their names are |
|
642 |
qualified by the parameter names of this locale. |
|
643 |
Hence the full name of associativity in @{text "semi"} is |
|
644 |
@{text "prod.assoc"}. Renaming parameters of a target also renames |
|
645 |
the qualifier of facts. Hence, associativity of @{text "sum"} is |
|
646 |
@{text "sum.assoc"}. Several parameters are separated by |
|
647 |
underscores in qualifiers. For example, the full name of the fact |
|
648 |
@{text "hom"} in the locale @{text "semi_hom"} is @{text |
|
649 |
"sum_prod_hom.hom"}. |
|
650 |
||
651 |
The following example is quite artificial, it illustrates the use of |
|
652 |
facts, though. |
|
653 |
*} |
|
654 |
||
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
655 |
theorem (in semi_hom) "hom x \<cdot> (hom y \<cdot> hom z) = hom (x \<oplus> (y \<oplus> z))" |
14586 | 656 |
proof - |
657 |
have "hom x \<cdot> (hom y \<cdot> hom z) = hom y \<cdot> (hom x \<cdot> hom z)" |
|
658 |
by (simp add: prod.lcomm) |
|
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
659 |
also have "\<dots> = hom (y \<oplus> (x \<oplus> z))" by (simp add: hom) |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
660 |
also have "\<dots> = hom (x \<oplus> (y \<oplus> z))" by (simp add: sum.lcomm) |
14586 | 661 |
finally show ?thesis . |
662 |
qed |
|
663 |
||
664 |
text {* |
|
665 |
Importing via a locale expression imports all facts of |
|
666 |
the imported locales, hence both @{text "sum.lcomm"} and @{text |
|
667 |
"prod.lcomm"} are |
|
668 |
available in @{text "hom_semi"}. The import is dynamic --- that is, |
|
669 |
whenever facts are added to a locale, they automatically |
|
670 |
become available in subsequent \textbf{theorem} commands that use |
|
671 |
the locale as a target, or a locale importing the locale. |
|
672 |
*} |
|
673 |
||
674 |
subsection {* Normal Forms *} |
|
675 |
||
676 |
text_raw {* |
|
677 |
\label{sec-normal-forms} |
|
678 |
\newcommand{\I}{\mathcal{I}} |
|
679 |
\newcommand{\F}{\mathcal{F}} |
|
680 |
\newcommand{\N}{\mathcal{N}} |
|
681 |
\newcommand{\C}{\mathcal{C}} |
|
682 |
\newcommand{\App}{\mathbin{\overline{@}}} |
|
683 |
*} |
|
684 |
||
685 |
text {* |
|
686 |
Locale expressions are interpreted in a two-step process. First, an |
|
687 |
expression is normalised, then it is converted to a list of context |
|
688 |
elements. |
|
689 |
||
690 |
Normal forms are based on \textbf{locale} declarations. These |
|
691 |
consist of an import section followed by a list of context |
|
692 |
elements. Let $\I(l)$ denote the locale expression imported by |
|
693 |
locale $l$. If $l$ has no import then $\I(l) = \varepsilon$. |
|
694 |
Likewise, let $\F(l)$ denote the list of context elements, also |
|
695 |
called the \emph{context fragment} of $l$. Note that $\F(l)$ |
|
696 |
contains only those context elements that are stated in the |
|
697 |
declaration of $l$, not imported ones. |
|
698 |
||
699 |
\paragraph{Example 1.} Consider the locales @{text semi} and @{text |
|
700 |
"comm_semi"}. We have $\I(@{text semi}) = \varepsilon$ and |
|
701 |
$\I(@{text "comm_semi"}) = @{text semi}$, and the context fragments |
|
702 |
are |
|
703 |
\begin{align*% |
|
704 |
} |
|
705 |
\F(@{text semi}) & = \left[ |
|
706 |
\begin{array}{ll} |
|
707 |
\textbf{fixes} & @{text prod} :: @{typ [quotes] "['a, 'a] \<Rightarrow> 'a"} |
|
708 |
~(\textbf{infixl}~@{text [quotes] "\<cdot>"}~70) \\ |
|
709 |
\textbf{assumes} & @{text [quotes] "semi prod"} \\ |
|
710 |
\textbf{notes} & @{text assoc}: @{text [quotes]"?x \<cdot> ?y \<cdot> ?z = ?x \<cdot> (?y \<cdot> |
|
711 |
?z)"} |
|
712 |
\end{array} \right], \\ |
|
713 |
\F(@{text "comm_semi"}) & = \left[ |
|
714 |
\begin{array}{ll} |
|
715 |
\textbf{assumes} & @{text [quotes] "comm_semi_axioms prod"} \\ |
|
716 |
\textbf{notes} & @{text comm}: @{text [quotes] "?x \<cdot> ?y = ?y \<cdot> ?x"} \\ |
|
717 |
\textbf{notes} & @{text lcomm}: @{text [quotes] "?x \<cdot> (?y \<cdot> ?z) = ?y \<cdot> (?x \<cdot> |
|
718 |
?z)"} |
|
719 |
\end{array} \right]. |
|
720 |
\end{align*% |
|
721 |
} |
|
722 |
Let $\pi_0(\F(l))$ denote the list of parameters defined in the |
|
723 |
\textbf{fixes} elements of $\F(l)$ in the order of their occurrence. |
|
724 |
The list of parameters of a locale expression $\pi(e)$ is defined as |
|
725 |
follows: |
|
726 |
\begin{align*% |
|
727 |
} |
|
728 |
\pi(l) & = \pi(\I(l)) \App \pi_0(\F(l)) \text{, for named locale $l$.} \\ |
|
729 |
\pi(e\: q_1 \ldots q_n) & = \text{$[q_1, \ldots, q_n, p_{n+1}, \ldots, |
|
730 |
p_{m}]$, where $\pi(e) = [p_1, \ldots, p_m]$.} \\ |
|
731 |
\pi(e_1 + e_2) & = \pi(e_1) \App \pi(e_2) |
|
732 |
\end{align*% |
|
733 |
} |
|
734 |
The operation $\App$ concatenates two lists but omits elements from |
|
735 |
the second list that are also present in the first list. |
|
736 |
It is not possible to rename more parameters than there |
|
737 |
are present in an expression --- that is, $n \le m$ --- otherwise |
|
738 |
the renaming is illegal. If $q_i |
|
739 |
= \_$ then the $i$th entry of the resulting list is $p_i$. |
|
740 |
||
741 |
In the normalisation phase, imports of named locales are unfolded, and |
|
742 |
renames and merges are recursively propagated to the imported locale |
|
743 |
expressions. The result is a list of locale names, each with a full |
|
744 |
list of parameters, where locale names occurring with the same parameter |
|
745 |
list twice are removed. Let $\N$ denote normalisation. It is defined |
|
746 |
by these equations: |
|
747 |
\begin{align*% |
|
748 |
} |
|
749 |
\N(l) & = \N(\I(l)) \App [l\:\pi(l)] \text{, for named locale $l$.} \\ |
|
750 |
\N(e\: q_1 \ldots q_n) & = \N(e)\:[q_1 \ldots q_n/\pi(e)] \\ |
|
751 |
\N(e_1 + e_2) & = \N(e_1) \App \N(e_2) |
|
752 |
\end{align*% |
|
753 |
} |
|
754 |
Normalisation yields a list of \emph{identifiers}. An identifier |
|
755 |
consists of a locale name and a (possibly empty) list of parameters. |
|
756 |
||
757 |
In the second phase, the list of identifiers $\N(e)$ is converted to |
|
758 |
a list of context elements $\C(e)$ by converting each identifier to |
|
759 |
a list of context elements, and flattening the obtained list. |
|
760 |
Conversion of the identifier $l\:q_1 \ldots q_n$ yields the list of |
|
761 |
context elements $\F(l)$, but with the following modifications: |
|
762 |
\begin{itemize} |
|
763 |
\item |
|
764 |
Rename the parameter in the $i$th \textbf{fixes} element of $\F(l)$ |
|
765 |
to $q_i$, $i = 1, \ldots, n$. If the parameter name is actually |
|
766 |
changed then delete the syntax annotation. Renaming a parameter may |
|
767 |
also change its type. |
|
768 |
\item |
|
769 |
Perform the same renamings on all occurrences of parameters (fixed |
|
770 |
variables) in \textbf{assumes}, \textbf{defines} and \textbf{notes} |
|
771 |
elements. |
|
772 |
\item |
|
773 |
Qualify names of facts by $q_1\_\ldots\_q_n$. |
|
774 |
\end{itemize} |
|
775 |
The locale expression is \emph{well-formed} if it contains no |
|
776 |
illegal renamings and the following conditions on $\C(e)$ hold, |
|
777 |
otherwise the expression is rejected: |
|
778 |
\begin{itemize} |
|
779 |
\item Parameters in \textbf{fixes} are distinct; |
|
780 |
\item Free variables in \textbf{assumes} and |
|
781 |
\textbf{defines} occur in preceding \textbf{fixes};% |
|
782 |
\footnote{This restriction is relaxed for contexts obtained with |
|
783 |
\textbf{includes}, see Section~\ref{sec-includes}.} |
|
784 |
\item Parameters defined in \textbf{defines} must neither occur in |
|
785 |
preceding \textbf{assumes} nor \textbf{defines}. |
|
786 |
\end{itemize} |
|
787 |
*} |
|
788 |
||
789 |
subsection {* Examples *} |
|
790 |
||
791 |
text {* |
|
792 |
\paragraph{Example 2.} |
|
793 |
We obtain the context fragment $\C(@{text "comm_semi"})$ of the |
|
794 |
locale @{text "comm_semi"}. First, the parameters are computed. |
|
795 |
\begin{align*% |
|
796 |
} |
|
797 |
\pi(@{text "semi"}) & = [@{text prod}] \\ |
|
798 |
\pi(@{text "comm_semi"}) & = \pi(@{text "semi"}) \App [] |
|
799 |
= [@{text prod}] |
|
800 |
\end{align*% |
|
801 |
} |
|
802 |
Next, the normal form of the locale expression |
|
803 |
@{text "comm_semi"} is obtained. |
|
804 |
\begin{align*% |
|
805 |
} |
|
806 |
\N(@{text "semi"}) & = [@{text semi} @{text prod}] \\ |
|
807 |
\N(@{text "comm_semi"}) & = \N(@{text "semi"}) \App |
|
808 |
[@{text "comm_semi prod"}] |
|
809 |
= [@{text "semi prod"}, @{text "comm_semi prod"}] |
|
810 |
\end{align*% |
|
811 |
} |
|
812 |
Converting this to a list of context elements leads to the |
|
813 |
list~(\ref{eq-comm-semi}) shown in |
|
814 |
Section~\ref{sec-locale-predicates}, but with fact names qualified |
|
815 |
by @{text prod} --- for example, @{text "prod.assoc"}. |
|
816 |
Qualification was omitted to keep the presentation simple. |
|
817 |
Isabelle's scoping rules identify the most recent fact with |
|
818 |
qualified name $x.a$ when a fact with name $a$ is requested. |
|
819 |
||
820 |
\paragraph{Example 3.} |
|
821 |
The locale expression @{text "comm_semi sum"} involves |
|
822 |
renaming. Computing parameters yields $\pi(@{text "comm_semi sum"}) |
|
823 |
= [@{text sum}]$, the normal form is |
|
824 |
\begin{align*% |
|
825 |
} |
|
826 |
\N(@{text "comm_semi sum"}) & = |
|
827 |
\N(@{text "comm_semi"})[@{text sum}/@{text prod}] = |
|
828 |
[@{text "semi sum"}, @{text "comm_semi sum"}] |
|
829 |
\end{align*% |
|
830 |
} |
|
831 |
and the list of context elements |
|
832 |
\[ |
|
833 |
\begin{array}{ll} |
|
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
834 |
\textbf{fixes} & @{text sum} :: @{typ [quotes] "['a, 'a] \<Rightarrow> 'a"} |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
835 |
~(\textbf{infixl}~@{text [quotes] "\<oplus>"}~65) \\ |
14586 | 836 |
\textbf{assumes} & @{text [quotes] "semi sum"} \\ |
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
837 |
\textbf{notes} & @{text sum.assoc}: @{text [quotes] "(?x \<oplus> ?y) \<oplus> ?z |
14586 | 838 |
= sum ?x (sum ?y ?z)"} \\ |
839 |
\textbf{assumes} & @{text [quotes] "comm_semi_axioms sum"} \\ |
|
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
840 |
\textbf{notes} & @{text sum.comm}: @{text [quotes] "?x \<oplus> ?y = |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
841 |
?y \<oplus> ?x"} \\ |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
842 |
\textbf{notes} & @{text sum.lcomm}: @{text [quotes] "?x \<oplus> (?y \<oplus> ?z) |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
843 |
= ?y \<oplus> (?x \<oplus> ?z)"} |
14586 | 844 |
\end{array} |
845 |
\] |
|
846 |
||
847 |
\paragraph{Example 4.} |
|
848 |
The context defined by the locale @{text "semi_hom"} involves |
|
849 |
merging two copies of @{text "comm_semi"}. We obtain the parameter |
|
850 |
list and normal form: |
|
851 |
\begin{align*% |
|
852 |
} |
|
853 |
\pi(@{text "semi_hom"}) & = \pi(@{text "comm_semi sum"} + |
|
854 |
@{text "comm_semi"}) \App [@{text hom}] \\ |
|
855 |
& = (\pi(@{text "comm_semi sum"}) \App \pi(@{text "comm_semi"})) |
|
856 |
\App [@{text hom}] \\ |
|
857 |
& = ([@{text sum}] \App [@{text prod}]) \App [@{text hom}] |
|
858 |
= [@{text sum}, @{text prod}, @{text hom}] \\ |
|
859 |
\N(@{text "semi_hom"}) & = |
|
860 |
\N(@{text "comm_semi sum"} + @{text "comm_semi"}) \App \\ |
|
861 |
& \phantom{==} |
|
862 |
[@{text "semi_hom sum prod hom"}] \\ |
|
863 |
& = (\N(@{text "comm_semi sum"}) \App \N(@{text "comm_semi"})) \App \\ |
|
864 |
& \phantom{==} |
|
865 |
[@{text "semi_hom sum prod hom"}] \\ |
|
866 |
& = ([@{text "semi sum"}, @{text "comm_semi sum"}] \App %\\ |
|
867 |
% & \phantom{==} |
|
868 |
[@{text "semi prod"}, @{text "comm_semi prod"}]) \App \\ |
|
869 |
& \phantom{==} |
|
870 |
[@{text "semi_hom sum prod hom"}] \\ |
|
871 |
& = [@{text "semi sum"}, @{text "comm_semi sum"}, |
|
872 |
@{text "semi prod"}, @{text "comm_semi prod"}, \\ |
|
873 |
& \phantom{==} |
|
874 |
@{text "semi_hom sum prod hom"}]. |
|
875 |
\end{align*% |
|
876 |
} |
|
877 |
Hence $\C(@{text "semi_hom"})$, shown below, is again well-formed. |
|
878 |
\[ |
|
879 |
\begin{array}{ll} |
|
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
880 |
\textbf{fixes} & @{text sum} :: @{typ [quotes] "['a, 'a] \<Rightarrow> 'a"} |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
881 |
~(\textbf{infixl}~@{text [quotes] "\<oplus>"}~65) \\ |
14586 | 882 |
\textbf{assumes} & @{text [quotes] "semi sum"} \\ |
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
883 |
\textbf{notes} & @{text sum.assoc}: @{text [quotes] "(?x \<oplus> ?y) \<oplus> ?z |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
884 |
= ?x \<oplus> (?y \<oplus> ?z)"} \\ |
14586 | 885 |
\textbf{assumes} & @{text [quotes] "comm_semi_axioms sum"} \\ |
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
886 |
\textbf{notes} & @{text sum.comm}: @{text [quotes] "?x \<oplus> ?y = ?y \<oplus> ?x"} \\ |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
887 |
\textbf{notes} & @{text sum.lcomm}: @{text [quotes] "?x \<oplus> (?y \<oplus> ?z) |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
888 |
= ?y \<oplus> (?x \<oplus> ?z)"} \\ |
14586 | 889 |
\textbf{fixes} & @{text prod} :: @{typ [quotes] "['b, 'b] \<Rightarrow> 'b"} |
890 |
~(\textbf{infixl}~@{text [quotes] "\<cdot>"}~70) \\ |
|
891 |
\textbf{assumes} & @{text [quotes] "semi prod"} \\ |
|
892 |
\textbf{notes} & @{text prod.assoc}: @{text [quotes] "?x \<cdot> ?y \<cdot> ?z = ?x \<cdot> (?y \<cdot> |
|
893 |
?z)"} \\ |
|
894 |
\textbf{assumes} & @{text [quotes] "comm_semi_axioms prod"} \\ |
|
895 |
\textbf{notes} & @{text prod.comm}: @{text [quotes] "?x \<cdot> ?y = ?y \<cdot> ?x"} \\ |
|
896 |
\textbf{notes} & @{text prod.lcomm}: @{text [quotes] "?x \<cdot> (?y \<cdot> ?z) = ?y \<cdot> (?x \<cdot> |
|
897 |
?z)"} \\ |
|
898 |
\textbf{fixes} & @{text hom} :: @{typ [quotes] "'a \<Rightarrow> 'b"} \\ |
|
899 |
\textbf{assumes} & @{text [quotes] "semi_hom_axioms sum"} \\ |
|
900 |
\textbf{notes} & @{text "sum_prod_hom.hom"}: |
|
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
901 |
@{text "hom (x \<oplus> y) = hom x \<cdot> hom y"} |
14586 | 902 |
\end{array} |
903 |
\] |
|
904 |
||
905 |
\paragraph{Example 5.} |
|
906 |
In this example, a locale expression leading to a list of context |
|
907 |
elements that is not well-defined is encountered, and it is illustrated |
|
908 |
how normalisation deals with multiple inheritance. |
|
909 |
Consider the specification of monads (in the algebraic sense) |
|
910 |
and monoids. |
|
911 |
*} |
|
912 |
||
913 |
locale monad = |
|
914 |
fixes prod :: "['a, 'a] \<Rightarrow> 'a" (infixl "\<cdot>" 70) and one :: 'a ("\<one>" 100) |
|
915 |
assumes l_one: "\<one> \<cdot> x = x" and r_one: "x \<cdot> \<one> = x" |
|
916 |
||
917 |
text {* |
|
918 |
Monoids are both semigroups and monads and one would want to |
|
919 |
specify them as locale expression @{text "semi + monad"}. |
|
920 |
Unfortunately, this expression is not well-formed. Its normal form |
|
921 |
\begin{align*% |
|
922 |
} |
|
923 |
\N(@{text "monad"}) & = [@{text "monad prod"}] \\ |
|
924 |
\N(@{text "semi"}+@{text "monad"}) & = |
|
925 |
\N(@{text "semi"}) \App \N(@{text "monad"}) |
|
926 |
= [@{text "semi prod"}, @{text "monad prod"}] |
|
927 |
\end{align*% |
|
928 |
} |
|
929 |
leads to a list containing the context element |
|
930 |
\[ |
|
931 |
\textbf{fixes}~@{text prod} :: @{typ [quotes] "['a, 'a] \<Rightarrow> 'a"} |
|
932 |
~(\textbf{infixl}~@{text [quotes] "\<cdot>"}~70) |
|
933 |
\] |
|
934 |
twice and thus violating the first criterion of well-formedness. To |
|
935 |
avoid this problem, one can |
|
936 |
introduce a new locale @{text magma} with the sole purpose of fixing the |
|
937 |
parameter and defining its syntax. The specifications of semigroup |
|
938 |
and monad are changed so that they import @{text magma}. |
|
939 |
*} |
|
940 |
||
941 |
locale magma = fixes prod (infixl "\<cdot>" 70) |
|
942 |
||
943 |
locale semi' = magma + assumes assoc: "(x \<cdot> y) \<cdot> z = x \<cdot> (y \<cdot> z)" |
|
944 |
locale monad' = magma + fixes one ("\<one>" 100) |
|
945 |
assumes l_one: "\<one> \<cdot> x = x" and r_one: "x \<cdot> \<one> = x" |
|
946 |
||
947 |
text {* |
|
948 |
Normalisation now yields |
|
949 |
\begin{align*% |
|
950 |
} |
|
951 |
\N(@{text "semi' + monad'"}) & = |
|
952 |
\N(@{text "semi'"}) \App \N(@{text "monad'"}) \\ |
|
953 |
& = (\N(@{text magma}) \App [@{text "semi' prod"}]) \App |
|
954 |
(\N(@{text magma}) \App [@{text "monad' prod"}]) \\ |
|
955 |
& = [@{text "magma prod"}, @{text "semi' prod"}] \App |
|
956 |
[@{text "magma prod"}, @{text "monad' prod"}]) \\ |
|
957 |
& = [@{text "magma prod"}, @{text "semi' prod"}, |
|
958 |
@{text "monad' prod"}] |
|
959 |
\end{align*% |
|
960 |
} |
|
961 |
where the second occurrence of @{text "magma prod"} is eliminated. |
|
962 |
The reader is encouraged to check, using the \textbf{print\_locale} |
|
963 |
command, that the list of context elements generated from this is |
|
964 |
indeed well-formed. |
|
965 |
||
966 |
It follows that importing |
|
967 |
parameters is more flexible than fixing them using a context element. |
|
968 |
The Locale package provides the predefined locale @{term var} that |
|
969 |
can be used to import parameters if no |
|
970 |
particular mixfix syntax is required. |
|
971 |
Its definition is |
|
972 |
\begin{center} |
|
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
973 |
\textbf{locale} @{text var} = \textbf{fixes} @{text "x_"} |
14586 | 974 |
\end{center} |
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
975 |
The use of the internal variable @{text "x_"} |
14586 | 976 |
enforces that the parameter is renamed before being used, because |
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
977 |
internal variables may not occur in the input syntax. Using |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
978 |
@{text var}, the locale @{text magma} could have been replaced by |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
979 |
the locale expression |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
980 |
\begin{center} |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
981 |
@{text var} @{text prod} (\textbf{infixl} @{text [quotes] \<cdot>} 70) |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
982 |
\end{center} |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
983 |
in the above locale declarations. |
14586 | 984 |
*} |
985 |
||
986 |
subsection {* Includes *} |
|
987 |
||
988 |
text {* |
|
989 |
\label{sec-includes} |
|
990 |
The context element \textbf{includes} takes a locale expression $e$ |
|
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
991 |
as argument. It can only occur in long goals, where it |
14586 | 992 |
adds, like a target, locale context to the proof context. Unlike |
993 |
with targets, the proved theorem is not stored |
|
994 |
in the locale. Instead, it is exported immediately. |
|
995 |
*} |
|
996 |
||
997 |
theorem lcomm2: |
|
998 |
includes comm_semi shows "x \<cdot> (y \<cdot> z) = y \<cdot> (x \<cdot> z)" |
|
999 |
proof - |
|
1000 |
have "x \<cdot> (y \<cdot> z) = (x \<cdot> y) \<cdot> z" by (simp add: assoc) |
|
1001 |
also have "\<dots> = (y \<cdot> x) \<cdot> z" by (simp add: comm) |
|
1002 |
also have "\<dots> = y \<cdot> (x \<cdot> z)" by (simp add: assoc) |
|
1003 |
finally show ?thesis . |
|
1004 |
qed |
|
1005 |
||
1006 |
text {* |
|
1007 |
This proof is identical to the proof of @{text lcomm}. The use of |
|
1008 |
\textbf{includes} provides the same context and facts as when using |
|
1009 |
@{term comm_semi} as target. On the other hand, @{thm [source] |
|
1010 |
lcomm2} is not added as a fact to the locale @{term comm_semi}, but |
|
1011 |
is directly visible in the theory. The theorem is |
|
1012 |
\[ |
|
1013 |
@{thm lcomm2}. |
|
1014 |
\] |
|
1015 |
Note that it is possible to |
|
1016 |
combine a target and (several) \textbf{includes} in a goal statement, thus |
|
1017 |
using contexts of several locales but storing the theorem in only |
|
1018 |
one of them. |
|
1019 |
*} |
|
1020 |
||
1021 |
section {* Structures *} |
|
1022 |
||
1023 |
text {* |
|
1024 |
\label{sec-structures} |
|
1025 |
The specifications of semigroups and monoids that served as examples |
|
1026 |
in previous sections modelled each operation of an algebraic |
|
1027 |
structure as a single parameter. This is rather inconvenient for |
|
1028 |
structures with many operations, and also unnatural. In accordance |
|
1029 |
to mathematical texts, one would rather fix two groups instead of |
|
1030 |
two sets of operations. |
|
1031 |
||
1032 |
The approach taken in Isabelle is to encode algebraic structures |
|
1033 |
with suitable types (in Isabelle/HOL usually records). An issue to |
|
1034 |
be addressed by |
|
1035 |
locales is syntax for algebraic structures. This is the purpose of |
|
1036 |
the \textbf{(structure)} annotation in \textbf{fixes}, introduced by |
|
1037 |
Wenzel. We illustrate this, independently of record types, with a |
|
1038 |
different formalisation of semigroups. |
|
1039 |
||
1040 |
Let @{text "'a semi_type"} be a not further specified type that |
|
1041 |
represents semigroups over the carrier type @{typ "'a"}. Let @{text |
|
1042 |
"s_op"} be an operation that maps an object of @{text "'a semi_type"} to |
|
1043 |
a binary operation. |
|
1044 |
*} |
|
1045 |
||
1046 |
typedecl 'a semi_type |
|
1047 |
consts s_op :: "['a semi_type, 'a, 'a] \<Rightarrow> 'a" (infixl "\<star>\<index>" 70) |
|
1048 |
||
1049 |
text {* |
|
1050 |
Although @{text "s_op"} is a ternary operation, it is declared |
|
1051 |
infix. The syntax annotation contains the token @{text \<index>} |
|
1052 |
(\verb.\<index>.), which refers to the first argument. This syntax is only |
|
1053 |
effective in the context of a locale, and only if the first argument |
|
1054 |
is a |
|
1055 |
\emph{structural} parameter --- that is, a parameter with annotation |
|
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
1056 |
\textbf{(structure)}. The token has the effect of subscripting the |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
1057 |
parameter --- by bracketing it between \verb.\<^bsub>. and \verb.\<^esub>.. |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
1058 |
Additionally, the subscript of the first structural parameter may be |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
1059 |
omitted, as in this specification of semigroups with structures: |
14586 | 1060 |
*} |
1061 |
||
1062 |
locale comm_semi' = |
|
1063 |
fixes G :: "'a semi_type" (structure) |
|
1064 |
assumes assoc: "(x \<star> y) \<star> z = x \<star> (y \<star> z)" and comm: "x \<star> y = y \<star> x" |
|
1065 |
||
1066 |
text {* |
|
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
1067 |
Here @{text "x \<star> y"} is equivalent to @{text "x \<star>\<^bsub>G\<^esub> y"} and |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
1068 |
abbreviates @{text "s_op G x y"}. A specification of homomorphisms |
14586 | 1069 |
requires a second structural parameter. |
1070 |
*} |
|
1071 |
||
1072 |
locale semi'_hom = comm_semi' + comm_semi' H + |
|
1073 |
fixes hom |
|
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
1074 |
assumes hom: "hom (x \<star> y) = hom x \<star>\<^bsub>H\<^esub> hom y" |
14586 | 1075 |
|
1076 |
text {* |
|
1077 |
The parameter @{text H} is defined in the second \textbf{fixes} |
|
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
1078 |
element of $\C(@{term "semi'_comm"})$. Hence @{text "\<star>\<^bsub>H\<^esub>"} |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
1079 |
abbreviates @{text "s_op H x y"}. The same construction can be done |
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
1080 |
with records instead of an \textit{ad-hoc} type. *} |
14586 | 1081 |
|
1082 |
record 'a semi = prod :: "['a, 'a] \<Rightarrow> 'a" (infixl "\<bullet>\<index>" 70) |
|
1083 |
||
1084 |
text {* |
|
1085 |
This declares the types @{typ "'a semi"} and @{typ "('a, 'b) |
|
1086 |
semi_scheme"}. The latter is an extensible record, where the second |
|
1087 |
type argument is the type of the extension field. For details on |
|
1088 |
records, see \cite{NipkowEtAl2002} Chapter~8.3. |
|
1089 |
*} |
|
1090 |
||
1091 |
locale semi_w_records = struct G + |
|
1092 |
assumes assoc: "(x \<bullet> y) \<bullet> z = x \<bullet> (y \<bullet> z)" |
|
1093 |
||
1094 |
text {* |
|
1095 |
The type @{typ "('a, 'b) semi_scheme"} is inferred for the parameter @{term |
|
1096 |
G}. Using subtyping on records, the specification can be extended |
|
1097 |
to groups easily. |
|
1098 |
*} |
|
1099 |
||
1100 |
record 'a group = "'a semi" + |
|
1101 |
one :: "'a" ("l\<index>" 100) |
|
1102 |
inv :: "'a \<Rightarrow> 'a" ("inv\<index> _" [81] 80) |
|
1103 |
locale group_w_records = semi_w_records + |
|
1104 |
assumes l_one: "l \<bullet> x = x" and l_inv: "inv x \<bullet> x = l" |
|
1105 |
||
1106 |
text {* |
|
1107 |
Finally, the predefined locale |
|
1108 |
\begin{center} |
|
1109 |
\textbf{locale} \textit{struct} = \textbf{fixes} @{text "S_"} |
|
1110 |
\textbf{(structure)}. |
|
1111 |
\end{center} |
|
1112 |
is analogous to @{text "var"}. |
|
1113 |
More examples on the use of structures, including groups, rings and |
|
1114 |
polynomials can be found in the Isabelle distribution in the |
|
1115 |
session HOL-Algebra. |
|
1116 |
*} |
|
1117 |
||
1118 |
section {* Conclusions and Outlook *} |
|
1119 |
||
1120 |
text {* |
|
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
1121 |
Locales provide a simple means of modular reasoning. They enable to |
14586 | 1122 |
abbreviate frequently occurring context statements and maintain facts |
1123 |
valid in these contexts. Importantly, using structures, they allow syntax to be |
|
1124 |
effective only in certain contexts, and thus to mimic common |
|
1125 |
practice in mathematics, where notation is chosen very flexibly. |
|
1126 |
This is also known as literate formalisation \cite{Bailey1998}. |
|
1127 |
Locale expressions allow to duplicate and merge |
|
1128 |
specifications. This is a necessity, for example, when reasoning about |
|
1129 |
homomorphisms. Normalisation makes it possible to deal with |
|
1130 |
diamond-shaped inheritance structures, and generally with directed |
|
1131 |
acyclic graphs. The combination of locales with record |
|
1132 |
types in higher-order logic provides an effective means for |
|
1133 |
specifying algebraic structures: locale import and record subtyping |
|
1134 |
provide independent hierarchies for specifications and structure |
|
1135 |
elements. Rich examples for this can be found in |
|
1136 |
the Isabelle distribution in the session HOL-Algebra. |
|
1137 |
||
1138 |
The primary reason for writing this report was to provide a better |
|
1139 |
understanding of locales in Isar. Wenzel provided hardly any |
|
1140 |
documentation, with the exception of \cite{Wenzel2002b}. The |
|
1141 |
present report should make it easier for users of Isabelle to take |
|
1142 |
advantage of locales. |
|
1143 |
||
1144 |
The report is also a base for future extensions. These include |
|
1145 |
improved syntax for structures. Identifying them by numbers seems |
|
1146 |
unnatural and can be confusing if more than two structures are |
|
1147 |
involved --- for example, when reasoning about universal |
|
1148 |
properties --- and numbering them by order of occurrence seems |
|
1149 |
arbitrary. Another desirable feature is \emph{instantiation}. One |
|
1150 |
may, in the course of a theory development, construct objects that |
|
1151 |
fulfil the specification of a locale. These objects are possibly |
|
1152 |
defined in the context of another locale. Instantiation should make it |
|
1153 |
simple to specialise abstract facts for the object under |
|
1154 |
consideration and to use the specified facts. |
|
1155 |
||
1156 |
A detailed comparison of locales with module systems in type theory |
|
1157 |
has not been undertaken yet, but could be beneficial. For example, |
|
1158 |
a module system for Coq has recently been presented by Chrzaszcz |
|
16168
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
14586
diff
changeset
|
1159 |
\cite{Chrzaszcz2003,Chrzaszcz2004}. While the |
14586 | 1160 |
latter usually constitute extensions of the calculus, locales are |
1161 |
a rather thin layer that does not change Isabelle's meta logic. |
|
1162 |
Locales mainly manage specifications and facts. Functors, like |
|
1163 |
the constructor for polynomial rings, remain objects of the |
|
1164 |
logic. |
|
1165 |
||
1166 |
\textbf{Acknowledgements.} Lawrence C.\ Paulson and Norbert |
|
1167 |
Schirmer made useful comments on a draft of this paper. |
|
1168 |
*} |
|
1169 |
||
1170 |
(*<*) |
|
1171 |
end |
|
1172 |
(*>*) |