| author | haftmann | 
| Fri, 29 Apr 2005 08:03:01 +0200 | |
| changeset 15880 | d6aa6c707acf | 
| parent 15687 | 8fa8872cdc49 | 
| child 15960 | 9bd6550dc004 | 
| permissions | -rw-r--r-- | 
| 7046 | 1 | |
| 13048 | 2 | \chapter{Syntax primitives}
 | 
| 7046 | 3 | |
| 12618 | 4 | The rather generic framework of Isabelle/Isar syntax emerges from three main | 
| 5 | syntactic categories: \emph{commands} of the top-level Isar engine (covering
 | |
| 6 | theory and proof elements), \emph{methods} for general goal refinements
 | |
| 7 | (analogous to traditional ``tactics''), and \emph{attributes} for operations
 | |
| 8 | on facts (within a certain context). Here we give a reference of basic | |
| 9 | syntactic entities underlying Isabelle/Isar syntax in a bottom-up manner. | |
| 10 | Concrete theory and proof language elements will be introduced later on. | |
| 7315 | 11 | |
| 12 | \medskip | |
| 13 | ||
| 14 | In order to get started with writing well-formed Isabelle/Isar documents, the | |
| 15 | most important aspect to be noted is the difference of \emph{inner} versus
 | |
| 16 | \emph{outer} syntax.  Inner syntax is that of Isabelle types and terms of the
 | |
| 12618 | 17 | logic, while outer syntax is that of Isabelle/Isar theory sources (including | 
| 7895 | 18 | proofs). As a general rule, inner syntax entities may occur only as | 
| 19 | \emph{atomic entities} within outer syntax.  For example, the string
 | |
| 20 | \texttt{"x + y"} and identifier \texttt{z} are legal term specifications
 | |
| 21 | within a theory, while \texttt{x + y} is not.
 | |
| 7315 | 22 | |
| 23 | \begin{warn}
 | |
| 12618 | 24 | Old-style Isabelle theories used to fake parts of the inner syntax of types, | 
| 25 | with rather complicated rules when quotes may be omitted. Despite the minor | |
| 26 | drawback of requiring quotes more often, the syntax of Isabelle/Isar is | |
| 27 | somewhat simpler and more robust in that respect. | |
| 7315 | 28 | \end{warn}
 | 
| 29 | ||
| 12618 | 30 | Printed theory documents usually omit quotes to gain readability (this is a | 
| 31 | matter of {\LaTeX} macro setup, say via \verb,\isabellestyle,, see also
 | |
| 32 | \cite{isabelle-sys}).  Experienced users of Isabelle/Isar may easily
 | |
| 33 | reconstruct the lost technical information, while mere readers need not care | |
| 34 | about quotes at all. | |
| 35 | ||
| 7466 | 36 | \medskip | 
| 37 | ||
| 9601 | 38 | Isabelle/Isar input may contain any number of input termination characters | 
| 12618 | 39 | ``\texttt{;}'' (semicolon) to separate commands explicitly.  This is
 | 
| 9601 | 40 | particularly useful in interactive shell sessions to make clear where the | 
| 12618 | 41 | current command is intended to end. Otherwise, the interpreter loop will | 
| 42 | continue to issue a secondary prompt ``\verb,#,'' until an end-of-command is | |
| 13039 | 43 | clearly recognized from the input syntax, e.g.\ encounter of the next command | 
| 12618 | 44 | keyword. | 
| 9601 | 45 | |
| 46 | Advanced interfaces such as Proof~General \cite{proofgeneral} do not require
 | |
| 47 | explicit semicolons, the amount of input text is determined automatically by | |
| 12618 | 48 | inspecting the present content of the Emacs text buffer. In the printed | 
| 49 | presentation of Isabelle/Isar documents semicolons are omitted altogether for | |
| 7981 | 50 | readability. | 
| 7466 | 51 | |
| 12618 | 52 | \begin{warn}
 | 
| 53 | Proof~General requires certain syntax classification tables in order to | |
| 54 | achieve properly synchronized interaction with the Isabelle/Isar process. | |
| 55 | These tables need to be consistent with the Isabelle version and particular | |
| 56 | logic image to be used in a running session (common object-logics may well | |
| 57 | change the outer syntax). The standard setup should work correctly with any | |
| 58 | of the ``official'' logic images derived from Isabelle/HOL (including HOLCF | |
| 59 | etc.). Users of alternative logics may need to tell Proof~General | |
| 60 | explicitly, e.g.\ by giving an option \verb,-k ZF, (in conjunction with | |
| 61 | \verb,-l ZF, to specify the default logic image). | |
| 62 | \end{warn}
 | |
| 7315 | 63 | |
| 64 | \section{Lexical matters}\label{sec:lex-syntax}
 | |
| 65 | ||
| 14955 | 66 | The Isabelle/Isar outer syntax provides token classes as presented below; most | 
| 67 | of these coincide with the inner lexical syntax as presented in | |
| 68 | \cite{isabelle-ref}.
 | |
| 7134 | 69 | |
| 9617 | 70 | \indexoutertoken{ident}\indexoutertoken{longident}\indexoutertoken{symident}
 | 
| 71 | \indexoutertoken{nat}\indexoutertoken{var}\indexoutertoken{typefree}
 | |
| 72 | \indexoutertoken{typevar}\indexoutertoken{string}\indexoutertoken{verbatim}
 | |
