doc-src/IsarRef/Thy/document/Outer_Syntax.tex
author wenzelm
Thu, 02 Feb 2012 18:11:42 +0100
changeset 46282 83864b045a72
parent 42705 528a2ba8fa74
child 47114 7c9e31ffcd9e
permissions -rw-r--r--
clarified syntax section structure;

%
\begin{isabellebody}%
\def\isabellecontext{Outer{\isaliteral{5F}{\isacharunderscore}}Syntax}%
%
\isadelimtheory
%
\endisadelimtheory
%
\isatagtheory
\isacommand{theory}\isamarkupfalse%
\ Outer{\isaliteral{5F}{\isacharunderscore}}Syntax\isanewline
\isakeyword{imports}\ Base\ Main\isanewline
\isakeyword{begin}%
\endisatagtheory
{\isafoldtheory}%
%
\isadelimtheory
%
\endisadelimtheory
%
\isamarkupchapter{Outer syntax --- the theory language%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
The rather generic framework of Isabelle/Isar syntax emerges from
  three main syntactic categories: \emph{commands} of the top-level
  Isar engine (covering theory and proof elements), \emph{methods} for
  general goal refinements (analogous to traditional ``tactics''), and
  \emph{attributes} for operations on facts (within a certain
  context).  Subsequently we give a reference of basic syntactic
  entities underlying Isabelle/Isar syntax in a bottom-up manner.
  Concrete theory and proof language elements will be introduced later
  on.

  \medskip In order to get started with writing well-formed
  Isabelle/Isar documents, the most important aspect to be noted is
  the difference of \emph{inner} versus \emph{outer} syntax.  Inner
  syntax is that of Isabelle types and terms of the logic, while outer
  syntax is that of Isabelle/Isar theory sources (specifications and
  proofs).  As a general rule, inner syntax entities may occur only as
  \emph{atomic entities} within outer syntax.  For example, the string
  \verb|"x + y"| and identifier \verb|z| are legal term
  specifications within a theory, while \verb|x + y| without
  quotes is not.

  Printed theory documents usually omit quotes to gain readability
  (this is a matter of {\LaTeX} macro setup, say via \verb|\isabellestyle|, see also \cite{isabelle-sys}).  Experienced
  users of Isabelle/Isar may easily reconstruct the lost technical
  information, while mere readers need not care about quotes at all.

  \medskip Isabelle/Isar input may contain any number of input
  termination characters ``\verb|;|'' (semicolon) to separate
  commands explicitly.  This is particularly useful in interactive
  shell sessions to make clear where the current command is intended
  to end.  Otherwise, the interpreter loop will continue to issue a
  secondary prompt ``\verb|#|'' until an end-of-command is
  clearly recognized from the input syntax, e.g.\ encounter of the
  next command keyword.

  More advanced interfaces such as Proof~General \cite{proofgeneral}
  do not require explicit semicolons, the amount of input text is
  determined automatically by inspecting the present content of the
  Emacs text buffer.  In the printed presentation of Isabelle/Isar
  documents semicolons are omitted altogether for readability.

  \begin{warn}
    Proof~General requires certain syntax classification tables in
    order to achieve properly synchronized interaction with the
    Isabelle/Isar process.  These tables need to be consistent with
    the Isabelle version and particular logic image to be used in a
    running session (common object-logics may well change the outer
    syntax).  The standard setup should work correctly with any of the
    ``official'' logic images derived from Isabelle/HOL (including
    HOLCF etc.).  Users of alternative logics may need to tell
    Proof~General explicitly, e.g.\ by giving an option \verb|-k ZF|
    (in conjunction with \verb|-l ZF|, to specify the default
    logic image).  Note that option \verb|-L| does both
    of this at the same time.
  \end{warn}%
\end{isamarkuptext}%
\isamarkuptrue%
%
\isamarkupsection{Lexical matters \label{sec:outer-lex}%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
The outer lexical syntax consists of three main categories of
  syntax tokens:

  \begin{enumerate}

  \item \emph{major keywords} --- the command names that are available
  in the present logic session;

  \item \emph{minor keywords} --- additional literal tokens required
  by the syntax of commands;

  \item \emph{named tokens} --- various categories of identifiers etc.

  \end{enumerate}

  Major keywords and minor keywords are guaranteed to be disjoint.
  This helps user-interfaces to determine the overall structure of a
  theory text, without knowing the full details of command syntax.
  Internally, there is some additional information about the kind of
  major keywords, which approximates the command type (theory command,
  proof command etc.).

  Keywords override named tokens.  For example, the presence of a
  command called \verb|term| inhibits the identifier \verb|term|, but the string \verb|"term"| can be used instead.
  By convention, the outer syntax always allows quoted strings in
  addition to identifiers, wherever a named entity is expected.

  When tokenizing a given input sequence, the lexer repeatedly takes
  the longest prefix of the input that forms a valid token.  Spaces,
  tabs, newlines and formfeeds between tokens serve as explicit
  separators.

  \medskip The categories for named tokens are defined once and for
  all as follows.

  \begin{center}
  \begin{supertabular}{rcl}
    \indexdef{}{syntax}{ident}\hypertarget{syntax.ident}{\hyperlink{syntax.ident}{\mbox{\isa{ident}}}} & = & \isa{{\isaliteral{22}{\isachardoublequote}}letter\ quasiletter\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} \\
    \indexdef{}{syntax}{longident}\hypertarget{syntax.longident}{\hyperlink{syntax.longident}{\mbox{\isa{longident}}}} & = & \isa{{\isaliteral{22}{\isachardoublequote}}ident{\isaliteral{28}{\isacharparenleft}}{\isaliteral{22}{\isachardoublequote}}}\verb|.|\isa{{\isaliteral{22}{\isachardoublequote}}ident{\isaliteral{29}{\isacharparenright}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2B}{\isacharplus}}{\isaliteral{22}{\isachardoublequote}}} \\
    \indexdef{}{syntax}{symident}\hypertarget{syntax.symident}{\hyperlink{syntax.symident}{\mbox{\isa{symident}}}} & = & \isa{{\isaliteral{22}{\isachardoublequote}}sym\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2B}{\isacharplus}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\|\verb|<|\isa{ident}\verb|>| \\
    \indexdef{}{syntax}{nat}\hypertarget{syntax.nat}{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}} & = & \isa{{\isaliteral{22}{\isachardoublequote}}digit\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2B}{\isacharplus}}{\isaliteral{22}{\isachardoublequote}}} \\
    \indexdef{}{syntax}{float}\hypertarget{syntax.float}{\hyperlink{syntax.float}{\mbox{\isa{float}}}} & = & \indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}}\verb|.|\indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}}\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|-|\indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}}\verb|.|\indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}} \\
    \indexdef{}{syntax}{var}\hypertarget{syntax.var}{\hyperlink{syntax.var}{\mbox{\isa{var}}}} & = & \verb|?|\isa{{\isaliteral{22}{\isachardoublequote}}ident\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|?|\isa{ident}\verb|.|\isa{nat} \\
    \indexdef{}{syntax}{typefree}\hypertarget{syntax.typefree}{\hyperlink{syntax.typefree}{\mbox{\isa{typefree}}}} & = & \verb|'|\isa{ident} \\
    \indexdef{}{syntax}{typevar}\hypertarget{syntax.typevar}{\hyperlink{syntax.typevar}{\mbox{\isa{typevar}}}} & = & \verb|?|\isa{{\isaliteral{22}{\isachardoublequote}}typefree\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|?|\isa{typefree}\verb|.|\isa{nat} \\
    \indexdef{}{syntax}{string}\hypertarget{syntax.string}{\hyperlink{syntax.string}{\mbox{\isa{string}}}} & = & \verb|"| \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} \verb|"| \\
    \indexdef{}{syntax}{altstring}\hypertarget{syntax.altstring}{\hyperlink{syntax.altstring}{\mbox{\isa{altstring}}}} & = & \verb|`| \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} \verb|`| \\
    \indexdef{}{syntax}{verbatim}\hypertarget{syntax.verbatim}{\hyperlink{syntax.verbatim}{\mbox{\isa{verbatim}}}} & = & \verb|{*| \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} \verb|*|\verb|}| \\[1ex]

    \isa{letter} & = & \isa{{\isaliteral{22}{\isachardoublequote}}latin\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\|\verb|<|\isa{latin}\verb|>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\|\verb|<|\isa{{\isaliteral{22}{\isachardoublequote}}latin\ latin{\isaliteral{22}{\isachardoublequote}}}\verb|>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ greek\ \ {\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} \\
          &   & \verb|\<^isub>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\<^isup>| \\
    \isa{quasiletter} & = & \isa{{\isaliteral{22}{\isachardoublequote}}letter\ \ {\isaliteral{7C}{\isacharbar}}\ \ digit\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|_|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|'| \\
    \isa{latin} & = & \verb|a|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|z|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|A|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|Z| \\
    \isa{digit} & = & \verb|0|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|9| \\
    \isa{sym} & = & \verb|!|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|#|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|$|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|%|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|&|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|*|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|+|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|-|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|/|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} \\
    & & \verb|<|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|=|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|?|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|@|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|^|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|_|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb||\verb,|,\verb||\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|~| \\
    \isa{greek} & = & \verb|\<alpha>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\<beta>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\<gamma>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\<delta>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} \\
          &   & \verb|\<epsilon>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\<zeta>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\<eta>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\<theta>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} \\
          &   & \verb|\<iota>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\<kappa>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\<mu>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\<nu>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} \\
          &   & \verb|\<xi>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\<pi>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\<rho>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\<sigma>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\<tau>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} \\
          &   & \verb|\<upsilon>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\<phi>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\<chi>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\<psi>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} \\
          &   & \verb|\<omega>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\<Gamma>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\<Delta>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\<Theta>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} \\
          &   & \verb|\<Lambda>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\<Xi>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\<Pi>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\<Sigma>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} \\
          &   & \verb|\<Upsilon>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\<Phi>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\<Psi>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\<Omega>| \\
  \end{supertabular}
  \end{center}

  A \indexref{}{syntax}{var}\hyperlink{syntax.var}{\mbox{\isa{var}}} or \indexref{}{syntax}{typevar}\hyperlink{syntax.typevar}{\mbox{\isa{typevar}}} describes an unknown,
  which is internally a pair of base name and index (ML type \verb|indexname|).  These components are either separated by a dot as in
  \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}x{\isaliteral{2E}{\isachardot}}{\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} or \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}x{\isadigit{7}}{\isaliteral{2E}{\isachardot}}{\isadigit{3}}{\isaliteral{22}{\isachardoublequote}}} or run together as in \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}x{\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}}.  The latter form is possible if the base name does not end
  with digits.  If the index is 0, it may be dropped altogether:
  \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}x{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}x{\isadigit{0}}{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}x{\isaliteral{2E}{\isachardot}}{\isadigit{0}}{\isaliteral{22}{\isachardoublequote}}} all refer to the
  same unknown, with basename \isa{{\isaliteral{22}{\isachardoublequote}}x{\isaliteral{22}{\isachardoublequote}}} and index 0.

  The syntax of \indexref{}{syntax}{string}\hyperlink{syntax.string}{\mbox{\isa{string}}} admits any characters, including
  newlines; ``\verb|"|'' (double-quote) and ``\verb|\|'' (backslash) need to be escaped by a backslash; arbitrary
  character codes may be specified as ``\verb|\|\isa{ddd}'',
  with three decimal digits.  Alternative strings according to
  \indexref{}{syntax}{altstring}\hyperlink{syntax.altstring}{\mbox{\isa{altstring}}} are analogous, using single back-quotes
  instead.

  The body of \indexref{}{syntax}{verbatim}\hyperlink{syntax.verbatim}{\mbox{\isa{verbatim}}} may consist of any text not
  containing ``\verb|*|\verb|}|''; this allows
  convenient inclusion of quotes without further escapes.  There is no
  way to escape ``\verb|*|\verb|}|''.  If the quoted
  text is {\LaTeX} source, one may usually add some blank or comment
  to avoid the critical character sequence.

  Source comments take the form \verb|(*|~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}}~\verb|*)| and may be nested, although the user-interface
  might prevent this.  Note that this form indicates source comments
  only, which are stripped after lexical analysis of the input.  The
  Isar syntax also provides proper \emph{document comments} that are
  considered as part of the text (see \secref{sec:comments}).

  Common mathematical symbols such as \isa{{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}} are represented in
  Isabelle as \verb|\<forall>|.  There are infinitely many Isabelle
  symbols like this, although proper presentation is left to front-end
  tools such as {\LaTeX} or Proof~General with the X-Symbol package.
  A list of predefined Isabelle symbols that work well with these
  tools is given in \appref{app:symbols}.  Note that \verb|\<lambda>|
  does not belong to the \isa{letter} category, since it is already
  used differently in the Pure term language.%
