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 |
|
|
28 |
Another notable point is proper input termination. Proof~General demands any
|
|
29 |
command to be terminated by ``\texttt{;}''
|
|
30 |
(semicolon)\index{semicolon}\index{*;}. As far as plain Isabelle/Isar is
|
7981
|
31 |
concerned, commands may be directly run together, though. In the presentation
|
|
32 |
of Isabelle/Isar documents, semicolons are omitted in order to gain
|
|
33 |
readability.
|
7466
|
34 |
|
7315
|
35 |
|
|
36 |
\section{Lexical matters}\label{sec:lex-syntax}
|
|
37 |
|
|
38 |
The Isabelle/Isar outer syntax provides token classes as presented below.
|
7895
|
39 |
Note that some of these coincide (by full intention) with the inner lexical
|
|
40 |
syntax as presented in \cite{isabelle-ref}. These different levels of syntax
|
|
41 |
should not be confused, though.
|
7134
|
42 |
|
7335
|
43 |
%FIXME keyword, command
|
7315
|
44 |
\begin{matharray}{rcl}
|
|
45 |
ident & = & letter~quasiletter^* \\
|
|
46 |
longident & = & ident\verb,.,ident~\dots~ident \\
|
8548
|
47 |
symident & = & sym^+ ~|~ symbol \\
|
7315
|
48 |
nat & = & digit^+ \\
|
|
49 |
var & = & \verb,?,ident ~|~ \verb,?,ident\verb,.,nat \\
|
|
50 |
typefree & = & \verb,',ident \\
|
|
51 |
typevar & = & \verb,?,typefree ~|~ \verb,?,typefree\verb,.,nat \\
|
|
52 |
string & = & \verb,", ~\dots~ \verb,", \\
|
7319
|
53 |
verbatim & = & \verb,{*, ~\dots~ \verb,*}, \\
|
|
54 |
\end{matharray}
|
|
55 |
\begin{matharray}{rcl}
|
7315
|
56 |
letter & = & \verb,a, ~|~ \dots ~|~ \verb,z, ~|~ \verb,A, ~|~ \dots ~|~ \verb,Z, \\
|
|
57 |
digit & = & \verb,0, ~|~ \dots ~|~ \verb,9, \\
|
|
58 |
quasiletter & = & letter ~|~ digit ~|~ \verb,_, ~|~ \verb,', \\
|
|
59 |
sym & = & \verb,!, ~|~ \verb,#, ~|~ \verb,$, ~|~ \verb,%, ~|~ \verb,&, ~|~ %$
|
7319
|
60 |
\verb,*, ~|~ \verb,+, ~|~ \verb,-, ~|~ \verb,/, ~|~ \verb,:, ~|~
|
|
61 |
\verb,<, ~|~ \verb,=, ~|~ \verb,>, ~|~ \verb,?, ~|~ \mathtt{\at} ~|~ \\
|
|
62 |
& & \verb,^, ~|~ \verb,_, ~|~ \verb,`, ~|~ \verb,|, ~|~ \verb,~, \\
|
8548
|
63 |
symbol & = & {\forall} ~|~ {\exists} ~|~ \dots
|
7315
|
64 |
\end{matharray}
|
|
65 |
|
|
66 |
The syntax of \texttt{string} admits any characters, including newlines;
|
7895
|
67 |
``\verb|"|'' (double-quote) and ``\verb|\|'' (backslash) have to be escaped by
|
7981
|
68 |
a backslash. Note that ML-style control characters are \emph{not} supported.
|
|
69 |
The body of \texttt{verbatim} may consist of any text not containing
|
|
70 |
``\verb|*}|''.
|
7315
|
71 |
|
7895
|
72 |
Comments take the form \texttt{(*~\dots~*)} and may be
|
8378
|
73 |
nested\footnote{Proof~General may occasionally get confused by nested
|
|
74 |
comments.}, just as in ML. Note that these are \emph{source} comments only,
|
|
75 |
which are stripped after lexical analysis of the input. The Isar document
|
|
76 |
syntax also provides \emph{formal comments} that are actually part of the text
|
|
77 |
(see \S\ref{sec:comments}).
|
7315
|
78 |
|
7046
|
79 |
|
|
80 |
\section{Common syntax entities}
|
|
81 |
|
7335
|
82 |
Subsequently, we introduce several basic syntactic entities, such as names,
|
7895
|
83 |
terms, and theorem specifications, which have been factored out of the actual
|
|
84 |
Isar language elements to be described later.
|
7134
|
85 |
|
7981
|
86 |
Note that some of the basic syntactic entities introduced below (e.g.\
|
7895
|
87 |
\railqtoken{name}) act much like tokens rather than plain nonterminals (e.g.\
|
|
88 |
\railnonterm{sort}), especially for the sake of error messages. E.g.\ syntax
|
|
89 |
elements such as $\CONSTS$ referring to \railqtoken{name} or \railqtoken{type}
|
|
90 |
would really report a missing name or type rather than any of the constituent
|
|
91 |
primitive tokens such as \railtoken{ident} or \railtoken{string}.
|
7046
|
92 |
|
7050
|
93 |
|
|
94 |
\subsection{Names}
|
|
95 |
|
7134
|
96 |
Entity \railqtoken{name} usually refers to any name of types, constants,
|
7167
|
97 |
theorems etc.\ that are to be \emph{declared} or \emph{defined} (so qualified
|
8548
|
98 |
identifiers are excluded here). Quoted strings provide an escape for
|
7134
|
99 |
non-identifier names or those ruled out by outer syntax keywords (e.g.\
|
|
100 |
\verb|"let"|). Already existing objects are usually referenced by
|
|
101 |
\railqtoken{nameref}.
|
7050
|
102 |
|
7141
|
103 |
\indexoutertoken{name}\indexoutertoken{parname}\indexoutertoken{nameref}
|
7046
|
104 |
\begin{rail}
|
8145
|
105 |
name: ident | symident | string | nat
|
7046
|
106 |
;
|
7167
|
107 |
parname: '(' name ')'
|
7141
|
108 |
;
|
7167
|
109 |
nameref: name | longident
|
7046
|
110 |
;
|
|
111 |
\end{rail}
|
|
112 |
|
7050
|
113 |
|
7315
|
114 |
\subsection{Comments}\label{sec:comments}
|
7046
|
115 |
|
7167
|
116 |
Large chunks of plain \railqtoken{text} are usually given
|
7895
|
117 |
\railtoken{verbatim}, i.e.\ enclosed in \verb|{*|~\dots~\verb|*}|. For
|
7175
|
118 |
convenience, any of the smaller text units conforming to \railqtoken{nameref}
|
8102
|
119 |
are admitted as well. Almost any of the Isar commands may be annotated by
|
7466
|
120 |
marginal \railnonterm{comment} of the form \texttt{--} \railqtoken{text}.
|
|
121 |
Note that the latter kind of comment is actually part of the language, while
|
7895
|
122 |
source level comments \verb|(*|~\dots~\verb|*)| are stripped at the lexical
|
7466
|
123 |
level. A few commands such as $\PROOFNAME$ admit additional markup with a
|
|
124 |
``level of interest'': \texttt{\%} followed by an optional number $n$ (default
|
|
125 |
$n = 1$) indicates that the respective part of the document becomes $n$ levels
|
|
126 |
more obscure; \texttt{\%\%} means that interest drops by $\infty$ --- abandon
|
|
127 |
every hope, who enter here.
|
7050
|
128 |
|
|
129 |
\indexoutertoken{text}\indexouternonterm{comment}\indexouternonterm{interest}
|
7046
|
130 |
\begin{rail}
|
7167
|
131 |
text: verbatim | nameref
|
7050
|
132 |
;
|
8102
|
133 |
comment: ('--' text +)
|
7046
|
134 |
;
|
7167
|
135 |
interest: percent nat? | ppercent
|
7046
|
136 |
;
|
|
137 |
\end{rail}
|
|
138 |
|
|
139 |
|
7335
|
140 |
\subsection{Type classes, sorts and arities}
|
7046
|
141 |
|
7050
|
142 |
The syntax of sorts and arities is given directly at the outer level. Note
|
7335
|
143 |
that this is in contrast to types and terms (see \ref{sec:types-terms}).
|
7050
|
144 |
|
|
145 |
\indexouternonterm{sort}\indexouternonterm{arity}\indexouternonterm{simplearity}
|
7135
|
146 |
\indexouternonterm{classdecl}
|
7046
|
147 |
\begin{rail}
|
7321
|
148 |
classdecl: name ('<' (nameref + ','))?
|
7046
|
149 |
;
|
7167
|
150 |
sort: nameref | lbrace (nameref * ',') rbrace
|
7046
|
151 |
;
|
7167
|
152 |
arity: ('(' (sort + ',') ')')? sort
|
7050
|
153 |
;
|
7167
|
154 |
simplearity: ('(' (sort + ',') ')')? nameref
|
|
155 |
;
|
7050
|
156 |
\end{rail}
|
|
157 |
|
|
158 |
|
7167
|
159 |
\subsection{Types and terms}\label{sec:types-terms}
|
7050
|
160 |
|
7167
|
161 |
The actual inner Isabelle syntax, that of types and terms of the logic, is far
|
7895
|
162 |
too sophisticated in order to be modelled explicitly at the outer theory
|
8548
|
163 |
level. Basically, any such entity has to be quoted to turn it into a single
|
|
164 |
token (the parsing and type-checking is performed internally later). For
|
|
165 |
convenience, a slightly more liberal convention is adopted: quotes may be
|
7895
|
166 |
omitted for any type or term that is already \emph{atomic} at the outer level.
|
|
167 |
For example, one may write just \texttt{x} instead of \texttt{"x"}. Note that
|
8548
|
168 |
symbolic identifiers (e.g.\ \texttt{++} or $\forall$) are available as well,
|
|
169 |
provided these are not superseded by commands or keywords (e.g.\ \texttt{+}).
|
7050
|
170 |
|
|
171 |
\indexoutertoken{type}\indexoutertoken{term}\indexoutertoken{prop}
|
|
172 |
\begin{rail}
|
7167
|
173 |
type: nameref | typefree | typevar
|
7050
|
174 |
;
|
8593
|
175 |
term: nameref | var
|
7050
|
176 |
;
|
7167
|
177 |
prop: term
|
7050
|
178 |
;
|
|
179 |
\end{rail}
|
|
180 |
|
7167
|
181 |
Type declarations and definitions usually refer to \railnonterm{typespec} on
|
|
182 |
the left-hand side. This models basic type constructor application at the
|
|
183 |
outer syntax level. Note that only plain postfix notation is available here,
|
|
184 |
but no infixes.
|
7050
|
185 |
|
|
186 |
\indexouternonterm{typespec}
|
|
187 |
\begin{rail}
|
7167
|
188 |
typespec: (() | typefree | '(' ( typefree + ',' ) ')') name
|
7050
|
189 |
;
|
|
190 |
\end{rail}
|
|
191 |
|
|
192 |
|
7315
|
193 |
\subsection{Term patterns}\label{sec:term-pats}
|
7050
|
194 |
|
7895
|
195 |
Assumptions and goal statements usually admit casual binding of schematic term
|
7981
|
196 |
variables by giving (optional) patterns of the form $\ISS{p@1\;\dots}{p@n}$.
|
7167
|
197 |
There are separate versions available for \railqtoken{term}s and
|
|
198 |
\railqtoken{prop}s. The latter provides a $\CONCLNAME$ part with patterns
|
|
199 |
referring the (atomic) conclusion of a rule.
|
7050
|
200 |
|
|
201 |
\indexouternonterm{termpat}\indexouternonterm{proppat}
|
|
202 |
\begin{rail}
|
7167
|
203 |
termpat: '(' ('is' term +) ')'
|
7050
|
204 |
;
|
7167
|
205 |
proppat: '(' (('is' prop +) | 'concl' ('is' prop +) | ('is' prop +) 'concl' ('is' prop +)) ')'
|
7046
|
206 |
;
|
|
207 |
\end{rail}
|
|
208 |
|
|
209 |
|
7050
|
210 |
\subsection{Mixfix annotations}
|
|
211 |
|
7134
|
212 |
Mixfix annotations specify concrete \emph{inner} syntax of Isabelle types and
|
8548
|
213 |
terms (see also \cite{isabelle-ref}). Some commands such as $\TYPES$ (see
|
|
214 |
\S\ref{sec:types-pure}) admit infixes only, while $\CONSTS$ (see
|
|
215 |
\S\ref{sec:consts}) and $\isarkeyword{syntax}$ (see \S\ref{sec:syn-trans})
|
|
216 |
support the full range of general mixfixes and binders.
|
7046
|
217 |
|
7050
|
218 |
\indexouternonterm{infix}\indexouternonterm{mixfix}
|
7046
|
219 |
\begin{rail}
|
7167
|
220 |
infix: '(' ('infixl' | 'infixr') string? nat ')'
|
|
221 |
;
|
7175
|
222 |
mixfix: infix | '(' string prios? nat? ')' | '(' 'binder' string prios? nat ')'
|
7050
|
223 |
;
|
|
224 |
|
7175
|
225 |
prios: '[' (nat + ',') ']'
|
7050
|
226 |
;
|
7046
|
227 |
\end{rail}
|
|
228 |
|
7050
|
229 |
|
7134
|
230 |
\subsection{Attributes and theorems}\label{sec:syn-att}
|
7050
|
231 |
|
|
232 |
Attributes (and proof methods, see \S\ref{sec:syn-meth}) have their own
|
7335
|
233 |
``semi-inner'' syntax, in the sense that input conforming to
|
|
234 |
\railnonterm{args} below is parsed by the attribute a second time. The
|
|
235 |
attribute argument specifications may be any sequence of atomic entities
|
|
236 |
(identifiers, strings etc.), or properly bracketed argument lists. Below
|
7981
|
237 |
\railqtoken{atom} refers to any atomic entity, including any
|
|
238 |
\railtoken{keyword} conforming to \railtoken{symident}.
|
7050
|
239 |
|
|
240 |
\indexoutertoken{atom}\indexouternonterm{args}\indexouternonterm{attributes}
|
|
241 |
\begin{rail}
|
7466
|
242 |
atom: nameref | typefree | typevar | var | nat | keyword
|
7134
|
243 |
;
|
7167
|
244 |
arg: atom | '(' args ')' | '[' args ']' | lbrace args rbrace
|
7050
|
245 |
;
|
7167
|
246 |
args: arg *
|
7134
|
247 |
;
|
7167
|
248 |
attributes: '[' (nameref args * ',') ']'
|
7050
|
249 |
;
|
|
250 |
\end{rail}
|
|
251 |
|
7895
|
252 |
Theorem specifications come in several flavors: \railnonterm{axmdecl} and
|
7175
|
253 |
\railnonterm{thmdecl} usually refer to axioms, assumptions or results of goal
|
7981
|
254 |
statements, while \railnonterm{thmdef} collects lists of existing theorems.
|
|
255 |
Existing theorems are given by \railnonterm{thmref} and \railnonterm{thmrefs},
|
|
256 |
the former requires an actual singleton result. Any of these theorem
|
7175
|
257 |
specifications may include lists of attributes both on the left and right hand
|
7466
|
258 |
sides; attributes are applied to any immediately preceding theorem. If names
|
7981
|
259 |
are omitted, the theorems are not stored within the theorem database of the
|
|
260 |
theory or proof context; any given attributes are still applied, though.
|
7050
|
261 |
|
7135
|
262 |
\indexouternonterm{thmdecl}\indexouternonterm{axmdecl}
|
|
263 |
\indexouternonterm{thmdef}\indexouternonterm{thmrefs}
|
7050
|
264 |
\begin{rail}
|
7167
|
265 |
axmdecl: name attributes? ':'
|
7050
|
266 |
;
|
7167
|
267 |
thmdecl: thmname ':'
|
7135
|
268 |
;
|
7167
|
269 |
thmdef: thmname '='
|
7050
|
270 |
;
|
7175
|
271 |
thmref: nameref attributes?
|
|
272 |
;
|
|
273 |
thmrefs: thmref +
|
7134
|
274 |
;
|
7167
|
275 |
|
|
276 |
thmname: name attributes | name | attributes
|
7050
|
277 |
;
|
|
278 |
\end{rail}
|
7046
|
279 |
|
|
280 |
|
7050
|
281 |
\subsection{Proof methods}\label{sec:syn-meth}
|
7046
|
282 |
|
7050
|
283 |
Proof methods are either basic ones, or expressions composed of methods via
|
7175
|
284 |
``\texttt{,}'' (sequential composition), ``\texttt{|}'' (alternative choices),
|
7981
|
285 |
``\texttt{?}'' (try), ``\texttt{+}'' (repeat at least once). In practice,
|
|
286 |
proof methods are usually just a comma separated list of
|
|
287 |
\railqtoken{nameref}~\railnonterm{args} specifications. Note that parentheses
|
|
288 |
may be dropped for single method specifications (with no arguments).
|
7046
|
289 |
|
7050
|
290 |
\indexouternonterm{method}
|
|
291 |
\begin{rail}
|
7430
|
292 |
method: (nameref | '(' methods ')') (() | '?' | '+')
|
7134
|
293 |
;
|
7167
|
294 |
methods: (nameref args | method) + (',' | '|')
|
7050
|
295 |
;
|
|
296 |
\end{rail}
|
7046
|
297 |
|
8532
|
298 |
Proper use of Isar proof methods does \emph{not} involve goal addressing.
|
|
299 |
Nevertheless, specifying goal ranges may occasionally come in handy in
|
|
300 |
emulating tactic scripts. Note that $[n-]$ refers to all goals, starting from
|
8548
|
301 |
$n$. All goals may be specified by $[!]$, which is the same as $[1-]$.
|
8532
|
302 |
|
|
303 |
\indexouternonterm{goalspec}
|
|
304 |
\begin{rail}
|
8548
|
305 |
goalspec: '[' (nat '-' nat | nat '-' | nat | '!' ) ']'
|
8532
|
306 |
;
|
|
307 |
\end{rail}
|
|
308 |
|
7046
|
309 |
|
|
310 |
%%% Local Variables:
|
|
311 |
%%% mode: latex
|
|
312 |
%%% TeX-master: "isar-ref"
|
|
313 |
%%% End:
|