| 26767 |      1 | (* $Id$ *)
 | 
| 26754 |      2 | 
 | 
|  |      3 | theory "syntax"
 | 
|  |      4 | imports CPure
 | 
|  |      5 | begin
 | 
|  |      6 | 
 | 
|  |      7 | chapter {* Syntax primitives *}
 | 
|  |      8 | 
 | 
|  |      9 | text {*
 | 
|  |     10 |   The rather generic framework of Isabelle/Isar syntax emerges from
 | 
|  |     11 |   three main syntactic categories: \emph{commands} of the top-level
 | 
|  |     12 |   Isar engine (covering theory and proof elements), \emph{methods} for
 | 
|  |     13 |   general goal refinements (analogous to traditional ``tactics''), and
 | 
|  |     14 |   \emph{attributes} for operations on facts (within a certain
 | 
|  |     15 |   context).  Subsequently we give a reference of basic syntactic
 | 
|  |     16 |   entities underlying Isabelle/Isar syntax in a bottom-up manner.
 | 
|  |     17 |   Concrete theory and proof language elements will be introduced later
 | 
|  |     18 |   on.
 | 
|  |     19 | 
 | 
|  |     20 |   \medskip In order to get started with writing well-formed
 | 
|  |     21 |   Isabelle/Isar documents, the most important aspect to be noted is
 | 
|  |     22 |   the difference of \emph{inner} versus \emph{outer} syntax.  Inner
 | 
|  |     23 |   syntax is that of Isabelle types and terms of the logic, while outer
 | 
|  |     24 |   syntax is that of Isabelle/Isar theory sources (specifications and
 | 
|  |     25 |   proofs).  As a general rule, inner syntax entities may occur only as
 | 
|  |     26 |   \emph{atomic entities} within outer syntax.  For example, the string
 | 
| 26760 |     27 |   @{verbatim "\"x + y\""} and identifier @{verbatim z} are legal term
 | 
|  |     28 |   specifications within a theory, while @{verbatim "x + y"} without
 | 
|  |     29 |   quotes is not.
 | 
| 26754 |     30 | 
 | 
|  |     31 |   Printed theory documents usually omit quotes to gain readability
 | 
| 26760 |     32 |   (this is a matter of {\LaTeX} macro setup, say via @{verbatim
 | 
|  |     33 |   "\\isabellestyle"}, see also \cite{isabelle-sys}).  Experienced
 | 
| 26754 |     34 |   users of Isabelle/Isar may easily reconstruct the lost technical
 | 
|  |     35 |   information, while mere readers need not care about quotes at all.
 | 
|  |     36 | 
 | 
|  |     37 |   \medskip Isabelle/Isar input may contain any number of input
 | 
| 26760 |     38 |   termination characters ``@{verbatim ";"}'' (semicolon) to separate
 | 
| 26754 |     39 |   commands explicitly.  This is particularly useful in interactive
 | 
|  |     40 |   shell sessions to make clear where the current command is intended
 | 
|  |     41 |   to end.  Otherwise, the interpreter loop will continue to issue a
 | 
| 26760 |     42 |   secondary prompt ``@{verbatim "#"}'' until an end-of-command is
 | 
|  |     43 |   clearly recognized from the input syntax, e.g.\ encounter of the
 | 
|  |     44 |   next command keyword.
 | 
| 26754 |     45 | 
 | 
|  |     46 |   More advanced interfaces such as Proof~General \cite{proofgeneral}
 | 
|  |     47 |   do not require explicit semicolons, the amount of input text is
 | 
|  |     48 |   determined automatically by inspecting the present content of the
 | 
|  |     49 |   Emacs text buffer.  In the printed presentation of Isabelle/Isar
 | 
|  |     50 |   documents semicolons are omitted altogether for readability.
 | 
|  |     51 | 
 | 
|  |     52 |   \begin{warn}
 | 
|  |     53 |     Proof~General requires certain syntax classification tables in
 | 
|  |     54 |     order to achieve properly synchronized interaction with the
 | 
|  |     55 |     Isabelle/Isar process.  These tables need to be consistent with
 | 
|  |     56 |     the Isabelle version and particular logic image to be used in a
 | 
|  |     57 |     running session (common object-logics may well change the outer
 | 
|  |     58 |     syntax).  The standard setup should work correctly with any of the
 | 
|  |     59 |     ``official'' logic images derived from Isabelle/HOL (including
 | 
|  |     60 |     HOLCF etc.).  Users of alternative logics may need to tell
 | 
| 26760 |     61 |     Proof~General explicitly, e.g.\ by giving an option @{verbatim "-k ZF"}
 | 
|  |     62 |     (in conjunction with @{verbatim "-l ZF"}, to specify the default
 | 
|  |     63 |     logic image).  Note that option @{verbatim "-L"} does both
 | 
|  |     64 |     of this at the same time.
 | 
| 26754 |     65 |   \end{warn}
 | 
|  |     66 | *}
 | 
|  |     67 | 
 | 
|  |     68 | 
 | 
|  |     69 | section {* Lexical matters \label{sec:lex-syntax} *}
 | 
|  |     70 | 
 | 
