author | wenzelm |
Wed, 03 Sep 2008 17:47:38 +0200 | |
changeset 28114 | 2637fb838f74 |
parent 28085 | 914183e229e9 |
child 28281 | 132456af0731 |
permissions | -rw-r--r-- |
26869 | 1 |
(* $Id$ *) |
2 |
||
3 |
theory Spec |
|
4 |
imports Main |
|
5 |
begin |
|
6 |
||
27046 | 7 |
chapter {* Theory specifications *} |
26869 | 8 |
|
26870 | 9 |
section {* Defining theories \label{sec:begin-thy} *} |
10 |
||
11 |
text {* |
|
12 |
\begin{matharray}{rcl} |
|
13 |
@{command_def "theory"} & : & \isartrans{toplevel}{theory} \\ |
|
27040 | 14 |
@{command_def (global) "end"} & : & \isartrans{theory}{toplevel} \\ |
26870 | 15 |
\end{matharray} |
16 |
||
27040 | 17 |
Isabelle/Isar theories are defined via theory file, which contain |
18 |
both specifications and proofs; occasionally definitional mechanisms |
|
19 |
also require some explicit proof. The theory body may be |
|
20 |
sub-structered by means of \emph{local theory} target mechanisms, |
|
21 |
notably @{command "locale"} and @{command "class"}. |
|
26870 | 22 |
|
23 |
The first ``real'' command of any theory has to be @{command |
|
24 |
"theory"}, which starts a new theory based on the merge of existing |
|
25 |
ones. Just preceding the @{command "theory"} keyword, there may be |
|
26 |
an optional @{command "header"} declaration, which is relevant to |
|
27 |
document preparation only; it acts very much like a special |
|
27051 | 28 |
pre-theory markup command (cf.\ \secref{sec:markup}). The @{command |
29 |
(global) "end"} command concludes a theory development; it has to be |
|
30 |
the very last command of any theory file loaded in batch-mode. |
|
26870 | 31 |
|
32 |
\begin{rail} |
|
33 |
'theory' name 'imports' (name +) uses? 'begin' |
|
34 |
; |
|
35 |
||
36 |
uses: 'uses' ((name | parname) +); |
|
37 |
\end{rail} |
|
38 |
||
39 |
\begin{descr} |
|
40 |
||
41 |
\item [@{command "theory"}~@{text "A \<IMPORTS> B\<^sub>1 \<dots> |
|
42 |
B\<^sub>n \<BEGIN>"}] starts a new theory @{text A} based on the |
|
43 |
merge of existing theories @{text "B\<^sub>1 \<dots> B\<^sub>n"}. |
|
44 |
||
45 |
Due to inclusion of several ancestors, the overall theory structure |
|
46 |
emerging in an Isabelle session forms a directed acyclic graph |
|
47 |
(DAG). Isabelle's theory loader ensures that the sources |
|
48 |
contributing to the development graph are always up-to-date. |
|
49 |
Changed files are automatically reloaded when processing theory |
|
50 |
headers. |
|
51 |
||
52 |
The optional @{keyword_def "uses"} specification declares additional |
|
53 |
dependencies on extra files (usually ML sources). Files will be |
|
54 |
loaded immediately (as ML), unless the name is put in parentheses, |
|
55 |
which merely documents the dependency to be resolved later in the |
|
56 |
text (typically via explicit @{command_ref "use"} in the body text, |
|
57 |
see \secref{sec:ML}). |
|
58 |
||
27040 | 59 |
\item [@{command (global) "end"}] concludes the current theory |
60 |
definition. |
|
61 |
||
62 |
\end{descr} |
|
63 |
*} |
|
64 |
||
65 |
||
66 |
section {* Local theory targets \label{sec:target} *} |
|
67 |
||
68 |
text {* |
|
69 |
A local theory target is a context managed separately within the |
|
70 |
enclosing theory. Contexts may introduce parameters (fixed |
|
71 |
variables) and assumptions (hypotheses). Definitions and theorems |
|
72 |
depending on the context may be added incrementally later on. Named |
|
73 |
contexts refer to locales (cf.\ \secref{sec:locale}) or type classes |
|
74 |
(cf.\ \secref{sec:class}); the name ``@{text "-"}'' signifies the |
|
75 |
global theory context. |
|
76 |
||
77 |
\begin{matharray}{rcll} |
|
78 |
@{command_def "context"} & : & \isartrans{theory}{local{\dsh}theory} \\ |
|
79 |
@{command_def (local) "end"} & : & \isartrans{local{\dsh}theory}{theory} \\ |
|
80 |
\end{matharray} |
|
81 |
||
82 |
\indexouternonterm{target} |
|
83 |
\begin{rail} |
|
84 |
'context' name 'begin' |
|
85 |
; |
|
86 |
||
87 |
target: '(' 'in' name ')' |
|
88 |
; |
|
89 |
\end{rail} |
|
90 |
||
91 |
\begin{descr} |
|
92 |
||
93 |
\item [@{command "context"}~@{text "c \<BEGIN>"}] recommences an |
|
94 |
existing locale or class context @{text c}. Note that locale and |
|
27051 | 95 |
class definitions allow to include the @{keyword "begin"} keyword as |
96 |
well, in order to continue the local theory immediately after the |
|
97 |
initial specification. |
|
27040 | 98 |
|
99 |
\item [@{command (local) "end"}] concludes the current local theory |
|
100 |
and continues the enclosing global theory. Note that a global |
|
101 |
@{command (global) "end"} has a different meaning: it concludes the |
|
102 |
theory itself (\secref{sec:begin-thy}). |
|
103 |
||
104 |
\item [@{text "(\<IN> c)"}] given after any local theory command |
|
105 |
specifies an immediate target, e.g.\ ``@{command |
|
106 |
"definition"}~@{text "(\<IN> c) \<dots>"}'' or ``@{command |
|
107 |
"theorem"}~@{text "(\<IN> c) \<dots>"}''. This works both in a local or |
|
108 |
global theory context; the current target context will be suspended |
|
109 |
for this command only. Note that ``@{text "(\<IN> -)"}'' will |
|
110 |
always produce a global result independently of the current target |
|
111 |
context. |
|
112 |
||
113 |
\end{descr} |
|
114 |
||
115 |
The exact meaning of results produced within a local theory context |
|
116 |
depends on the underlying target infrastructure (locale, type class |
|
117 |
etc.). The general idea is as follows, considering a context named |
|
118 |
@{text c} with parameter @{text x} and assumption @{text "A[x]"}. |
|
119 |
||
120 |
Definitions are exported by introducing a global version with |
|
121 |
additional arguments; a syntactic abbreviation links the long form |
|
122 |
with the abstract version of the target context. For example, |
|
123 |
@{text "a \<equiv> t[x]"} becomes @{text "c.a ?x \<equiv> t[?x]"} at the theory |
|
124 |
level (for arbitrary @{text "?x"}), together with a local |
|
125 |
abbreviation @{text "c \<equiv> c.a x"} in the target context (for the |
|
126 |
fixed parameter @{text x}). |
|
127 |
||
128 |
Theorems are exported by discharging the assumptions and |
|
129 |
generalizing the parameters of the context. For example, @{text "a: |
|
130 |
B[x]"} becomes @{text "c.a: A[?x] \<Longrightarrow> B[?x]"}, again for arbitrary |
|
131 |
@{text "?x"}. |
|
132 |
*} |
|
133 |
||
134 |
||
135 |
section {* Basic specification elements *} |
|
136 |
||
137 |
text {* |
|
138 |
\begin{matharray}{rcll} |
|
28114 | 139 |
@{command_def "axiomatization"} & : & \isarkeep{theory} & (axiomatic!)\\ |
27040 | 140 |
@{command_def "definition"} & : & \isarkeep{local{\dsh}theory} \\ |
141 |
@{attribute_def "defn"} & : & \isaratt \\ |
|
142 |
@{command_def "abbreviation"} & : & \isarkeep{local{\dsh}theory} \\ |
|
143 |
@{command_def "print_abbrevs"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\ |
|
144 |
@{command_def "notation"} & : & \isarkeep{local{\dsh}theory} \\ |
|
145 |
@{command_def "no_notation"} & : & \isarkeep{local{\dsh}theory} \\ |
|
146 |
\end{matharray} |
|
147 |
||
148 |
These specification mechanisms provide a slightly more abstract view |
|
149 |
than the underlying primitives of @{command "consts"}, @{command |
|
150 |
"defs"} (see \secref{sec:consts}), and @{command "axioms"} (see |
|
151 |
\secref{sec:axms-thms}). In particular, type-inference is commonly |
|
152 |
available, and result names need not be given. |
|
153 |
||
154 |
\begin{rail} |
|
155 |
'axiomatization' target? fixes? ('where' specs)? |
|
156 |
; |
|
157 |
'definition' target? (decl 'where')? thmdecl? prop |
|
158 |
; |
|
159 |
'abbreviation' target? mode? (decl 'where')? prop |
|
160 |
; |
|
161 |
('notation' | 'no\_notation') target? mode? (nameref structmixfix + 'and') |
|
162 |
; |
|
163 |
||
164 |
fixes: ((name ('::' type)? mixfix? | vars) + 'and') |
|
165 |
; |
|
166 |
specs: (thmdecl? props + 'and') |
|
167 |
; |
|
168 |
decl: name ('::' type)? mixfix? |
|
169 |
; |
|
170 |
\end{rail} |
|
171 |
||
172 |
\begin{descr} |
|
173 |
||
174 |
\item [@{command "axiomatization"}~@{text "c\<^sub>1 \<dots> c\<^sub>m |
|
175 |
\<WHERE> \<phi>\<^sub>1 \<dots> \<phi>\<^sub>n"}] introduces several constants |
|
176 |
simultaneously and states axiomatic properties for these. The |
|
177 |
constants are marked as being specified once and for all, which |
|
178 |
prevents additional specifications being issued later on. |
|
179 |
||
180 |
Note that axiomatic specifications are only appropriate when |
|
28114 | 181 |
declaring a new logical system; axiomatic specifications are |
182 |
restricted to global theory contexts. Normal applications should |
|
183 |
only use definitional mechanisms! |
|
27040 | 184 |
|
185 |
\item [@{command "definition"}~@{text "c \<WHERE> eq"}] produces an |
|
186 |
internal definition @{text "c \<equiv> t"} according to the specification |
|
187 |
given as @{text eq}, which is then turned into a proven fact. The |
|
188 |
given proposition may deviate from internal meta-level equality |
|
189 |
according to the rewrite rules declared as @{attribute defn} by the |
|
190 |
object-logic. This usually covers object-level equality @{text "x = |
|
191 |
y"} and equivalence @{text "A \<leftrightarrow> B"}. End-users normally need not |
|
192 |
change the @{attribute defn} setup. |
|
193 |
||
194 |
Definitions may be presented with explicit arguments on the LHS, as |
|
195 |
well as additional conditions, e.g.\ @{text "f x y = t"} instead of |
|
196 |
@{text "f \<equiv> \<lambda>x y. t"} and @{text "y \<noteq> 0 \<Longrightarrow> g x y = u"} instead of an |
|
197 |
unrestricted @{text "g \<equiv> \<lambda>x y. u"}. |
|
198 |
||
199 |
\item [@{command "abbreviation"}~@{text "c \<WHERE> eq"}] introduces |
|
200 |
a syntactic constant which is associated with a certain term |
|
201 |
according to the meta-level equality @{text eq}. |
|
202 |
||
203 |
Abbreviations participate in the usual type-inference process, but |
|
204 |
are expanded before the logic ever sees them. Pretty printing of |
|
205 |
terms involves higher-order rewriting with rules stemming from |
|
206 |
reverted abbreviations. This needs some care to avoid overlapping |
|
207 |
or looping syntactic replacements! |
|
208 |
||
209 |
The optional @{text mode} specification restricts output to a |
|
210 |
particular print mode; using ``@{text input}'' here achieves the |
|
211 |
effect of one-way abbreviations. The mode may also include an |
|
212 |
``@{keyword "output"}'' qualifier that affects the concrete syntax |
|
213 |
declared for abbreviations, cf.\ @{command "syntax"} in |
|
214 |
\secref{sec:syn-trans}. |
|
215 |
||
216 |
\item [@{command "print_abbrevs"}] prints all constant abbreviations |
|
217 |
of the current context. |
|
218 |
||
219 |
\item [@{command "notation"}~@{text "c (mx)"}] associates mixfix |
|
220 |
syntax with an existing constant or fixed variable. This is a |
|
221 |
robust interface to the underlying @{command "syntax"} primitive |
|
222 |
(\secref{sec:syn-trans}). Type declaration and internal syntactic |
|
223 |
representation of the given entity is retrieved from the context. |
|
224 |
||
225 |
\item [@{command "no_notation"}] is similar to @{command |
|
226 |
"notation"}, but removes the specified syntax annotation from the |
|
227 |
present context. |
|
228 |
||
229 |
\end{descr} |
|
230 |
||
231 |
All of these specifications support local theory targets (cf.\ |
|
232 |
\secref{sec:target}). |
|
233 |
*} |
|
234 |
||
235 |
||
236 |
section {* Generic declarations *} |
|
237 |
||
238 |
text {* |
|
239 |
Arbitrary operations on the background context may be wrapped-up as |
|
240 |
generic declaration elements. Since the underlying concept of local |
|
241 |
theories may be subject to later re-interpretation, there is an |
|
242 |
additional dependency on a morphism that tells the difference of the |
|
243 |
original declaration context wrt.\ the application context |
|
244 |
encountered later on. A fact declaration is an important special |
|
245 |
case: it consists of a theorem which is applied to the context by |
|
246 |
means of an attribute. |
|
247 |
||
248 |
\begin{matharray}{rcl} |
|
249 |
@{command_def "declaration"} & : & \isarkeep{local{\dsh}theory} \\ |
|
250 |
@{command_def "declare"} & : & \isarkeep{local{\dsh}theory} \\ |
|
251 |
\end{matharray} |
|
252 |
||
253 |
\begin{rail} |
|
254 |
'declaration' target? text |
|
255 |
; |
|
256 |
'declare' target? (thmrefs + 'and') |
|
257 |
; |
|
258 |
\end{rail} |
|
259 |
||
260 |
\begin{descr} |
|
261 |
||
262 |
\item [@{command "declaration"}~@{text d}] adds the declaration |
|
263 |
function @{text d} of ML type @{ML_type declaration}, to the current |
|
264 |
local theory under construction. In later application contexts, the |
|
265 |
function is transformed according to the morphisms being involved in |
|
266 |
the interpretation hierarchy. |
|
267 |
||
268 |
\item [@{command "declare"}~@{text thms}] declares theorems to the |
|
269 |
current local theory context. No theorem binding is involved here, |
|
270 |
unlike @{command "theorems"} or @{command "lemmas"} (cf.\ |
|
271 |
\secref{sec:axms-thms}), so @{command "declare"} only has the effect |
|
272 |
of applying attributes as included in the theorem specification. |
|
273 |
||
274 |
\end{descr} |
|
275 |
*} |
|
276 |
||
277 |
||
278 |
section {* Locales \label{sec:locale} *} |
|
279 |
||
280 |
text {* |
|
281 |
Locales are named local contexts, consisting of a list of |
|
282 |
declaration elements that are modeled after the Isar proof context |
|
283 |
commands (cf.\ \secref{sec:proof-context}). |
|
284 |
*} |
|
285 |
||
286 |
||
287 |
subsection {* Locale specifications *} |
|
288 |
||
289 |
text {* |
|
290 |
\begin{matharray}{rcl} |
|
291 |
@{command_def "locale"} & : & \isartrans{theory}{local{\dsh}theory} \\ |
|
292 |
@{command_def "print_locale"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\ |
|
293 |
@{command_def "print_locales"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\ |
|
294 |
@{method_def intro_locales} & : & \isarmeth \\ |
|
295 |
@{method_def unfold_locales} & : & \isarmeth \\ |
|
296 |
\end{matharray} |
|
297 |
||
298 |
\indexouternonterm{contextexpr}\indexouternonterm{contextelem} |
|
299 |
\indexisarelem{fixes}\indexisarelem{constrains}\indexisarelem{assumes} |
|
300 |
\indexisarelem{defines}\indexisarelem{notes}\indexisarelem{includes} |
|
301 |
\begin{rail} |
|
27681 | 302 |
'locale' name ('=' localeexpr)? 'begin'? |
27040 | 303 |
; |
304 |
'print\_locale' '!'? localeexpr |
|
305 |
; |
|
306 |
localeexpr: ((contextexpr '+' (contextelem+)) | contextexpr | (contextelem+)) |
|
307 |
; |
|
308 |
||
309 |
contextexpr: nameref | '(' contextexpr ')' | |
|
310 |
(contextexpr (name mixfix? +)) | (contextexpr + '+') |
|
311 |
; |
|
312 |
contextelem: fixes | constrains | assumes | defines | notes |
|
313 |
; |
|
314 |
fixes: 'fixes' ((name ('::' type)? structmixfix? | vars) + 'and') |
|
315 |
; |
|
316 |
constrains: 'constrains' (name '::' type + 'and') |
|
317 |
; |
|
318 |
assumes: 'assumes' (thmdecl? props + 'and') |
|
319 |
; |
|
320 |
defines: 'defines' (thmdecl? prop proppat? + 'and') |
|
321 |
; |
|
322 |
notes: 'notes' (thmdef? thmrefs + 'and') |
|
323 |
; |
|
324 |
includes: 'includes' contextexpr |
|
325 |
; |
|
326 |
\end{rail} |
|
327 |
||
328 |
\begin{descr} |
|
329 |
||
330 |
\item [@{command "locale"}~@{text "loc = import + body"}] defines a |
|
331 |
new locale @{text loc} as a context consisting of a certain view of |
|
332 |
existing locales (@{text import}) plus some additional elements |
|
333 |
(@{text body}). Both @{text import} and @{text body} are optional; |
|
334 |
the degenerate form @{command "locale"}~@{text loc} defines an empty |
|
335 |
locale, which may still be useful to collect declarations of facts |
|
336 |
later on. Type-inference on locale expressions automatically takes |
|
337 |
care of the most general typing that the combined context elements |
|
338 |
may acquire. |
|
339 |
||
340 |
The @{text import} consists of a structured context expression, |
|
341 |
consisting of references to existing locales, renamed contexts, or |
|
342 |
merged contexts. Renaming uses positional notation: @{text "c |
|
343 |
x\<^sub>1 \<dots> x\<^sub>n"} means that (a prefix of) the fixed |
|
344 |
parameters of context @{text c} are named @{text "x\<^sub>1, \<dots>, |
|
345 |
x\<^sub>n"}; a ``@{text _}'' (underscore) means to skip that |
|
346 |
position. Renaming by default deletes concrete syntax, but new |
|
347 |
syntax may by specified with a mixfix annotation. An exeption of |
|
348 |
this rule is the special syntax declared with ``@{text |
|
349 |
"(\<STRUCTURE>)"}'' (see below), which is neither deleted nor can it |
|
350 |
be changed. Merging proceeds from left-to-right, suppressing any |
|
351 |
duplicates stemming from different paths through the import |
|
352 |
hierarchy. |
|
353 |
||
354 |
The @{text body} consists of basic context elements, further context |
|
355 |
expressions may be included as well. |
|
356 |
||
357 |
\begin{descr} |
|
358 |
||
359 |
\item [@{element "fixes"}~@{text "x :: \<tau> (mx)"}] declares a local |
|
360 |
parameter of type @{text \<tau>} and mixfix annotation @{text mx} (both |
|
361 |
are optional). The special syntax declaration ``@{text |
|
362 |
"(\<STRUCTURE>)"}'' means that @{text x} may be referenced |
|
363 |
implicitly in this context. |
|
364 |
||
365 |
\item [@{element "constrains"}~@{text "x :: \<tau>"}] introduces a type |
|
366 |
constraint @{text \<tau>} on the local parameter @{text x}. |
|
367 |
||
368 |
\item [@{element "assumes"}~@{text "a: \<phi>\<^sub>1 \<dots> \<phi>\<^sub>n"}] |
|
369 |
introduces local premises, similar to @{command "assume"} within a |
|
370 |
proof (cf.\ \secref{sec:proof-context}). |
|
371 |
||
372 |
\item [@{element "defines"}~@{text "a: x \<equiv> t"}] defines a previously |
|
373 |
declared parameter. This is similar to @{command "def"} within a |
|
374 |
proof (cf.\ \secref{sec:proof-context}), but @{element "defines"} |
|
375 |
takes an equational proposition instead of variable-term pair. The |
|
376 |
left-hand side of the equation may have additional arguments, e.g.\ |
|
377 |
``@{element "defines"}~@{text "f x\<^sub>1 \<dots> x\<^sub>n \<equiv> t"}''. |
|
378 |
||
379 |
\item [@{element "notes"}~@{text "a = b\<^sub>1 \<dots> b\<^sub>n"}] |
|
380 |
reconsiders facts within a local context. Most notably, this may |
|
381 |
include arbitrary declarations in any attribute specifications |
|
382 |
included here, e.g.\ a local @{attribute simp} rule. |
|
383 |
||
384 |
\item [@{element "includes"}~@{text c}] copies the specified context |
|
385 |
in a statically scoped manner. Only available in the long goal |
|
386 |
format of \secref{sec:goals}. |
|
387 |
||
388 |
In contrast, the initial @{text import} specification of a locale |
|
389 |
expression maintains a dynamic relation to the locales being |
|
390 |
referenced (benefiting from any later fact declarations in the |
|
391 |
obvious manner). |
|
392 |
||
393 |
\end{descr} |
|
394 |
||
395 |
Note that ``@{text "(\<IS> p\<^sub>1 \<dots> p\<^sub>n)"}'' patterns given |
|
396 |
in the syntax of @{element "assumes"} and @{element "defines"} above |
|
397 |
are illegal in locale definitions. In the long goal format of |
|
398 |
\secref{sec:goals}, term bindings may be included as expected, |
|
399 |
though. |
|
400 |
||
401 |
\medskip By default, locale specifications are ``closed up'' by |
|
402 |
turning the given text into a predicate definition @{text |
|
403 |
loc_axioms} and deriving the original assumptions as local lemmas |
|
404 |
(modulo local definitions). The predicate statement covers only the |
|
405 |
newly specified assumptions, omitting the content of included locale |
|
406 |
expressions. The full cumulative view is only provided on export, |
|
407 |
involving another predicate @{text loc} that refers to the complete |
|
408 |
specification text. |
|
409 |
||
410 |
In any case, the predicate arguments are those locale parameters |
|
411 |
that actually occur in the respective piece of text. Also note that |
|
412 |
these predicates operate at the meta-level in theory, but the locale |
|
413 |
packages attempts to internalize statements according to the |
|
414 |
object-logic setup (e.g.\ replacing @{text \<And>} by @{text \<forall>}, and |
|
415 |
@{text "\<Longrightarrow>"} by @{text "\<longrightarrow>"} in HOL; see also |
|
416 |
\secref{sec:object-logic}). Separate introduction rules @{text |
|
417 |
loc_axioms.intro} and @{text loc.intro} are provided as well. |
|
418 |
||
419 |
\item [@{command "print_locale"}~@{text "import + body"}] prints the |
|
420 |
specified locale expression in a flattened form. The notable |
|
421 |
special case @{command "print_locale"}~@{text loc} just prints the |
|
422 |
contents of the named locale, but keep in mind that type-inference |
|
423 |
will normalize type variables according to the usual alphabetical |
|
424 |
order. The command omits @{element "notes"} elements by default. |
|
425 |
Use @{command "print_locale"}@{text "!"} to get them included. |
|
426 |
||
427 |
\item [@{command "print_locales"}] prints the names of all locales |
|
428 |
of the current theory. |
|
429 |
||
430 |
\item [@{method intro_locales} and @{method unfold_locales}] |
|
431 |
repeatedly expand all introduction rules of locale predicates of the |
|
432 |
theory. While @{method intro_locales} only applies the @{text |
|
433 |
loc.intro} introduction rules and therefore does not decend to |
|
434 |
assumptions, @{method unfold_locales} is more aggressive and applies |
|
435 |
@{text loc_axioms.intro} as well. Both methods are aware of locale |
|
436 |
specifications entailed by the context, both from target and |
|
437 |
@{element "includes"} statements, and from interpretations (see |
|
438 |
below). New goals that are entailed by the current context are |
|
439 |
discharged automatically. |
|
440 |
||
441 |
\end{descr} |
|
442 |
*} |
|
443 |
||
444 |
||
445 |
subsection {* Interpretation of locales *} |
|
446 |
||
447 |
text {* |
|
448 |
Locale expressions (more precisely, \emph{context expressions}) may |
|
449 |
be instantiated, and the instantiated facts added to the current |
|
450 |
context. This requires a proof of the instantiated specification |
|
451 |
and is called \emph{locale interpretation}. Interpretation is |
|
452 |
possible in theories and locales (command @{command |
|
453 |
"interpretation"}) and also within a proof body (command @{command |
|
454 |
"interpret"}). |
|
455 |
||
456 |
\begin{matharray}{rcl} |
|
457 |
@{command_def "interpretation"} & : & \isartrans{theory}{proof(prove)} \\ |
|
458 |
@{command_def "interpret"} & : & \isartrans{proof(state) ~|~ proof(chain)}{proof(prove)} \\ |
|
459 |
@{command_def "print_interps"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\ |
|
460 |
\end{matharray} |
|
461 |
||
462 |
\indexouternonterm{interp} |
|
463 |
\begin{rail} |
|
464 |
'interpretation' (interp | name ('<' | subseteq) contextexpr) |
|
465 |
; |
|
466 |
'interpret' interp |
|
467 |
; |
|
468 |
'print\_interps' '!'? name |
|
469 |
; |
|
470 |
instantiation: ('[' (inst+) ']')? |
|
471 |
; |
|
28085
914183e229e9
Interpretation commands no longer accept interpretation attributes.
ballarin
parents:
27681
diff
changeset
|
472 |
interp: (name ':')? \\ (contextexpr instantiation | |
27040 | 473 |
name instantiation 'where' (thmdecl? prop + 'and')) |
474 |
; |
|
475 |
\end{rail} |
|
476 |
||
477 |
\begin{descr} |
|
478 |
||
479 |
\item [@{command "interpretation"}~@{text "expr insts \<WHERE> eqns"}] |
|
480 |
||
481 |
The first form of @{command "interpretation"} interprets @{text |
|
482 |
expr} in the theory. The instantiation is given as a list of terms |
|
483 |
@{text insts} and is positional. All parameters must receive an |
|
484 |
instantiation term --- with the exception of defined parameters. |
|
485 |
These are, if omitted, derived from the defining equation and other |
|
486 |
instantiations. Use ``@{text _}'' to omit an instantiation term. |
|
487 |
||
488 |
The command generates proof obligations for the instantiated |
|
489 |
specifications (assumes and defines elements). Once these are |
|
490 |
discharged by the user, instantiated facts are added to the theory |
|
491 |
in a post-processing phase. |
|
492 |
||
493 |
Additional equations, which are unfolded in facts during |
|
494 |
post-processing, may be given after the keyword @{keyword "where"}. |
|
495 |
This is useful for interpreting concepts introduced through |
|
496 |
definition specification elements. The equations must be proved. |
|
497 |
Note that if equations are present, the context expression is |
|
498 |
restricted to a locale name. |
|
499 |
||
500 |
The command is aware of interpretations already active in the |
|
28085
914183e229e9
Interpretation commands no longer accept interpretation attributes.
ballarin
parents:
27681
diff
changeset
|
501 |
theory, but does not simplify the goal automatically. In order to |
914183e229e9
Interpretation commands no longer accept interpretation attributes.
ballarin
parents:
27681
diff
changeset
|
502 |
simplify the proof obligations use methods @{method intro_locales} |
914183e229e9
Interpretation commands no longer accept interpretation attributes.
ballarin
parents:
27681
diff
changeset
|
503 |
or @{method unfold_locales}. Post-processing is not applied to |
914183e229e9
Interpretation commands no longer accept interpretation attributes.
ballarin
parents:
27681
diff
changeset
|
504 |
facts of interpretations that are already active. This avoids |
914183e229e9
Interpretation commands no longer accept interpretation attributes.
ballarin
parents:
27681
diff
changeset
|
505 |
duplication of interpreted facts, in particular. Note that, in the |
914183e229e9
Interpretation commands no longer accept interpretation attributes.
ballarin
parents:
27681
diff
changeset
|
506 |
case of a locale with import, parts of the interpretation may |
914183e229e9
Interpretation commands no longer accept interpretation attributes.
ballarin
parents:
27681
diff
changeset
|
507 |
already be active. The command will only process facts for new |
914183e229e9
Interpretation commands no longer accept interpretation attributes.
ballarin
parents:
27681
diff
changeset
|
508 |
parts. |
27040 | 509 |
|
28085
914183e229e9
Interpretation commands no longer accept interpretation attributes.
ballarin
parents:
27681
diff
changeset
|
510 |
The context expression may be preceded by a name, which takes effect |
914183e229e9
Interpretation commands no longer accept interpretation attributes.
ballarin
parents:
27681
diff
changeset
|
511 |
in the post-processing of facts. It is used to prefix fact names, |
914183e229e9
Interpretation commands no longer accept interpretation attributes.
ballarin
parents:
27681
diff
changeset
|
512 |
for example to avoid accidental hiding of other facts. |
27040 | 513 |
|
514 |
Adding facts to locales has the effect of adding interpreted facts |
|
515 |
to the theory for all active interpretations also. That is, |
|
516 |
interpretations dynamically participate in any facts added to |
|
517 |
locales. |
|
518 |
||
519 |
\item [@{command "interpretation"}~@{text "name \<subseteq> expr"}] |
|
520 |
||
521 |
This form of the command interprets @{text expr} in the locale |
|
522 |
@{text name}. It requires a proof that the specification of @{text |
|
523 |
name} implies the specification of @{text expr}. As in the |
|
524 |
localized version of the theorem command, the proof is in the |
|
525 |
context of @{text name}. After the proof obligation has been |
|
526 |
dischared, the facts of @{text expr} become part of locale @{text |
|
527 |
name} as \emph{derived} context elements and are available when the |
|
528 |
context @{text name} is subsequently entered. Note that, like |
|
529 |
import, this is dynamic: facts added to a locale part of @{text |
|
530 |
expr} after interpretation become also available in @{text name}. |
|
531 |
Like facts of renamed context elements, facts obtained by |
|
532 |
interpretation may be accessed by prefixing with the parameter |
|
533 |
renaming (where the parameters are separated by ``@{text _}''). |
|
534 |
||
535 |
Unlike interpretation in theories, instantiation is confined to the |
|
536 |
renaming of parameters, which may be specified as part of the |
|
537 |
context expression @{text expr}. Using defined parameters in @{text |
|
538 |
name} one may achieve an effect similar to instantiation, though. |
|
539 |
||
540 |
Only specification fragments of @{text expr} that are not already |
|
541 |
part of @{text name} (be it imported, derived or a derived fragment |
|
542 |
of the import) are considered by interpretation. This enables |
|
543 |
circular interpretations. |
|
544 |
||
545 |
If interpretations of @{text name} exist in the current theory, the |
|
546 |
command adds interpretations for @{text expr} as well, with the same |
|
547 |
prefix and attributes, although only for fragments of @{text expr} |
|
548 |
that are not interpreted in the theory already. |
|
549 |
||
550 |
\item [@{command "interpret"}~@{text "expr insts \<WHERE> eqns"}] |
|
551 |
interprets @{text expr} in the proof context and is otherwise |
|
552 |
similar to interpretation in theories. |
|
553 |
||
554 |
\item [@{command "print_interps"}~@{text loc}] prints the |
|
555 |
interpretations of a particular locale @{text loc} that are active |
|
556 |
in the current context, either theory or proof context. The |
|
557 |
exclamation point argument triggers printing of \emph{witness} |
|
558 |
theorems justifying interpretations. These are normally omitted |
|
559 |
from the output. |
|
560 |
||
561 |
\end{descr} |
|
562 |
||
563 |
\begin{warn} |
|
564 |
Since attributes are applied to interpreted theorems, |
|
565 |
interpretation may modify the context of common proof tools, e.g.\ |
|
566 |
the Simplifier or Classical Reasoner. Since the behavior of such |
|
567 |
automated reasoning tools is \emph{not} stable under |
|
568 |
interpretation morphisms, manual declarations might have to be |
|
569 |
issued. |
|
570 |
\end{warn} |
|
571 |
||
572 |
\begin{warn} |
|
573 |
An interpretation in a theory may subsume previous |
|
574 |
interpretations. This happens if the same specification fragment |
|
575 |
is interpreted twice and the instantiation of the second |
|
576 |
interpretation is more general than the interpretation of the |
|
577 |
first. A warning is issued, since it is likely that these could |
|
578 |
have been generalized in the first place. The locale package does |
|
579 |
not attempt to remove subsumed interpretations. |
|
580 |
\end{warn} |
|
581 |
*} |
|
582 |
||
583 |
||
584 |
section {* Classes \label{sec:class} *} |
|
585 |
||
586 |
text {* |
|
587 |
A class is a particular locale with \emph{exactly one} type variable |
|
588 |
@{text \<alpha>}. Beyond the underlying locale, a corresponding type class |
|
589 |
is established which is interpreted logically as axiomatic type |
|
590 |
class \cite{Wenzel:1997:TPHOL} whose logical content are the |
|
591 |
assumptions of the locale. Thus, classes provide the full |
|
592 |
generality of locales combined with the commodity of type classes |
|
593 |
(notably type-inference). See \cite{isabelle-classes} for a short |
|
594 |
tutorial. |
|
595 |
||
596 |
\begin{matharray}{rcl} |
|
597 |
@{command_def "class"} & : & \isartrans{theory}{local{\dsh}theory} \\ |
|
598 |
@{command_def "instantiation"} & : & \isartrans{theory}{local{\dsh}theory} \\ |
|
599 |
@{command_def "instance"} & : & \isartrans{local{\dsh}theory}{local{\dsh}theory} \\ |
|
600 |
@{command_def "subclass"} & : & \isartrans{local{\dsh}theory}{local{\dsh}theory} \\ |
|
601 |
@{command_def "print_classes"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\ |
|
602 |
@{method_def intro_classes} & : & \isarmeth \\ |
|
603 |
\end{matharray} |
|
604 |
||
605 |
\begin{rail} |
|
606 |
'class' name '=' ((superclassexpr '+' (contextelem+)) | superclassexpr | (contextelem+)) \\ |
|
607 |
'begin'? |
|
608 |
; |
|
609 |
'instantiation' (nameref + 'and') '::' arity 'begin' |
|
610 |
; |
|
611 |
'instance' |
|
612 |
; |
|
613 |
'subclass' target? nameref |
|
614 |
; |
|
615 |
'print\_classes' |
|
616 |
; |
|
617 |
||
618 |
superclassexpr: nameref | (nameref '+' superclassexpr) |
|
619 |
; |
|
620 |
\end{rail} |
|
621 |
||
622 |
\begin{descr} |
|
623 |
||
624 |
\item [@{command "class"}~@{text "c = superclasses + body"}] defines |
|
625 |
a new class @{text c}, inheriting from @{text superclasses}. This |
|
626 |
introduces a locale @{text c} with import of all locales @{text |
|
627 |
superclasses}. |
|
628 |
||
629 |
Any @{element "fixes"} in @{text body} are lifted to the global |
|
630 |
theory level (\emph{class operations} @{text "f\<^sub>1, \<dots>, |
|
631 |
f\<^sub>n"} of class @{text c}), mapping the local type parameter |
|
632 |
@{text \<alpha>} to a schematic type variable @{text "?\<alpha> :: c"}. |
|
633 |
||
634 |
Likewise, @{element "assumes"} in @{text body} are also lifted, |
|
635 |
mapping each local parameter @{text "f :: \<tau>[\<alpha>]"} to its |
|
636 |
corresponding global constant @{text "f :: \<tau>[?\<alpha> :: c]"}. The |
|
637 |
corresponding introduction rule is provided as @{text |
|
638 |
c_class_axioms.intro}. This rule should be rarely needed directly |
|
639 |
--- the @{method intro_classes} method takes care of the details of |
|
640 |
class membership proofs. |
|
641 |
||
642 |
\item [@{command "instantiation"}~@{text "t :: (s\<^sub>1, \<dots>, |
|
643 |
s\<^sub>n) s \<BEGIN>"}] opens a theory target (cf.\ |
|
644 |
\secref{sec:target}) which allows to specify class operations @{text |
|
645 |
"f\<^sub>1, \<dots>, f\<^sub>n"} corresponding to sort @{text s} at the |
|
646 |
particular type instance @{text "(\<alpha>\<^sub>1 :: s\<^sub>1, \<dots>, |
|
647 |
\<alpha>\<^sub>n :: s\<^sub>n) t"}. A plain @{command "instance"} command |
|
648 |
in the target body poses a goal stating these type arities. The |
|
649 |
target is concluded by an @{command_ref (local) "end"} command. |
|
650 |
||
651 |
Note that a list of simultaneous type constructors may be given; |
|
652 |
this corresponds nicely to mutual recursive type definitions, e.g.\ |
|
653 |
in Isabelle/HOL. |
|
654 |
||
655 |
\item [@{command "instance"}] in an instantiation target body sets |
|
656 |
up a goal stating the type arities claimed at the opening @{command |
|
657 |
"instantiation"}. The proof would usually proceed by @{method |
|
658 |
intro_classes}, and then establish the characteristic theorems of |
|
659 |
the type classes involved. After finishing the proof, the |
|
660 |
background theory will be augmented by the proven type arities. |
|
661 |
||
662 |
\item [@{command "subclass"}~@{text c}] in a class context for class |
|
663 |
@{text d} sets up a goal stating that class @{text c} is logically |
|
664 |
contained in class @{text d}. After finishing the proof, class |
|
665 |
@{text d} is proven to be subclass @{text c} and the locale @{text |
|
666 |
c} is interpreted into @{text d} simultaneously. |
|
667 |
||
668 |
\item [@{command "print_classes"}] prints all classes in the current |
|
669 |
theory. |
|
670 |
||
671 |
\item [@{method intro_classes}] repeatedly expands all class |
|
672 |
introduction rules of this theory. Note that this method usually |
|
673 |
needs not be named explicitly, as it is already included in the |
|
674 |
default proof step (e.g.\ of @{command "proof"}). In particular, |
|
675 |
instantiation of trivial (syntactic) classes may be performed by a |
|
676 |
single ``@{command ".."}'' proof step. |
|
26870 | 677 |
|
678 |
\end{descr} |
|
679 |
*} |
|
680 |
||
27040 | 681 |
|
682 |
subsection {* The class target *} |
|
683 |
||
684 |
text {* |
|
685 |
%FIXME check |
|
686 |
||
687 |
A named context may refer to a locale (cf.\ \secref{sec:target}). |
|
688 |
If this locale is also a class @{text c}, apart from the common |
|
689 |
locale target behaviour the following happens. |
|
690 |
||
691 |
\begin{itemize} |
|
692 |
||
693 |
\item Local constant declarations @{text "g[\<alpha>]"} referring to the |
|
694 |
local type parameter @{text \<alpha>} and local parameters @{text "f[\<alpha>]"} |
|
695 |
are accompanied by theory-level constants @{text "g[?\<alpha> :: c]"} |
|
696 |
referring to theory-level class operations @{text "f[?\<alpha> :: c]"}. |
|
697 |
||
698 |
\item Local theorem bindings are lifted as are assumptions. |
|
699 |
||
700 |
\item Local syntax refers to local operations @{text "g[\<alpha>]"} and |
|
701 |
global operations @{text "g[?\<alpha> :: c]"} uniformly. Type inference |
|
702 |
resolves ambiguities. In rare cases, manual type annotations are |
|
703 |
needed. |
|
704 |
||
705 |
\end{itemize} |
|
706 |
*} |
|
707 |
||
708 |
||
27053 | 709 |
subsection {* Old-style axiomatic type classes \label{sec:axclass} *} |
27040 | 710 |
|
711 |
text {* |
|
712 |
\begin{matharray}{rcl} |
|
713 |
@{command_def "axclass"} & : & \isartrans{theory}{theory} \\ |
|
714 |
@{command_def "instance"} & : & \isartrans{theory}{proof(prove)} \\ |
|
715 |
\end{matharray} |
|
716 |
||
717 |
Axiomatic type classes are Isabelle/Pure's primitive |
|
718 |
\emph{definitional} interface to type classes. For practical |
|
719 |
applications, you should consider using classes |
|
720 |
(cf.~\secref{sec:classes}) which provide high level interface. |
|
721 |
||
722 |
\begin{rail} |
|
723 |
'axclass' classdecl (axmdecl prop +) |
|
724 |
; |
|
725 |
'instance' (nameref ('<' | subseteq) nameref | nameref '::' arity) |
|
726 |
; |
|
727 |
\end{rail} |
|
728 |
||
729 |
\begin{descr} |
|
730 |
||
731 |
\item [@{command "axclass"}~@{text "c \<subseteq> c\<^sub>1, \<dots>, c\<^sub>n |
|
732 |
axms"}] defines an axiomatic type class as the intersection of |
|
733 |
existing classes, with additional axioms holding. Class axioms may |
|
734 |
not contain more than one type variable. The class axioms (with |
|
735 |
implicit sort constraints added) are bound to the given names. |
|
736 |
Furthermore a class introduction rule is generated (being bound as |
|
737 |
@{text c_class.intro}); this rule is employed by method @{method |
|
738 |
intro_classes} to support instantiation proofs of this class. |
|
739 |
||
740 |
The ``class axioms'' are stored as theorems according to the given |
|
741 |
name specifications, adding @{text "c_class"} as name space prefix; |
|
742 |
the same facts are also stored collectively as @{text |
|
743 |
c_class.axioms}. |
|
744 |
||
745 |
\item [@{command "instance"}~@{text "c\<^sub>1 \<subseteq> c\<^sub>2"} and |
|
746 |
@{command "instance"}~@{text "t :: (s\<^sub>1, \<dots>, s\<^sub>n) s"}] |
|
747 |
setup a goal stating a class relation or type arity. The proof |
|
748 |
would usually proceed by @{method intro_classes}, and then establish |
|
749 |
the characteristic theorems of the type classes involved. After |
|
750 |
finishing the proof, the theory will be augmented by a type |
|
751 |
signature declaration corresponding to the resulting theorem. |
|
752 |
||
753 |
\end{descr} |
|
754 |
*} |
|
755 |
||
756 |
||
757 |
section {* Unrestricted overloading *} |
|
758 |
||
759 |
text {* |
|
760 |
Isabelle/Pure's definitional schemes support certain forms of |
|
761 |
overloading (see \secref{sec:consts}). At most occassions |
|
762 |
overloading will be used in a Haskell-like fashion together with |
|
763 |
type classes by means of @{command "instantiation"} (see |
|
764 |
\secref{sec:class}). Sometimes low-level overloading is desirable. |
|
765 |
The @{command "overloading"} target provides a convenient view for |
|
766 |
end-users. |
|
767 |
||
768 |
\begin{matharray}{rcl} |
|
769 |
@{command_def "overloading"} & : & \isartrans{theory}{local{\dsh}theory} \\ |
|
770 |
\end{matharray} |
|
771 |
||
772 |
\begin{rail} |
|
773 |
'overloading' \\ |
|
774 |
( string ( '==' | equiv ) term ( '(' 'unchecked' ')' )? + ) 'begin' |
|
775 |
\end{rail} |
|
776 |
||
777 |
\begin{descr} |
|
778 |
||
779 |
\item [@{command "overloading"}~@{text "x\<^sub>1 \<equiv> c\<^sub>1 :: |
|
780 |
\<tau>\<^sub>1 \<AND> \<dots> x\<^sub>n \<equiv> c\<^sub>n :: \<tau>\<^sub>n \<BEGIN>"}] |
|
781 |
opens a theory target (cf.\ \secref{sec:target}) which allows to |
|
782 |
specify constants with overloaded definitions. These are identified |
|
783 |
by an explicitly given mapping from variable names @{text |
|
784 |
"x\<^sub>i"} to constants @{text "c\<^sub>i"} at particular type |
|
785 |
instances. The definitions themselves are established using common |
|
786 |
specification tools, using the names @{text "x\<^sub>i"} as |
|
787 |
reference to the corresponding constants. The target is concluded |
|
788 |
by @{command (local) "end"}. |
|
789 |
||
790 |
A @{text "(unchecked)"} option disables global dependency checks for |
|
791 |
the corresponding definition, which is occasionally useful for |
|
792 |
exotic overloading. It is at the discretion of the user to avoid |
|
793 |
malformed theory specifications! |
|
794 |
||
795 |
\end{descr} |
|
796 |
*} |
|
797 |
||
798 |
||
799 |
section {* Incorporating ML code \label{sec:ML} *} |
|
800 |
||
801 |
text {* |
|
802 |
\begin{matharray}{rcl} |
|
803 |
@{command_def "use"} & : & \isarkeep{theory~|~local{\dsh}theory} \\ |
|
804 |
@{command_def "ML"} & : & \isarkeep{theory~|~local{\dsh}theory} \\ |
|
805 |
@{command_def "ML_val"} & : & \isartrans{\cdot}{\cdot} \\ |
|
806 |
@{command_def "ML_command"} & : & \isartrans{\cdot}{\cdot} \\ |
|
807 |
@{command_def "setup"} & : & \isartrans{theory}{theory} \\ |
|
808 |
@{command_def "method_setup"} & : & \isartrans{theory}{theory} \\ |
|
809 |
\end{matharray} |
|
810 |
||
811 |
\begin{rail} |
|
812 |
'use' name |
|
813 |
; |
|
814 |
('ML' | 'ML\_val' | 'ML\_command' | 'setup') text |
|
815 |
; |
|
816 |
'method\_setup' name '=' text text |
|
817 |
; |
|
818 |
\end{rail} |
|
819 |
||
820 |
\begin{descr} |
|
821 |
||
822 |
\item [@{command "use"}~@{text "file"}] reads and executes ML |
|
823 |
commands from @{text "file"}. The current theory context is passed |
|
824 |
down to the ML toplevel and may be modified, using @{ML |
|
825 |
"Context.>>"} or derived ML commands. The file name is checked with |
|
826 |
the @{keyword_ref "uses"} dependency declaration given in the theory |
|
827 |
header (see also \secref{sec:begin-thy}). |
|
828 |
||
829 |
\item [@{command "ML"}~@{text "text"}] is similar to @{command |
|
830 |
"use"}, but executes ML commands directly from the given @{text |
|
831 |
"text"}. |
|
832 |
||
833 |
\item [@{command "ML_val"} and @{command "ML_command"}] are |
|
834 |
diagnostic versions of @{command "ML"}, which means that the context |
|
835 |
may not be updated. @{command "ML_val"} echos the bindings produced |
|
836 |
at the ML toplevel, but @{command "ML_command"} is silent. |
|
837 |
||
838 |
\item [@{command "setup"}~@{text "text"}] changes the current theory |
|
839 |
context by applying @{text "text"}, which refers to an ML expression |
|
840 |
of type @{ML_type "theory -> theory"}. This enables to initialize |
|
841 |
any object-logic specific tools and packages written in ML, for |
|
842 |
example. |
|
843 |
||
844 |
\item [@{command "method_setup"}~@{text "name = text description"}] |
|
845 |
defines a proof method in the current theory. The given @{text |
|
846 |
"text"} has to be an ML expression of type @{ML_type "Args.src -> |
|
847 |
Proof.context -> Proof.method"}. Parsing concrete method syntax |
|
848 |
from @{ML_type Args.src} input can be quite tedious in general. The |
|
849 |
following simple examples are for methods without any explicit |
|
850 |
arguments, or a list of theorems, respectively. |
|
851 |
||
852 |
%FIXME proper antiquotations |
|
853 |
{\footnotesize |
|
854 |
\begin{verbatim} |
|
855 |
Method.no_args (Method.METHOD (fn facts => foobar_tac)) |
|
856 |
Method.thms_args (fn thms => Method.METHOD (fn facts => foobar_tac)) |
|
857 |
Method.ctxt_args (fn ctxt => Method.METHOD (fn facts => foobar_tac)) |
|
858 |
Method.thms_ctxt_args (fn thms => fn ctxt => |
|
859 |
Method.METHOD (fn facts => foobar_tac)) |
|
860 |
\end{verbatim} |
|
861 |
} |
|
862 |
||
863 |
Note that mere tactic emulations may ignore the @{text facts} |
|
864 |
parameter above. Proper proof methods would do something |
|
865 |
appropriate with the list of current facts, though. Single-rule |
|
866 |
methods usually do strict forward-chaining (e.g.\ by using @{ML |
|
867 |
Drule.multi_resolves}), while automatic ones just insert the facts |
|
868 |
using @{ML Method.insert_tac} before applying the main tactic. |
|
869 |
||
870 |
\end{descr} |
|
871 |
*} |
|
872 |
||
873 |
||
874 |
section {* Primitive specification elements *} |
|
875 |
||
876 |
subsection {* Type classes and sorts \label{sec:classes} *} |
|
877 |
||
878 |
text {* |
|
879 |
\begin{matharray}{rcll} |
|
880 |
@{command_def "classes"} & : & \isartrans{theory}{theory} \\ |
|
881 |
@{command_def "classrel"} & : & \isartrans{theory}{theory} & (axiomatic!) \\ |
|
882 |
@{command_def "defaultsort"} & : & \isartrans{theory}{theory} \\ |
|
27059 | 883 |
@{command_def "class_deps"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\ |
27040 | 884 |
\end{matharray} |
885 |
||
886 |
\begin{rail} |
|
887 |
'classes' (classdecl +) |
|
888 |
; |
|
889 |
'classrel' (nameref ('<' | subseteq) nameref + 'and') |
|
890 |
; |
|
891 |
'defaultsort' sort |
|
892 |
; |
|
893 |
\end{rail} |
|
894 |
||
895 |
\begin{descr} |
|
896 |
||
897 |
\item [@{command "classes"}~@{text "c \<subseteq> c\<^sub>1, \<dots>, c\<^sub>n"}] |
|
898 |
declares class @{text c} to be a subclass of existing classes @{text |
|
899 |
"c\<^sub>1, \<dots>, c\<^sub>n"}. Cyclic class structures are not permitted. |
|
900 |
||
901 |
\item [@{command "classrel"}~@{text "c\<^sub>1 \<subseteq> c\<^sub>2"}] states |
|
902 |
subclass relations between existing classes @{text "c\<^sub>1"} and |
|
903 |
@{text "c\<^sub>2"}. This is done axiomatically! The @{command_ref |
|
904 |
"instance"} command (see \secref{sec:axclass}) provides a way to |
|
905 |
introduce proven class relations. |
|
906 |
||
907 |
\item [@{command "defaultsort"}~@{text s}] makes sort @{text s} the |
|
908 |
new default sort for any type variables given without sort |
|
909 |
constraints. Usually, the default sort would be only changed when |
|
910 |
defining a new object-logic. |
|
911 |
||
912 |
\item [@{command "class_deps"}] visualizes the subclass relation, |
|
913 |
using Isabelle's graph browser tool (see also \cite{isabelle-sys}). |
|
914 |
||
915 |
\end{descr} |
|
916 |
*} |
|
917 |
||
918 |
||
919 |
subsection {* Types and type abbreviations \label{sec:types-pure} *} |
|
920 |
||
921 |
text {* |
|
922 |
\begin{matharray}{rcll} |
|
923 |
@{command_def "types"} & : & \isartrans{theory}{theory} \\ |
|
924 |
@{command_def "typedecl"} & : & \isartrans{theory}{theory} \\ |
|
925 |
@{command_def "nonterminals"} & : & \isartrans{theory}{theory} \\ |
|
926 |
@{command_def "arities"} & : & \isartrans{theory}{theory} & (axiomatic!) \\ |
|
927 |
\end{matharray} |
|
928 |
||
929 |
\begin{rail} |
|
930 |
'types' (typespec '=' type infix? +) |
|
931 |
; |
|
932 |
'typedecl' typespec infix? |
|
933 |
; |
|
934 |
'nonterminals' (name +) |
|
935 |
; |
|
936 |
'arities' (nameref '::' arity +) |
|
937 |
; |
|
938 |
\end{rail} |
|
939 |
||
940 |
\begin{descr} |
|
941 |
||
942 |
\item [@{command "types"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t = \<tau>"}] |
|
943 |
introduces \emph{type synonym} @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t"} |
|
944 |
for existing type @{text "\<tau>"}. Unlike actual type definitions, as |
|
945 |
are available in Isabelle/HOL for example, type synonyms are just |
|
946 |
purely syntactic abbreviations without any logical significance. |
|
947 |
Internally, type synonyms are fully expanded. |
|
948 |
||
949 |
\item [@{command "typedecl"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t"}] |
|
950 |
declares a new type constructor @{text t}, intended as an actual |
|
951 |
logical type (of the object-logic, if available). |
|
952 |
||
953 |
\item [@{command "nonterminals"}~@{text c}] declares type |
|
954 |
constructors @{text c} (without arguments) to act as purely |
|
955 |
syntactic types, i.e.\ nonterminal symbols of Isabelle's inner |
|
956 |
syntax of terms or types. |
|
957 |
||
958 |
\item [@{command "arities"}~@{text "t :: (s\<^sub>1, \<dots>, s\<^sub>n) |
|
959 |
s"}] augments Isabelle's order-sorted signature of types by new type |
|
960 |
constructor arities. This is done axiomatically! The @{command_ref |
|
961 |
"instance"} command (see \S\ref{sec:axclass}) provides a way to |
|
962 |
introduce proven type arities. |
|
963 |
||
964 |
\end{descr} |
|
965 |
*} |
|
966 |
||
967 |
||
968 |
subsection {* Constants and definitions \label{sec:consts} *} |
|
969 |
||
970 |
text {* |
|
971 |
Definitions essentially express abbreviations within the logic. The |
|
972 |
simplest form of a definition is @{text "c :: \<sigma> \<equiv> t"}, where @{text |
|
973 |
c} is a newly declared constant. Isabelle also allows derived forms |
|
974 |
where the arguments of @{text c} appear on the left, abbreviating a |
|
975 |
prefix of @{text \<lambda>}-abstractions, e.g.\ @{text "c \<equiv> \<lambda>x y. t"} may be |
|
976 |
written more conveniently as @{text "c x y \<equiv> t"}. Moreover, |
|
977 |
definitions may be weakened by adding arbitrary pre-conditions: |
|
978 |
@{text "A \<Longrightarrow> c x y \<equiv> t"}. |
|
979 |
||
980 |
\medskip The built-in well-formedness conditions for definitional |
|
981 |
specifications are: |
|
982 |
||
983 |
\begin{itemize} |
|
984 |
||
985 |
\item Arguments (on the left-hand side) must be distinct variables. |
|
986 |
||
987 |
\item All variables on the right-hand side must also appear on the |
|
988 |
left-hand side. |
|
989 |
||
990 |
\item All type variables on the right-hand side must also appear on |
|
991 |
the left-hand side; this prohibits @{text "0 :: nat \<equiv> length ([] :: |
|
992 |
\<alpha> list)"} for example. |
|
993 |
||
994 |
\item The definition must not be recursive. Most object-logics |
|
995 |
provide definitional principles that can be used to express |
|
996 |
recursion safely. |
|
997 |
||
998 |
\end{itemize} |
|
999 |
||
1000 |
Overloading means that a constant being declared as @{text "c :: \<alpha> |
|
1001 |
decl"} may be defined separately on type instances @{text "c :: |
|
1002 |
(\<beta>\<^sub>1, \<dots>, \<beta>\<^sub>n) t decl"} for each type constructor @{text |
|
1003 |
t}. The right-hand side may mention overloaded constants |
|
1004 |
recursively at type instances corresponding to the immediate |
|
1005 |
argument types @{text "\<beta>\<^sub>1, \<dots>, \<beta>\<^sub>n"}. Incomplete |
|
1006 |
specification patterns impose global constraints on all occurrences, |
|
1007 |
e.g.\ @{text "d :: \<alpha> \<times> \<alpha>"} on the left-hand side means that all |
|
1008 |
corresponding occurrences on some right-hand side need to be an |
|
1009 |
instance of this, general @{text "d :: \<alpha> \<times> \<beta>"} will be disallowed. |
|
1010 |
||
1011 |
\begin{matharray}{rcl} |
|
1012 |
@{command_def "consts"} & : & \isartrans{theory}{theory} \\ |
|
1013 |
@{command_def "defs"} & : & \isartrans{theory}{theory} \\ |
|
1014 |
@{command_def "constdefs"} & : & \isartrans{theory}{theory} \\ |
|
1015 |
\end{matharray} |
|
1016 |
||
1017 |
\begin{rail} |
|
1018 |
'consts' ((name '::' type mixfix?) +) |
|
1019 |
; |
|
1020 |
'defs' ('(' 'unchecked'? 'overloaded'? ')')? \\ (axmdecl prop +) |
|
1021 |
; |
|
1022 |
\end{rail} |
|
1023 |
||
1024 |
\begin{rail} |
|
1025 |
'constdefs' structs? (constdecl? constdef +) |
|
1026 |
; |
|
1027 |
||
1028 |
structs: '(' 'structure' (vars + 'and') ')' |
|
1029 |
; |
|
1030 |
constdecl: ((name '::' type mixfix | name '::' type | name mixfix) 'where'?) | name 'where' |
|
1031 |
; |
|
1032 |
constdef: thmdecl? prop |
|
1033 |
; |
|
1034 |
\end{rail} |
|
1035 |
||
1036 |
\begin{descr} |
|
1037 |
||
1038 |
\item [@{command "consts"}~@{text "c :: \<sigma>"}] declares constant |
|
1039 |
@{text c} to have any instance of type scheme @{text \<sigma>}. The |
|
1040 |
optional mixfix annotations may attach concrete syntax to the |
|
1041 |
constants declared. |
|
1042 |
||
1043 |
\item [@{command "defs"}~@{text "name: eqn"}] introduces @{text eqn} |
|
1044 |
as a definitional axiom for some existing constant. |
|
1045 |
||
1046 |
The @{text "(unchecked)"} option disables global dependency checks |
|
1047 |
for this definition, which is occasionally useful for exotic |
|
1048 |
overloading. It is at the discretion of the user to avoid malformed |
|
1049 |
theory specifications! |
|
1050 |
||
1051 |
The @{text "(overloaded)"} option declares definitions to be |
|
1052 |
potentially overloaded. Unless this option is given, a warning |
|
1053 |
message would be issued for any definitional equation with a more |
|
1054 |
special type than that of the corresponding constant declaration. |
|
1055 |
||
1056 |
\item [@{command "constdefs"}] provides a streamlined combination of |
|
1057 |
constants declarations and definitions: type-inference takes care of |
|
1058 |
the most general typing of the given specification (the optional |
|
1059 |
type constraint may refer to type-inference dummies ``@{text |
|
1060 |
_}'' as usual). The resulting type declaration needs to agree with |
|
1061 |
that of the specification; overloading is \emph{not} supported here! |
|
1062 |
||
1063 |
The constant name may be omitted altogether, if neither type nor |
|
1064 |
syntax declarations are given. The canonical name of the |
|
1065 |
definitional axiom for constant @{text c} will be @{text c_def}, |
|
1066 |
unless specified otherwise. Also note that the given list of |
|
1067 |
specifications is processed in a strictly sequential manner, with |
|
1068 |
type-checking being performed independently. |
|
1069 |
||
1070 |
An optional initial context of @{text "(structure)"} declarations |
|
1071 |
admits use of indexed syntax, using the special symbol @{verbatim |
|
1072 |
"\<index>"} (printed as ``@{text "\<index>"}''). The latter concept is |
|
1073 |
particularly useful with locales (see also \S\ref{sec:locale}). |
|
1074 |
||
1075 |
\end{descr} |
|
1076 |
*} |
|
1077 |
||
1078 |
||
1079 |
section {* Axioms and theorems \label{sec:axms-thms} *} |
|
1080 |
||
1081 |
text {* |
|
1082 |
\begin{matharray}{rcll} |
|
1083 |
@{command_def "axioms"} & : & \isartrans{theory}{theory} & (axiomatic!) \\ |
|
1084 |
@{command_def "lemmas"} & : & \isarkeep{local{\dsh}theory} \\ |
|
27046 | 1085 |
@{command_def "theorems"} & : & \isarkeep{local{\dsh}theory} \\ |
27040 | 1086 |
\end{matharray} |
1087 |
||
1088 |
\begin{rail} |
|
1089 |
'axioms' (axmdecl prop +) |
|
1090 |
; |
|
1091 |
('lemmas' | 'theorems') target? (thmdef? thmrefs + 'and') |
|
1092 |
; |
|
1093 |
\end{rail} |
|
1094 |
||
1095 |
\begin{descr} |
|
1096 |
||
1097 |
\item [@{command "axioms"}~@{text "a: \<phi>"}] introduces arbitrary |
|
1098 |
statements as axioms of the meta-logic. In fact, axioms are |
|
1099 |
``axiomatic theorems'', and may be referred later just as any other |
|
1100 |
theorem. |
|
1101 |
||
1102 |
Axioms are usually only introduced when declaring new logical |
|
1103 |
systems. Everyday work is typically done the hard way, with proper |
|
1104 |
definitions and proven theorems. |
|
1105 |
||
1106 |
\item [@{command "lemmas"}~@{text "a = b\<^sub>1 \<dots> b\<^sub>n"}] |
|
1107 |
retrieves and stores existing facts in the theory context, or the |
|
1108 |
specified target context (see also \secref{sec:target}). Typical |
|
1109 |
applications would also involve attributes, to declare Simplifier |
|
1110 |
rules, for example. |
|
1111 |
||
1112 |
\item [@{command "theorems"}] is essentially the same as @{command |
|
1113 |
"lemmas"}, but marks the result as a different kind of facts. |
|
1114 |
||
1115 |
\end{descr} |
|
1116 |
*} |
|
1117 |
||
1118 |
||
1119 |
section {* Oracles *} |
|
1120 |
||
1121 |
text {* |
|
1122 |
\begin{matharray}{rcl} |
|
1123 |
@{command_def "oracle"} & : & \isartrans{theory}{theory} \\ |
|
1124 |
\end{matharray} |
|
1125 |
||
1126 |
The oracle interface promotes a given ML function @{ML_text |
|
1127 |
"theory -> T -> term"} to @{ML_text "theory -> T -> thm"}, for some |
|
1128 |
type @{ML_text T} given by the user. This acts like an infinitary |
|
1129 |
specification of axioms -- there is no internal check of the |
|
1130 |
correctness of the results! The inference kernel records oracle |
|
1131 |
invocations within the internal derivation object of theorems, and |
|
1132 |
the pretty printer attaches ``@{text "[!]"}'' to indicate results |
|
1133 |
that are not fully checked by Isabelle inferences. |
|
1134 |
||
1135 |
\begin{rail} |
|
1136 |
'oracle' name '(' type ')' '=' text |
|
1137 |
; |
|
1138 |
\end{rail} |
|
1139 |
||
1140 |
\begin{descr} |
|
1141 |
||
1142 |
\item [@{command "oracle"}~@{text "name (type) = text"}] turns the |
|
1143 |
given ML expression @{text "text"} of type |
|
1144 |
@{ML_text "theory ->"}~@{text "type"}~@{ML_text "-> term"} into an |
|
1145 |
ML function of type |
|
1146 |
@{ML_text "theory ->"}~@{text "type"}~@{ML_text "-> thm"}, which is |
|
1147 |
bound to the global identifier @{ML_text name}. |
|
1148 |
||
1149 |
\end{descr} |
|
1150 |
*} |
|
1151 |
||
1152 |
||
1153 |
section {* Name spaces *} |
|
1154 |
||
1155 |
text {* |
|
1156 |
\begin{matharray}{rcl} |
|
1157 |
@{command_def "global"} & : & \isartrans{theory}{theory} \\ |
|
1158 |
@{command_def "local"} & : & \isartrans{theory}{theory} \\ |
|
1159 |
@{command_def "hide"} & : & \isartrans{theory}{theory} \\ |
|
1160 |
\end{matharray} |
|
1161 |
||
1162 |
\begin{rail} |
|
1163 |
'hide' ('(open)')? name (nameref + ) |
|
1164 |
; |
|
1165 |
\end{rail} |
|
1166 |
||
1167 |
Isabelle organizes any kind of name declarations (of types, |
|
1168 |
constants, theorems etc.) by separate hierarchically structured name |
|
1169 |
spaces. Normally the user does not have to control the behavior of |
|
1170 |
name spaces by hand, yet the following commands provide some way to |
|
1171 |
do so. |
|
1172 |
||
1173 |
\begin{descr} |
|
1174 |
||
1175 |
\item [@{command "global"} and @{command "local"}] change the |
|
1176 |
current name declaration mode. Initially, theories start in |
|
1177 |
@{command "local"} mode, causing all names to be automatically |
|
1178 |
qualified by the theory name. Changing this to @{command "global"} |
|
1179 |
causes all names to be declared without the theory prefix, until |
|
1180 |
@{command "local"} is declared again. |
|
1181 |
||
1182 |
Note that global names are prone to get hidden accidently later, |
|
1183 |
when qualified names of the same base name are introduced. |
|
1184 |
||
1185 |
\item [@{command "hide"}~@{text "space names"}] fully removes |
|
1186 |
declarations from a given name space (which may be @{text "class"}, |
|
1187 |
@{text "type"}, @{text "const"}, or @{text "fact"}); with the @{text |
|
1188 |
"(open)"} option, only the base name is hidden. Global |
|
1189 |
(unqualified) names may never be hidden. |
|
1190 |
||
1191 |
Note that hiding name space accesses has no impact on logical |
|
1192 |
declarations -- they remain valid internally. Entities that are no |
|
1193 |
longer accessible to the user are printed with the special qualifier |
|
1194 |
``@{text "??"}'' prefixed to the full internal name. |
|
1195 |
||
1196 |
\end{descr} |
|
1197 |
*} |
|
1198 |
||
1199 |
||
1200 |
section {* Syntax and translations \label{sec:syn-trans} *} |
|
1201 |
||
1202 |
text {* |
|
1203 |
\begin{matharray}{rcl} |
|
1204 |
@{command_def "syntax"} & : & \isartrans{theory}{theory} \\ |
|
1205 |
@{command_def "no_syntax"} & : & \isartrans{theory}{theory} \\ |
|
1206 |
@{command_def "translations"} & : & \isartrans{theory}{theory} \\ |
|
1207 |
@{command_def "no_translations"} & : & \isartrans{theory}{theory} \\ |
|
1208 |
\end{matharray} |
|
1209 |
||
1210 |
\begin{rail} |
|
1211 |
('syntax' | 'no\_syntax') mode? (constdecl +) |
|
1212 |
; |
|
1213 |
('translations' | 'no\_translations') (transpat ('==' | '=>' | '<=' | rightleftharpoons | rightharpoonup | leftharpoondown) transpat +) |
|
1214 |
; |
|
1215 |
||
1216 |
mode: ('(' ( name | 'output' | name 'output' ) ')') |
|
1217 |
; |
|
1218 |
transpat: ('(' nameref ')')? string |
|
1219 |
; |
|
1220 |
\end{rail} |
|
1221 |
||
1222 |
\begin{descr} |
|
1223 |
||
1224 |
\item [@{command "syntax"}~@{text "(mode) decls"}] is similar to |
|
1225 |
@{command "consts"}~@{text decls}, except that the actual logical |
|
1226 |
signature extension is omitted. Thus the context free grammar of |
|
1227 |
Isabelle's inner syntax may be augmented in arbitrary ways, |
|
1228 |
independently of the logic. The @{text mode} argument refers to the |
|
1229 |
print mode that the grammar rules belong; unless the @{keyword_ref |
|
1230 |
"output"} indicator is given, all productions are added both to the |
|
1231 |
input and output grammar. |
|
1232 |
||
1233 |
\item [@{command "no_syntax"}~@{text "(mode) decls"}] removes |
|
1234 |
grammar declarations (and translations) resulting from @{text |
|
1235 |
decls}, which are interpreted in the same manner as for @{command |
|
1236 |
"syntax"} above. |
|
1237 |
||
1238 |
\item [@{command "translations"}~@{text rules}] specifies syntactic |
|
1239 |
translation rules (i.e.\ macros): parse~/ print rules (@{text "\<rightleftharpoons>"}), |
|
1240 |
parse rules (@{text "\<rightharpoonup>"}), or print rules (@{text "\<leftharpoondown>"}). |
|
1241 |
Translation patterns may be prefixed by the syntactic category to be |
|
1242 |
used for parsing; the default is @{text logic}. |
|
1243 |
||
1244 |
\item [@{command "no_translations"}~@{text rules}] removes syntactic |
|
1245 |
translation rules, which are interpreted in the same manner as for |
|
1246 |
@{command "translations"} above. |
|
1247 |
||
1248 |
\end{descr} |
|
1249 |
*} |
|
1250 |
||
1251 |
||
1252 |
section {* Syntax translation functions *} |
|
1253 |
||
1254 |
text {* |
|
1255 |
\begin{matharray}{rcl} |
|
1256 |
@{command_def "parse_ast_translation"} & : & \isartrans{theory}{theory} \\ |
|
1257 |
@{command_def "parse_translation"} & : & \isartrans{theory}{theory} \\ |
|
1258 |
@{command_def "print_translation"} & : & \isartrans{theory}{theory} \\ |
|
1259 |
@{command_def "typed_print_translation"} & : & \isartrans{theory}{theory} \\ |
|
1260 |
@{command_def "print_ast_translation"} & : & \isartrans{theory}{theory} \\ |
|
1261 |
\end{matharray} |
|
1262 |
||
1263 |
\begin{rail} |
|
1264 |
( 'parse\_ast\_translation' | 'parse\_translation' | 'print\_translation' | |
|
1265 |
'typed\_print\_translation' | 'print\_ast\_translation' ) ('(advanced)')? text |
|
1266 |
; |
|
1267 |
\end{rail} |
|
1268 |
||
1269 |
Syntax translation functions written in ML admit almost arbitrary |
|
1270 |
manipulations of Isabelle's inner syntax. Any of the above commands |
|
1271 |
have a single \railqtok{text} argument that refers to an ML |
|
1272 |
expression of appropriate type, which are as follows by default: |
|
1273 |
||
1274 |
%FIXME proper antiquotations |
|
1275 |
\begin{ttbox} |
|
1276 |
val parse_ast_translation : (string * (ast list -> ast)) list |
|
1277 |
val parse_translation : (string * (term list -> term)) list |
|
1278 |
val print_translation : (string * (term list -> term)) list |
|
1279 |
val typed_print_translation : |
|
1280 |
(string * (bool -> typ -> term list -> term)) list |
|
1281 |
val print_ast_translation : (string * (ast list -> ast)) list |
|
1282 |
\end{ttbox} |
|
1283 |
||
1284 |
If the @{text "(advanced)"} option is given, the corresponding |
|
1285 |
translation functions may depend on the current theory or proof |
|
1286 |
context. This allows to implement advanced syntax mechanisms, as |
|
1287 |
translations functions may refer to specific theory declarations or |
|
1288 |
auxiliary proof data. |
|
1289 |
||
1290 |
See also \cite[\S8]{isabelle-ref} for more information on the |
|
1291 |
general concept of syntax transformations in Isabelle. |
|
1292 |
||
1293 |
%FIXME proper antiquotations |
|
1294 |
\begin{ttbox} |
|
1295 |
val parse_ast_translation: |
|
27046 | 1296 |
(string * (Proof.context -> ast list -> ast)) list |
27040 | 1297 |
val parse_translation: |
27046 | 1298 |
(string * (Proof.context -> term list -> term)) list |
27040 | 1299 |
val print_translation: |
27046 | 1300 |
(string * (Proof.context -> term list -> term)) list |
27040 | 1301 |
val typed_print_translation: |
27046 | 1302 |
(string * (Proof.context -> bool -> typ -> term list -> term)) list |
27040 | 1303 |
val print_ast_translation: |
27046 | 1304 |
(string * (Proof.context -> ast list -> ast)) list |
27040 | 1305 |
\end{ttbox} |
1306 |
*} |
|
1307 |
||
26869 | 1308 |
end |