| 7315 | 73 | \begin{matharray}{rcl}
 | 
| 14955 | 74 | ident & = & letter\,quasiletter^* \\ | 
| 75 | longident & = & ident (\verb,.,ident)^+ \\ | |
| 76 | symident & = & sym^+ ~|~ \verb,\<,ident\verb,>, \\ | |
| 7315 | 77 | nat & = & digit^+ \\ | 
| 14212 | 78 | var & = & ident ~|~ \verb,?,ident ~|~ \verb,?,ident\verb,.,nat \\ | 
| 7315 | 79 | typefree & = & \verb,',ident \\ | 
| 14212 | 80 | typevar & = & typefree ~|~ \verb,?,typefree ~|~ \verb,?,typefree\verb,.,nat \\ | 
| 7315 | 81 | string & = & \verb,", ~\dots~ \verb,", \\ | 
| 14483 | 82 |   verbatim & = & \verb,{*, ~\dots~ \verb,*}, \\[1ex]
 | 
| 83 | ||
| 14960 | 84 | letter & = & latin ~|~ \verb,\<,latin\verb,>, ~|~ \verb,\<,latin\,latin\verb,>, ~|~ greek ~|~ \\ | 
| 85 | & & \verb,\<^isub>, ~|~ \verb,\<^isup>, \\ | |
| 86 | quasiletter & = & letter ~|~ digit ~|~ \verb,_, ~|~ \verb,', \\ | |
| 14895 | 87 | latin & = & \verb,a, ~|~ \dots ~|~ \verb,z, ~|~ \verb,A, ~|~ \dots ~|~ \verb,Z, \\ | 
| 7315 | 88 | digit & = & \verb,0, ~|~ \dots ~|~ \verb,9, \\ | 
| 89 | sym & = & \verb,!, ~|~ \verb,#, ~|~ \verb,$, ~|~ \verb,%, ~|~ \verb,&, ~|~ %$ | |
| 10160 | 90 | \verb,*, ~|~ \verb,+, ~|~ \verb,-, ~|~ \verb,/, ~|~ \verb,:, ~|~ \\ | 
| 91 |   & & \verb,<, ~|~ \verb,=, ~|~ \verb,>, ~|~ \verb,?, ~|~ \texttt{\at} ~|~
 | |