\end{isamarkuptext}%
\isamarkuptrue%
%
\isamarkupsection{Common syntax entities%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
We now introduce several basic syntactic entities, such as names,
  terms, and theorem specifications, which are factored out of the
  actual Isar language elements to be described later.%
\end{isamarkuptext}%
\isamarkuptrue%
%
\isamarkupsubsection{Names%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
Entity \hyperlink{syntax.name}{\mbox{\isa{name}}} usually refers to any name of types,
  constants, theorems etc.\ that are to be \emph{declared} or
  \emph{defined} (so qualified identifiers are excluded here).  Quoted
  strings provide an escape for non-identifier names or those ruled
  out by outer syntax keywords (e.g.\ quoted \verb|"let"|).
  Already existing objects are usually referenced by \hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}.

  \begin{railoutput}
\rail@begin{4}{\indexdef{}{syntax}{name}\hypertarget{syntax.name}{\hyperlink{syntax.name}{\mbox{\isa{name}}}}}
\rail@bar
\rail@nont{\hyperlink{syntax.ident}{\mbox{\isa{ident}}}}[]
\rail@nextbar{1}
\rail@nont{\hyperlink{syntax.symident}{\mbox{\isa{symident}}}}[]
\rail@nextbar{2}
\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[]
\rail@nextbar{3}
\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[]
\rail@endbar
\rail@end
\rail@begin{1}{\indexdef{}{syntax}{parname}\hypertarget{syntax.parname}{\hyperlink{syntax.parname}{\mbox{\isa{parname}}}}}
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
\rail@end
\rail@begin{2}{\indexdef{}{syntax}{nameref}\hypertarget{syntax.nameref}{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}}
\rail@bar
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
\rail@nextbar{1}
\rail@nont{\hyperlink{syntax.longident}{\mbox{\isa{longident}}}}[]
\rail@endbar
\rail@end
\end{railoutput}%
\end{isamarkuptext}%
\isamarkuptrue%
%
\isamarkupsubsection{Numbers%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
The outer lexical syntax (\secref{sec:outer-lex}) admits
  natural numbers and floating point numbers.  These are combined as
  \hyperlink{syntax.int}{\mbox{\isa{int}}} and \hyperlink{syntax.real}{\mbox{\isa{real}}} as follows.

  \begin{railoutput}
\rail@begin{2}{\indexdef{}{syntax}{int}\hypertarget{syntax.int}{\hyperlink{syntax.int}{\mbox{\isa{int}}}}}
\rail@bar
\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[]
\rail@nextbar{1}
\rail@term{\isa{{\isaliteral{2D}{\isacharminus}}}}[]
\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[]
\rail@endbar
\rail@end
\rail@begin{2}{\indexdef{}{syntax}{real}\hypertarget{syntax.real}{\hyperlink{syntax.real}{\mbox{\isa{real}}}}}
\rail@bar
\rail@nont{\hyperlink{syntax.float}{\mbox{\isa{float}}}}[]
\rail@nextbar{1}
\rail@nont{\hyperlink{syntax.int}{\mbox{\isa{int}}}}[]
\rail@endbar
\rail@end
\end{railoutput}


  Note that there is an overlap with the category \hyperlink{syntax.name}{\mbox{\isa{name}}},
  which also includes \hyperlink{syntax.nat}{\mbox{\isa{nat}}}.%
\end{isamarkuptext}%
\isamarkuptrue%
%
\isamarkupsubsection{Comments \label{sec:comments}%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
Large chunks of plain \hyperlink{syntax.text}{\mbox{\isa{text}}} are usually given
  \hyperlink{syntax.verbatim}{\mbox{\isa{verbatim}}}, i.e.\ enclosed in \verb|{|\verb|*|~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}}~\verb|*|\verb|}|.  For convenience,
  any of the smaller text units conforming to \hyperlink{syntax.nameref}{\mbox{\isa{nameref}}} are
  admitted as well.  A marginal \hyperlink{syntax.comment}{\mbox{\isa{comment}}} is of the form
  \verb|--|~\hyperlink{syntax.text}{\mbox{\isa{text}}}.  Any number of these may occur
  within Isabelle/Isar commands.

  \begin{railoutput}
\rail@begin{2}{\indexdef{}{syntax}{text}\hypertarget{syntax.text}{\hyperlink{syntax.text}{\mbox{\isa{text}}}}}
\rail@bar
\rail@nont{\hyperlink{syntax.verbatim}{\mbox{\isa{verbatim}}}}[]
\rail@nextbar{1}
\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[]
\rail@endbar
\rail@end
\rail@begin{1}{\indexdef{}{syntax}{comment}\hypertarget{syntax.comment}{\hyperlink{syntax.comment}{\mbox{\isa{comment}}}}}
\rail@term{\isa{{\isaliteral{2D}{\isacharminus}}{\isaliteral{2D}{\isacharminus}}}}[]
\rail@nont{\hyperlink{syntax.text}{\mbox{\isa{text}}}}[]
\rail@end
\end{railoutput}%
\end{isamarkuptext}%
\isamarkuptrue%
%
\isamarkupsubsection{Type classes, sorts and arities%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
Classes are specified by plain names.  Sorts have a very simple
  inner syntax, which is either a single class name \isa{c} or a
  list \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7B}{\isacharbraceleft}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}} referring to the
  intersection of these classes.  The syntax of type arities is given
  directly at the outer level.

  \begin{railoutput}