|  |     71 | text {*
 | 
|  |     72 |   The Isabelle/Isar outer syntax provides token classes as presented
 | 
|  |     73 |   below; most of these coincide with the inner lexical syntax as
 | 
|  |     74 |   presented in \cite{isabelle-ref}.
 | 
|  |     75 | 
 | 
|  |     76 |   \begin{matharray}{rcl}
 | 
|  |     77 |     @{syntax_def ident} & = & letter\,quasiletter^* \\
 | 
|  |     78 |     @{syntax_def longident} & = & ident (\verb,.,ident)^+ \\
 | 
|  |     79 |     @{syntax_def symident} & = & sym^+ ~|~ \verb,\,\verb,<,ident\verb,>, \\
 | 
|  |     80 |     @{syntax_def nat} & = & digit^+ \\
 | 
|  |     81 |     @{syntax_def var} & = & ident ~|~ \verb,?,ident ~|~ \verb,?,ident\verb,.,nat \\
 | 
|  |     82 |     @{syntax_def typefree} & = & \verb,',ident \\
 | 
|  |     83 |     @{syntax_def typevar} & = & typefree ~|~ \verb,?,typefree ~|~ \verb,?,typefree\verb,.,nat \\
 | 
|  |     84 |     @{syntax_def string} & = & \verb,", ~\dots~ \verb,", \\
 | 
|  |     85 |     @{syntax_def altstring} & = & \backquote ~\dots~ \backquote \\
 | 
|  |     86 |     @{syntax_def verbatim} & = & \verb,{*, ~\dots~ \verb,*,\verb,}, \\[1ex]
 | 
|  |     87 | 
 | 
|  |     88 |     letter & = & latin ~|~ \verb,\,\verb,<,latin\verb,>, ~|~ \verb,\,\verb,<,latin\,latin\verb,>, ~|~ greek ~|~ \\
 | 
|  |     89 |            &   & \verb,\<^isub>, ~|~ \verb,\<^isup>, \\
 | 
|  |     90 |     quasiletter & = & letter ~|~ digit ~|~ \verb,_, ~|~ \verb,', \\
 | 
|  |     91 |     latin & = & \verb,a, ~|~ \dots ~|~ \verb,z, ~|~ \verb,A, ~|~ \dots ~|~ \verb,Z, \\
 | 
|  |     92 |     digit & = & \verb,0, ~|~ \dots ~|~ \verb,9, \\
 | 
|  |     93 |     sym & = & \verb,!, ~|~ \verb,#, ~|~ \verb,$, ~|~ \verb,%, ~|~ \verb,&, ~|~
 | 
|  |     94 |      \verb,*, ~|~ \verb,+, ~|~ \verb,-, ~|~ \verb,/, ~|~ \\
 | 
|  |     95 |     & & \verb,<, ~|~ \verb,=, ~|~ \verb,>, ~|~ \verb,?, ~|~ \texttt{\at} ~|~
 | 
|  |     96 |     \verb,^, ~|~ \verb,_, ~|~ \verb,|, ~|~ \verb,~, \\
 | 
|  |     97 |     greek & = & \verb,\<alpha>, ~|~ \verb,\<beta>, ~|~ \verb,\<gamma>, ~|~ \verb,\<delta>, ~| \\
 | 
|  |     98 |           &   & \verb,\<epsilon>, ~|~ \verb,\<zeta>, ~|~ \verb,\<eta>, ~|~ \verb,\<theta>, ~| \\
 | 
|  |     99 |           &   & \verb,\<iota>, ~|~ \verb,\<kappa>, ~|~ \verb,\<mu>, ~|~ \verb,\<nu>, ~| \\
 | 
|  |    100 |           &   & \verb,\<xi>, ~|~ \verb,\<pi>, ~|~ \verb,\<rho>, ~|~ \verb,\<sigma>, ~|~ \verb,\<tau>, ~| \\
 | 
|  |    101 |           &   & \verb,\<upsilon>, ~|~ \verb,\<phi>, ~|~ \verb,\<chi>, ~|~ \verb,\<psi>, ~| \\
 | 
|  |    102 |           &   & \verb,\<omega>, ~|~ \verb,\<Gamma>, ~|~ \verb,\<Delta>, ~|~ \verb,\<Theta>, ~| \\
 | 
|  |    103 |           &   & \verb,\<Lambda>, ~|~ \verb,\<Xi>, ~|~ \verb,\<Pi>, ~|~ \verb,\<Sigma>, ~| \\
 | 
|  |    104 |           &   & \verb,\<Upsilon>, ~|~ \verb,\<Phi>, ~|~ \verb,\<Psi>, ~|~ \verb,\<Omega>, \\
 | 
|  |    105 |   \end{matharray}
 | 
|  |    106 | 
 | 
|  |    107 |   The syntax of @{syntax string} admits any characters, including
 | 
| 26760 |    108 |   newlines; ``@{verbatim "\""}'' (double-quote) and ``@{verbatim
 | 
|  |    109 |   "\\"}'' (backslash) need to be escaped by a backslash; arbitrary
 | 
|  |    110 |   character codes may be specified as ``@{verbatim "\\"}@{text ddd}'',
 | 
|  |    111 |   with three decimal digits.  Alternative strings according to
 | 
|  |    112 |   @{syntax altstring} are analogous, using single back-quotes instead.
 | 
|  |    113 |   The body of @{syntax verbatim} may consist of any text not
 | 
|  |    114 |   containing ``@{verbatim "*"}@{verbatim "}"}''; this allows
 | 
|  |    115 |   convenient inclusion of quotes without further escapes.  The greek
 | 
|  |    116 |   letters do \emph{not} include @{verbatim "\<lambda>"}, which is already used
 | 
|  |    117 |   differently in the meta-logic.
 | 
| 26754 |    118 | 
 | 
|  |    119 |   Common mathematical symbols such as @{text \<forall>} are represented in
 | 
| 26760 |    120 |   Isabelle as @{verbatim \<forall>}.  There are infinitely many Isabelle
 | 
|  |    121 |   symbols like this, although proper presentation is left to front-end
 | 
|  |    122 |   tools such as {\LaTeX} or Proof~General with the X-Symbol package.
 | 
|  |    123 |   A list of standard Isabelle symbols that work well with these tools
 | 
|  |    124 |   is given in \cite[appendix~A]{isabelle-sys}.
 | 
| 26754 |    125 |   
 | 
| 26760 |    126 |   Source comments take the form @{verbatim "(*"}~@{text
 | 
|  |    127 |   "\<dots>"}~@{verbatim "*)"} and may be nested, although user-interface
 | 
|  |    128 |   tools might prevent this.  Note that this form indicates source
 | 
|  |    129 |   comments only, which are stripped after lexical analysis of the
 | 
|  |    130 |   input.  The Isar document syntax also provides formal comments that
 | 
|  |    131 |   are considered as part of the text (see \secref{sec:comments}).
 | 
| 26754 |    132 | *}
 | 
|  |    133 | 
 | 
|  |    134 | 
 | 
|  |    135 | section {* Common syntax entities *}
 | 
|  |    136 | 
 | 
|  |    137 | text {*
 | 
|  |    138 |   We now introduce several basic syntactic entities, such as names,
 | 
|  |    139 |   terms, and theorem specifications, which are factored out of the
 | 
|  |    140 |   actual Isar language elements to be described later.
 | 
|  |    141 | *}
 | 
|  |    142 | 
 | 
|  |    143 | 
 | 
|  |    144 | subsection {* Names *}
 | 
|  |    145 | 
 | 
|  |    146 | text {*
 | 
|  |    147 |   Entity \railqtok{name} usually refers to any name of types,
 | 
|  |    148 |   constants, theorems etc.\ that are to be \emph{declared} or
 | 
|  |    149 |   \emph{defined} (so qualified identifiers are excluded here).  Quoted
 | 
|  |    150 |   strings provide an escape for non-identifier names or those ruled
 | 
| 26760 |    151 |   out by outer syntax keywords (e.g.\ quoted @{verbatim "\"let\""}).
 | 
|  |    152 |   Already existing objects are usually referenced by
 | 
|  |    153 |   \railqtok{nameref}.
 | 
| 26754 |    154 | 
 | 
|  |    155 |   \indexoutertoken{name}\indexoutertoken{parname}\indexoutertoken{nameref}
 | 
|  |    156 |   \indexoutertoken{int}
 | 
|  |    157 |   \begin{rail}
 | 
|  |    158 |     name: ident | symident | string | nat
 | 
|  |    159 |     ;
 | 
|  |    160 |     parname: '(' name ')'
 | 
|  |    161 |     ;
 | 
|  |    162 |     nameref: name | longident
 | 
|  |    163 |     ;
 | 
|  |    164 |     int: nat | '-' nat
 | 
|  |    165 |     ;
 | 
|  |    166 |   \end{rail}
 | 
|  |    167 | *}
 | 