| 92 | \verb,^, ~|~ \verb,_, ~|~ \verb,`, ~|~ \verb,|, ~|~ \verb,~, \\ | |
| 14955 | 93 | greek & = & \verb,\<alpha>, ~|~ \verb,\<beta>, ~|~ \verb,\<gamma>, ~|~ \verb,\<delta>, ~| \\ | 
| 94 | & & \verb,\<epsilon>, ~|~ \verb,\<zeta>, ~|~ \verb,\<eta>, ~|~ \verb,\<theta>, ~| \\ | |
| 95 | & & \verb,\<iota>, ~|~ \verb,\<kappa>, ~|~ \verb,\<mu>, ~|~ \verb,\<nu>, ~| \\ | |
| 96 | & & \verb,\<xi>, ~|~ \verb,\<pi>, ~|~ \verb,\<rho>, ~|~ \verb,\<sigma>, ~| \\ | |
| 97 | & & \verb,\<tau>, ~|~ \verb,\<upsilon>, ~|~ \verb,\<phi>, ~|~ \verb,\<psi>, ~| \\ | |
| 98 | & & \verb,\<omega>, ~|~ \verb,\<Gamma>, ~|~ \verb,\<Delta>, ~|~ \verb,\<Theta>, ~| \\ | |
| 99 | & & \verb,\<Lambda>, ~|~ \verb,\<Xi>, ~|~ \verb,\<Pi>, ~|~ \verb,\<Sigma>, ~| \\ | |
| 100 | & & \verb,\<Upsilon>, ~|~ \verb,\<Phi>, ~|~ \verb,\<Psi>, ~|~ \verb,\<Omega>, \\ | |
| 7315 | 101 | \end{matharray}
 | 
| 102 | ||
| 13039 | 103 | The syntax of $string$ admits any characters, including newlines; ``\verb|"|'' | 
| 104 | (double-quote) and ``\verb|\|'' (backslash) need to be escaped by a backslash. | |
| 14955 | 105 | The body of $verbatim$ may consist of any text not containing ``\verb|*}|''; | 
| 106 | this allows convenient inclusion of quotes without further escapes. The greek | |
| 107 | letters do \emph{not} include \verb,\<lambda>,, which is already used
 | |
| 108 | differently in the meta-logic. | |
| 7315 | 109 | |
| 14955 | 110 | Common mathematical symbols such as $\forall$ are represented in Isabelle as | 
| 111 | \verb,\<forall>,. There are infinitely many legal symbols like this, although | |
| 112 | proper presentation is left to front-end tools such as {\LaTeX} or
 | |
| 113 | Proof~General with the X-Symbol package. A list of standard Isabelle symbols | |
| 114 | that work well with these tools is given in \cite[appendix~A]{isabelle-sys}.
 | |
| 115 | ||
| 116 | Comments take the form \texttt{(*~\dots~*)} and may be nested, although
 | |
| 117 | user-interface tools may prevent this.  Note that \texttt{(*~\dots~*)}
 | |
| 118 | indicate source comments only, which are stripped after lexical analysis of | |
| 119 | the input. The Isar document syntax also provides formal comments that are | |
| 120 | considered as part of the text (see \S\ref{sec:comments}).
 | |
| 12618 | 121 | |
| 122 | \begin{warn}
 | |
| 123 | Proof~General does not handle nested comments properly; it is also unable to | |
| 124 |   keep \verb,(*,\,/\,\verb,{*, and \verb,*),\,/\,\verb,*}, apart, despite
 | |
| 125 | their rather different meaning. These are inherent problems of Emacs | |
| 13039 | 126 | legacy. Users should not be overly aggressive about nesting or alternating | 
| 127 | these delimiters. | |
| 12618 | 128 | \end{warn}
 | 
| 129 | ||
| 7046 | 130 | |
| 131 | \section{Common syntax entities}
 | |
| 132 | ||
| 7335 | 133 | Subsequently, we introduce several basic syntactic entities, such as names, | 
| 7895 | 134 | terms, and theorem specifications, which have been factored out of the actual | 
| 135 | Isar language elements to be described later. | |
| 7134 | 136 | |
| 14895 | 137 | Note that some of the basic syntactic entities introduced below (e.g.\ | 
| 138 | \railqtok{name}) act much like tokens rather than plain nonterminals (e.g.\
 | |
| 7895 | 139 | \railnonterm{sort}), especially for the sake of error messages.  E.g.\ syntax
 | 
| 13048 | 140 | elements like $\CONSTS$ referring to \railqtok{name} or \railqtok{type} would
 | 
| 141 | really report a missing name or type rather than any of the constituent | |
| 142 | primitive tokens such as \railtok{ident} or \railtok{string}.
 | |
| 7046 | 143 | |
| 7050 | 144 | |
| 145 | \subsection{Names}
 | |
| 146 | ||
| 13048 | 147 | Entity \railqtok{name} usually refers to any name of types, constants,
 | 
| 7167 | 148 | theorems etc.\ that are to be \emph{declared} or \emph{defined} (so qualified
 | 
| 8548 | 149 | identifiers are excluded here). Quoted strings provide an escape for | 
| 14895 | 150 | non-identifier names or those ruled out by outer syntax keywords (e.g.\ | 
| 7134 | 151 | \verb|"let"|). Already existing objects are usually referenced by | 
| 13048 | 152 | \railqtok{nameref}.
 | 
| 7050 | 153 | |
| 7141 | 154 | \indexoutertoken{name}\indexoutertoken{parname}\indexoutertoken{nameref}
 | 
| 9617 | 155 | \indexoutertoken{int}
 | 
| 7046 | 156 | \begin{rail}
 | 
| 8145 | 157 | name: ident | symident | string | nat | 
| 7046 | 158 | ; | 
| 7167 | 159 |   parname: '(' name ')'
 | 
| 7141 | 160 | ; | 
| 7167 | 161 | nameref: name | longident | 
| 7046 | 162 | ; | 
| 9617 | 163 | int: nat | '-' nat | 
| 164 | ; | |
| 7046 | 165 | \end{rail}
 | 
| 166 | ||
| 7050 | 167 | |
| 7315 | 168 | \subsection{Comments}\label{sec:comments}
 | 
| 7046 | 169 | |
| 13048 | 170 | Large chunks of plain \railqtok{text} are usually given \railtok{verbatim},
 | 
| 171 | i.e.\ enclosed in \verb|{*|~\dots~\verb|*}|.  For convenience, any of the
 | |
| 172 | smaller text units conforming to \railqtok{nameref} are admitted as well.  A
 | |
| 173 | marginal \railnonterm{comment} is of the form \texttt{--} \railqtok{text}.
 | |
| 174 | Any number of these may occur within Isabelle/Isar commands. | |
| 12618 | 175 | |
| 12879 | 176 | \indexoutertoken{text}\indexouternonterm{comment}
 | 
| 7046 | 177 | \begin{rail}
 | 
| 7167 | 178 | text: verbatim | nameref | 
| 7050 | 179 | ; | 
| 12879 | 180 | comment: '--' text | 
| 7046 | 181 | ; | 
| 182 | \end{rail}
 | |
| 183 | ||
| 184 | ||
| 7335 | 185 | \subsection{Type classes, sorts and arities}
 | 
| 7046 | 186 | |
| 8896 | 187 | Classes are specified by plain names. Sorts have a very simple inner syntax, | 
| 188 | which is either a single class name $c$ or a list $\{c@1, \dots, c@n\}$
 | |
| 189 | referring to the intersection of these classes. The syntax of type arities is | |
| 190 | given directly at the outer level. | |
| 7050 | 191 | |
| 11100 
34d58b1818f4
\<subseteq> syntax for classes/classrel/axclass/instance;
 wenzelm parents: 
10858diff
changeset | 192 | \railalias{subseteq}{\isasymsubseteq}
 | 
| 
34d58b1818f4
\<subseteq> syntax for classes/classrel/axclass/instance;
 wenzelm parents: 
10858diff
changeset | 193 | \railterm{subseteq}
 | 
| 
34d58b1818f4
\<subseteq> syntax for classes/classrel/axclass/instance;
 wenzelm parents: 
10858diff
changeset | 194 | |
| 14605 
9de4d64eee3b
'instance' and intro_classes now handle general sorts;
 wenzelm parents: 
14483diff
changeset | 195 | \indexouternonterm{sort}\indexouternonterm{arity}
 | 
| 7135 | 196 | \indexouternonterm{classdecl}
 | 
| 7046 | 197 | \begin{rail}
 | 
| 11100 
34d58b1818f4
\<subseteq> syntax for classes/classrel/axclass/instance;
 wenzelm parents: 
10858diff
changeset | 198 |   classdecl: name (('<' | subseteq) (nameref + ','))?
 | 
| 7046 | 199 | ; | 
| 8896 | 200 | sort: nameref | 
| 7046 | 201 | ; | 
| 7167 | 202 |   arity: ('(' (sort + ',') ')')? sort
 | 
| 7050 | 203 | ; | 
| 204 | \end{rail}
 | |
| 205 | ||
| 206 | ||
| 7167 | 207 | \subsection{Types and terms}\label{sec:types-terms}
 | 
| 7050 | 208 | |
| 7167 | 209 | The actual inner Isabelle syntax, that of types and terms of the logic, is far | 
| 7895 | 210 | too sophisticated in order to be modelled explicitly at the outer theory | 
| 8548 | 211 | level. Basically, any such entity has to be quoted to turn it into a single | 
| 212 | token (the parsing and type-checking is performed internally later). For | |
| 213 | convenience, a slightly more liberal convention is adopted: quotes may be | |
| 13039 | 214 | omitted for any type or term that is already atomic at the outer level. For | 
| 215 | example, one may just write \texttt{x} instead of \texttt{"x"}.  Note that
 | |
| 8548 | 216 | symbolic identifiers (e.g.\ \texttt{++} or $\forall$) are available as well,
 | 
| 12618 | 217 | provided these have not been superseded by commands or other keywords already | 
| 218 | (e.g.\ \texttt{=} or \texttt{+}).
 | |
| 7050 | 219 | |
| 220 | \indexoutertoken{type}\indexoutertoken{term}\indexoutertoken{prop}
 | |
| 221 | \begin{rail}
 | |
| 7167 | 222 | type: nameref | typefree | typevar | 
| 7050 | 223 | ; | 
| 8593 | 224 | term: nameref | var | 
| 7050 | 225 | ; | 
| 7167 | 226 | prop: term | 
| 7050 | 227 | ; | 
| 228 | \end{rail}
 | |
| 229 | ||
| 8690 | 230 | Positional instantiations are indicated by giving a sequence of terms, or the | 
| 231 | placeholder ``$\_$'' (underscore), which means to skip a position. | |
| 232 | ||
| 233 | \indexoutertoken{inst}\indexoutertoken{insts}
 | |
| 234 | \begin{rail}
 | |
| 235 | inst: underscore | term | |
| 236 | ; | |
| 237 | insts: (inst *) | |
| 238 | ; | |
| 239 | \end{rail}
 | |
| 240 | ||
| 7167 | 241 | Type declarations and definitions usually refer to \railnonterm{typespec} on
 | 
| 242 | the left-hand side. This models basic type constructor application at the | |
| 243 | outer syntax level. Note that only plain postfix notation is available here, | |
| 244 | but no infixes. | |
| 7050 | 245 | |
| 246 | \indexouternonterm{typespec}
 | |
| 247 | \begin{rail}
 | |
| 7167 | 248 |   typespec: (() | typefree | '(' ( typefree + ',' ) ')') name
 | 
| 7050 | 249 | ; | 
| 250 | \end{rail}
 | |
| 251 | ||
| 252 | ||
| 253 | \subsection{Mixfix annotations}
 | |
| 254 | ||
| 7134 | 255 | Mixfix annotations specify concrete \emph{inner} syntax of Isabelle types and
 | 
| 12618 | 256 | terms.  Some commands such as $\TYPES$ (see \S\ref{sec:types-pure}) admit
 | 
| 257 | infixes only, while $\CONSTS$ (see \S\ref{sec:consts}) and
 | |
| 258 | $\isarkeyword{syntax}$ (see \S\ref{sec:syn-trans}) support the full range of
 | |
| 259 | general mixfixes and binders. | |
| 7046 | 260 | |
| 12976 | 261 | \indexouternonterm{infix}\indexouternonterm{mixfix}\indexouternonterm{structmixfix}
 | 
| 7046 | 262 | \begin{rail}
 | 
| 11651 | 263 |   infix: '(' ('infix' | 'infixl' | 'infixr') string? nat ')'
 | 
| 7167 | 264 | ; | 
| 7175 | 265 |   mixfix: infix | '(' string prios? nat? ')' | '(' 'binder' string prios? nat ')'
 | 
| 7050 | 266 | ; | 
| 12976 | 267 |   structmixfix: mixfix | '(' 'structure' ')'
 | 
| 268 | ; | |
| 7050 | 269 | |
| 7175 | 270 | prios: '[' (nat + ',') ']' | 
| 7050 | 271 | ; | 
| 7046 | 272 | \end{rail}
 | 
| 273 | ||
| 13048 | 274 | Here the \railtok{string} specifications refer to the actual mixfix template
 | 
| 12618 | 275 | (see also \cite{isabelle-ref}), which may include literal text, spacing,
 | 
| 276 | blocks, and arguments (denoted by ``$_$''); the special symbol \verb,\<index>, | |
| 277 | (printed as ``\i'') represents an index argument that specifies an implicit | |
| 278 | structure reference (see also \S\ref{sec:locale}).  Infix and binder
 | |
| 279 | declarations provide common abbreviations for particular mixfix declarations. | |
| 280 | So in practice, mixfix templates mostly degenerate to literal text for | |
| 281 | concrete syntax, such as ``\verb,++,'' for an infix symbol, or ``\verb,++,\i'' | |
| 282 | for an infix of an implicit structure. | |
| 283 | ||
| 284 | ||
| 285 | ||
| 286 | \subsection{Proof methods}\label{sec:syn-meth}
 | |
| 287 | ||
| 288 | Proof methods are either basic ones, or expressions composed of methods via | |
| 289 | ``\texttt{,}'' (sequential composition), ``\texttt{|}'' (alternative choices),
 | |
| 290 | ``\texttt{?}'' (try), ``\texttt{+}'' (repeat at least once).  In practice,
 | |
| 291 | proof methods are usually just a comma separated list of | |
| 13048 | 292 | \railqtok{nameref}~\railnonterm{args} specifications.  Note that parentheses
 | 
| 12618 | 293 | may be dropped for single method specifications (with no arguments). | 
| 294 | ||
| 295 | \indexouternonterm{method}
 | |
| 296 | \begin{rail}
 | |
| 297 |   method: (nameref | '(' methods ')') (() | '?' | '+')
 | |
| 298 | ; | |
| 299 |   methods: (nameref args | method) + (',' | '|')
 | |
| 300 | ; | |
| 301 | \end{rail}
 | |
| 302 | ||
| 303 | Proper use of Isar proof methods does \emph{not} involve goal addressing.
 | |
| 304 | Nevertheless, specifying goal ranges may occasionally come in handy in | |
| 305 | emulating tactic scripts. Note that $[n-]$ refers to all goals, starting from | |
| 306 | $n$. All goals may be specified by $[!]$, which is the same as $[1-]$. | |
| 307 | ||
| 308 | \indexouternonterm{goalspec}
 | |
| 309 | \begin{rail}
 | |
| 310 | goalspec: '[' (nat '-' nat | nat '-' | nat | '!' ) ']' | |
| 311 | ; | |
| 312 | \end{rail}
 | |
| 313 | ||
| 7050 | 314 | |
| 7134 | 315 | \subsection{Attributes and theorems}\label{sec:syn-att}
 | 
| 7050 | 316 | |
| 317 | Attributes (and proof methods, see \S\ref{sec:syn-meth}) have their own
 | |
| 7335 | 318 | ``semi-inner'' syntax, in the sense that input conforming to | 
| 319 | \railnonterm{args} below is parsed by the attribute a second time.  The
 | |
| 320 | attribute argument specifications may be any sequence of atomic entities | |
| 321 | (identifiers, strings etc.), or properly bracketed argument lists. Below | |
| 13048 | 322 | \railqtok{atom} refers to any atomic entity, including any \railtok{keyword}
 | 
| 323 | conforming to \railtok{symident}.
 | |
| 7050 | 324 | |
| 325 | \indexoutertoken{atom}\indexouternonterm{args}\indexouternonterm{attributes}
 | |
| 326 | \begin{rail}
 | |
| 7466 | 327 | atom: nameref | typefree | typevar | var | nat | keyword | 
| 7134 | 328 | ; | 
| 8896 | 329 |   arg: atom | '(' args ')' | '[' args ']'
 | 
| 7050 | 330 | ; | 
| 7167 | 331 | args: arg * | 
| 7134 | 332 | ; | 
| 7167 | 333 | attributes: '[' (nameref args * ',') ']' | 
| 7050 | 334 | ; | 
| 335 | \end{rail}
 | |
| 336 | ||
| 7895 | 337 | Theorem specifications come in several flavors: \railnonterm{axmdecl} and
 | 
| 7175 | 338 | \railnonterm{thmdecl} usually refer to axioms, assumptions or results of goal
 | 
| 7981 | 339 | statements, while \railnonterm{thmdef} collects lists of existing theorems.
 | 
| 340 | Existing theorems are given by \railnonterm{thmref} and \railnonterm{thmrefs},
 | |
| 15687 | 341 | the former requires an actual singleton result. An optional index selection | 
| 342 | specifies the individual theorems to be picked out of a given fact list. Any | |
| 343 | kind of theorem specification may include lists of attributes both on the left | |
| 344 | and right hand sides; attributes are applied to any immediately preceding | |
| 345 | fact. If names are omitted, the theorems are not stored within the theorem | |
| 346 | database of the theory or proof context, but any given attributes are applied | |
| 347 | nonetheless. | |
| 7050 | 348 | |
| 15687 | 349 | \indexouternonterm{axmdecl}\indexouternonterm{thmdecl}
 | 
| 350 | \indexouternonterm{thmdef}\indexouternonterm{thmref}
 | |
| 351 | \indexouternonterm{thmrefs}\indexouternonterm{selection}
 | |
| 7050 | 352 | \begin{rail}
 | 
| 7167 | 353 | axmdecl: name attributes? ':' | 
| 7050 | 354 | ; | 
| 9200 | 355 | thmdecl: thmbind ':' | 
| 7135 | 356 | ; | 
| 9200 | 357 | thmdef: thmbind '=' | 
| 7050 | 358 | ; | 
| 15687 | 359 | thmref: nameref selection? attributes? | 
| 7175 | 360 | ; | 
| 361 | thmrefs: thmref + | |
| 7134 | 362 | ; | 
| 7167 | 363 | |
| 9200 | 364 | thmbind: name attributes | name | attributes | 
| 7050 | 365 | ; | 
| 15687 | 366 |   selection: '(' ((nat | nat '-' nat?) + ',') ')'
 | 
| 367 | ; | |
| 7050 | 368 | \end{rail}
 | 
| 7046 | 369 | |
| 370 | ||
| 12618 | 371 | \subsection{Term patterns and declarations}\label{sec:term-decls}
 | 
| 7046 | 372 | |
| 12618 | 373 | Wherever explicit propositions (or term fragments) occur in a proof text, | 
| 374 | casual binding of schematic term variables may be given specified via patterns | |
| 13039 | 375 | of the form ``$\ISS{p@1\;\dots}{p@n}$''.  There are separate versions
 | 
| 13048 | 376 | available for \railqtok{term}s and \railqtok{prop}s.  The latter provides a
 | 
| 377 | $\CONCLNAME$ part with patterns referring the (atomic) conclusion of a rule. | |
| 7046 | 378 | |
| 12618 | 379 | \indexouternonterm{termpat}\indexouternonterm{proppat}
 | 
| 7050 | 380 | \begin{rail}
 | 
| 12618 | 381 |   termpat: '(' ('is' term +) ')'
 | 
| 7134 | 382 | ; | 
| 12618 | 383 |   proppat: '(' (('is' prop +) | 'concl' ('is' prop +) | ('is' prop +) 'concl' ('is' prop +)) ')'
 | 
| 7050 | 384 | ; | 
| 385 | \end{rail}
 | |
| 7046 | 386 | |
| 12618 | 387 | Declarations of local variables $x :: \tau$ and logical propositions $a : | 
| 388 | \phi$ represent different views on the same principle of introducing a local | |
| 389 | scope. In practice, one may usually omit the typing of $vars$ (due to | |
| 13039 | 390 | type-inference), and the naming of propositions (due to implicit references of | 
| 391 | current facts). In any case, Isar proof elements usually admit to introduce | |
| 12618 | 392 | multiple such items simultaneously. | 
| 8532 | 393 | |
| 12618 | 394 | \indexouternonterm{vars}\indexouternonterm{props}
 | 
| 8532 | 395 | \begin{rail}
 | 
| 12618 | 396 |   vars: (name+) ('::' type)?
 | 
| 397 | ; | |
| 398 | props: thmdecl? (prop proppat? +) | |
| 8532 | 399 | ; | 
| 400 | \end{rail}
 | |
| 401 | ||
| 12618 | 402 | The treatment of multiple declarations corresponds to the complementary focus | 
| 403 | of $vars$ versus $props$: in ``$x@1~\dots~x@n :: \tau$'' the typing refers to | |
| 404 | all variables, while in $a\colon \phi@1~\dots~\phi@n$ the naming refers to all | |
| 405 | propositions collectively. Isar language elements that refer to $vars$ or | |
| 406 | $props$ typically admit separate typings or namings via another level of | |
| 407 | iteration, with explicit $\AND$ separators; e.g.\ see $\FIXNAME$ and | |
| 408 | $\ASSUMENAME$ in \S\ref{sec:proof-context}.
 | |
| 409 | ||
| 7046 | 410 | |
| 9200 | 411 | \subsection{Antiquotations}\label{sec:antiq}
 | 
| 412 | ||
| 15357 | 413 | \begin{matharray}{rcl@{\hspace*{2cm}}rcl}
 | 
| 414 | thm & : & \isarantiq & text & : & \isarantiq \\ | |
| 415 | lhs & : & \isarantiq & goals & : & \isarantiq \\ | |
| 416 | rhs & : & \isarantiq & subgoals & : & \isarantiq \\ | |
| 417 | prop & : & \isarantiq & prf & : & \isarantiq \\ | |
| 418 | term & : & \isarantiq & full_prf & : & \isarantiq \\ | |
| 419 | typ & : & \isarantiq \\ | |
| 10336 | 420 | \end{matharray}
 | 
| 421 | ||
| 9200 | 422 | The text body of formal comments (see also \S\ref{sec:comments}) may contain
 | 
| 423 | antiquotations of logical entities, such as theorems, terms and types, which | |
| 424 | are to be presented in the final output produced by the Isabelle document | |
| 425 | preparation system (see also \S\ref{sec:document-prep}).
 | |
| 426 | ||
| 9601 | 427 | Thus embedding of | 
| 13039 | 428 | ``\texttt{{\at}{\ttlbrace}term~[show_types]~"f(x)~=~a~+~x"{\ttrbrace}}''
 | 
| 429 | within a text block would cause | |
| 9200 | 430 | \isa{(f{\isasymColon}'a~{\isasymRightarrow}~'a)~(x{\isasymColon}'a)~=~(a{\isasymColon}'a)~+~x}
 | 
| 10160 | 431 | to appear in the final {\LaTeX} document.  Also note that theorem
 | 
| 432 | antiquotations may involve attributes as well. For example, | |
| 433 | \texttt{{\at}{\ttlbrace}thm~sym~[no_vars]{\ttrbrace}} would print the
 | |
| 434 | statement where all schematic variables have been replaced by fixed ones, | |
| 12618 | 435 | which are easier to read. | 
| 9200 | 436 | |
| 15357 | 437 | \indexisarant{thm}\indexisarant{lhs}\indexisarant{rhs}\indexisarant{prop}\indexisarant{term}
 | 
| 10355 | 438 | \indexisarant{typ}\indexisarant{text}\indexisarant{goals}\indexisarant{subgoals}
 | 
| 9200 | 439 | \begin{rail}
 | 
| 440 | atsign lbrace antiquotation rbrace | |
| 441 | ; | |
| 442 | ||
| 443 | antiquotation: | |
| 444 | 'thm' options thmrefs | | |
| 15357 | 445 | 'lhs' options thmrefs | | 
| 446 | 'rhs' options thmrefs | | |
| 9200 | 447 | 'prop' options prop | | 
| 448 | 'term' options term | | |
| 9728 | 449 | 'typ' options type | | 
| 10319 
02463775cafb
added antiquotation "goals" and option "goals_limit";
 wenzelm parents: 
10160diff
changeset | 450 | 'text' options name | | 
| 10355 | 451 | 'goals' options | | 
| 13827 
c690cb885db4
Documented prf / full_prf commands and antiquotations.
 berghofe parents: 
13048diff
changeset | 452 | 'subgoals' options | | 
| 
c690cb885db4
Documented prf / full_prf commands and antiquotations.
 berghofe parents: 
13048diff
changeset | 453 | 'prf' options thmrefs | | 
| 
c690cb885db4
Documented prf / full_prf commands and antiquotations.
 berghofe parents: 
13048diff
changeset | 454 | 'full\_prf' options thmrefs | 
| 9200 | 455 | ; | 
| 456 | options: '[' (option * ',') ']' | |
| 457 | ; | |
| 458 | option: name | name '=' name | |
| 459 | ; | |
| 460 | \end{rail}
 | |
| 461 | ||
| 462 | Note that the syntax of antiquotations may \emph{not} include source comments
 | |
| 463 | \texttt{(*~\dots~*)} or verbatim text \verb|{*|~\dots~\verb|*}|.
 | |
| 464 | ||
| 10319 
02463775cafb
added antiquotation "goals" and option "goals_limit";
 wenzelm parents: 
10160diff
changeset | 465 | \begin{descr}
 | 
| 14895 | 466 | |
| 10336 | 467 | \item [$\at\{thm~\vec a\}$] prints theorems $\vec a$. Note that attribute
 | 
| 468 |   specifications may be included as well (see also \S\ref{sec:syn-att}); the
 | |
| 12618 | 469 |   $no_vars$ operation (see \S\ref{sec:misc-meth-att}) would be particularly
 | 
| 10336 | 470 | useful to suppress printing of schematic variables. | 
| 13039 | 471 | |
| 15357 | 472 | \item [$\at\{lhs~\vec a\}$] prints the left hand sides of theorems $\vec a$.
 | 
| 473 | The antiquotation only works for $\vec a$ that are definitions, | |
| 474 | equations or other binary operators. It understands the same | |
| 475 |   options and attributes as $\at\{thm~\vec a\}$.
 | |
| 476 | ||
| 477 | \item [$\at\{rhs~\vec a\}$] prints the right hand sides of theorems $\vec a$.
 | |
| 478 |   As for $\at\{lhs~\vec a\}$, $\vec a$ must be definitions, equations or
 | |
| 479 | other binary operators. | |
| 480 | ||
| 10336 | 481 | \item [$\at\{prop~\phi\}$] prints a well-typed proposition $\phi$.
 | 
| 13039 | 482 | |
| 10336 | 483 | \item [$\at\{term~t\}$] prints a well-typed term $t$.
 | 
| 13039 | 484 | |
| 10336 | 485 | \item [$\at\{typ~\tau\}$] prints a well-formed type $\tau$.
 | 
| 14895 | 486 | |
| 10336 | 487 | \item [$\at\{text~s\}$] prints uninterpreted source text $s$.  This is
 | 
| 488 | particularly useful to print portions of text according to the Isabelle | |
| 489 |   {\LaTeX} output style, without demanding well-formedness (e.g.\ small pieces
 | |
| 13039 | 490 | of terms that should not be parsed or type-checked yet). | 
| 14895 | 491 | |
| 10336 | 492 | \item [$\at\{goals\}$] prints the current \emph{dynamic} goal state.  This is
 | 
| 13039 | 493 | mainly for support of tactic-emulation scripts within Isar --- presentation | 
| 494 | of goal states does not conform to actual human-readable proof documents. | |
| 10319 
02463775cafb
added antiquotation "goals" and option "goals_limit";
 wenzelm parents: 
10160diff
changeset | 495 | Please do not include goal states into document output unless you really | 
| 
02463775cafb
added antiquotation "goals" and option "goals_limit";
 wenzelm parents: 
10160diff
changeset | 496 | know what you are doing! | 
| 13039 | 497 | |
| 10355 | 498 | \item [$\at\{subgoals\}$] behaves almost like $goals$, except that it does not
 | 
| 499 | print the main goal. | |
| 13039 | 500 | |
| 13827 
c690cb885db4
Documented prf / full_prf commands and antiquotations.
 berghofe parents: 
13048diff
changeset | 501 | \item [$\at\{prf~\vec a\}$] prints the (compact) proof terms corresponding to
 | 
| 
c690cb885db4
Documented prf / full_prf commands and antiquotations.
 berghofe parents: 
13048diff
changeset | 502 | the theorems $\vec a$. Note that this | 
| 
c690cb885db4
Documented prf / full_prf commands and antiquotations.
 berghofe parents: 
13048diff
changeset | 503 | requires proof terms to be switched on for the current object logic | 
| 
c690cb885db4
Documented prf / full_prf commands and antiquotations.
 berghofe parents: 
13048diff
changeset | 504 | (see the ``Proof terms'' section of the Isabelle reference manual | 
| 
c690cb885db4
Documented prf / full_prf commands and antiquotations.
 berghofe parents: 
13048diff
changeset | 505 | for information on how to do this). | 
| 
c690cb885db4
Documented prf / full_prf commands and antiquotations.
 berghofe parents: 
13048diff
changeset | 506 | |
| 
c690cb885db4
Documented prf / full_prf commands and antiquotations.
 berghofe parents: 
13048diff
changeset | 507 | \item [$\at\{full_prf~\vec a\}$] is like $\at\{prf~\vec a\}$, but displays
 | 
| 
c690cb885db4
Documented prf / full_prf commands and antiquotations.
 berghofe parents: 
13048diff
changeset | 508 | the full proof terms, i.e.\ also displays information omitted in | 
| 
c690cb885db4
Documented prf / full_prf commands and antiquotations.
 berghofe parents: 
13048diff
changeset | 509 | the compact proof term, which is denoted by ``$_$'' placeholders there. | 
| 
c690cb885db4
Documented prf / full_prf commands and antiquotations.
 berghofe parents: 
13048diff
changeset | 510 | |
| 10319 
02463775cafb
added antiquotation "goals" and option "goals_limit";
 wenzelm parents: 
10160diff
changeset | 511 | \end{descr}
 | 
| 
02463775cafb
added antiquotation "goals" and option "goals_limit";
 wenzelm parents: 
10160diff
changeset | 512 | |
| 9200 | 513 | \medskip | 
| 514 | ||
| 10336 | 515 | The following options are available to tune the output. Note that most of | 
| 9233 | 516 | these coincide with ML flags of the same names (see also \cite{isabelle-ref}).
 | 
| 9200 | 517 | \begin{descr}
 | 
| 9233 | 518 | \item[$show_types = bool$ and $show_sorts = bool$] control printing of | 
| 9234 | 519 | explicit type and sort constraints. | 
| 14707 | 520 | \item[$show_structs = bool$] controls printing of implicit structures. | 
| 9233 | 521 | \item[$long_names = bool$] forces names of types and constants etc.\ to be | 
| 522 | printed in their fully qualified internal form. | |
| 523 | \item[$eta_contract = bool$] prints terms in $\eta$-contracted form. | |
| 9200 | 524 | \item[$display = bool$] indicates if the text is to be output as multi-line | 
| 525 | ``display material'', rather than a small piece of text without line breaks | |
| 526 | (which is the default). | |
| 14689 | 527 | \item[$breaks = bool$] controls line breaks in non-display material. | 
| 9200 | 528 | \item[$quotes = bool$] indicates if the output should be enclosed in double | 
| 529 | quotes. | |
| 9233 | 530 | \item[$mode = name$] adds $name$ to the print mode to be used for presentation | 
| 531 |   (see also \cite{isabelle-ref}).  Note that the standard setup for {\LaTeX}
 | |
| 532 | output is already present by default, including the modes ``$latex$'', | |
| 533 | ``$xsymbols$'', ``$symbols$''. | |
| 9728 | 534 | \item[$margin = nat$ and $indent = nat$] change the margin or indentation for | 
| 535 | pretty printing of display material. | |
| 9752 | 536 | \item[$source = bool$] prints the source text of the antiquotation arguments, | 
| 537 | rather than the actual value. Note that this does not affect | |
| 538 | well-formedness checks of $thm$, $term$, etc. (only the $text$ antiquotation | |
| 539 | admits arbitrary output). | |
| 10319 
02463775cafb
added antiquotation "goals" and option "goals_limit";
 wenzelm parents: 
10160diff
changeset | 540 | \item[$goals_limit = nat$] determines the maximum number of goals to be | 
| 
02463775cafb
added antiquotation "goals" and option "goals_limit";
 wenzelm parents: 
10160diff
changeset | 541 | printed. | 
| 14895 | 542 | \item[$locale = name$] specifies an alternative context used for evaluating | 
| 14919 | 543 | and printing the subsequent argument. | 
| 9200 | 544 | \end{descr}
 | 
| 545 | ||
| 546 | For boolean flags, ``$name = true$'' may be abbreviated as ``$name$''. All of | |
| 547 | the above flags are disabled by default, unless changed from ML. | |
| 548 | ||
| 10336 | 549 | \medskip Note that antiquotations do not only spare the author from tedious | 
| 13039 | 550 | typing of logical entities, but also achieve some degree of | 
| 551 | consistency-checking of informal explanations with formal developments: | |
| 552 | well-formedness of terms and types with respect to the current theory or proof | |
| 553 | context is ensured here. | |
| 9200 | 554 | |
| 14895 | 555 | %%% Local Variables: | 
| 7046 | 556 | %%% mode: latex | 
| 557 | %%% TeX-master: "isar-ref" | |
| 14895 | 558 | %%% End: |