| author | wenzelm | 
| Wed, 10 Jan 2001 20:18:55 +0100 | |
| changeset 10858 | 479dad7b3b41 | 
| parent 10355 | aef4f587a0e4 | 
| child 11100 | 34d58b1818f4 | 
| permissions | -rw-r--r-- | 
| 7046 | 1  | 
|
| 7895 | 2  | 
\chapter{Isar Syntax Primitives}
 | 
| 7046 | 3  | 
|
| 7315 | 4  | 
We give a complete reference of all basic syntactic entities underlying the  | 
| 7335 | 5  | 
Isabelle/Isar document syntax. Actual theory and proof commands will be  | 
6  | 
introduced later on.  | 
|
| 7315 | 7  | 
|
8  | 
\medskip  | 
|
9  | 
||
10  | 
In order to get started with writing well-formed Isabelle/Isar documents, the  | 
|
11  | 
most important aspect to be noted is the difference of \emph{inner} versus
 | 
|
12  | 
\emph{outer} syntax.  Inner syntax is that of Isabelle types and terms of the
 | 
|
| 7895 | 13  | 
logic, while outer syntax is that of Isabelle/Isar theories (including  | 
14  | 
proofs). As a general rule, inner syntax entities may occur only as  | 
|
15  | 
\emph{atomic entities} within outer syntax.  For example, the string
 | 
|
16  | 
\texttt{"x + y"} and identifier \texttt{z} are legal term specifications
 | 
|
17  | 
within a theory, while \texttt{x + y} is not.
 | 
|
| 7315 | 18  | 
|
19  | 
\begin{warn}
 | 
|
| 8378 | 20  | 
Note that classic Isabelle theories used to fake parts of the inner syntax  | 
21  | 
of types, with rather complicated rules when quotes may be omitted. Despite  | 
|
| 7981 | 22  | 
the minor drawback of requiring quotes more often, the syntax of  | 
| 8548 | 23  | 
Isabelle/Isar is much simpler and more robust in that respect.  | 
| 7315 | 24  | 
\end{warn}
 | 
25  | 
||
| 7466 | 26  | 
\medskip  | 
27  | 
||
| 9601 | 28  | 
Isabelle/Isar input may contain any number of input termination characters  | 
29  | 
``\texttt{;}'' (semicolon) to separate commands explicitly.  This may be
 | 
|
30  | 
particularly useful in interactive shell sessions to make clear where the  | 
|
31  | 
current command is intended to end. Otherwise, the interactive loop will  | 
|
| 10858 | 32  | 
continue until end-of-command is clearly indicated from the input syntax,  | 
| 9601 | 33  | 
e.g.\ encounter of the next command keyword.  | 
34  | 
||
35  | 
Advanced interfaces such as Proof~General \cite{proofgeneral} do not require
 | 
|
36  | 
explicit semicolons, the amount of input text is determined automatically by  | 
|
37  | 
inspecting the Emacs text buffer. Also note that in the presentation of  | 
|
38  | 
Isabelle/Isar documents, semicolons are omitted in any case to gain  | 
|
| 7981 | 39  | 
readability.  | 
| 7466 | 40  | 
|
| 7315 | 41  | 
|
42  | 
\section{Lexical matters}\label{sec:lex-syntax}
 | 
|
43  | 
||
44  | 
The Isabelle/Isar outer syntax provides token classes as presented below.  | 
|
| 7895 | 45  | 
Note that some of these coincide (by full intention) with the inner lexical  | 
46  | 
syntax as presented in \cite{isabelle-ref}.  These different levels of syntax
 | 
|
47  | 
should not be confused, though.  | 
|
| 7134 | 48  | 
|
| 7335 | 49  | 
%FIXME keyword, command  | 
| 9617 | 50  | 
\indexoutertoken{ident}\indexoutertoken{longident}\indexoutertoken{symident}
 | 
51  | 
\indexoutertoken{nat}\indexoutertoken{var}\indexoutertoken{typefree}
 | 
|
52  | 
\indexoutertoken{typevar}\indexoutertoken{string}\indexoutertoken{verbatim}
 | 
|
| 7315 | 53  | 
\begin{matharray}{rcl}
 | 