\rail@begin{3}{\indexdef{}{syntax}{classdecl}\hypertarget{syntax.classdecl}{\hyperlink{syntax.classdecl}{\mbox{\isa{classdecl}}}}}
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
\rail@bar
\rail@nextbar{1}
\rail@bar
\rail@term{\isa{{\isaliteral{3C}{\isacharless}}}}[]
\rail@nextbar{2}
\rail@term{\isa{{\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}}}[]
\rail@endbar
\rail@plus
\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[]
\rail@nextplus{2}
\rail@cterm{\isa{{\isaliteral{2C}{\isacharcomma}}}}[]
\rail@endplus
\rail@endbar
\rail@end
\rail@begin{1}{\indexdef{}{syntax}{sort}\hypertarget{syntax.sort}{\hyperlink{syntax.sort}{\mbox{\isa{sort}}}}}
\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[]
\rail@end
\rail@begin{3}{\indexdef{}{syntax}{arity}\hypertarget{syntax.arity}{\hyperlink{syntax.arity}{\mbox{\isa{arity}}}}}
\rail@bar
\rail@nextbar{1}
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
\rail@plus
\rail@nont{\hyperlink{syntax.sort}{\mbox{\isa{sort}}}}[]
\rail@nextplus{2}
\rail@cterm{\isa{{\isaliteral{2C}{\isacharcomma}}}}[]
\rail@endplus
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
\rail@endbar
\rail@nont{\hyperlink{syntax.sort}{\mbox{\isa{sort}}}}[]
\rail@end
\end{railoutput}%
\end{isamarkuptext}%
\isamarkuptrue%
%
\isamarkupsubsection{Types and terms \label{sec:types-terms}%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
The actual inner Isabelle syntax, that of types and terms of the
  logic, is far too sophisticated in order to be modelled explicitly
  at the outer theory level.  Basically, any such entity has to be
  quoted to turn it into a single token (the parsing and type-checking
  is performed internally later).  For convenience, a slightly more
  liberal convention is adopted: quotes may be omitted for any type or
  term that is already atomic at the outer level.  For example, one
  may just write \verb|x| instead of quoted \verb|"x"|.
  Note that symbolic identifiers (e.g.\ \verb|++| or \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}{\isaliteral{22}{\isachardoublequote}}} are available as well, provided these have not been superseded
  by commands or other keywords already (such as \verb|=| or
  \verb|+|).

  \begin{railoutput}