|  |    168 | 
 | 
|  |    169 | 
 | 
|  |    170 | subsection {* Comments \label{sec:comments} *}
 | 
|  |    171 | 
 | 
|  |    172 | text {*
 | 
|  |    173 |   Large chunks of plain \railqtok{text} are usually given
 | 
| 26760 |    174 |   \railtok{verbatim}, i.e.\ enclosed in @{verbatim "{"}@{verbatim
 | 
|  |    175 |   "*"}~@{text "\<dots>"}~@{verbatim "*"}@{verbatim "}"}.  For convenience,
 | 
|  |    176 |   any of the smaller text units conforming to \railqtok{nameref} are
 | 
|  |    177 |   admitted as well.  A marginal \railnonterm{comment} is of the form
 | 
|  |    178 |   @{verbatim "--"} \railqtok{text}.  Any number of these may occur
 | 
|  |    179 |   within Isabelle/Isar commands.
 | 
| 26754 |    180 | 
 | 
|  |    181 |   \indexoutertoken{text}\indexouternonterm{comment}
 | 
|  |    182 |   \begin{rail}
 | 
|  |    183 |     text: verbatim | nameref
 | 
|  |    184 |     ;
 | 
|  |    185 |     comment: '--' text
 | 
|  |    186 |     ;
 | 
|  |    187 |   \end{rail}
 | 
|  |    188 | *}
 | 
|  |    189 | 
 | 
|  |    190 | 
 | 
|  |    191 | subsection {* Type classes, sorts and arities *}
 | 
|  |    192 | 
 | 
|  |    193 | text {*
 | 
|  |    194 |   Classes are specified by plain names.  Sorts have a very simple
 | 
|  |    195 |   inner syntax, which is either a single class name @{text c} or a
 | 
|  |    196 |   list @{text "{c\<^sub>1, \<dots>, c\<^sub>n}"} referring to the
 | 
|  |    197 |   intersection of these classes.  The syntax of type arities is given
 | 
|  |    198 |   directly at the outer level.
 | 
|  |    199 | 
 | 
|  |    200 |   \indexouternonterm{sort}\indexouternonterm{arity}
 | 
|  |    201 |   \indexouternonterm{classdecl}
 | 
|  |    202 |   \begin{rail}
 | 
|  |    203 |     classdecl: name (('<' | subseteq) (nameref + ','))?
 | 
|  |    204 |     ;
 | 
|  |    205 |     sort: nameref
 | 
|  |    206 |     ;
 | 
|  |    207 |     arity: ('(' (sort + ',') ')')? sort
 | 
|  |    208 |     ;
 | 
|  |    209 |   \end{rail}
 | 
|  |    210 | *}
 | 
|  |    211 | 
 | 
|  |    212 | 
 | 
|  |    213 | subsection {* Types and terms \label{sec:types-terms} *}
 | 
|  |    214 | 
 | 