54  | 
ident & = & letter~quasiletter^* \\  | 
|
55  | 
longident & = & ident\verb,.,ident~\dots~ident \\  | 
|
| 8548 | 56  | 
symident & = & sym^+ ~|~ symbol \\  | 
| 7315 | 57  | 
nat & = & digit^+ \\  | 
58  | 
var & = & \verb,?,ident ~|~ \verb,?,ident\verb,.,nat \\  | 
|
59  | 
typefree & = & \verb,',ident \\  | 
|
60  | 
typevar & = & \verb,?,typefree ~|~ \verb,?,typefree\verb,.,nat \\  | 
|
61  | 
string & = & \verb,", ~\dots~ \verb,", \\  | 
|
| 7319 | 62  | 
  verbatim & = & \verb,{*, ~\dots~ \verb,*}, \\
 | 
63  | 
\end{matharray}
 | 
|
64  | 
\begin{matharray}{rcl}
 | 
|
| 7315 | 65  | 
letter & = & \verb,a, ~|~ \dots ~|~ \verb,z, ~|~ \verb,A, ~|~ \dots ~|~ \verb,Z, \\  | 
66  | 
digit & = & \verb,0, ~|~ \dots ~|~ \verb,9, \\  | 
|
67  | 
quasiletter & = & letter ~|~ digit ~|~ \verb,_, ~|~ \verb,', \\  | 
|
68  | 
sym & = & \verb,!, ~|~ \verb,#, ~|~ \verb,$, ~|~ \verb,%, ~|~ \verb,&, ~|~ %$  | 
|
| 10160 | 69  | 
\verb,*, ~|~ \verb,+, ~|~ \verb,-, ~|~ \verb,/, ~|~ \verb,:, ~|~ \\  | 
70  | 
  & & \verb,<, ~|~ \verb,=, ~|~ \verb,>, ~|~ \verb,?, ~|~ \texttt{\at} ~|~
 | 