\rail@begin{3}{\indexdef{}{syntax}{type}\hypertarget{syntax.type}{\hyperlink{syntax.type}{\mbox{\isa{type}}}}}
\rail@bar
\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[]
\rail@nextbar{1}
\rail@nont{\hyperlink{syntax.typefree}{\mbox{\isa{typefree}}}}[]
\rail@nextbar{2}
\rail@nont{\hyperlink{syntax.typevar}{\mbox{\isa{typevar}}}}[]
\rail@endbar
\rail@end
\rail@begin{2}{\indexdef{}{syntax}{term}\hypertarget{syntax.term}{\hyperlink{syntax.term}{\mbox{\isa{term}}}}}
\rail@bar
\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[]
\rail@nextbar{1}
\rail@nont{\hyperlink{syntax.var}{\mbox{\isa{var}}}}[]
\rail@endbar
\rail@end
\rail@begin{1}{\indexdef{}{syntax}{prop}\hypertarget{syntax.prop}{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}}
\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[]
\rail@end
\end{railoutput}


  Positional instantiations are indicated by giving a sequence of
  terms, or the placeholder ``\isa{{\isaliteral{5F}{\isacharunderscore}}}'' (underscore), which means to
  skip a position.

  \begin{railoutput}
\rail@begin{2}{\indexdef{}{syntax}{inst}\hypertarget{syntax.inst}{\hyperlink{syntax.inst}{\mbox{\isa{inst}}}}}
\rail@bar
\rail@term{\isa{{\isaliteral{5F}{\isacharunderscore}}}}[]
\rail@nextbar{1}
\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[]
\rail@endbar
\rail@end
\rail@begin{2}{\indexdef{}{syntax}{insts}\hypertarget{syntax.insts}{\hyperlink{syntax.insts}{\mbox{\isa{insts}}}}}
\rail@plus
\rail@nextplus{1}
\rail@cnont{\hyperlink{syntax.inst}{\mbox{\isa{inst}}}}[]
\rail@endplus
\rail@end
\end{railoutput}


  Type declarations and definitions usually refer to \hyperlink{syntax.typespec}{\mbox{\isa{typespec}}} on the left-hand side.  This models basic type constructor
  application at the outer syntax level.  Note that only plain postfix
  notation is available here, but no infixes.

  \begin{railoutput}
