11647
|
1 |
(*<*)
|
|
2 |
theory Documents = Main:
|
|
3 |
(*>*)
|
|
4 |
|
12648
|
5 |
section {* Concrete Syntax \label{sec:concrete-syntax} *}
|
12629
|
6 |
|
|
7 |
text {*
|
|
8 |
Concerning Isabelle's ``inner'' language of simply-typed @{text
|
12653
|
9 |
\<lambda>}-calculus, the core concept of Isabelle's elaborate
|
|
10 |
infrastructure for concrete syntax is that of general
|
|
11 |
\bfindex{mixfix annotations}. Associated with any kind of constant
|
|
12 |
declaration, mixfixes affect both the grammar productions for the
|
|
13 |
parser and output templates for the pretty printer.
|
12629
|
14 |
|
|
15 |
In full generality, the whole affair of parser and pretty printer
|
12653
|
16 |
configuration is rather subtle, see \cite{isabelle-ref} for further
|
|
17 |
details. Any syntax specifications given by end-users need to
|
|
18 |
interact properly with the existing setup of Isabelle/Pure and
|
|
19 |
Isabelle/HOL. It is particularly important to get the precedence of
|
|
20 |
new syntactic constructs right, avoiding ambiguities with existing
|
12629
|
21 |
elements.
|
|
22 |
|
|
23 |
\medskip Subsequently we introduce a few simple declaration forms
|
|
24 |
that already cover the most common situations fairly well.
|
|
25 |
*}
|
|
26 |
|
|
27 |
|
12648
|
28 |
subsection {* Infix Annotations *}
|
12629
|
29 |
|
|
30 |
text {*
|
|
31 |
Syntax annotations may be included wherever constants are declared
|
|
32 |
directly or indirectly, including \isacommand{consts},
|
|
33 |
\isacommand{constdefs}, or \isacommand{datatype} (for the
|
|
34 |
constructor operations). Type-constructors may be annotated as
|
|
35 |
well, although this is less frequently encountered in practice
|
|
36 |
(@{text "*"} and @{text "+"} types may come to mind).
|
|
37 |
|
12645
|
38 |
Infix declarations\index{infix annotations} provide a useful special
|
|
39 |
case of mixfixes, where users need not care about the full details
|
|
40 |
of priorities, nesting, spacing, etc. The subsequent example of the
|
|
41 |
exclusive-or operation on boolean values illustrates typical infix
|
12653
|
42 |
declarations arising in practice.
|
12629
|
43 |
*}
|
|
44 |
|
|
45 |
constdefs
|
|
46 |
xor :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infixl "[+]" 60)
|
|
47 |
"A [+] B \<equiv> (A \<and> \<not> B) \<or> (\<not> A \<and> B)"
|
|
48 |
|
|
49 |
text {*
|
12653
|
50 |
\noindent Now @{text "xor A B"} and @{text "A [+] B"} refer to the
|
|
51 |
same expression internally. Any curried function with at least two
|
|
52 |
arguments may be associated with infix syntax. For partial
|
|
53 |
applications with less than two operands there is a special notation
|
|
54 |
with \isa{op} prefix: @{text xor} without arguments is represented
|
|
55 |
as @{text "op [+]"}; together with plain prefix application this
|
|
56 |
turns @{text "xor A"} into @{text "op [+] A"}.
|
12629
|
57 |
|
12653
|
58 |
\medskip The string @{text [source] "[+]"} in the above annotation
|
12629
|
59 |
refers to the bit of concrete syntax to represent the operator,
|
12653
|
60 |
while the number @{text 60} determines the precedence of the
|
|
61 |
construct (i.e.\ the syntactic priorities of the arguments and
|
|
62 |
result).
|
12629
|
63 |
|
|
64 |
As it happens, Isabelle/HOL already spends many popular combinations
|
|
65 |
of ASCII symbols for its own use, including both @{text "+"} and
|
|
66 |
@{text "++"}. Slightly more awkward combinations like the present
|
|
67 |
@{text "[+]"} tend to be available for user extensions. The current
|
|
68 |
arrangement of inner syntax may be inspected via
|
|
69 |
\commdx{print\protect\_syntax}, albeit its output is enormous.
|
|
70 |
|
|
71 |
Operator precedence also needs some special considerations. The
|
|
72 |
admissible range is 0--1000. Very low or high priorities are
|
|
73 |
basically reserved for the meta-logic. Syntax of Isabelle/HOL
|
|
74 |
mainly uses the range of 10--100: the equality infix @{text "="} is
|
|
75 |
centered at 50, logical connectives (like @{text "\<or>"} and @{text
|
|
76 |
"\<and>"}) are below 50, and algebraic ones (like @{text "+"} and @{text
|
|
77 |
"*"}) above 50. User syntax should strive to coexist with common
|
|
78 |
HOL forms, or use the mostly unused range 100--900.
|
|
79 |
|
|
80 |
\medskip The keyword \isakeyword{infixl} specifies an operator that
|
|
81 |
is nested to the \emph{left}: in iterated applications the more
|
|
82 |
complex expression appears on the left-hand side: @{term "A [+] B
|
|
83 |
[+] C"} stands for @{text "(A [+] B) [+] C"}. Similarly,
|
12653
|
84 |
\isakeyword{infixr} specifies to nesting to the \emph{right},
|
|
85 |
reading @{term "A [+] B [+] C"} as @{text "A [+] (B [+] C)"}. In
|
|
86 |
contrast, a \emph{non-oriented} declaration via \isakeyword{infix}
|
|
87 |
would have rendered @{term "A [+] B [+] C"} illegal, but demand
|
|
88 |
explicit parentheses about the intended grouping.
|
12629
|
89 |
*}
|
|
90 |
|
12635
|
91 |
|
12659
|
92 |
subsection {* Mathematical Symbols \label{sec:syntax-symbols} *}
|
12629
|
93 |
|
|
94 |
text {*
|
12635
|
95 |
Concrete syntax based on plain ASCII characters has its inherent
|
|
96 |
limitations. Rich mathematical notation demands a larger repertoire
|
|
97 |
of symbols. Several standards of extended character sets have been
|
|
98 |
proposed over decades, but none has become universally available so
|
12653
|
99 |
far, not even Unicode\index{Unicode}. Isabelle supports a generic
|
|
100 |
notion of \bfindex{symbols} as the smallest entities of source text,
|
|
101 |
without referring to internal encodings.
|
12635
|
102 |
|
12653
|
103 |
There are three kinds of such ``generalized characters'' available:
|
12635
|
104 |
|
|
105 |
\begin{enumerate}
|
|
106 |
|
12653
|
107 |
\item 7-bit ASCII characters
|
12635
|
108 |
|
12653
|
109 |
\item named symbols: \verb,\,\verb,<,$ident$\verb,>,
|
12629
|
110 |
|
12653
|
111 |
\item named control symbols: \verb,\,\verb,<^,$ident$\verb,>,
|
12635
|
112 |
|
|
113 |
\end{enumerate}
|
|
114 |
|
|
115 |
Here $ident$ may be any identifier according to the usual Isabelle
|
|
116 |
conventions. This results in an infinite store of symbols, whose
|
12653
|
117 |
interpretation is left to further front-end tools. For example,
|
|
118 |
both by the user-interface of Proof~General + X-Symbol and the
|
|
119 |
Isabelle document processor (see \S\ref{sec:document-preparation})
|
|
120 |
display the \verb,\,\verb,<forall>, symbol really as ``$\forall$''.
|
12635
|
121 |
|
|
122 |
A list of standard Isabelle symbols is given in
|
|
123 |
\cite[appendix~A]{isabelle-sys}. Users may introduce their own
|
|
124 |
interpretation of further symbols by configuring the appropriate
|
12653
|
125 |
front-end tool accordingly, e.g.\ by defining certain {\LaTeX}
|
|
126 |
macros (see also \S\ref{sec:doc-prep-symbols}). There are also a
|
|
127 |
few predefined control symbols, such as \verb,\,\verb,<^sub>, and
|
12635
|
128 |
\verb,\,\verb,<^sup>, for sub- and superscript of the subsequent
|
12653
|
129 |
(printable) symbol, respectively. For example, \verb,A\<^sup>\<star>, is
|
|
130 |
shown as ``@{text "A\<^sup>\<star>"}''.
|
12635
|
131 |
|
|
132 |
\medskip The following version of our @{text xor} definition uses a
|
|
133 |
standard Isabelle symbol to achieve typographically pleasing output.
|
12629
|
134 |
*}
|
|
135 |
|
|
136 |
(*<*)
|
|
137 |
hide const xor
|
|
138 |
ML_setup {* Context.>> (Theory.add_path "1") *}
|
|
139 |
(*>*)
|
|
140 |
constdefs
|
|
141 |
xor :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infixl "\<oplus>" 60)
|
|
142 |
"A \<oplus> B \<equiv> (A \<and> \<not> B) \<or> (\<not> A \<and> B)"
|
12635
|
143 |
(*<*)
|
|
144 |
local
|
|
145 |
(*>*)
|
12629
|
146 |
|
12635
|
147 |
text {*
|
12653
|
148 |
\noindent The X-Symbol package within Proof~General provides several
|
|
149 |
input methods to enter @{text \<oplus>} in the text. If all fails one may
|
|
150 |
just type \verb,\,\verb,<oplus>, by hand; the display will be
|
|
151 |
adapted immediately after continuing input.
|
12635
|
152 |
|
|
153 |
\medskip A slightly more refined scheme is to provide alternative
|
12645
|
154 |
syntax via the \bfindex{print mode} concept of Isabelle (see also
|
12653
|
155 |
\cite{isabelle-ref}). By convention, the mode of ``$xsymbols$'' is
|
|
156 |
enabled whenever Proof~General's X-Symbol mode (or {\LaTeX} output)
|
|
157 |
is active. Consider the following hybrid declaration of @{text
|
|
158 |
xor}.
|
12635
|
159 |
*}
|
|
160 |
|
|
161 |
(*<*)
|
|
162 |
hide const xor
|
|
163 |
ML_setup {* Context.>> (Theory.add_path "2") *}
|
|
164 |
(*>*)
|
|
165 |
constdefs
|
|
166 |
xor :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infixl "[+]\<ignore>" 60)
|
|
167 |
"A [+]\<ignore> B \<equiv> (A \<and> \<not> B) \<or> (\<not> A \<and> B)"
|
|
168 |
|
|
169 |
syntax (xsymbols)
|
|
170 |
xor :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infixl "\<oplus>\<ignore>" 60)
|
12629
|
171 |
(*<*)
|
|
172 |
local
|
|
173 |
(*>*)
|
|
174 |
|
12635
|
175 |
text {*
|
12653
|
176 |
The \commdx{syntax} command introduced here acts like
|
|
177 |
\isakeyword{consts}, but without declaring a logical constant; an
|
|
178 |
optional print mode specification may be given, too. Note that the
|
|
179 |
type declaration given here merely serves for syntactic purposes,
|
|
180 |
and is not checked for consistency with the real constant.
|
12635
|
181 |
|
12653
|
182 |
\medskip We may now write either @{text "[+]"} or @{text "\<oplus>"} in
|
12635
|
183 |
input, while output uses the nicer syntax of $xsymbols$, provided
|
12653
|
184 |
that print mode is presently active. Such an arrangement is
|
|
185 |
particularly useful for interactive development, where users may
|
|
186 |
type plain ASCII text, but gain improved visual feedback from the
|
|
187 |
system (say in current goal output).
|
12635
|
188 |
|
|
189 |
\begin{warn}
|
12653
|
190 |
Alternative syntax declarations are apt to result in varying
|
|
191 |
occurrences of concrete syntax in the input sources. Isabelle
|
|
192 |
provides no systematic way to convert alternative syntax expressions
|
|
193 |
back and forth; print modes only affect situations where formal
|
|
194 |
entities are pretty printed by the Isabelle process (e.g.\ output of
|
|
195 |
terms and types), but not the original theory text.
|
12635
|
196 |
\end{warn}
|
|
197 |
|
|
198 |
\medskip The following variant makes the alternative @{text \<oplus>}
|
|
199 |
notation only available for output. Thus we may enforce input
|
12653
|
200 |
sources to refer to plain ASCII only, but effectively disable
|
|
201 |
cut-and-paste from output as well.
|
12635
|
202 |
*}
|
|
203 |
|
|
204 |
syntax (xsymbols output)
|
|
205 |
xor :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infixl "\<oplus>\<ignore>" 60)
|
|
206 |
|
12629
|
207 |
|
12648
|
208 |
subsection {* Prefix Annotations *}
|
12629
|
209 |
|
|
210 |
text {*
|
12653
|
211 |
Prefix syntax annotations\index{prefix annotation} are just another
|
|
212 |
degenerate form of general mixfixes \cite{isabelle-ref}, without any
|
|
213 |
template arguments or priorities --- just some bits of literal
|
|
214 |
syntax. The following example illustrates this idea idea by
|
|
215 |
associating common symbols with the constructors of a datatype.
|
12629
|
216 |
*}
|
|
217 |
|
|
218 |
datatype currency =
|
|
219 |
Euro nat ("\<euro>")
|
|
220 |
| Pounds nat ("\<pounds>")
|
|
221 |
| Yen nat ("\<yen>")
|
|
222 |
| Dollar nat ("$")
|
|
223 |
|
|
224 |
text {*
|
12653
|
225 |
\noindent Here the mixfix annotations on the rightmost column happen
|
|
226 |
to consist of a single Isabelle symbol each: \verb,\,\verb,<euro>,,
|
|
227 |
\verb,\,\verb,<pounds>,, \verb,\,\verb,<yen>,, and \verb,$,. Recall
|
|
228 |
that a constructor like @{text Euro} actually is a function @{typ
|
|
229 |
"nat \<Rightarrow> currency"}. An expression like @{text "Euro 10"} will be
|
|
230 |
printed as @{term "\<euro> 10"}; only the head of the application is
|
12651
|
231 |
subject to our concrete syntax. This simple form already achieves
|
12645
|
232 |
conformance with notational standards of the European Commission.
|
12629
|
233 |
|
12635
|
234 |
\medskip Certainly, the same idea of prefix syntax also works for
|
12651
|
235 |
\isakeyword{consts}, \isakeyword{constdefs} etc. The slightly
|
|
236 |
unusual syntax declaration below decorates the existing @{typ
|
|
237 |
currency} type with the international currency symbol @{text \<currency>}
|
|
238 |
(\verb,\,\verb,<currency>,).
|
|
239 |
*}
|
|
240 |
|
|
241 |
syntax
|
|
242 |
currency :: type ("\<currency>")
|
|
243 |
|
|
244 |
text {*
|
|
245 |
\noindent Here @{typ type} refers to the builtin syntactic category
|
12653
|
246 |
of Isabelle types. We may now write down funny things like @{text
|
12651
|
247 |
"\<euro> :: nat \<Rightarrow> \<currency>"}, for example.
|
|
248 |
*}
|
|
249 |
|
|
250 |
|
|
251 |
subsection {* Syntax Translations \label{sec:syntax-translations} *}
|
|
252 |
|
|
253 |
text{*
|
12653
|
254 |
Mixfix syntax annotations work well for those situations where a
|
|
255 |
particular constant application forms need to be decorated by
|
|
256 |
concrete syntax; just reconsider @{text "xor A B"} versus @{text "A
|
|
257 |
\<oplus> B"} covered before. Occasionally, the relationship between some
|
|
258 |
piece of notation and its internal form is slightly more involved.
|
|
259 |
Here the concept of \bfindex{syntax translations} enters the scene.
|
12651
|
260 |
|
|
261 |
Using the raw \isakeyword{syntax}\index{syntax (command)} command we
|
12653
|
262 |
may introduce uninterpreted notational elements, while
|
|
263 |
\commdx{translations} relates the input forms with more complex
|
|
264 |
logical expressions. This essentially provides a simple mechanism
|
|
265 |
for for syntactic macros; even heavier transformations may be
|
|
266 |
programmed in ML \cite{isabelle-ref}.
|
12651
|
267 |
|
12653
|
268 |
\medskip A typical example of syntax translations is to decorate
|
|
269 |
relational expressions (set membership of tuples) with nice symbolic
|
|
270 |
notation, such as @{text "(x, y) \<in> sim"} versus @{text "x \<approx> y"}.
|
12635
|
271 |
*}
|
|
272 |
|
|
273 |
consts
|
12651
|
274 |
sim :: "('a \<times> 'a) set"
|
12629
|
275 |
|
12651
|
276 |
syntax
|
|
277 |
"_sim" :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<approx>" 50)
|
|
278 |
translations
|
|
279 |
"x \<approx> y" \<rightleftharpoons> "(x, y) \<in> sim"
|
12629
|
280 |
|
12651
|
281 |
text {*
|
|
282 |
\noindent Here the name of the dummy constant @{text "_sim"} does
|
12653
|
283 |
not really matter, as long as it is not used elsewhere. Prefixing
|
12651
|
284 |
an underscore is a common convention. The \isakeyword{translations}
|
|
285 |
declaration already uses concrete syntax on the left-hand side;
|
|
286 |
internally we relate a raw application @{text "_sim x y"} with
|
|
287 |
@{text "(x, y) \<in> sim"}.
|
|
288 |
|
12653
|
289 |
\medskip Another common application of syntax translations is to
|
12651
|
290 |
provide variant versions of fundamental relational expressions, such
|
|
291 |
as @{text \<noteq>} for negated equalities. The following declaration
|
|
292 |
stems from Isabelle/HOL itself:
|
12629
|
293 |
*}
|
|
294 |
|
12651
|
295 |
syntax "_not_equal" :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl "\<noteq>\<ignore>" 50)
|
|
296 |
translations "x \<noteq>\<ignore> y" \<rightleftharpoons> "\<not> (x = y)"
|
12629
|
297 |
|
12651
|
298 |
text {*
|
|
299 |
\noindent Normally one would introduce derived concepts like this
|
12653
|
300 |
within the logic, using \isakeyword{consts} + \isakeyword{defs}
|
|
301 |
instead of \isakeyword{syntax} + \isakeyword{translations}. The
|
12651
|
302 |
present formulation has the virtue that expressions are immediately
|
|
303 |
replaced by its ``definition'' upon parsing; the effect is reversed
|
|
304 |
upon printing. Internally, @{text"\<noteq>"} never appears.
|
|
305 |
|
|
306 |
Simulating definitions via translations is adequate for very basic
|
12653
|
307 |
logical concepts, where a new representation is a trivial variation
|
|
308 |
on an existing one. On the other hand, syntax translations do not
|
|
309 |
scale up well to large hierarchies of concepts built on each other.
|
12629
|
310 |
*}
|
|
311 |
|
|
312 |
|
12653
|
313 |
section {* Document Preparation \label{sec:document-preparation} *}
|
12629
|
314 |
|
12645
|
315 |
text {*
|
12653
|
316 |
Isabelle/Isar is centered around the concept of \bfindex{formal
|
|
317 |
proof documents}\index{documents|bold}. Ultimately the result of
|
|
318 |
the user's theory development efforts is meant to be a
|
|
319 |
human-readable record, presented as a browsable PDF file or printed
|
|
320 |
on paper. The overall document structure follows traditional
|
|
321 |
mathematical articles, with sectioning, intermediate explanations,
|
|
322 |
definitions, theorem statements and proofs.
|
12629
|
323 |
|
12645
|
324 |
The Isar proof language \cite{Wenzel-PhD}, which is not covered in
|
|
325 |
this book, admits to write formal proof texts that are acceptable
|
|
326 |
both to the machine and human readers at the same time. Thus
|
12653
|
327 |
marginal comments and explanations may be kept at a minimum. Even
|
|
328 |
without proper coverage of human-readable proofs, Isabelle document
|
|
329 |
is very useful to produce formally derived texts (elaborating on the
|
|
330 |
specifications etc.). Unstructured proof scripts given here may be
|
|
331 |
just ignored by readers, or intentionally suppressed from the text
|
|
332 |
by the writer (see also \S\ref{sec:doc-prep-suppress}).
|
12629
|
333 |
|
12645
|
334 |
\medskip The Isabelle document preparation system essentially acts
|
12653
|
335 |
like a formal front-end to {\LaTeX}. After checking specifications
|
|
336 |
and proofs, the theory sources are turned into typesetting
|
|
337 |
instructions in a well-defined manner. This enables users to write
|
|
338 |
authentic reports on formal theory developments with little
|
|
339 |
additional effort, the most tedious consistency checks are handled
|
|
340 |
by the system.
|
12645
|
341 |
*}
|
|
342 |
|
|
343 |
|
12648
|
344 |
subsection {* Isabelle Sessions *}
|
12629
|
345 |
|
|
346 |
text {*
|
12653
|
347 |
In contrast to the highly interactive mode of Isabelle/Isar theory
|
|
348 |
development, the document preparation stage essentially works in
|
|
349 |
batch-mode. An Isabelle \bfindex{session} essentially consists of a
|
|
350 |
collection of theory source files that contribute to a single output
|
|
351 |
document eventually. Session is derived from a single parent each
|
|
352 |
(usually an object-logic image like \texttt{HOL}), resulting in an
|
|
353 |
overall tree structure that is reflected in the output location
|
|
354 |
within the file system (usually rooted at
|
|
355 |
\verb,~/isabelle/browser_info,).
|
12645
|
356 |
|
12653
|
357 |
Here is the canonical arrangement of sources of a session called
|
|
358 |
\texttt{MySession}:
|
12645
|
359 |
|
|
360 |
\begin{itemize}
|
|
361 |
|
|
362 |
\item Directory \texttt{MySession} contains the required theory
|
12653
|
363 |
files ($A@1$\texttt{.thy}, \dots, $A@n$\texttt{.thy}).
|
12645
|
364 |
|
|
365 |
\item File \texttt{MySession/ROOT.ML} holds appropriate ML commands
|
|
366 |
for loading all wanted theories, usually just
|
12653
|
367 |
``\texttt{use_thy~"$A@i$";}'' for any $A@i$ in leaf position of the
|
12645
|
368 |
theory dependency graph.
|
|
369 |
|
|
370 |
\item Directory \texttt{MySession/document} contains everything
|
12653
|
371 |
required for the {\LaTeX} stage; only \texttt{root.tex} needs to be
|
|
372 |
provided initially.
|
12645
|
373 |
|
12653
|
374 |
The latter file holds appropriate {\LaTeX} code to commence a
|
|
375 |
document (\verb,\documentclass, etc.), and to include the generated
|
|
376 |
files $A@i$\texttt{.tex} for each theory. The generated
|
|
377 |
\texttt{session.tex} will hold {\LaTeX} commands to include all
|
|
378 |
theory output files in topologically sorted order, so
|
|
379 |
\verb,\input{session}, in \texttt{root.tex} will do it in most
|
|
380 |
situations.
|
|
381 |
|
|
382 |
\item In addition, \texttt{IsaMakefile} outside of the directory
|
12645
|
383 |
\texttt{MySession} holds appropriate dependencies and invocations of
|
12653
|
384 |
Isabelle tools to control the batch job. In fact, several sessions
|
|
385 |
may be controlled by the same \texttt{IsaMakefile}. See also
|
|
386 |
\cite{isabelle-sys} for further details, especially on
|
|
387 |
\texttt{isatool usedir} and \texttt{isatool make}.
|
12645
|
388 |
|
|
389 |
\end{itemize}
|
|
390 |
|
|
391 |
With everything put in its proper place, \texttt{isatool make}
|
|
392 |
should be sufficient to process the Isabelle session completely,
|
12653
|
393 |
with the generated document appearing in its proper place.
|
12645
|
394 |
|
12653
|
395 |
\medskip In practice, users may want to have \texttt{isatool mkdir}
|
|
396 |
generate an initial working setup without further ado. For example,
|
|
397 |
an empty session \texttt{MySession} derived from \texttt{HOL} may be
|
|
398 |
produced as follows:
|
12645
|
399 |
|
|
400 |
\begin{verbatim}
|
|
401 |
isatool mkdir HOL MySession
|
|
402 |
isatool make
|
|
403 |
\end{verbatim}
|
|
404 |
|
12653
|
405 |
This processes the session with sensible default options, including
|
|
406 |
verbose mode to tell the user where the ultimate results will
|
|
407 |
appear. The above dry run should produce should already be able to
|
|
408 |
produce a single page of output (with a dummy title, empty table of
|
|
409 |
contents etc.). Any failure at that stage is likely to indicate
|
|
410 |
technical problems with the user's {\LaTeX}
|
|
411 |
installation.\footnote{Especially make sure that \texttt{pdflatex}
|
|
412 |
is present; if all fails one may fall back on DVI output by changing
|
|
413 |
\texttt{usedir} options \cite{isabelle-sys}.}
|
12645
|
414 |
|
12653
|
415 |
\medskip One may now start to populate the directory
|
12645
|
416 |
\texttt{MySession}, and the file \texttt{MySession/ROOT.ML}
|
|
417 |
accordingly. \texttt{MySession/document/root.tex} should be also
|
12653
|
418 |
adapted at some point; the default version is mostly
|
12659
|
419 |
self-explanatory. Note that the \verb,\isabellestyle, enables
|
|
420 |
fine-tuning of the general appearance of characters and mathematical
|
|
421 |
symbols (see also \S\ref{sec:doc-prep-symbols}).
|
12653
|
422 |
|
|
423 |
Especially note the standard inclusion of {\LaTeX} packages
|
|
424 |
\texttt{isabelle} (mandatory), and \texttt{isabellesym} (required
|
|
425 |
for mathematical symbols), and the final \texttt{pdfsetup} (provides
|
|
426 |
handsome defaults for \texttt{hyperref}, including URL markup).
|
|
427 |
Further {\LaTeX} packages further packages may required in
|
|
428 |
particular applications, e.g.\ for unusual Isabelle symbols.
|
12645
|
429 |
|
12653
|
430 |
\medskip Further auxiliary files for the {\LaTeX} stage should be
|
|
431 |
included in the \texttt{MySession/document} directory, e.g.\
|
|
432 |
additional {\TeX} sources or graphics. In particular, adding
|
|
433 |
\texttt{root.bib} here (with that specific name) causes an automatic
|
|
434 |
run of \texttt{bibtex} to process a bibliographic database; see for
|
|
435 |
further commodities \texttt{isatool document} covered in
|
|
436 |
\cite{isabelle-sys}.
|
12645
|
437 |
|
12653
|
438 |
\medskip Any failure of the document preparation phase in an
|
|
439 |
Isabelle batch session leaves the generated sources in there target
|
|
440 |
location (as pointed out by the accompanied error message). In case
|
|
441 |
of {\LaTeX} errors, users may trace error messages at the file
|
|
442 |
position of the generated text.
|
12645
|
443 |
*}
|
|
444 |
|
|
445 |
|
12648
|
446 |
subsection {* Structure Markup *}
|
12645
|
447 |
|
12653
|
448 |
text {*
|
|
449 |
The large-scale structure of Isabelle documents follows existing
|
|
450 |
{\LaTeX} conventions, with chapters, sections, subsubsections etc.
|
|
451 |
The Isar language includes separate \bfindex{markup commands}, which
|
|
452 |
do not effect the formal content of a theory (or proof), but result
|
|
453 |
in corresponding {\LaTeX} elements issued to the output.
|
12645
|
454 |
|
12653
|
455 |
There are separate markup commands for different formal contexts: in
|
|
456 |
header position (just before a \isakeyword{theory} command), within
|
|
457 |
the theory body, or within a proof. Note that the header needs to
|
|
458 |
be treated specially, since ordinary theory and proof commands may
|
|
459 |
only occur \emph{after} the initial \isakeyword{theory}
|
12645
|
460 |
specification.
|
|
461 |
|
|
462 |
\smallskip
|
|
463 |
|
|
464 |
\begin{tabular}{llll}
|
|
465 |
header & theory & proof & default meaning \\\hline
|
|
466 |
& \commdx{chapter} & & \verb,\chapter, \\
|
|
467 |
\commdx{header} & \commdx{section} & \commdx{sect} & \verb,\section, \\
|
|
468 |
& \commdx{subsection} & \commdx{subsect} & \verb,\subsection, \\
|
|
469 |
& \commdx{subsubsection} & \commdx{subsubsect} & \verb,\subsubsection, \\
|
|
470 |
\end{tabular}
|
|
471 |
|
|
472 |
\medskip
|
|
473 |
|
|
474 |
From the Isabelle perspective, each markup command takes a single
|
12659
|
475 |
$text$ argument (delimited by \verb,",\dots\verb,", or
|
12653
|
476 |
\verb,{,\verb,*,~\dots~\verb,*,\verb,},). After stripping any
|
12645
|
477 |
surrounding white space, the argument is passed to a {\LaTeX} macro
|
12653
|
478 |
\verb,\isamarkupXYZ, for any command \isakeyword{XYZ}. These macros
|
|
479 |
are defined in \verb,isabelle.sty, according to the meaning given in
|
|
480 |
the rightmost column above.
|
12645
|
481 |
|
|
482 |
\medskip The following source fragment illustrates structure markup
|
12653
|
483 |
of a theory. Note that {\LaTeX} labels may be included inside of
|
|
484 |
section headings as well.
|
12645
|
485 |
|
|
486 |
\begin{ttbox}
|
|
487 |
header {\ttlbrace}* Some properties of Foo Bar elements *{\ttrbrace}
|
|
488 |
|
|
489 |
theory Foo_Bar = Main:
|
|
490 |
|
|
491 |
subsection {\ttlbrace}* Basic definitions *{\ttrbrace}
|
|
492 |
|
|
493 |
consts
|
|
494 |
foo :: \dots
|
|
495 |
bar :: \dots
|
12648
|
496 |
|
12645
|
497 |
defs \dots
|
12648
|
498 |
|
12645
|
499 |
subsection {\ttlbrace}* Derived rules *{\ttrbrace}
|
|
500 |
|
|
501 |
lemma fooI: \dots
|
|
502 |
lemma fooE: \dots
|
|
503 |
|
12648
|
504 |
subsection {\ttlbrace}* Main theorem {\ttback}label{\ttlbrace}sec:main-theorem{\ttrbrace} *{\ttrbrace}
|
12645
|
505 |
|
|
506 |
theorem main: \dots
|
|
507 |
|
|
508 |
end
|
|
509 |
\end{ttbox}
|
|
510 |
|
12653
|
511 |
Users may occasionally want to change the meaning of markup
|
|
512 |
commands, say via \verb,\renewcommand, in \texttt{root.tex}. The
|
|
513 |
\verb,\isamarkupheader, is a good candidate for some adaption, e.g.\
|
|
514 |
moving it up in the hierarchy to become \verb,\chapter,.
|
12645
|
515 |
|
|
516 |
\begin{verbatim}
|
|
517 |
\renewcommand{\isamarkupheader}[1]{\chapter{#1}}
|
|
518 |
\end{verbatim}
|
|
519 |
|
12653
|
520 |
\noindent Certainly, this requires to change the default
|
12645
|
521 |
\verb,\documentclass{article}, in \texttt{root.tex} to something
|
|
522 |
that supports the notion of chapters in the first place, e.g.\
|
12648
|
523 |
\verb,\documentclass{report},.
|
12645
|
524 |
|
12648
|
525 |
\medskip The {\LaTeX} macro \verb,\isabellecontext, is maintained to
|
|
526 |
hold the name of the current theory context. This is particularly
|
12653
|
527 |
useful for document headings:
|
12645
|
528 |
|
|
529 |
\begin{verbatim}
|
12653
|
530 |
\renewcommand{\isamarkupheader}[1]
|
12645
|
531 |
{\chapter{#1}\markright{THEORY~\isabellecontext}}
|
|
532 |
\end{verbatim}
|
|
533 |
|
|
534 |
\noindent Make sure to include something like
|
12648
|
535 |
\verb,\pagestyle{headings}, in \texttt{root.tex}; the document
|
|
536 |
should have more than 2 pages to show the effect.
|
12645
|
537 |
*}
|
|
538 |
|
|
539 |
|
12648
|
540 |
subsection {* Formal Comments and Antiquotations *}
|
12645
|
541 |
|
|
542 |
text {*
|
12659
|
543 |
FIXME check
|
|
544 |
|
|
545 |
Source comments of the form \verb,(,\verb,*,~\dots~\verb,*,\verb,),
|
|
546 |
essentially act like white space and do not contribute to the formal
|
|
547 |
text at all. They mainly serve technical purposes to mark certain
|
|
548 |
oddities or problems with the raw sources.
|
|
549 |
|
|
550 |
In contrast, \bfindex{formal comments} are portions of text that are
|
|
551 |
associated with formal Isabelle/Isar commands (\bfindex{marginal
|
|
552 |
comments}), or even as stanalone paragraphs positioned explicitly
|
|
553 |
within a theory or proof context (\bfindex{text blocks}).
|
|
554 |
|
|
555 |
\medskip Marginal comments are part of each command's concrete
|
|
556 |
syntax \cite{isabelle-ref}; the common form \verb,--,~text, where
|
|
557 |
$text$, delimited by \verb,",\dots\verb,", or
|
|
558 |
\verb,{,\verb,*,~\dots~\verb,*,\verb,}, as usual. Multiple marginal
|
|
559 |
comments may be given at the same time. Here is a simple example:
|
|
560 |
|
|
561 |
\begin{verbatim}
|
|
562 |
lemma "A --> A"
|
|
563 |
-- "a triviality of propositional logic"
|
|
564 |
-- "(should not really bother)"
|
|
565 |
by (rule impI) -- "implicit assumption step involved here"
|
|
566 |
\end{verbatim}
|
|
567 |
|
|
568 |
From the {\LaTeX} view, \verb,--, acts like a markup command, the
|
|
569 |
corresponding macro is \verb,\isamarkupcmt, (of a single argument).
|
|
570 |
|
|
571 |
\medskip The commands \bfindex{text} and \bfindex{txt} both
|
|
572 |
introduce a text block (for theory and proof contexts,
|
|
573 |
respectively), taking a single $text$ argument. The {\LaTeX} output
|
|
574 |
appears as a free-form text, surrounded with separate paragraphs and
|
|
575 |
additional vertical spacing. This behavior is determined by the
|
|
576 |
{\LaTeX} environment definitions \verb,isamarkuptext, and
|
|
577 |
\verb,isamarkuptxt,, respectively. This may be changed via
|
|
578 |
\verb,\renewenvironment,, but it is often sufficient to redefine
|
|
579 |
\verb,\isastyletext, or \verb,\isastyletxt, only, which determine
|
|
580 |
the body text style.
|
|
581 |
|
|
582 |
\medskip The $text$ part of each of the various markup commands
|
|
583 |
considered so far essentially inserts \emph{quoted} material within
|
|
584 |
a formal text, essentially for instruction of the reader only
|
|
585 |
(arbitrary {\LaTeX} macros may be included).
|
|
586 |
|
|
587 |
An \bfindex{antiquotation} is again a formal object that has been
|
|
588 |
embedded into such an informal portion of text. Typically, the
|
|
589 |
interpretation of an antiquotation is limited to well-formedness
|
|
590 |
checks, with the result being pretty printed into the resulting
|
|
591 |
document output. So quoted text blocks together with antiquotations
|
|
592 |
provide very handsome means to reference formal entities within
|
|
593 |
informal expositions, with a high level of confidence in the
|
|
594 |
technical details (especially syntax and types).
|
|
595 |
|
|
596 |
The general format of antiquotations is as follows:
|
|
597 |
\texttt{{\at}{\ttlbrace}$name$ $arguments${\ttrbrace}}, or
|
|
598 |
\texttt{{\at}{\ttlbrace}$name$ [$options$] $arguments${\ttrbrace}}
|
|
599 |
for a comma-separated list of $name$ or assignments
|
|
600 |
\texttt{$name$=$value$} of $options$ (see \cite{isabelle-isar-ref}
|
|
601 |
for details). The syntax of $arguments$ depends on the
|
|
602 |
antiquotation name, it generally follows the same conventions for
|
|
603 |
types, terms, or theorems as in the formal part of a theory.
|
12645
|
604 |
|
12659
|
605 |
\medskip Here is an example use of the quotation-antiquotation
|
|
606 |
technique: @{term "%x y. x"} is a well-typed term.
|
|
607 |
|
|
608 |
\medskip This output has been produced as follows:
|
|
609 |
\begin{ttbox}
|
|
610 |
text {\ttlbrace}*
|
|
611 |
Here is an example use of the quotation-antiquotation technique:
|
|
612 |
{\at}{\ttlbrace}term "%x y. x"{\ttrbrace} is a well-typed term.
|
|
613 |
*{\ttrbrace}
|
|
614 |
\end{ttbox}
|
|
615 |
|
|
616 |
From the notational change of the ASCII character \verb,%, to the
|
|
617 |
symbol @{text \<lambda>} we see that the term really got printed by the
|
|
618 |
system --- recall that the document preparation system enables
|
|
619 |
symbolic output by default.
|
|
620 |
|
|
621 |
\medskip In the following example we include an option to enable the
|
|
622 |
\verb,show_types, flag of Isabelle: the antiquotation
|
|
623 |
\texttt{{\at}}\verb,{term [show_types] "%x y. x"}, produces @{term
|
|
624 |
[show_types] "%x y. x"}. Here type-inference has figured out the
|
|
625 |
most general typings in the present (theory) context. Within a
|
|
626 |
proof, one may get different results according to typings that have
|
|
627 |
already been figured out in the text so far, say some fixed
|
|
628 |
variables in the main statement given before hand.
|
|
629 |
|
|
630 |
\medskip Several further kinds of antiquotations (and options) are
|
|
631 |
available \cite{isabelle-sys}. The most commonly used combinations
|
|
632 |
are as follows:
|
|
633 |
|
|
634 |
\medskip
|
12651
|
635 |
|
12659
|
636 |
\begin{tabular}{ll}
|
|
637 |
\texttt{\at}\verb,{typ,~$\tau$\verb,}, & print type $\tau$ \\
|
|
638 |
\texttt{\at}\verb,{term,~$t$\verb,}, & print term $t$ \\
|
|
639 |
\texttt{\at}\verb,{prop,~$\phi$\verb,}, & print proposition $\phi$ \\
|
|
640 |
\texttt{\at}\verb,{prop [source],~$\phi$\verb,}, & check proposition $\phi$, print its input \\
|
|
641 |
\texttt{\at}\verb,{thm,~$a$\verb,}, & print fact $a$ \\
|
|
642 |
\texttt{\at}\verb,{thm,~$a$~\verb,[no_vars]}, & print fact $a$, fixing schematic variables \\
|
|
643 |
\texttt{\at}\verb,{thm [source],~$a$\verb,}, & check availability of fact $a$, print its name \\
|
|
644 |
\texttt{\at}\verb,{text,~$s$\verb,}, & print uninterpreted text $s$ \\
|
|
645 |
\end{tabular}
|
|
646 |
|
|
647 |
\medskip
|
|
648 |
|
|
649 |
Note that \verb,no_vars, given above is \emph{not} an option, but
|
|
650 |
merely an attribute of the theorem argument given here.
|
|
651 |
|
|
652 |
The \texttt{\at}\verb,{text, $s$\verb,}, antiquotation is
|
|
653 |
particularly interesting. Embedding uninterpreted text within an
|
|
654 |
informal text body might appear useless at first sight. Here the
|
|
655 |
key virtue is that the string $s$ is processed as proper Isabelle
|
|
656 |
output, interpreting Isabelle symbols (\S\ref{sec:syntax-symbols})
|
|
657 |
appropriately.
|
|
658 |
|
|
659 |
For example, \texttt{\at}\verb,{text "\<forall>\<exists>"}, produces @{text "\<forall>\<exists>"},
|
|
660 |
according to the standard interpretation of these symbol (cf.\
|
|
661 |
\S\ref{sec:doc-prep-symbols}). Thus we achieve consistent
|
|
662 |
mathematical notation in both the formal and informal parts of the
|
|
663 |
document very easily. Manual {\LaTeX} code leaves more control over
|
|
664 |
the type-setting, but is slightly more tedious. Also note that
|
|
665 |
Isabelle symbols may be also displayed within the editor buffer of
|
|
666 |
Proof~General.
|
12645
|
667 |
*}
|
|
668 |
|
|
669 |
|
12659
|
670 |
subsection {* Interpretation of symbols \label{sec:doc-prep-symbols} *}
|
12645
|
671 |
|
|
672 |
text {*
|
12659
|
673 |
FIXME tune
|
|
674 |
|
|
675 |
According to \S\ref{sec:syntax-symbols}, the smalles syntactic
|
|
676 |
entities of Isabelle text are symbols, a straight-forward
|
|
677 |
generalization of ASCII characters. Concerning document
|
|
678 |
preperation, symbols are translated uniformly to {\LaTeX} code as
|
|
679 |
follows.
|
|
680 |
|
|
681 |
\begin{enumerate} \item 7-bit ASCII characters: letters
|
|
682 |
\texttt{A\dots Z} and \texttt{a\dots z} are output verbatim, digits
|
|
683 |
are passed as an argument to the \verb,\isadigit, macro, other
|
|
684 |
characters are replaced by specific macros \verb,\isacharXYZ, (see
|
|
685 |
also \texttt{isabelle.sty}).
|
|
686 |
|
|
687 |
\item Named symbols: \verb,\,\verb,<,$XYZ$\verb,>, become
|
|
688 |
\verb,\{\isasym,$XYZ$\verb,}, each (note the additional braces).
|
|
689 |
See \cite[appendix~A]{isabelle-sys} and \texttt{isabellesym.sty} for
|
|
690 |
the collection of predefined standard symbols.
|
|
691 |
|
|
692 |
\item Named control symbols: \verb,\,\verb,<^,$XYZ$\verb,>, become
|
|
693 |
\verb,\isactrl,$XYZ$; subsequent symbols may act as arguments, if
|
|
694 |
the corresponding macro is defined accordingly.
|
|
695 |
\end{enumerate}
|
12645
|
696 |
*}
|
|
697 |
|
|
698 |
|
12653
|
699 |
subsection {* Suppressing Output \label{sec:doc-prep-suppress} *}
|
12645
|
700 |
|
|
701 |
text {*
|
12648
|
702 |
By default Isabelle's document system generates a {\LaTeX} source
|
|
703 |
file for each theory that happens to get loaded during the session.
|
12653
|
704 |
The generated \texttt{session.tex} will include all of these in
|
12648
|
705 |
order of appearance, which in turn gets included by the standard
|
12653
|
706 |
\texttt{root.tex}. Certainly one may change the order of appearance
|
|
707 |
or suppress unwanted theories by ignoring \texttt{session.tex} and
|
|
708 |
include individual files in \texttt{root.tex} by hand. On the other
|
|
709 |
hand, such an arrangement requires additional maintenance chores
|
|
710 |
whenever the collection of theories changes.
|
12648
|
711 |
|
|
712 |
Alternatively, one may tune the theory loading process in
|
12653
|
713 |
\texttt{ROOT.ML} itself: traversal of the theory dependency graph
|
|
714 |
may be fine-tuned by adding further \verb,use_thy, invocations,
|
|
715 |
although topological sorting still has to be observed. Moreover,
|
|
716 |
the ML operator \verb,no_document, temporarily disables document
|
|
717 |
generation while executing a theory loader command; its usage is
|
|
718 |
like this:
|
12648
|
719 |
|
|
720 |
\begin{verbatim}
|
12653
|
721 |
no_document use_thy "A";
|
12648
|
722 |
\end{verbatim}
|
12645
|
723 |
|
12653
|
724 |
\medskip Theory output may be also suppressed in smaller portions as
|
|
725 |
well. For example, research papers or slides usually do not include
|
|
726 |
the formal content in full. In order to delimit \bfindex{ignored
|
|
727 |
material} special source comments
|
12648
|
728 |
\verb,(,\verb,*,\verb,<,\verb,*,\verb,), and
|
12653
|
729 |
\verb,(,\verb,*,\verb,>,\verb,*,\verb,), may be included in the
|
|
730 |
text. Only the document preparation system is affected, the formal
|
|
731 |
checking the theory is performed as before.
|
12648
|
732 |
|
|
733 |
In the following example we suppress the slightly formalistic
|
12653
|
734 |
\isakeyword{theory} + \isakeyword{end} surroundings a theory.
|
12648
|
735 |
|
|
736 |
\medskip
|
|
737 |
|
|
738 |
\begin{tabular}{l}
|
|
739 |
\verb,(,\verb,*,\verb,<,\verb,*,\verb,), \\
|
|
740 |
\texttt{theory A = Main:} \\
|
|
741 |
\verb,(,\verb,*,\verb,>,\verb,*,\verb,), \\
|
|
742 |
~~$\vdots$ \\
|
|
743 |
\verb,(,\verb,*,\verb,<,\verb,*,\verb,), \\
|
|
744 |
\texttt{end} \\
|
|
745 |
\verb,(,\verb,*,\verb,>,\verb,*,\verb,), \\
|
|
746 |
\end{tabular}
|
|
747 |
|
|
748 |
\medskip
|
|
749 |
|
12653
|
750 |
Text may be suppressed in a fine grained manner. For example, we
|
|
751 |
may even drop vital parts of a formal proof, pretending that things
|
|
752 |
have been simpler than in reality. For example, the following
|
|
753 |
``fully automatic'' proof is actually a fake:
|
12651
|
754 |
*}
|
|
755 |
|
|
756 |
lemma "x \<noteq> (0::int) \<Longrightarrow> 0 < x * x"
|
|
757 |
by (auto(*<*)simp add: int_less_le(*>*))
|
|
758 |
|
|
759 |
text {*
|
|
760 |
\noindent Here the real source of the proof has been as follows:
|
|
761 |
|
|
762 |
\begin{verbatim}
|
|
763 |
by (auto(*<*)simp add: int_less_le(*>*))
|
12659
|
764 |
\end{verbatim}
|
|
765 |
%(*
|
12651
|
766 |
|
12653
|
767 |
\medskip Ignoring portions of printed does demand some care by the
|
|
768 |
user. First of all, the writer is responsible not to obfuscate the
|
|
769 |
underlying formal development in an unduly manner. It is fairly
|
|
770 |
easy to invalidate the remaining visible text, e.g.\ by referencing
|
|
771 |
questionable formal items (strange definitions, arbitrary axioms
|
|
772 |
etc.) that have been hidden from sight beforehand.
|
12651
|
773 |
|
|
774 |
Some minor technical subtleties of the
|
12653
|
775 |
\verb,(,\verb,*,\verb,<,\verb,*,\verb,),-\verb,(,\verb,*,\verb,>,\verb,*,\verb,),
|
|
776 |
elements need to be kept in mind as well, since the system performs
|
|
777 |
little sanity checks here. Arguments of markup commands and formal
|
12651
|
778 |
comments must not be hidden, otherwise presentation fails. Open and
|
|
779 |
close parentheses need to be inserted carefully; it is fairly easy
|
12653
|
780 |
to hide the wrong parts, especially after rearranging the sources.
|
12651
|
781 |
|
|
782 |
\medskip Authentic reports of formal theories, say as part of a
|
12653
|
783 |
library, usually should refrain from suppressing parts of the text
|
12651
|
784 |
at all. Other users may need the full information for their own
|
12653
|
785 |
derivative work. If a particular formalization appears inadequate
|
|
786 |
for general public coverage, it is often more appropriate to think
|
|
787 |
of a better way in the first place.
|
12629
|
788 |
*}
|
|
789 |
|
11647
|
790 |
(*<*)
|
|
791 |
end
|
|
792 |
(*>*)
|