|
71  | 
\verb,^, ~|~ \verb,_, ~|~ \verb,`, ~|~ \verb,|, ~|~ \verb,~, \\  | 
|
| 8548 | 72  | 
  symbol & = & {\forall} ~|~ {\exists} ~|~ \dots
 | 
| 7315 | 73  | 
\end{matharray}
 | 
74  | 
||
75  | 
The syntax of \texttt{string} admits any characters, including newlines;
 | 
|
| 7895 | 76  | 
``\verb|"|'' (double-quote) and ``\verb|\|'' (backslash) have to be escaped by  | 
| 9051 | 77  | 
a backslash; newlines need not be escaped. Note that ML-style control  | 
78  | 
characters are \emph{not} supported.  The body of \texttt{verbatim} may
 | 
|
79  | 
consist of any text not containing ``\verb|*}|''.  | 
|
| 7315 | 80  | 
|
| 7895 | 81  | 
Comments take the form \texttt{(*~\dots~*)} and may be
 | 
| 8378 | 82  | 
nested\footnote{Proof~General may occasionally get confused by nested
 | 
83  | 
  comments.}, just as in ML. Note that these are \emph{source} comments only,
 | 
|
84  | 
which are stripped after lexical analysis of the input. The Isar document  | 
|
85  | 
syntax also provides \emph{formal comments} that are actually part of the text
 | 
|
86  | 
(see \S\ref{sec:comments}).
 | 
|
| 7315 | 87  | 
|
| 10160 | 88  | 
Mathematical symbols such as ``$\forall$'' are represented in plain ASCII as  | 
89  | 
``\verb,\<forall>,''.  | 
|
90  | 
||
| 7046 | 91  | 
|
92  | 
\section{Common syntax entities}
 | 
|
93  | 
||
| 7335 | 94  | 
Subsequently, we introduce several basic syntactic entities, such as names,  | 
| 7895 | 95  | 
terms, and theorem specifications, which have been factored out of the actual  | 
96  | 
Isar language elements to be described later.  | 
|
| 7134 | 97  | 
|
| 7981 | 98  | 
Note that some of the basic syntactic entities introduced below (e.g.\  | 
| 7895 | 99  | 
\railqtoken{name}) act much like tokens rather than plain nonterminals (e.g.\ 
 | 
100  | 
\railnonterm{sort}), especially for the sake of error messages.  E.g.\ syntax
 | 
|
101  | 
elements such as $\CONSTS$ referring to \railqtoken{name} or \railqtoken{type}
 | 
|
102  | 
would really report a missing name or type rather than any of the constituent  | 
|
103  | 
primitive tokens such as \railtoken{ident} or \railtoken{string}.
 | 
|
| 7046 | 104  | 
|
| 7050 | 105  | 
|
106  | 
\subsection{Names}
 | 
|
107  | 
||
| 7134 | 108  | 
Entity \railqtoken{name} usually refers to any name of types, constants,
 | 
| 7167 | 109  | 
theorems etc.\ that are to be \emph{declared} or \emph{defined} (so qualified
 | 
| 8548 | 110  | 
identifiers are excluded here). Quoted strings provide an escape for  | 
| 7134 | 111  | 
non-identifier names or those ruled out by outer syntax keywords (e.g.\  | 
112  | 
\verb|"let"|). Already existing objects are usually referenced by  | 
|
113  | 
\railqtoken{nameref}.
 | 
|
| 7050 | 114  | 
|
| 7141 | 115  | 
\indexoutertoken{name}\indexoutertoken{parname}\indexoutertoken{nameref}
 | 
| 9617 | 116  | 
\indexoutertoken{int}
 | 
| 7046 | 117  | 
\begin{rail}
 | 
| 8145 | 118  | 
name: ident | symident | string | nat  | 
| 7046 | 119  | 
;  | 
| 7167 | 120  | 
  parname: '(' name ')'
 | 
| 7141 | 121  | 
;  | 
| 7167 | 122  | 
nameref: name | longident  | 
| 7046 | 123  | 
;  | 
| 9617 | 124  | 
int: nat | '-' nat  | 
125  | 
;  | 
|
| 7046 | 126  | 
\end{rail}
 | 
127  | 
||
| 7050 | 128  | 
|
| 7315 | 129  | 
\subsection{Comments}\label{sec:comments}
 | 
| 7046 | 130  | 
|
| 7167 | 131  | 
Large chunks of plain \railqtoken{text} are usually given
 | 
| 7895 | 132  | 
\railtoken{verbatim}, i.e.\ enclosed in \verb|{*|~\dots~\verb|*}|.  For
 | 
| 7175 | 133  | 
convenience, any of the smaller text units conforming to \railqtoken{nameref}
 | 
| 8102 | 134  | 
are admitted as well. Almost any of the Isar commands may be annotated by  | 
| 7466 | 135  | 
marginal \railnonterm{comment} of the form \texttt{--} \railqtoken{text}.
 | 
136  | 
Note that the latter kind of comment is actually part of the language, while  | 
|
| 7895 | 137  | 
source level comments \verb|(*|~\dots~\verb|*)| are stripped at the lexical  | 
| 7466 | 138  | 
level. A few commands such as $\PROOFNAME$ admit additional markup with a  | 
139  | 
``level of interest'': \texttt{\%} followed by an optional number $n$ (default
 | 
|
140  | 
$n = 1$) indicates that the respective part of the document becomes $n$ levels  | 
|
141  | 
more obscure; \texttt{\%\%} means that interest drops by $\infty$ --- abandon
 | 
|
142  | 
every hope, who enter here.  | 
|
| 7050 | 143  | 
|
144  | 
\indexoutertoken{text}\indexouternonterm{comment}\indexouternonterm{interest}
 | 
|
| 7046 | 145  | 
\begin{rail}
 | 
| 7167 | 146  | 
text: verbatim | nameref  | 
| 7050 | 147  | 
;  | 
| 8102 | 148  | 
  comment: ('--' text +)
 | 
| 7046 | 149  | 
;  | 
| 7167 | 150  | 
interest: percent nat? | ppercent  | 
| 7046 | 151  | 
;  | 
152  | 
\end{rail}
 | 
|
153  | 
||
154  | 
||
| 7335 | 155  | 
\subsection{Type classes, sorts and arities}
 | 
| 7046 | 156  | 
|
| 8896 | 157  | 
Classes are specified by plain names. Sorts have a very simple inner syntax,  | 
158  | 
which is either a single class name $c$ or a list $\{c@1, \dots, c@n\}$
 | 
|
159  | 
referring to the intersection of these classes. The syntax of type arities is  | 
|
160  | 
given directly at the outer level.  | 
|
| 7050 | 161  | 
|
162  | 
\indexouternonterm{sort}\indexouternonterm{arity}\indexouternonterm{simplearity}
 | 
|
| 7135 | 163  | 
\indexouternonterm{classdecl}
 | 
| 7046 | 164  | 
\begin{rail}
 | 
| 7321 | 165  | 
  classdecl: name ('<' (nameref + ','))?
 | 
| 7046 | 166  | 
;  | 
| 8896 | 167  | 
sort: nameref  | 
| 7046 | 168  | 
;  | 
| 7167 | 169  | 
  arity: ('(' (sort + ',') ')')? sort
 | 
| 7050 | 170  | 
;  | 
| 7167 | 171  | 
  simplearity: ('(' (sort + ',') ')')? nameref
 | 
172  | 
;  | 
|
| 7050 | 173  | 
\end{rail}
 | 
174  | 
||
175  | 
||
| 7167 | 176  | 
\subsection{Types and terms}\label{sec:types-terms}
 | 
| 7050 | 177  | 
|
| 7167 | 178  | 
The actual inner Isabelle syntax, that of types and terms of the logic, is far  | 
| 7895 | 179  | 
too sophisticated in order to be modelled explicitly at the outer theory  | 
| 8548 | 180  | 
level. Basically, any such entity has to be quoted to turn it into a single  | 
181  | 
token (the parsing and type-checking is performed internally later). For  | 
|
182  | 
convenience, a slightly more liberal convention is adopted: quotes may be  | 
|
| 7895 | 183  | 
omitted for any type or term that is already \emph{atomic} at the outer level.
 | 
184  | 
For example, one may write just \texttt{x} instead of \texttt{"x"}.  Note that
 | 
|
| 8548 | 185  | 
symbolic identifiers (e.g.\ \texttt{++} or $\forall$) are available as well,
 | 
186  | 
provided these are not superseded by commands or keywords (e.g.\ \texttt{+}).
 | 
|
| 7050 | 187  | 
|
188  | 
\indexoutertoken{type}\indexoutertoken{term}\indexoutertoken{prop}
 | 
|
189  | 
\begin{rail}
 | 
|
| 7167 | 190  | 
type: nameref | typefree | typevar  | 
| 7050 | 191  | 
;  | 
| 8593 | 192  | 
term: nameref | var  | 
| 7050 | 193  | 
;  | 
| 7167 | 194  | 
prop: term  | 
| 7050 | 195  | 
;  | 
196  | 
\end{rail}
 | 
|
197  | 
||
| 8690 | 198  | 
Positional instantiations are indicated by giving a sequence of terms, or the  | 
199  | 
placeholder ``$\_$'' (underscore), which means to skip a position.  | 
|
200  | 
||
201  | 
\indexoutertoken{inst}\indexoutertoken{insts}
 | 
|
202  | 
\begin{rail}
 | 
|
203  | 
inst: underscore | term  | 
|
204  | 
;  | 
|
205  | 
insts: (inst *)  | 
|
206  | 
;  | 
|
207  | 
\end{rail}
 | 
|
208  | 
||
| 7167 | 209  | 
Type declarations and definitions usually refer to \railnonterm{typespec} on
 | 
210  | 
the left-hand side. This models basic type constructor application at the  | 
|
211  | 
outer syntax level. Note that only plain postfix notation is available here,  | 
|
212  | 
but no infixes.  | 
|
| 7050 | 213  | 
|
214  | 
\indexouternonterm{typespec}
 | 
|
215  | 
\begin{rail}
 | 
|
| 7167 | 216  | 
  typespec: (() | typefree | '(' ( typefree + ',' ) ')') name
 | 
| 7050 | 217  | 
;  | 
218  | 
\end{rail}
 | 
|
219  | 
||
220  | 
||
| 7315 | 221  | 
\subsection{Term patterns}\label{sec:term-pats}
 | 
| 7050 | 222  | 
|
| 7895 | 223  | 
Assumptions and goal statements usually admit casual binding of schematic term  | 
| 7981 | 224  | 
variables by giving (optional) patterns of the form $\ISS{p@1\;\dots}{p@n}$.
 | 
| 7167 | 225  | 
There are separate versions available for \railqtoken{term}s and
 | 
226  | 
\railqtoken{prop}s.  The latter provides a $\CONCLNAME$ part with patterns
 | 
|
227  | 
referring the (atomic) conclusion of a rule.  | 
|
| 7050 | 228  | 
|
229  | 
\indexouternonterm{termpat}\indexouternonterm{proppat}
 | 
|
230  | 
\begin{rail}
 | 
|
| 7167 | 231  | 
  termpat: '(' ('is' term +) ')'
 | 
| 7050 | 232  | 
;  | 
| 7167 | 233  | 
  proppat: '(' (('is' prop +) | 'concl' ('is' prop +) | ('is' prop +) 'concl' ('is' prop +)) ')'
 | 
| 7046 | 234  | 
;  | 
235  | 
\end{rail}
 | 
|
236  | 
||
237  | 
||
| 7050 | 238  | 
\subsection{Mixfix annotations}
 | 
239  | 
||
| 7134 | 240  | 
Mixfix annotations specify concrete \emph{inner} syntax of Isabelle types and
 | 
| 8548 | 241  | 
terms (see also \cite{isabelle-ref}).  Some commands such as $\TYPES$ (see
 | 
242  | 
\S\ref{sec:types-pure}) admit infixes only, while $\CONSTS$ (see
 | 
|
243  | 
\S\ref{sec:consts}) and $\isarkeyword{syntax}$ (see \S\ref{sec:syn-trans})
 | 
|
244  | 
support the full range of general mixfixes and binders.  | 
|
| 7046 | 245  | 
|
| 7050 | 246  | 
\indexouternonterm{infix}\indexouternonterm{mixfix}
 | 
| 7046 | 247  | 
\begin{rail}
 | 
| 7167 | 248  | 
  infix: '(' ('infixl' | 'infixr') string? nat ')'
 | 
249  | 
;  | 
|
| 7175 | 250  | 
  mixfix: infix | '(' string prios? nat? ')' | '(' 'binder' string prios? nat ')'
 | 
| 7050 | 251  | 
;  | 
252  | 
||
| 7175 | 253  | 
prios: '[' (nat + ',') ']'  | 
| 7050 | 254  | 
;  | 
| 7046 | 255  | 
\end{rail}
 | 
256  | 
||
| 7050 | 257  | 
|
| 7134 | 258  | 
\subsection{Attributes and theorems}\label{sec:syn-att}
 | 
| 7050 | 259  | 
|
260  | 
Attributes (and proof methods, see \S\ref{sec:syn-meth}) have their own
 | 
|
| 7335 | 261  | 
``semi-inner'' syntax, in the sense that input conforming to  | 
262  | 
\railnonterm{args} below is parsed by the attribute a second time.  The
 | 
|
263  | 
attribute argument specifications may be any sequence of atomic entities  | 
|
264  | 
(identifiers, strings etc.), or properly bracketed argument lists. Below  | 
|
| 7981 | 265  | 
\railqtoken{atom} refers to any atomic entity, including any
 | 
266  | 
\railtoken{keyword} conforming to \railtoken{symident}.
 | 
|
| 7050 | 267  | 
|
268  | 
\indexoutertoken{atom}\indexouternonterm{args}\indexouternonterm{attributes}
 | 
|
269  | 
\begin{rail}
 | 
|
| 7466 | 270  | 
atom: nameref | typefree | typevar | var | nat | keyword  | 
| 7134 | 271  | 
;  | 
| 8896 | 272  | 
  arg: atom | '(' args ')' | '[' args ']'
 | 
| 7050 | 273  | 
;  | 
| 7167 | 274  | 
args: arg *  | 
| 7134 | 275  | 
;  | 
| 7167 | 276  | 
attributes: '[' (nameref args * ',') ']'  | 
| 7050 | 277  | 
;  | 
278  | 
\end{rail}
 | 
|
279  | 
||
| 7895 | 280  | 
Theorem specifications come in several flavors: \railnonterm{axmdecl} and
 | 
| 7175 | 281  | 
\railnonterm{thmdecl} usually refer to axioms, assumptions or results of goal
 | 
| 7981 | 282  | 
statements, while \railnonterm{thmdef} collects lists of existing theorems.
 | 
283  | 
Existing theorems are given by \railnonterm{thmref} and \railnonterm{thmrefs},
 | 
|
284  | 
the former requires an actual singleton result. Any of these theorem  | 
|
| 7175 | 285  | 
specifications may include lists of attributes both on the left and right hand  | 
| 7466 | 286  | 
sides; attributes are applied to any immediately preceding theorem. If names  | 
| 7981 | 287  | 
are omitted, the theorems are not stored within the theorem database of the  | 
288  | 
theory or proof context; any given attributes are still applied, though.  | 
|
| 7050 | 289  | 
|
| 7135 | 290  | 
\indexouternonterm{thmdecl}\indexouternonterm{axmdecl}
 | 
291  | 
\indexouternonterm{thmdef}\indexouternonterm{thmrefs}
 | 
|
| 7050 | 292  | 
\begin{rail}
 | 
| 7167 | 293  | 
axmdecl: name attributes? ':'  | 
| 7050 | 294  | 
;  | 
| 9200 | 295  | 
thmdecl: thmbind ':'  | 
| 7135 | 296  | 
;  | 
| 9200 | 297  | 
thmdef: thmbind '='  | 
| 7050 | 298  | 
;  | 
| 7175 | 299  | 
thmref: nameref attributes?  | 
300  | 
;  | 
|
301  | 
thmrefs: thmref +  | 
|
| 7134 | 302  | 
;  | 
| 7167 | 303  | 
|
| 9200 | 304  | 
thmbind: name attributes | name | attributes  | 
| 7050 | 305  | 
;  | 
306  | 
\end{rail}
 | 
|
| 7046 | 307  | 
|
308  | 
||
| 7050 | 309  | 
\subsection{Proof methods}\label{sec:syn-meth}
 | 
| 7046 | 310  | 
|
| 7050 | 311  | 
Proof methods are either basic ones, or expressions composed of methods via  | 
| 7175 | 312  | 
``\texttt{,}'' (sequential composition), ``\texttt{|}'' (alternative choices),
 | 
| 7981 | 313  | 
``\texttt{?}'' (try), ``\texttt{+}'' (repeat at least once).  In practice,
 | 
314  | 
proof methods are usually just a comma separated list of  | 
|
315  | 
\railqtoken{nameref}~\railnonterm{args} specifications.  Note that parentheses
 | 
|
316  | 
may be dropped for single method specifications (with no arguments).  | 
|
| 7046 | 317  | 
|
| 7050 | 318  | 
\indexouternonterm{method}
 | 
319  | 
\begin{rail}
 | 
|
| 7430 | 320  | 
  method: (nameref | '(' methods ')') (() | '?' | '+')
 | 
| 7134 | 321  | 
;  | 
| 7167 | 322  | 
  methods: (nameref args | method) + (',' | '|')
 | 
| 7050 | 323  | 
;  | 
324  | 
\end{rail}
 | 
|
| 7046 | 325  | 
|
| 8532 | 326  | 
Proper use of Isar proof methods does \emph{not} involve goal addressing.
 | 
327  | 
Nevertheless, specifying goal ranges may occasionally come in handy in  | 
|
328  | 
emulating tactic scripts. Note that $[n-]$ refers to all goals, starting from  | 
|
| 8548 | 329  | 
$n$. All goals may be specified by $[!]$, which is the same as $[1-]$.  | 
| 8532 | 330  | 
|
331  | 
\indexouternonterm{goalspec}
 | 
|
332  | 
\begin{rail}
 | 
|
| 8548 | 333  | 
goalspec: '[' (nat '-' nat | nat '-' | nat | '!' ) ']'  | 
| 8532 | 334  | 
;  | 
335  | 
\end{rail}
 | 
|
336  | 
||
| 7046 | 337  | 
|
| 9200 | 338  | 
\subsection{Antiquotations}\label{sec:antiq}
 | 
339  | 
||
| 10336 | 340  | 
\begin{matharray}{rcl}
 | 
341  | 
thm & : & \isarantiq \\  | 
|
342  | 
prop & : & \isarantiq \\  | 
|
343  | 
term & : & \isarantiq \\  | 
|
344  | 
typ & : & \isarantiq \\  | 
|
345  | 
text & : & \isarantiq \\  | 
|
346  | 
goals & : & \isarantiq \\  | 
|
| 10351 | 347  | 
subgoals & : & \isarantiq \\  | 
| 10336 | 348  | 
\end{matharray}
 | 
349  | 
||
| 9200 | 350  | 
The text body of formal comments (see also \S\ref{sec:comments}) may contain
 | 
351  | 
antiquotations of logical entities, such as theorems, terms and types, which  | 
|
352  | 
are to be presented in the final output produced by the Isabelle document  | 
|
353  | 
preparation system (see also \S\ref{sec:document-prep}).
 | 
|
354  | 
||
| 9601 | 355  | 
Thus embedding of  | 
| 10160 | 356  | 
\texttt{{\at}{\ttlbrace}term~[show_types]~"f(x)~=~a~+~x"{\ttrbrace}} within a
 | 
| 9601 | 357  | 
text block would cause  | 
| 9200 | 358  | 
\isa{(f{\isasymColon}'a~{\isasymRightarrow}~'a)~(x{\isasymColon}'a)~=~(a{\isasymColon}'a)~+~x}
 | 
| 10160 | 359  | 
to appear in the final {\LaTeX} document.  Also note that theorem
 | 
360  | 
antiquotations may involve attributes as well. For example,  | 
|
361  | 
\texttt{{\at}{\ttlbrace}thm~sym~[no_vars]{\ttrbrace}} would print the
 | 
|
362  | 
statement where all schematic variables have been replaced by fixed ones,  | 
|
363  | 
which are better readable.  | 
|
| 9200 | 364  | 
|
| 9728 | 365  | 
\indexisarant{thm}\indexisarant{prop}\indexisarant{term}
 | 
| 10355 | 366  | 
\indexisarant{typ}\indexisarant{text}\indexisarant{goals}\indexisarant{subgoals}
 | 
| 9200 | 367  | 
\begin{rail}
 | 
368  | 
atsign lbrace antiquotation rbrace  | 
|
369  | 
;  | 
|
370  | 
||
371  | 
antiquotation:  | 
|
372  | 
'thm' options thmrefs |  | 
|
373  | 
'prop' options prop |  | 
|
374  | 
'term' options term |  | 
|
| 9728 | 375  | 
'typ' options type |  | 
| 
10319
 
02463775cafb
added antiquotation "goals" and option "goals_limit";
 
wenzelm 
parents: 
10160 
diff
changeset
 | 
376  | 
'text' options name |  | 
| 10355 | 377  | 
'goals' options |  | 
378  | 
'subgoals' options  | 
|
| 9200 | 379  | 
;  | 
380  | 
options: '[' (option * ',') ']'  | 
|
381  | 
;  | 
|
382  | 
option: name | name '=' name  | 
|
383  | 
;  | 
|
384  | 
\end{rail}
 | 
|
385  | 
||
386  | 
Note that the syntax of antiquotations may \emph{not} include source comments
 | 
|
387  | 
\texttt{(*~\dots~*)} or verbatim text \verb|{*|~\dots~\verb|*}|.
 | 
|
388  | 
||
| 
10319
 
02463775cafb
added antiquotation "goals" and option "goals_limit";
 
wenzelm 
parents: 
10160 
diff
changeset
 | 
389  | 
\begin{descr}
 | 
| 10336 | 390  | 
\item [$\at\{thm~\vec a\}$] prints theorems $\vec a$. Note that attribute
 | 
391  | 
  specifications may be included as well (see also \S\ref{sec:syn-att}); the
 | 
|
392  | 
  $no_vars$ operation (see \S\ref{sec:misc-methods}) would be particularly
 | 
|
393  | 
useful to suppress printing of schematic variables.  | 
|
394  | 
\item [$\at\{prop~\phi\}$] prints a well-typed proposition $\phi$.
 | 
|
395  | 
\item [$\at\{term~t\}$] prints a well-typed term $t$.
 | 
|
396  | 
\item [$\at\{typ~\tau\}$] prints a well-formed type $\tau$.
 | 
|
397  | 
\item [$\at\{text~s\}$] prints uninterpreted source text $s$.  This is
 | 
|
398  | 
particularly useful to print portions of text according to the Isabelle  | 
|
399  | 
  {\LaTeX} output style, without demanding well-formedness (e.g.\ small pieces
 | 
|
400  | 
of terms that cannot be parsed or type-checked yet).  | 
|
401  | 
\item [$\at\{goals\}$] prints the current \emph{dynamic} goal state.  This is
 | 
|
402  | 
only for support of tactic-emulation scripts within Isar --- presentation of  | 
|
403  | 
goal states does not conform to actual human-readable proof documents.  | 
|
| 
10319
 
02463775cafb
added antiquotation "goals" and option "goals_limit";
 
wenzelm 
parents: 
10160 
diff
changeset
 | 
404  | 
|
| 
 
02463775cafb
added antiquotation "goals" and option "goals_limit";
 
wenzelm 
parents: 
10160 
diff
changeset
 | 
405  | 
Please do not include goal states into document output unless you really  | 
| 
 
02463775cafb
added antiquotation "goals" and option "goals_limit";
 
wenzelm 
parents: 
10160 
diff
changeset
 | 
406  | 
know what you are doing!  | 
| 10355 | 407  | 
\item [$\at\{subgoals\}$] behaves almost like $goals$, except that it does not
 | 
408  | 
print the main goal.  | 
|
| 
10319
 
02463775cafb
added antiquotation "goals" and option "goals_limit";
 
wenzelm 
parents: 
10160 
diff
changeset
 | 
409  | 
\end{descr}
 | 
| 
 
02463775cafb
added antiquotation "goals" and option "goals_limit";
 
wenzelm 
parents: 
10160 
diff
changeset
 | 
410  | 
|
| 9200 | 411  | 
\medskip  | 
412  | 
||
| 10336 | 413  | 
The following options are available to tune the output. Note that most of  | 
| 9233 | 414  | 
these coincide with ML flags of the same names (see also \cite{isabelle-ref}).
 | 
| 9200 | 415  | 
\begin{descr}
 | 
| 9233 | 416  | 
\item[$show_types = bool$ and $show_sorts = bool$] control printing of  | 
| 9234 | 417  | 
explicit type and sort constraints.  | 
| 9233 | 418  | 
\item[$long_names = bool$] forces names of types and constants etc.\ to be  | 
419  | 
printed in their fully qualified internal form.  | 
|
420  | 
\item[$eta_contract = bool$] prints terms in $\eta$-contracted form.  | 
|
| 9200 | 421  | 
\item[$display = bool$] indicates if the text is to be output as multi-line  | 
422  | 
``display material'', rather than a small piece of text without line breaks  | 
|
423  | 
(which is the default).  | 
|
424  | 
\item[$quotes = bool$] indicates if the output should be enclosed in double  | 
|
425  | 
quotes.  | 
|
| 9233 | 426  | 
\item[$mode = name$] adds $name$ to the print mode to be used for presentation  | 
427  | 
  (see also \cite{isabelle-ref}).  Note that the standard setup for {\LaTeX}
 | 
|
428  | 
output is already present by default, including the modes ``$latex$'',  | 
|
429  | 
``$xsymbols$'', ``$symbols$''.  | 
|
| 9728 | 430  | 
\item[$margin = nat$ and $indent = nat$] change the margin or indentation for  | 
431  | 
pretty printing of display material.  | 
|
| 9752 | 432  | 
\item[$source = bool$] prints the source text of the antiquotation arguments,  | 
433  | 
rather than the actual value. Note that this does not affect  | 
|
434  | 
well-formedness checks of $thm$, $term$, etc. (only the $text$ antiquotation  | 
|
435  | 
admits arbitrary output).  | 
|
| 
10319
 
02463775cafb
added antiquotation "goals" and option "goals_limit";
 
wenzelm 
parents: 
10160 
diff
changeset
 | 
436  | 
\item[$goals_limit = nat$] determines the maximum number of goals to be  | 
| 
 
02463775cafb
added antiquotation "goals" and option "goals_limit";
 
wenzelm 
parents: 
10160 
diff
changeset
 | 
437  | 
printed.  | 
| 9200 | 438  | 
\end{descr}
 | 
439  | 
||
440  | 
For boolean flags, ``$name = true$'' may be abbreviated as ``$name$''. All of  | 
|
441  | 
the above flags are disabled by default, unless changed from ML.  | 
|
442  | 
||
| 10336 | 443  | 
\medskip Note that antiquotations do not only spare the author from tedious  | 
444  | 
typing, but also achieve some degree of consistency-checking of informal  | 
|
445  | 
explanations with formal developments, since well-formedness of terms and  | 
|
446  | 
types with respect to the current theory or proof context can be ensured.  | 
|
| 9200 | 447  | 
|
| 7046 | 448  | 
%%% Local Variables:  | 
449  | 
%%% mode: latex  | 
|
450  | 
%%% TeX-master: "isar-ref"  | 
|
451  | 
%%% End:  |