\rail@begin{4}{\indexdef{}{syntax}{typespec}\hypertarget{syntax.typespec}{\hyperlink{syntax.typespec}{\mbox{\isa{typespec}}}}}
\rail@bar
\rail@nextbar{1}
\rail@nont{\hyperlink{syntax.typefree}{\mbox{\isa{typefree}}}}[]
\rail@nextbar{2}
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
\rail@plus
\rail@nont{\hyperlink{syntax.typefree}{\mbox{\isa{typefree}}}}[]
\rail@nextplus{3}
\rail@cterm{\isa{{\isaliteral{2C}{\isacharcomma}}}}[]
\rail@endplus
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
\rail@endbar
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
\rail@end
\rail@begin{6}{\indexdef{}{syntax}{typespec\_sorts}\hypertarget{syntax.typespec-sorts}{\hyperlink{syntax.typespec-sorts}{\mbox{\isa{typespec{\isaliteral{5F}{\isacharunderscore}}sorts}}}}}
\rail@bar
\rail@nextbar{1}
\rail@nont{\hyperlink{syntax.typefree}{\mbox{\isa{typefree}}}}[]
\rail@bar
\rail@nextbar{2}
\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[]
\rail@nont{\hyperlink{syntax.sort}{\mbox{\isa{sort}}}}[]
\rail@endbar
\rail@nextbar{3}
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
\rail@plus
\rail@nont{\hyperlink{syntax.typefree}{\mbox{\isa{typefree}}}}[]
\rail@bar
\rail@nextbar{4}
\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[]
\rail@nont{\hyperlink{syntax.sort}{\mbox{\isa{sort}}}}[]
\rail@endbar
\rail@nextplus{5}
\rail@cterm{\isa{{\isaliteral{2C}{\isacharcomma}}}}[]
\rail@endplus
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
\rail@endbar
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
\rail@end
\end{railoutput}%
\end{isamarkuptext}%
\isamarkuptrue%
%
\isamarkupsubsection{Term patterns and declarations \label{sec:term-decls}%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
Wherever explicit propositions (or term fragments) occur in a
  proof text, casual binding of schematic term variables may be given
  specified via patterns of the form ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C49533E}{\isasymIS}}\ p\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ p\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}''.
  This works both for \hyperlink{syntax.term}{\mbox{\isa{term}}} and \hyperlink{syntax.prop}{\mbox{\isa{prop}}}.

  \begin{railoutput}