|  |    215 | text {*
 | 
|  |    216 |   The actual inner Isabelle syntax, that of types and terms of the
 | 
|  |    217 |   logic, is far too sophisticated in order to be modelled explicitly
 | 
|  |    218 |   at the outer theory level.  Basically, any such entity has to be
 | 
|  |    219 |   quoted to turn it into a single token (the parsing and type-checking
 | 
|  |    220 |   is performed internally later).  For convenience, a slightly more
 | 
|  |    221 |   liberal convention is adopted: quotes may be omitted for any type or
 | 
|  |    222 |   term that is already atomic at the outer level.  For example, one
 | 
| 26760 |    223 |   may just write @{verbatim x} instead of quoted @{verbatim "\"x\""}.
 | 
|  |    224 |   Note that symbolic identifiers (e.g.\ @{verbatim "++"} or @{text
 | 
|  |    225 |   "\<forall>"} are available as well, provided these have not been superseded
 | 
|  |    226 |   by commands or other keywords already (such as @{verbatim "="} or
 | 
|  |    227 |   @{verbatim "+"}).
 | 
| 26754 |    228 | 
 | 
|  |    229 |   \indexoutertoken{type}\indexoutertoken{term}\indexoutertoken{prop}
 | 
|  |    230 |   \begin{rail}
 | 
|  |    231 |     type: nameref | typefree | typevar
 | 
|  |    232 |     ;
 | 
|  |    233 |     term: nameref | var
 | 
|  |    234 |     ;
 | 
|  |    235 |     prop: term
 | 
|  |    236 |     ;
 | 
|  |    237 |   \end{rail}
 | 
|  |    238 | 
 | 
|  |    239 |   Positional instantiations are indicated by giving a sequence of
 | 
| 26777 |    240 |   terms, or the placeholder ``@{text _}'' (underscore), which means to
 | 
|  |    241 |   skip a position.
 | 
| 26754 |    242 | 
 | 
|  |    243 |   \indexoutertoken{inst}\indexoutertoken{insts}
 | 
|  |    244 |   \begin{rail}
 | 
|  |    245 |     inst: underscore | term
 | 
|  |    246 |     ;
 | 
|  |    247 |     insts: (inst *)
 | 
|  |    248 |     ;
 | 
|  |    249 |   \end{rail}
 | 
|  |    250 | 
 | 
|  |    251 |   Type declarations and definitions usually refer to
 | 
|  |    252 |   \railnonterm{typespec} on the left-hand side.  This models basic
 | 
|  |    253 |   type constructor application at the outer syntax level.  Note that
 | 
|  |    254 |   only plain postfix notation is available here, but no infixes.
 | 
|  |    255 | 
 | 
|  |    256 |   \indexouternonterm{typespec}
 | 
|  |    257 |   \begin{rail}
 | 
|  |    258 |     typespec: (() | typefree | '(' ( typefree + ',' ) ')') name
 | 
|  |    259 |     ;
 | 
|  |    260 |   \end{rail}
 | 
|  |    261 | *}
 | 
|  |    262 | 
 | 
|  |    263 | 
 | 
|  |    264 | subsection {* Mixfix annotations *}
 | 
|  |    265 | 
 | 
|  |    266 | text {*
 | 
|  |    267 |   Mixfix annotations specify concrete \emph{inner} syntax of Isabelle
 | 
|  |    268 |   types and terms.  Some commands such as @{command "types"} (see
 | 
| 26760 |    269 |   \secref{sec:types-pure}) admit infixes only, while @{command
 | 
|  |    270 |   "consts"} (see \secref{sec:consts}) and @{command "syntax"} (see
 | 
|  |    271 |   \secref{sec:syn-trans}) support the full range of general mixfixes
 | 
| 26754 |    272 |   and binders.
 | 
|  |    273 | 
 | 
|  |    274 |   \indexouternonterm{infix}\indexouternonterm{mixfix}\indexouternonterm{structmixfix}
 | 
|  |    275 |   \begin{rail}
 | 
|  |    276 |     infix: '(' ('infix' | 'infixl' | 'infixr') string? nat ')'
 | 
|  |    277 |     ;
 | 
|  |    278 |     mixfix: infix | '(' string prios? nat? ')' | '(' 'binder' string prios? nat ')'
 | 
|  |    279 |     ;
 | 
|  |    280 |     structmixfix: mixfix | '(' 'structure' ')'
 | 
|  |    281 |     ;
 | 
|  |    282 | 
 | 
|  |    283 |     prios: '[' (nat + ',') ']'
 | 
|  |    284 |     ;
 | 
|  |    285 |   \end{rail}
 | 
|  |    286 | 
 | 
|  |    287 |   Here the \railtok{string} specifications refer to the actual mixfix
 | 
|  |    288 |   template (see also \cite{isabelle-ref}), which may include literal
 | 
| 26777 |    289 |   text, spacing, blocks, and arguments (denoted by ``@{text _}''); the
 | 
|  |    290 |   special symbol ``@{verbatim "\<index>"}'' (printed as ``@{text "\<index>"}'')
 | 
| 26760 |    291 |   represents an index argument that specifies an implicit structure
 | 
|  |    292 |   reference (see also \secref{sec:locale}).  Infix and binder
 | 
|  |    293 |   declarations provide common abbreviations for particular mixfix
 | 
|  |    294 |   declarations.  So in practice, mixfix templates mostly degenerate to
 | 
|  |    295 |   literal text for concrete syntax, such as ``@{verbatim "++"}'' for
 | 
|  |    296 |   an infix symbol, or ``@{verbatim "++"}@{text "\<index>"}'' for an infix of
 | 
|  |    297 |   an implicit structure.
 | 
| 26754 |    298 | *}
 | 
|  |    299 | 
 | 
|  |    300 | 
 | 
|  |    301 | subsection {* Proof methods \label{sec:syn-meth} *}
 | 
|  |    302 | 
 | 
|  |    303 | text {*
 | 
|  |    304 |   Proof methods are either basic ones, or expressions composed of
 | 
| 26760 |    305 |   methods via ``@{verbatim ","}'' (sequential composition),
 | 
|  |    306 |   ``@{verbatim "|"}'' (alternative choices), ``@{verbatim "?"}'' 
 | 
|  |    307 |   (try), ``@{verbatim "+"}'' (repeat at least once), ``@{verbatim
 | 
|  |    308 |   "["}@{text n}@{verbatim "]"}'' (restriction to first @{text n}
 | 
|  |    309 |   sub-goals, with default @{text "n = 1"}).  In practice, proof
 | 
|  |    310 |   methods are usually just a comma separated list of
 | 
|  |    311 |   \railqtok{nameref}~\railnonterm{args} specifications.  Note that
 | 
|  |    312 |   parentheses may be dropped for single method specifications (with no
 | 
|  |    313 |   arguments).
 | 
| 26754 |    314 | 
 | 
|  |    315 |   \indexouternonterm{method}
 | 
|  |    316 |   \begin{rail}
 | 
|  |    317 |     method: (nameref | '(' methods ')') (() | '?' | '+' | '[' nat? ']')
 | 
|  |    318 |     ;
 | 
|  |    319 |     methods: (nameref args | method) + (',' | '|')
 | 
|  |    320 |     ;
 | 
|  |    321 |   \end{rail}
 | 
|  |    322 | 
 | 
|  |    323 |   Proper Isar proof methods do \emph{not} admit arbitrary goal
 | 
|  |    324 |   addressing, but refer either to the first sub-goal or all sub-goals
 | 
| 26760 |    325 |   uniformly.  The goal restriction operator ``@{text "[n]"}''
 | 
| 26754 |    326 |   evaluates a method expression within a sandbox consisting of the
 | 
| 26760 |    327 |   first @{text n} sub-goals (which need to exist).  For example, the
 | 
|  |    328 |   method ``@{text "simp_all[3]"}'' simplifies the first three
 | 
|  |    329 |   sub-goals, while ``@{text "(rule foo, simp_all)[]"}'' simplifies all
 | 
|  |    330 |   new goals that emerge from applying rule @{text "foo"} to the
 | 
|  |    331 |   originally first one.
 | 
| 26754 |    332 | 
 | 
|  |    333 |   Improper methods, notably tactic emulations, offer a separate
 | 
|  |    334 |   low-level goal addressing scheme as explicit argument to the
 | 
| 26760 |    335 |   individual tactic being involved.  Here ``@{text "[!]"}'' refers to
 | 
|  |    336 |   all goals, and ``@{text "[n-]"}'' to all goals starting from @{text
 | 
|  |    337 |   "n"}.
 | 
| 26754 |    338 | 
 | 
|  |    339 |   \indexouternonterm{goalspec}
 | 
|  |    340 |   \begin{rail}
 | 
|  |    341 |     goalspec: '[' (nat '-' nat | nat '-' | nat | '!' ) ']'
 | 
|  |    342 |     ;
 | 
|  |    343 |   \end{rail}
 | 
|  |    344 | *}
 | 
|  |    345 | 
 | 
|  |    346 | 
 | 
|  |    347 | subsection {* Attributes and theorems \label{sec:syn-att} *}
 | 
|  |    348 | 
 | 
|  |    349 | text {*
 | 
| 26760 |    350 |   Attributes (and proof methods, see \secref{sec:syn-meth}) have their
 | 
| 26754 |    351 |   own ``semi-inner'' syntax, in the sense that input conforming to
 | 
|  |    352 |   \railnonterm{args} below is parsed by the attribute a second time.
 | 
|  |    353 |   The attribute argument specifications may be any sequence of atomic
 | 
|  |    354 |   entities (identifiers, strings etc.), or properly bracketed argument
 | 
|  |    355 |   lists.  Below \railqtok{atom} refers to any atomic entity, including
 | 
|  |    356 |   any \railtok{keyword} conforming to \railtok{symident}.
 | 
|  |    357 | 
 | 
|  |    358 |   \indexoutertoken{atom}\indexouternonterm{args}\indexouternonterm{attributes}
 | 
|  |    359 |   \begin{rail}
 | 
|  |    360 |     atom: nameref | typefree | typevar | var | nat | keyword
 | 
|  |    361 |     ;
 | 
|  |    362 |     arg: atom | '(' args ')' | '[' args ']'
 | 
|  |    363 |     ;
 | 
|  |    364 |     args: arg *
 | 
|  |    365 |     ;
 | 
|  |    366 |     attributes: '[' (nameref args * ',') ']'
 | 
|  |    367 |     ;
 | 
|  |    368 |   \end{rail}
 | 
|  |    369 | 
 | 
|  |    370 |   Theorem specifications come in several flavors:
 | 
|  |    371 |   \railnonterm{axmdecl} and \railnonterm{thmdecl} usually refer to
 | 
|  |    372 |   axioms, assumptions or results of goal statements, while
 | 
|  |    373 |   \railnonterm{thmdef} collects lists of existing theorems.  Existing
 | 
|  |    374 |   theorems are given by \railnonterm{thmref} and
 | 
|  |    375 |   \railnonterm{thmrefs}, the former requires an actual singleton
 | 
|  |    376 |   result.
 | 
|  |    377 | 
 | 
|  |    378 |   There are three forms of theorem references:
 | 
|  |    379 |   \begin{enumerate}
 | 
|  |    380 |   
 | 
| 26760 |    381 |   \item named facts @{text "a"},
 | 
| 26754 |    382 | 
 | 
| 26760 |    383 |   \item selections from named facts @{text "a(i)"} or @{text "a(j - k)"},
 | 
| 26754 |    384 | 
 | 
|  |    385 |   \item literal fact propositions using @{syntax_ref altstring} syntax
 | 
| 26760 |    386 |   @{verbatim "`"}@{text "\<phi>"}@{verbatim "`"} (see also method
 | 
|  |    387 |   @{method_ref fact} in \secref{sec:pure-meth-att}).
 | 
| 26754 |    388 | 
 | 
|  |    389 |   \end{enumerate}
 | 
|  |    390 | 
 | 
|  |    391 |   Any kind of theorem specification may include lists of attributes
 | 
|  |    392 |   both on the left and right hand sides; attributes are applied to any
 | 
|  |    393 |   immediately preceding fact.  If names are omitted, the theorems are
 | 
|  |    394 |   not stored within the theorem database of the theory or proof
 | 
|  |    395 |   context, but any given attributes are applied nonetheless.
 | 
|  |    396 | 
 | 
| 26760 |    397 |   An extra pair of brackets around attributes (like ``@{text
 | 
|  |    398 |   "[[simproc a]]"}'') abbreviates a theorem reference involving an
 | 
|  |    399 |   internal dummy fact, which will be ignored later on.  So only the
 | 
|  |    400 |   effect of the attribute on the background context will persist.
 | 
|  |    401 |   This form of in-place declarations is particularly useful with
 | 
|  |    402 |   commands like @{command "declare"} and @{command "using"}.
 | 
| 26754 |    403 | 
 | 
|  |    404 |   \indexouternonterm{axmdecl}\indexouternonterm{thmdecl}
 | 
|  |    405 |   \indexouternonterm{thmdef}\indexouternonterm{thmref}
 | 
|  |    406 |   \indexouternonterm{thmrefs}\indexouternonterm{selection}
 | 
|  |    407 |   \begin{rail}
 | 
|  |    408 |     axmdecl: name attributes? ':'
 | 
|  |    409 |     ;
 | 
|  |    410 |     thmdecl: thmbind ':'
 | 
|  |    411 |     ;
 | 
|  |    412 |     thmdef: thmbind '='
 | 
|  |    413 |     ;
 | 
|  |    414 |     thmref: (nameref selection? | altstring) attributes? | '[' attributes ']'
 | 
|  |    415 |     ;
 | 
|  |    416 |     thmrefs: thmref +
 | 
|  |    417 |     ;
 | 
|  |    418 | 
 | 
|  |    419 |     thmbind: name attributes | name | attributes
 | 
|  |    420 |     ;
 | 
|  |    421 |     selection: '(' ((nat | nat '-' nat?) + ',') ')'
 | 
|  |    422 |     ;
 | 
|  |    423 |   \end{rail}
 | 
|  |    424 | *}
 | 
|  |    425 | 
 | 
|  |    426 | 
 | 
|  |    427 | subsection {* Term patterns and declarations \label{sec:term-decls} *}
 | 
|  |    428 | 
 | 
|  |    429 | text {*
 | 
|  |    430 |   Wherever explicit propositions (or term fragments) occur in a proof
 | 
|  |    431 |   text, casual binding of schematic term variables may be given
 | 
|  |    432 |   specified via patterns of the form ``@{text "(\<IS> p\<^sub>1 \<dots>
 | 
|  |    433 |   p\<^sub>n)"}''.  This works both for \railqtok{term} and \railqtok{prop}.
 | 
|  |    434 | 
 | 
|  |    435 |   \indexouternonterm{termpat}\indexouternonterm{proppat}
 | 
|  |    436 |   \begin{rail}
 | 
|  |    437 |     termpat: '(' ('is' term +) ')'
 | 
|  |    438 |     ;
 | 
|  |    439 |     proppat: '(' ('is' prop +) ')'
 | 
|  |    440 |     ;
 | 
|  |    441 |   \end{rail}
 | 
|  |    442 | 
 | 
|  |    443 |   \medskip Declarations of local variables @{text "x :: \<tau>"} and
 | 
|  |    444 |   logical propositions @{text "a : \<phi>"} represent different views on
 | 
|  |    445 |   the same principle of introducing a local scope.  In practice, one
 | 
|  |    446 |   may usually omit the typing of \railnonterm{vars} (due to
 | 
|  |    447 |   type-inference), and the naming of propositions (due to implicit
 | 
|  |    448 |   references of current facts).  In any case, Isar proof elements
 | 
|  |    449 |   usually admit to introduce multiple such items simultaneously.
 | 
|  |    450 | 
 | 
|  |    451 |   \indexouternonterm{vars}\indexouternonterm{props}
 | 
|  |    452 |   \begin{rail}
 | 
|  |    453 |     vars: (name+) ('::' type)?
 | 
|  |    454 |     ;
 | 
|  |    455 |     props: thmdecl? (prop proppat? +)
 | 
|  |    456 |     ;
 | 
|  |    457 |   \end{rail}
 | 
|  |    458 | 
 | 
|  |    459 |   The treatment of multiple declarations corresponds to the
 | 
|  |    460 |   complementary focus of \railnonterm{vars} versus
 | 
|  |    461 |   \railnonterm{props}.  In ``@{text "x\<^sub>1 \<dots> x\<^sub>n :: \<tau>"}''
 | 
|  |    462 |   the typing refers to all variables, while in @{text "a: \<phi>\<^sub>1 \<dots>
 | 
|  |    463 |   \<phi>\<^sub>n"} the naming refers to all propositions collectively.
 | 
|  |    464 |   Isar language elements that refer to \railnonterm{vars} or
 | 
|  |    465 |   \railnonterm{props} typically admit separate typings or namings via
 | 
|  |    466 |   another level of iteration, with explicit @{keyword_ref "and"}
 | 
|  |    467 |   separators; e.g.\ see @{command "fix"} and @{command "assume"} in
 | 
| 26760 |    468 |   \secref{sec:proof-context}.
 | 
| 26754 |    469 | *}
 | 
|  |    470 | 
 | 
|  |    471 | 
 | 
|  |    472 | subsection {* Antiquotations \label{sec:antiq} *}
 | 
|  |    473 | 
 | 
|  |    474 | text {*
 | 
|  |    475 |   \begin{matharray}{rcl}
 | 
|  |    476 |     @{antiquotation_def "theory"} & : & \isarantiq \\
 | 
|  |    477 |     @{antiquotation_def "thm"} & : & \isarantiq \\
 | 
|  |    478 |     @{antiquotation_def "prop"} & : & \isarantiq \\
 | 
|  |    479 |     @{antiquotation_def "term"} & : & \isarantiq \\
 | 
|  |    480 |     @{antiquotation_def const} & : & \isarantiq \\
 | 
|  |    481 |     @{antiquotation_def abbrev} & : & \isarantiq \\
 | 
|  |    482 |     @{antiquotation_def typeof} & : & \isarantiq \\
 | 
|  |    483 |     @{antiquotation_def typ} & : & \isarantiq \\
 | 
|  |    484 |     @{antiquotation_def thm_style} & : & \isarantiq \\
 | 
|  |    485 |     @{antiquotation_def term_style} & : & \isarantiq \\
 | 
|  |    486 |     @{antiquotation_def "text"} & : & \isarantiq \\
 | 
|  |    487 |     @{antiquotation_def goals} & : & \isarantiq \\
 | 
|  |    488 |     @{antiquotation_def subgoals} & : & \isarantiq \\
 | 
|  |    489 |     @{antiquotation_def prf} & : & \isarantiq \\
 | 
|  |    490 |     @{antiquotation_def full_prf} & : & \isarantiq \\
 | 
|  |    491 |     @{antiquotation_def ML} & : & \isarantiq \\
 | 
|  |    492 |     @{antiquotation_def ML_type} & : & \isarantiq \\
 | 
|  |    493 |     @{antiquotation_def ML_struct} & : & \isarantiq \\
 | 
|  |    494 |   \end{matharray}
 | 
|  |    495 | 
 | 
| 26760 |    496 |   The text body of formal comments (see also \secref{sec:comments})
 | 
|  |    497 |   may contain antiquotations of logical entities, such as theorems,
 | 
|  |    498 |   terms and types, which are to be presented in the final output
 | 
|  |    499 |   produced by the Isabelle document preparation system (see also
 | 
|  |    500 |   \secref{sec:document-prep}).
 | 
| 26754 |    501 | 
 | 
|  |    502 |   Thus embedding of ``@{text "@{term [show_types] \"f x = a + x\"}"}''
 | 
|  |    503 |   within a text block would cause
 | 
|  |    504 |   \isa{{\isacharparenleft}f{\isasymColon}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}\ {\isacharparenleft}x{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}a{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isacharplus}\ x} to appear in the final {\LaTeX} document.  Also note that theorem
 | 
|  |    505 |   antiquotations may involve attributes as well.  For example,
 | 
| 26760 |    506 |   @{text "@{thm sym [no_vars]}"} would print the theorem's
 | 
|  |    507 |   statement where all schematic variables have been replaced by fixed
 | 
|  |    508 |   ones, which are easier to read.
 | 
| 26754 |    509 | 
 | 
|  |    510 |   \begin{rail}
 | 
|  |    511 |     atsign lbrace antiquotation rbrace
 | 
|  |    512 |     ;
 | 
|  |    513 | 
 | 
|  |    514 |     antiquotation:
 | 
|  |    515 |       'theory' options name |
 | 
|  |    516 |       'thm' options thmrefs |
 | 
|  |    517 |       'prop' options prop |
 | 
|  |    518 |       'term' options term |
 | 
|  |    519 |       'const' options term |
 | 
|  |    520 |       'abbrev' options term |
 | 
|  |    521 |       'typeof' options term |
 | 
|  |    522 |       'typ' options type |
 | 
|  |    523 |       'thm\_style' options name thmref |
 | 
|  |    524 |       'term\_style' options name term |
 | 
|  |    525 |       'text' options name |
 | 
|  |    526 |       'goals' options |
 | 
|  |    527 |       'subgoals' options |
 | 
|  |    528 |       'prf' options thmrefs |
 | 
|  |    529 |       'full\_prf' options thmrefs |
 | 
|  |    530 |       'ML' options name |
 | 
|  |    531 |       'ML\_type' options name |
 | 
|  |    532 |       'ML\_struct' options name
 | 
|  |    533 |     ;
 | 
|  |    534 |     options: '[' (option * ',') ']'
 | 
|  |    535 |     ;
 | 
|  |    536 |     option: name | name '=' name
 | 
|  |    537 |     ;
 | 
|  |    538 |   \end{rail}
 | 
|  |    539 | 
 | 
|  |    540 |   Note that the syntax of antiquotations may \emph{not} include source
 | 
| 26760 |    541 |   comments @{verbatim "(*"}~@{text "\<dots>"}~@{verbatim "*)"} or verbatim
 | 
|  |    542 |   text @{verbatim "{"}@{verbatim "*"}~@{text "\<dots>"}~@{verbatim
 | 
|  |    543 |   "*"}@{verbatim "}"}.
 | 
| 26754 |    544 | 
 | 
|  |    545 |   \begin{descr}
 | 
|  |    546 |   
 | 
|  |    547 |   \item [@{text "@{theory A}"}] prints the name @{text "A"}, which is
 | 
|  |    548 |   guaranteed to refer to a valid ancestor theory in the current
 | 
|  |    549 |   context.
 | 
|  |    550 | 
 | 
| 26760 |    551 |   \item [@{text "@{thm a\<^sub>1 \<dots> a\<^sub>n}"}] prints theorems
 | 
|  |    552 |   @{text "a\<^sub>1 \<dots> a\<^sub>n"}.  Note that attribute specifications
 | 
|  |    553 |   may be included as well (see also \secref{sec:syn-att}); the
 | 
|  |    554 |   @{attribute_ref no_vars} rule (see \secref{sec:misc-meth-att}) would
 | 
|  |    555 |   be particularly useful to suppress printing of schematic variables.
 | 
| 26754 |    556 | 
 | 
|  |    557 |   \item [@{text "@{prop \<phi>}"}] prints a well-typed proposition @{text
 | 
|  |    558 |   "\<phi>"}.
 | 
|  |    559 | 
 | 
|  |    560 |   \item [@{text "@{term t}"}] prints a well-typed term @{text "t"}.
 | 
|  |    561 | 
 | 
|  |    562 |   \item [@{text "@{const c}"}] prints a logical or syntactic constant
 | 
|  |    563 |   @{text "c"}.
 | 
|  |    564 |   
 | 
|  |    565 |   \item [@{text "@{abbrev c x\<^sub>1 \<dots> x\<^sub>n}"}] prints a constant
 | 
|  |    566 |   abbreviation @{text "c x\<^sub>1 \<dots> x\<^sub>n \<equiv> rhs"} as defined in
 | 
|  |    567 |   the current context.
 | 
|  |    568 | 
 | 
|  |    569 |   \item [@{text "@{typeof t}"}] prints the type of a well-typed term
 | 
|  |    570 |   @{text "t"}.
 | 
|  |    571 | 
 | 
|  |    572 |   \item [@{text "@{typ \<tau>}"}] prints a well-formed type @{text "\<tau>"}.
 | 
|  |    573 |   
 | 
|  |    574 |   \item [@{text "@{thm_style s a}"}] prints theorem @{text a},
 | 
|  |    575 |   previously applying a style @{text s} to it (see below).
 | 
|  |    576 |   
 | 
|  |    577 |   \item [@{text "@{term_style s t}"}] prints a well-typed term @{text
 | 
|  |    578 |   t} after applying a style @{text s} to it (see below).
 | 
|  |    579 | 
 | 
|  |    580 |   \item [@{text "@{text s}"}] prints uninterpreted source text @{text
 | 
|  |    581 |   s}.  This is particularly useful to print portions of text according
 | 
|  |    582 |   to the Isabelle {\LaTeX} output style, without demanding
 | 
|  |    583 |   well-formedness (e.g.\ small pieces of terms that should not be
 | 
|  |    584 |   parsed or type-checked yet).
 | 
|  |    585 | 
 | 
|  |    586 |   \item [@{text "@{goals}"}] prints the current \emph{dynamic} goal
 | 
|  |    587 |   state.  This is mainly for support of tactic-emulation scripts
 | 
|  |    588 |   within Isar --- presentation of goal states does not conform to
 | 
|  |    589 |   actual human-readable proof documents.
 | 
|  |    590 | 
 | 
|  |    591 |   Please do not include goal states into document output unless you
 | 
|  |    592 |   really know what you are doing!
 | 
|  |    593 |   
 | 
|  |    594 |   \item [@{text "@{subgoals}"}] is similar to @{text "@{goals}"}, but
 | 
|  |    595 |   does not print the main goal.
 | 
|  |    596 |   
 | 
|  |    597 |   \item [@{text "@{prf a\<^sub>1 \<dots> a\<^sub>n}"}] prints the (compact)
 | 
|  |    598 |   proof terms corresponding to the theorems @{text "a\<^sub>1 \<dots>
 | 
|  |    599 |   a\<^sub>n"}. Note that this requires proof terms to be switched on
 | 
|  |    600 |   for the current object logic (see the ``Proof terms'' section of the
 | 
|  |    601 |   Isabelle reference manual for information on how to do this).
 | 
|  |    602 |   
 | 
|  |    603 |   \item [@{text "@{full_prf a\<^sub>1 \<dots> a\<^sub>n}"}] is like @{text
 | 
|  |    604 |   "@{prf a\<^sub>1 \<dots> a\<^sub>n}"}, but displays the full proof terms,
 | 
|  |    605 |   i.e.\ also displays information omitted in the compact proof term,
 | 
| 26777 |    606 |   which is denoted by ``@{text _}'' placeholders there.
 | 
| 26754 |    607 |   
 | 
|  |    608 |   \item [@{text "@{ML s}"}, @{text "@{ML_type s}"}, and @{text
 | 
|  |    609 |   "@{ML_struct s}"}] check text @{text s} as ML value, type, and
 | 
|  |    610 |   structure, respectively.  The source is displayed verbatim.
 | 
|  |    611 | 
 | 
|  |    612 |   \end{descr}
 | 
|  |    613 | 
 | 
|  |    614 |   \medskip The following standard styles for use with @{text
 | 
|  |    615 |   thm_style} and @{text term_style} are available:
 | 
|  |    616 | 
 | 
|  |    617 |   \begin{descr}
 | 
|  |    618 |   
 | 
|  |    619 |   \item [@{text lhs}] extracts the first argument of any application
 | 
|  |    620 |   form with at least two arguments -- typically meta-level or
 | 
|  |    621 |   object-level equality, or any other binary relation.
 | 
|  |    622 |   
 | 
|  |    623 |   \item [@{text rhs}] is like @{text lhs}, but extracts the second
 | 
|  |    624 |   argument.
 | 
|  |    625 |   
 | 
|  |    626 |   \item [@{text "concl"}] extracts the conclusion @{text C} from a rule
 | 
|  |    627 |   in Horn-clause normal form @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C"}.
 | 
|  |    628 |   
 | 
|  |    629 |   \item [@{text "prem1"}, \dots, @{text "prem9"}] extract premise
 | 
| 26760 |    630 |   number @{text "1, \<dots>, 9"}, respectively, from from a rule in
 | 
| 26754 |    631 |   Horn-clause normal form @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C"}
 | 
|  |    632 | 
 | 
|  |    633 |   \end{descr}
 | 
|  |    634 | 
 | 
|  |    635 |   \medskip
 | 
|  |    636 |   The following options are available to tune the output.  Note that most of
 | 
|  |    637 |   these coincide with ML flags of the same names (see also \cite{isabelle-ref}).
 | 
|  |    638 | 
 | 
|  |    639 |   \begin{descr}
 | 
|  |    640 | 
 | 
|  |    641 |   \item[@{text "show_types = bool"} and @{text "show_sorts = bool"}]
 | 
|  |    642 |   control printing of explicit type and sort constraints.
 | 
|  |    643 | 
 | 
|  |    644 |   \item[@{text "show_structs = bool"}] controls printing of implicit
 | 
|  |    645 |   structures.
 | 
|  |    646 | 
 | 
|  |    647 |   \item[@{text "long_names = bool"}] forces names of types and
 | 
|  |    648 |   constants etc.\ to be printed in their fully qualified internal
 | 
|  |    649 |   form.
 | 
|  |    650 | 
 | 
|  |    651 |   \item[@{text "short_names = bool"}] forces names of types and
 | 
|  |    652 |   constants etc.\ to be printed unqualified.  Note that internalizing
 | 
|  |    653 |   the output again in the current context may well yield a different
 | 
|  |    654 |   result.
 | 
|  |    655 | 
 | 
|  |    656 |   \item[@{text "unique_names = bool"}] determines whether the printed
 | 
|  |    657 |   version of qualified names should be made sufficiently long to avoid
 | 
|  |    658 |   overlap with names declared further back.  Set to @{text false} for
 | 
|  |    659 |   more concise output.
 | 
|  |    660 | 
 | 
|  |    661 |   \item[@{text "eta_contract = bool"}] prints terms in @{text
 | 
|  |    662 |   \<eta>}-contracted form.
 | 
|  |    663 | 
 | 
|  |    664 |   \item[@{text "display = bool"}] indicates if the text is to be
 | 
|  |    665 |   output as multi-line ``display material'', rather than a small piece
 | 
|  |    666 |   of text without line breaks (which is the default).
 | 
|  |    667 | 
 | 
|  |    668 |   \item[@{text "break = bool"}] controls line breaks in non-display
 | 
|  |    669 |   material.
 | 
|  |    670 | 
 | 
|  |    671 |   \item[@{text "quotes = bool"}] indicates if the output should be
 | 
|  |    672 |   enclosed in double quotes.
 | 
|  |    673 | 
 | 
|  |    674 |   \item[@{text "mode = name"}] adds @{text name} to the print mode to
 | 
|  |    675 |   be used for presentation (see also \cite{isabelle-ref}).  Note that
 | 
|  |    676 |   the standard setup for {\LaTeX} output is already present by
 | 
|  |    677 |   default, including the modes @{text latex} and @{text xsymbols}.
 | 
|  |    678 | 
 | 
|  |    679 |   \item[@{text "margin = nat"} and @{text "indent = nat"}] change the
 | 
|  |    680 |   margin or indentation for pretty printing of display material.
 | 
|  |    681 | 
 | 
|  |    682 |   \item[@{text "source = bool"}] prints the source text of the
 | 
|  |    683 |   antiquotation arguments, rather than the actual value.  Note that
 | 
|  |    684 |   this does not affect well-formedness checks of @{antiquotation
 | 
|  |    685 |   "thm"}, @{antiquotation "term"}, etc. (only the @{antiquotation
 | 
|  |    686 |   "text"} antiquotation admits arbitrary output).
 | 
|  |    687 | 
 | 
|  |    688 |   \item[@{text "goals_limit = nat"}] determines the maximum number of
 | 
|  |    689 |   goals to be printed.
 | 
|  |    690 | 
 | 
|  |    691 |   \item[@{text "locale = name"}] specifies an alternative locale
 | 
|  |    692 |   context used for evaluating and printing the subsequent argument.
 | 
|  |    693 | 
 | 
|  |    694 |   \end{descr}
 | 
|  |    695 | 
 | 
|  |    696 |   For boolean flags, ``@{text "name = true"}'' may be abbreviated as
 | 
|  |    697 |   ``@{text name}''.  All of the above flags are disabled by default,
 | 
|  |    698 |   unless changed from ML.
 | 
|  |    699 | 
 | 
|  |    700 |   \medskip Note that antiquotations do not only spare the author from
 | 
|  |    701 |   tedious typing of logical entities, but also achieve some degree of
 | 
|  |    702 |   consistency-checking of informal explanations with formal
 | 
|  |    703 |   developments: well-formedness of terms and types with respect to the
 | 
|  |    704 |   current theory or proof context is ensured here.
 | 
|  |    705 | *}
 | 
|  |    706 | 
 | 
|  |    707 | 
 | 
|  |    708 | subsection {* Tagged commands \label{sec:tags} *}
 | 
|  |    709 | 
 | 
|  |    710 | text {*
 | 
|  |    711 |   Each Isabelle/Isar command may be decorated by presentation tags:
 | 
|  |    712 | 
 | 
|  |    713 |   \indexouternonterm{tags}
 | 
|  |    714 |   \begin{rail}
 | 
|  |    715 |     tags: ( tag * )
 | 
|  |    716 |     ;
 | 
|  |    717 |     tag: '\%' (ident | string)
 | 
|  |    718 |   \end{rail}
 | 
|  |    719 | 
 | 
|  |    720 |   The tags @{text "theory"}, @{text "proof"}, @{text "ML"} are already
 | 
|  |    721 |   pre-declared for certain classes of commands:
 | 
|  |    722 | 
 | 
|  |    723 |  \medskip
 | 
|  |    724 | 
 | 
|  |    725 |   \begin{tabular}{ll}
 | 
|  |    726 |     @{text "theory"} & theory begin/end \\
 | 
|  |    727 |     @{text "proof"} & all proof commands \\
 | 
|  |    728 |     @{text "ML"} & all commands involving ML code \\
 | 
|  |    729 |   \end{tabular}
 | 
|  |    730 | 
 | 
|  |    731 |   \medskip The Isabelle document preparation system (see also
 | 
|  |    732 |   \cite{isabelle-sys}) allows tagged command regions to be presented
 | 
|  |    733 |   specifically, e.g.\ to fold proof texts, or drop parts of the text
 | 
|  |    734 |   completely.
 | 
|  |    735 | 
 | 
|  |    736 |   For example ``@{command "by"}~@{text "%invisible auto"}'' would
 | 
|  |    737 |   cause that piece of proof to be treated as @{text invisible} instead
 | 
|  |    738 |   of @{text "proof"} (the default), which may be either show or hidden
 | 
|  |    739 |   depending on the document setup.  In contrast, ``@{command
 | 
|  |    740 |   "by"}~@{text "%visible auto"}'' would force this text to be shown
 | 
|  |    741 |   invariably.
 | 
|  |    742 | 
 | 
|  |    743 |   Explicit tag specifications within a proof apply to all subsequent
 | 
|  |    744 |   commands of the same level of nesting.  For example, ``@{command
 | 
|  |    745 |   "proof"}~@{text "%visible \<dots>"}~@{command "qed"}'' would force the
 | 
|  |    746 |   whole sub-proof to be typeset as @{text visible} (unless some of its
 | 
|  |    747 |   parts are tagged differently).
 | 
|  |    748 | *}
 | 
|  |    749 | 
 | 
|  |    750 | end
 |