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