\rail@begin{2}{\indexdef{}{syntax}{term\_pat}\hypertarget{syntax.term-pat}{\hyperlink{syntax.term-pat}{\mbox{\isa{term{\isaliteral{5F}{\isacharunderscore}}pat}}}}}
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
\rail@plus
\rail@term{\isa{\isakeyword{is}}}[]
\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[]
\rail@nextplus{1}
\rail@endplus
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
\rail@end
\rail@begin{2}{\indexdef{}{syntax}{prop\_pat}\hypertarget{syntax.prop-pat}{\hyperlink{syntax.prop-pat}{\mbox{\isa{prop{\isaliteral{5F}{\isacharunderscore}}pat}}}}}
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
\rail@plus
\rail@term{\isa{\isakeyword{is}}}[]
\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[]
\rail@nextplus{1}
\rail@endplus
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
\rail@end
\end{railoutput}


  \medskip Declarations of local variables \isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{22}{\isachardoublequote}}} and
  logical propositions \isa{{\isaliteral{22}{\isachardoublequote}}a\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} represent different views on
  the same principle of introducing a local scope.  In practice, one
  may usually omit the typing of \hyperlink{syntax.vars}{\mbox{\isa{vars}}} (due to
  type-inference), and the naming of propositions (due to implicit
  references of current facts).  In any case, Isar proof elements
  usually admit to introduce multiple such items simultaneously.

  \begin{railoutput}
\rail@begin{2}{\indexdef{}{syntax}{vars}\hypertarget{syntax.vars}{\hyperlink{syntax.vars}{\mbox{\isa{vars}}}}}
\rail@plus
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
\rail@nextplus{1}
\rail@endplus
\rail@bar
\rail@nextbar{1}
\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[]
\rail@nont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[]
\rail@endbar
\rail@end
\rail@begin{3}{\indexdef{}{syntax}{props}\hypertarget{syntax.props}{\hyperlink{syntax.props}{\mbox{\isa{props}}}}}
\rail@bar
\rail@nextbar{1}
\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[]
\rail@endbar
\rail@plus
\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[]
\rail@bar
\rail@nextbar{1}
\rail@nont{\hyperlink{syntax.prop-pat}{\mbox{\isa{prop{\isaliteral{5F}{\isacharunderscore}}pat}}}}[]
\rail@endbar
\rail@nextplus{2}
\rail@endplus
\rail@end
\end{railoutput}


  The treatment of multiple declarations corresponds to the
  complementary focus of \hyperlink{syntax.vars}{\mbox{\isa{vars}}} versus \hyperlink{syntax.props}{\mbox{\isa{props}}}.  In
  ``\isa{{\isaliteral{22}{\isachardoublequote}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{22}{\isachardoublequote}}}'' the typing refers to all variables, while
  in \isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} the naming refers to all propositions
  collectively.  Isar language elements that refer to \hyperlink{syntax.vars}{\mbox{\isa{vars}}}
  or \hyperlink{syntax.props}{\mbox{\isa{props}}} typically admit separate typings or namings via
  another level of iteration, with explicit \indexref{}{keyword}{and}\hyperlink{keyword.and}{\mbox{\isa{\isakeyword{and}}}}
  separators; e.g.\ see \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}} and \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}} in
  \secref{sec:proof-context}.%
\end{isamarkuptext}%
\isamarkuptrue%
%
\isamarkupsubsection{Attributes and theorems \label{sec:syn-att}%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
Attributes have their own ``semi-inner'' syntax, in the sense
  that input conforming to \hyperlink{syntax.args}{\mbox{\isa{args}}} below is parsed by the
  attribute a second time.  The attribute argument specifications may
  be any sequence of atomic entities (identifiers, strings etc.), or
  properly bracketed argument lists.  Below \hyperlink{syntax.atom}{\mbox{\isa{atom}}} refers to
  any atomic entity, including any \hyperlink{syntax.keyword}{\mbox{\isa{keyword}}} conforming to
  \hyperlink{syntax.symident}{\mbox{\isa{symident}}}.

  \begin{railoutput}
\rail@begin{7}{\indexdef{}{syntax}{atom}\hypertarget{syntax.atom}{\hyperlink{syntax.atom}{\mbox{\isa{atom}}}}}
\rail@bar
\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[]
\rail@nextbar{1}
\rail@nont{\hyperlink{syntax.typefree}{\mbox{\isa{typefree}}}}[]
\rail@nextbar{2}
\rail@nont{\hyperlink{syntax.typevar}{\mbox{\isa{typevar}}}}[]
\rail@nextbar{3}
\rail@nont{\hyperlink{syntax.var}{\mbox{\isa{var}}}}[]
\rail@nextbar{4}
\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[]
\rail@nextbar{5}
\rail@nont{\hyperlink{syntax.float}{\mbox{\isa{float}}}}[]
\rail@nextbar{6}
\rail@nont{\hyperlink{syntax.keyword}{\mbox{\isa{keyword}}}}[]
\rail@endbar
\rail@end
\rail@begin{3}{\isa{arg}}
\rail@bar
\rail@nont{\hyperlink{syntax.atom}{\mbox{\isa{atom}}}}[]
\rail@nextbar{1}
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
\rail@nont{\hyperlink{syntax.args}{\mbox{\isa{args}}}}[]
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
\rail@nextbar{2}
\rail@term{\isa{{\isaliteral{5B}{\isacharbrackleft}}}}[]
\rail@nont{\hyperlink{syntax.args}{\mbox{\isa{args}}}}[]
\rail@term{\isa{{\isaliteral{5D}{\isacharbrackright}}}}[]
\rail@endbar
\rail@end
\rail@begin{2}{\indexdef{}{syntax}{args}\hypertarget{syntax.args}{\hyperlink{syntax.args}{\mbox{\isa{args}}}}}
\rail@plus
\rail@nextplus{1}
\rail@cnont{\isa{arg}}[]
\rail@endplus
\rail@end
\rail@begin{3}{\indexdef{}{syntax}{attributes}\hypertarget{syntax.attributes}{\hyperlink{syntax.attributes}{\mbox{\isa{attributes}}}}}
\rail@term{\isa{{\isaliteral{5B}{\isacharbrackleft}}}}[]
\rail@bar
\rail@nextbar{1}
\rail@plus
\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[]
\rail@nont{\hyperlink{syntax.args}{\mbox{\isa{args}}}}[]
\rail@nextplus{2}
\rail@cterm{\isa{{\isaliteral{2C}{\isacharcomma}}}}[]
\rail@endplus
\rail@endbar
\rail@term{\isa{{\isaliteral{5D}{\isacharbrackright}}}}[]
\rail@end
\end{railoutput}


  Theorem specifications come in several flavors: \hyperlink{syntax.axmdecl}{\mbox{\isa{axmdecl}}}
  and \hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}} usually refer to axioms, assumptions or
  results of goal statements, while \hyperlink{syntax.thmdef}{\mbox{\isa{thmdef}}} collects lists of
  existing theorems.  Existing theorems are given by \hyperlink{syntax.thmref}{\mbox{\isa{thmref}}}
  and \hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}, the former requires an actual singleton
  result.

  There are three forms of theorem references:
  \begin{enumerate}
  
  \item named facts \isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{22}{\isachardoublequote}}},

  \item selections from named facts \isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{28}{\isacharparenleft}}i{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} or \isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{28}{\isacharparenleft}}j\ {\isaliteral{2D}{\isacharminus}}\ k{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}},

  \item literal fact propositions using \indexref{}{syntax}{altstring}\hyperlink{syntax.altstring}{\mbox{\isa{altstring}}} syntax
  \verb|`|\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}}\verb|`| (see also method
  \indexref{}{method}{fact}\hyperlink{method.fact}{\mbox{\isa{fact}}}).

  \end{enumerate}

  Any kind of theorem specification may include lists of attributes
  both on the left and right hand sides; attributes are applied to any
  immediately preceding fact.  If names are omitted, the theorems are
  not stored within the theorem database of the theory or proof
  context, but any given attributes are applied nonetheless.

  An extra pair of brackets around attributes (like ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5B}{\isacharbrackleft}}simproc\ a{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}}'') abbreviates a theorem reference involving an
  internal dummy fact, which will be ignored later on.  So only the
  effect of the attribute on the background context will persist.
  This form of in-place declarations is particularly useful with
  commands like \hyperlink{command.declare}{\mbox{\isa{\isacommand{declare}}}} and \hyperlink{command.using}{\mbox{\isa{\isacommand{using}}}}.

  \begin{railoutput}
\rail@begin{2}{\indexdef{}{syntax}{axmdecl}\hypertarget{syntax.axmdecl}{\hyperlink{syntax.axmdecl}{\mbox{\isa{axmdecl}}}}}
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
\rail@bar
\rail@nextbar{1}
\rail@nont{\hyperlink{syntax.attributes}{\mbox{\isa{attributes}}}}[]
\rail@endbar
\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[]
\rail@end
\rail@begin{1}{\indexdef{}{syntax}{thmdecl}\hypertarget{syntax.thmdecl}{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}}
\rail@nont{\isa{thmbind}}[]
\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[]
\rail@end
\rail@begin{1}{\indexdef{}{syntax}{thmdef}\hypertarget{syntax.thmdef}{\hyperlink{syntax.thmdef}{\mbox{\isa{thmdef}}}}}
\rail@nont{\isa{thmbind}}[]
\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[]
\rail@end
\rail@begin{4}{\indexdef{}{syntax}{thmref}\hypertarget{syntax.thmref}{\hyperlink{syntax.thmref}{\mbox{\isa{thmref}}}}}
\rail@bar
\rail@bar
\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[]
\rail@bar
\rail@nextbar{1}
\rail@nont{\isa{selection}}[]
\rail@endbar
\rail@nextbar{2}
\rail@nont{\hyperlink{syntax.altstring}{\mbox{\isa{altstring}}}}[]
\rail@endbar
\rail@bar
\rail@nextbar{1}
\rail@nont{\hyperlink{syntax.attributes}{\mbox{\isa{attributes}}}}[]
\rail@endbar
\rail@nextbar{3}
\rail@term{\isa{{\isaliteral{5B}{\isacharbrackleft}}}}[]
\rail@nont{\hyperlink{syntax.attributes}{\mbox{\isa{attributes}}}}[]
\rail@term{\isa{{\isaliteral{5D}{\isacharbrackright}}}}[]
\rail@endbar
\rail@end
\rail@begin{2}{\indexdef{}{syntax}{thmrefs}\hypertarget{syntax.thmrefs}{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}}
\rail@plus
\rail@nont{\hyperlink{syntax.thmref}{\mbox{\isa{thmref}}}}[]
\rail@nextplus{1}
\rail@endplus
\rail@end
\rail@begin{3}{\isa{thmbind}}
\rail@bar
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
\rail@nont{\hyperlink{syntax.attributes}{\mbox{\isa{attributes}}}}[]
\rail@nextbar{1}
\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
\rail@nextbar{2}
\rail@nont{\hyperlink{syntax.attributes}{\mbox{\isa{attributes}}}}[]
\rail@endbar
\rail@end
\rail@begin{4}{\isa{selection}}
\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
\rail@plus
\rail@bar
\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[]
\rail@nextbar{1}
\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[]
\rail@term{\isa{{\isaliteral{2D}{\isacharminus}}}}[]
\rail@bar
\rail@nextbar{2}
\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[]
\rail@endbar
\rail@endbar
\rail@nextplus{3}
\rail@cterm{\isa{{\isaliteral{2C}{\isacharcomma}}}}[]
\rail@endplus
\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
\rail@end
\end{railoutput}%
\end{isamarkuptext}%
\isamarkuptrue%
%
\isadelimtheory
%
\endisadelimtheory
%
\isatagtheory
\isacommand{end}\isamarkupfalse%
%
\endisatagtheory
{\isafoldtheory}%
%
\isadelimtheory
%
\endisadelimtheory
\isanewline
\end{isabellebody}%
%%% Local Variables:
%%% mode: latex
%%% TeX-master: "root"
%%% End: