| 
28762
 | 
     1  | 
%
  | 
| 
 | 
     2  | 
\begin{isabellebody}%
 | 
| 
 | 
     3  | 
\def\isabellecontext{Inner{\isacharunderscore}Syntax}%
 | 
| 
 | 
     4  | 
%
  | 
| 
 | 
     5  | 
\isadelimtheory
  | 
| 
 | 
     6  | 
\isanewline
  | 
| 
 | 
     7  | 
\isanewline
  | 
| 
 | 
     8  | 
%
  | 
| 
 | 
     9  | 
\endisadelimtheory
  | 
| 
 | 
    10  | 
%
  | 
| 
 | 
    11  | 
\isatagtheory
  | 
| 
 | 
    12  | 
\isacommand{theory}\isamarkupfalse%
 | 
| 
 | 
    13  | 
\ Inner{\isacharunderscore}Syntax\isanewline
 | 
| 
 | 
    14  | 
\isakeyword{imports}\ Main\isanewline
 | 
| 
 | 
    15  | 
\isakeyword{begin}%
 | 
| 
 | 
    16  | 
\endisatagtheory
  | 
| 
 | 
    17  | 
{\isafoldtheory}%
 | 
| 
 | 
    18  | 
%
  | 
| 
 | 
    19  | 
\isadelimtheory
  | 
| 
 | 
    20  | 
%
  | 
| 
 | 
    21  | 
\endisadelimtheory
  | 
| 
 | 
    22  | 
%
  | 
| 
 | 
    23  | 
\isamarkupchapter{Inner syntax --- the term language \label{ch:inner-syntax}%
 | 
| 
 | 
    24  | 
}
  | 
| 
 | 
    25  | 
\isamarkuptrue%
  | 
| 
 | 
    26  | 
%
  | 
| 
 | 
    27  | 
\isamarkupsection{Printing logical entities%
 | 
| 
 | 
    28  | 
}
  | 
| 
 | 
    29  | 
\isamarkuptrue%
  | 
| 
 | 
    30  | 
%
  | 
| 
 | 
    31  | 
\isamarkupsubsection{Diagnostic commands%
 | 
| 
 | 
    32  | 
}
  | 
| 
 | 
    33  | 
\isamarkuptrue%
  | 
| 
 | 
    34  | 
%
  | 
| 
 | 
    35  | 
\begin{isamarkuptext}%
 | 
| 
 | 
    36  | 
\begin{matharray}{rcl}
 | 
| 
 | 
    37  | 
    \indexdef{}{command}{typ}\hypertarget{command.typ}{\hyperlink{command.typ}{\mbox{\isa{\isacommand{typ}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}context\ {\isasymrightarrow}{\isachardoublequote}} \\
 | 
| 
 | 
    38  | 
    \indexdef{}{command}{term}\hypertarget{command.term}{\hyperlink{command.term}{\mbox{\isa{\isacommand{term}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}context\ {\isasymrightarrow}{\isachardoublequote}} \\
 | 
| 
 | 
    39  | 
    \indexdef{}{command}{prop}\hypertarget{command.prop}{\hyperlink{command.prop}{\mbox{\isa{\isacommand{prop}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}context\ {\isasymrightarrow}{\isachardoublequote}} \\
 | 
| 
 | 
    40  | 
    \indexdef{}{command}{thm}\hypertarget{command.thm}{\hyperlink{command.thm}{\mbox{\isa{\isacommand{thm}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}context\ {\isasymrightarrow}{\isachardoublequote}} \\
 | 
| 
 | 
    41  | 
    \indexdef{}{command}{prf}\hypertarget{command.prf}{\hyperlink{command.prf}{\mbox{\isa{\isacommand{prf}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}context\ {\isasymrightarrow}{\isachardoublequote}} \\
 | 
| 
 | 
    42  | 
    \indexdef{}{command}{full\_prf}\hypertarget{command.full-prf}{\hyperlink{command.full-prf}{\mbox{\isa{\isacommand{full{\isacharunderscore}prf}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}context\ {\isasymrightarrow}{\isachardoublequote}} \\
 | 
| 
 | 
    43  | 
    \indexdef{}{command}{pr}\hypertarget{command.pr}{\hyperlink{command.pr}{\mbox{\isa{\isacommand{pr}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}any\ {\isasymrightarrow}{\isachardoublequote}} \\
 | 
| 
 | 
    44  | 
  \end{matharray}
 | 
| 
 | 
    45  | 
  | 
| 
 | 
    46  | 
  These diagnostic commands assist interactive development by printing
  | 
| 
 | 
    47  | 
  internal logical entities in a human-readable fashion.
  | 
| 
 | 
    48  | 
  | 
| 
 | 
    49  | 
  \begin{rail}
 | 
| 
 | 
    50  | 
    'typ' modes? type
  | 
| 
 | 
    51  | 
    ;
  | 
| 
 | 
    52  | 
    'term' modes? term
  | 
| 
 | 
    53  | 
    ;
  | 
| 
 | 
    54  | 
    'prop' modes? prop
  | 
| 
 | 
    55  | 
    ;
  | 
| 
 | 
    56  | 
    'thm' modes? thmrefs
  | 
| 
 | 
    57  | 
    ;
  | 
| 
 | 
    58  | 
    ( 'prf' | 'full\_prf' ) modes? thmrefs?
  | 
| 
 | 
    59  | 
    ;
  | 
| 
 | 
    60  | 
    'pr' modes? nat? (',' nat)?
 | 
| 
 | 
    61  | 
    ;
  | 
| 
 | 
    62  | 
  | 
| 
 | 
    63  | 
    modes: '(' (name + ) ')'
 | 
| 
 | 
    64  | 
    ;
  | 
| 
 | 
    65  | 
  \end{rail}
 | 
| 
 | 
    66  | 
  | 
| 
 | 
    67  | 
  \begin{description}
 | 
| 
 | 
    68  | 
  | 
| 
 | 
    69  | 
  \item \hyperlink{command.typ}{\mbox{\isa{\isacommand{typ}}}}~\isa{{\isasymtau}} reads and prints types of the
 | 
| 
 | 
    70  | 
  meta-logic according to the current theory or proof context.
  | 
| 
 | 
    71  | 
  | 
| 
 | 
    72  | 
  \item \hyperlink{command.term}{\mbox{\isa{\isacommand{term}}}}~\isa{t} and \hyperlink{command.prop}{\mbox{\isa{\isacommand{prop}}}}~\isa{{\isasymphi}}
 | 
| 
 | 
    73  | 
  read, type-check and print terms or propositions according to the
  | 
| 
 | 
    74  | 
  current theory or proof context; the inferred type of \isa{t} is
 | 
| 
 | 
    75  | 
  output as well.  Note that these commands are also useful in
  | 
| 
 | 
    76  | 
  inspecting the current environment of term abbreviations.
  | 
| 
 | 
    77  | 
  | 
| 
 | 
    78  | 
  \item \hyperlink{command.thm}{\mbox{\isa{\isacommand{thm}}}}~\isa{{\isachardoublequote}a\isactrlsub {\isadigit{1}}\ {\isasymdots}\ a\isactrlsub n{\isachardoublequote}} retrieves
 | 
| 
 | 
    79  | 
  theorems from the current theory or proof context.  Note that any
  | 
| 
 | 
    80  | 
  attributes included in the theorem specifications are applied to a
  | 
| 
 | 
    81  | 
  temporary context derived from the current theory or proof; the
  | 
| 
 | 
    82  | 
  result is discarded, i.e.\ attributes involved in \isa{{\isachardoublequote}a\isactrlsub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ a\isactrlsub n{\isachardoublequote}} do not have any permanent effect.
 | 
| 
 | 
    83  | 
  | 
| 
 | 
    84  | 
  \item \hyperlink{command.prf}{\mbox{\isa{\isacommand{prf}}}} displays the (compact) proof term of the
 | 
| 
 | 
    85  | 
  current proof state (if present), or of the given theorems. Note
  | 
| 
 | 
    86  | 
  that this requires proof terms to be switched on for the current
  | 
| 
 | 
    87  | 
  object logic (see the ``Proof terms'' section of the Isabelle
  | 
| 
 | 
    88  | 
  reference manual for information on how to do this).
  | 
| 
 | 
    89  | 
  | 
| 
 | 
    90  | 
  \item \hyperlink{command.full-prf}{\mbox{\isa{\isacommand{full{\isacharunderscore}prf}}}} is like \hyperlink{command.prf}{\mbox{\isa{\isacommand{prf}}}}, but displays
 | 
| 
 | 
    91  | 
  the full proof term, i.e.\ also displays information omitted in the
  | 
| 
 | 
    92  | 
  compact proof term, which is denoted by ``\isa{{\isacharunderscore}}'' placeholders
 | 
| 
 | 
    93  | 
  there.
  | 
| 
 | 
    94  | 
  | 
| 
 | 
    95  | 
  \item \hyperlink{command.pr}{\mbox{\isa{\isacommand{pr}}}}~\isa{{\isachardoublequote}goals{\isacharcomma}\ prems{\isachardoublequote}} prints the current
 | 
| 
 | 
    96  | 
  proof state (if present), including the proof context, current facts
  | 
| 
 | 
    97  | 
  and goals.  The optional limit arguments affect the number of goals
  | 
| 
 | 
    98  | 
  and premises to be displayed, which is initially 10 for both.
  | 
| 
 | 
    99  | 
  Omitting limit values leaves the current setting unchanged.
  | 
| 
 | 
   100  | 
  | 
| 
 | 
   101  | 
  \end{description}
 | 
| 
 | 
   102  | 
  | 
| 
 | 
   103  | 
  All of the diagnostic commands above admit a list of \isa{modes}
 | 
| 
 | 
   104  | 
  to be specified, which is appended to the current print mode (see
  | 
| 
 | 
   105  | 
  also \cite{isabelle-ref}).  Thus the output behavior may be modified
 | 
| 
 | 
   106  | 
  according particular print mode features.  For example, \hyperlink{command.pr}{\mbox{\isa{\isacommand{pr}}}}~\isa{{\isachardoublequote}{\isacharparenleft}latex\ xsymbols{\isacharparenright}{\isachardoublequote}} would print the current proof state
 | 
| 
 | 
   107  | 
  with mathematical symbols and special characters represented in
  | 
| 
 | 
   108  | 
  {\LaTeX} source, according to the Isabelle style
 | 
| 
 | 
   109  | 
  \cite{isabelle-sys}.
 | 
| 
 | 
   110  | 
  | 
| 
 | 
   111  | 
  Note that antiquotations (cf.\ \secref{sec:antiq}) provide a more
 | 
| 
 | 
   112  | 
  systematic way to include formal items into the printed text
  | 
| 
 | 
   113  | 
  document.%
  | 
| 
 | 
   114  | 
\end{isamarkuptext}%
 | 
| 
 | 
   115  | 
\isamarkuptrue%
  | 
| 
 | 
   116  | 
%
  | 
| 
 | 
   117  | 
\isamarkupsubsection{Details of printed content%
 | 
| 
 | 
   118  | 
}
  | 
| 
 | 
   119  | 
\isamarkuptrue%
  | 
| 
 | 
   120  | 
%
  | 
| 
 | 
   121  | 
\begin{isamarkuptext}%
 | 
| 
 | 
   122  | 
\begin{mldecls} 
 | 
| 
 | 
   123  | 
    \indexml{show\_types}\verb|show_types: bool ref| & default \verb|false| \\
 | 
| 
 | 
   124  | 
    \indexml{show\_sorts}\verb|show_sorts: bool ref| & default \verb|false| \\
 | 
| 
 | 
   125  | 
    \indexml{show\_consts}\verb|show_consts: bool ref| & default \verb|false| \\
 | 
| 
 | 
   126  | 
    \indexml{long\_names}\verb|long_names: bool ref| & default \verb|false| \\
 | 
| 
 | 
   127  | 
    \indexml{short\_names}\verb|short_names: bool ref| & default \verb|false| \\
 | 
| 
 | 
   128  | 
    \indexml{unique\_names}\verb|unique_names: bool ref| & default \verb|true| \\
 | 
| 
 | 
   129  | 
    \indexml{show\_brackets}\verb|show_brackets: bool ref| & default \verb|false| \\
 | 
| 
 | 
   130  | 
    \indexml{eta\_contract}\verb|eta_contract: bool ref| & default \verb|true| \\
 | 
| 
 | 
   131  | 
    \indexml{goals\_limit}\verb|goals_limit: int ref| & default \verb|10| \\
 | 
| 
 | 
   132  | 
    \indexml{Proof.show\_main\_goal}\verb|Proof.show_main_goal: bool ref| & default \verb|false| \\
 | 
| 
 | 
   133  | 
    \indexml{show\_hyps}\verb|show_hyps: bool ref| & default \verb|false| \\
 | 
| 
 | 
   134  | 
    \indexml{show\_tags}\verb|show_tags: bool ref| & default \verb|false| \\
 | 
| 
 | 
   135  | 
    \indexml{show\_question\_marks}\verb|show_question_marks: bool ref| & default \verb|true| \\
 | 
| 
 | 
   136  | 
  \end{mldecls}
 | 
| 
 | 
   137  | 
  | 
| 
 | 
   138  | 
  These global ML variables control the detail of information that is
  | 
| 
 | 
   139  | 
  displayed for types, terms, theorems, goals etc.
  | 
| 
 | 
   140  | 
  | 
| 
 | 
   141  | 
  In interactive sessions, the user interface usually manages these
  | 
| 
 | 
   142  | 
  global parameters of the Isabelle process, even with some concept of
  | 
| 
 | 
   143  | 
  persistence.  Nonetheless it is occasionally useful to manipulate ML
  | 
| 
 | 
   144  | 
  variables directly, e.g.\ using \hyperlink{command.ML-val}{\mbox{\isa{\isacommand{ML{\isacharunderscore}val}}}} or \hyperlink{command.ML-command}{\mbox{\isa{\isacommand{ML{\isacharunderscore}command}}}}.
 | 
| 
 | 
   145  | 
  | 
| 
 | 
   146  | 
  Batch-mode logic sessions may be configured by putting appropriate
  | 
| 
 | 
   147  | 
  ML text directly into the \verb|ROOT.ML| file.
  | 
| 
 | 
   148  | 
  | 
| 
 | 
   149  | 
  \begin{description}
 | 
| 
 | 
   150  | 
  | 
| 
 | 
   151  | 
  \item \verb|show_types| and \verb|show_sorts| control printing of type
  | 
| 
 | 
   152  | 
  constraints for term variables, and sort constraints for type
  | 
| 
 | 
   153  | 
  variables.  By default, neither of these are shown in output.  If
  | 
| 
 | 
   154  | 
  \verb|show_sorts| is set to \verb|true|, types are always shown as
  | 
| 
 | 
   155  | 
  well.
  | 
| 
 | 
   156  | 
  | 
| 
 | 
   157  | 
  Note that displaying types and sorts may explain why a polymorphic
  | 
| 
 | 
   158  | 
  inference rule fails to resolve with some goal, or why a rewrite
  | 
| 
 | 
   159  | 
  rule does not apply as expected.
  | 
| 
 | 
   160  | 
  | 
| 
 | 
   161  | 
  \item \verb|show_consts| controls printing of types of constants when
  | 
| 
 | 
   162  | 
  displaying a goal state.
  | 
| 
 | 
   163  | 
  | 
| 
 | 
   164  | 
  Note that the output can be enormous, because polymorphic constants
  | 
| 
 | 
   165  | 
  often occur at several different type instances.
  | 
| 
 | 
   166  | 
  | 
| 
 | 
   167  | 
  \item \verb|long_names|, \verb|short_names|, and \verb|unique_names|
  | 
| 
 | 
   168  | 
  control the way of printing fully qualified internal names in
  | 
| 
 | 
   169  | 
  external form.  See also \secref{sec:antiq} for the document
 | 
| 
 | 
   170  | 
  antiquotation options of the same names.
  | 
| 
 | 
   171  | 
  | 
| 
 | 
   172  | 
  \item \verb|show_brackets| controls bracketing in pretty printed
  | 
| 
 | 
   173  | 
  output.  If set to \verb|true|, all sub-expressions of the pretty
  | 
| 
 | 
   174  | 
  printing tree will be parenthesized, even if this produces malformed
  | 
| 
 | 
   175  | 
  term syntax!  This crude way of showing the internal structure of
  | 
| 
 | 
   176  | 
  pretty printed entities may occasionally help to diagnose problems
  | 
| 
 | 
   177  | 
  with operator priorities, for example.
  | 
| 
 | 
   178  | 
  | 
| 
 | 
   179  | 
  \item \verb|eta_contract| controls \isa{{\isachardoublequote}{\isasymeta}{\isachardoublequote}}-contracted printing of
 | 
| 
 | 
   180  | 
  terms.
  | 
| 
 | 
   181  | 
  | 
| 
 | 
   182  | 
  The \isa{{\isasymeta}}-contraction law asserts \isa{{\isachardoublequote}{\isacharparenleft}{\isasymlambda}x{\isachardot}\ f\ x{\isacharparenright}\ {\isasymequiv}\ f{\isachardoublequote}},
 | 
| 
 | 
   183  | 
  provided \isa{x} is not free in \isa{f}.  It asserts
 | 
| 
 | 
   184  | 
  \emph{extensionality} of functions: \isa{{\isachardoublequote}f\ {\isasymequiv}\ g{\isachardoublequote}} if \isa{{\isachardoublequote}f\ x\ {\isasymequiv}\ g\ x{\isachardoublequote}} for all \isa{x}.  Higher-order unification frequently puts
 | 
| 
 | 
   185  | 
  terms into a fully \isa{{\isasymeta}}-expanded form.  For example, if \isa{F} has type \isa{{\isachardoublequote}{\isacharparenleft}{\isasymtau}\ {\isasymRightarrow}\ {\isasymtau}{\isacharparenright}\ {\isasymRightarrow}\ {\isasymtau}{\isachardoublequote}} then its expanded form is \isa{{\isachardoublequote}{\isasymlambda}h{\isachardot}\ F\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ h\ x{\isacharparenright}{\isachardoublequote}}.
 | 
| 
 | 
   186  | 
  | 
| 
 | 
   187  | 
  Setting \verb|eta_contract| makes Isabelle perform \isa{{\isasymeta}}-contractions before printing, so that \isa{{\isachardoublequote}{\isasymlambda}h{\isachardot}\ F\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ h\ x{\isacharparenright}{\isachardoublequote}}
 | 
| 
 | 
   188  | 
  appears simply as \isa{F}.
 | 
| 
 | 
   189  | 
  | 
| 
 | 
   190  | 
  Note that the distinction between a term and its \isa{{\isasymeta}}-expanded
 | 
| 
 | 
   191  | 
  form occasionally matters.  While higher-order resolution and
  | 
| 
 | 
   192  | 
  rewriting operate modulo \isa{{\isachardoublequote}{\isasymalpha}{\isasymbeta}{\isasymeta}{\isachardoublequote}}-conversion, some other tools
 | 
| 
 | 
   193  | 
  might look at terms more discretely.
  | 
| 
 | 
   194  | 
  | 
| 
 | 
   195  | 
  \item \verb|goals_limit| controls the maximum number of subgoals to
  | 
| 
 | 
   196  | 
  be shown in goal output.
  | 
| 
 | 
   197  | 
  | 
| 
 | 
   198  | 
  \item \verb|Proof.show_main_goal| controls whether the main result to
  | 
| 
 | 
   199  | 
  be proven should be displayed.  This information might be relevant
  | 
| 
 | 
   200  | 
  for schematic goals, to inspect the current claim that has been
  | 
| 
 | 
   201  | 
  synthesized so far.
  | 
| 
 | 
   202  | 
  | 
| 
 | 
   203  | 
  \item \verb|show_hyps| controls printing of implicit hypotheses of
  | 
| 
 | 
   204  | 
  local facts.  Normally, only those hypotheses are displayed that are
  | 
| 
 | 
   205  | 
  \emph{not} covered by the assumptions of the current context: this
 | 
| 
 | 
   206  | 
  situation indicates a fault in some tool being used.
  | 
| 
 | 
   207  | 
  | 
| 
 | 
   208  | 
  By setting \verb|show_hyps| to \verb|true|, output of \emph{all}
 | 
| 
 | 
   209  | 
  hypotheses can be enforced, which is occasionally useful for
  | 
| 
 | 
   210  | 
  diagnostic purposes.
  | 
| 
 | 
   211  | 
  | 
| 
 | 
   212  | 
  \item \verb|show_tags| controls printing of extra annotations within
  | 
| 
 | 
   213  | 
  theorems, such as internal position information, or the case names
  | 
| 
 | 
   214  | 
  being attached by the attribute \hyperlink{attribute.case-names}{\mbox{\isa{case{\isacharunderscore}names}}}.
 | 
| 
 | 
   215  | 
  | 
| 
 | 
   216  | 
  Note that the \hyperlink{attribute.tagged}{\mbox{\isa{tagged}}} and \hyperlink{attribute.untagged}{\mbox{\isa{untagged}}}
 | 
| 
 | 
   217  | 
  attributes provide low-level access to the collection of tags
  | 
| 
 | 
   218  | 
  associated with a theorem.
  | 
| 
 | 
   219  | 
  | 
| 
 | 
   220  | 
  \item \verb|show_question_marks| controls printing of question marks
  | 
| 
 | 
   221  | 
  for schematic variables, such as \isa{{\isacharquery}x}.  Only the leading
 | 
| 
 | 
   222  | 
  question mark is affected, the remaining text is unchanged
  | 
| 
 | 
   223  | 
  (including proper markup for schematic variables that might be
  | 
| 
 | 
   224  | 
  relevant for user interfaces).
  | 
| 
 | 
   225  | 
  | 
| 
 | 
   226  | 
  \end{description}%
 | 
| 
 | 
   227  | 
\end{isamarkuptext}%
 | 
| 
 | 
   228  | 
\isamarkuptrue%
  | 
| 
 | 
   229  | 
%
  | 
| 
 | 
   230  | 
\isamarkupsubsection{Printing limits%
 | 
| 
 | 
   231  | 
}
  | 
| 
 | 
   232  | 
\isamarkuptrue%
  | 
| 
 | 
   233  | 
%
  | 
| 
 | 
   234  | 
\begin{isamarkuptext}%
 | 
| 
 | 
   235  | 
\begin{mldecls}
 | 
| 
 | 
   236  | 
    \indexml{Pretty.setdepth}\verb|Pretty.setdepth: int -> unit| \\
 | 
| 
 | 
   237  | 
    \indexml{Pretty.setmargin}\verb|Pretty.setmargin: int -> unit| \\
 | 
| 
 | 
   238  | 
    \indexml{print\_depth}\verb|print_depth: int -> unit| \\
 | 
| 
 | 
   239  | 
  \end{mldecls}
 | 
| 
 | 
   240  | 
  | 
| 
 | 
   241  | 
  These ML functions set limits for pretty printed text.
  | 
| 
 | 
   242  | 
  | 
| 
 | 
   243  | 
  \begin{description}
 | 
| 
 | 
   244  | 
  | 
| 
 | 
   245  | 
  \item \verb|Pretty.setdepth|~\isa{d} tells the pretty printer to
 | 
| 
 | 
   246  | 
  limit the printing depth to \isa{d}.  This affects the display of
 | 
| 
 | 
   247  | 
  types, terms, theorems etc.  The default value is 0, which permits
  | 
| 
 | 
   248  | 
  printing to an arbitrary depth.  Other useful values for \isa{d}
 | 
| 
 | 
   249  | 
  are 10 and 20.
  | 
| 
 | 
   250  | 
  | 
| 
 | 
   251  | 
  \item \verb|Pretty.setmargin|~\isa{m} tells the pretty printer to
 | 
| 
 | 
   252  | 
  assume a right margin (page width) of \isa{m}.  The initial margin
 | 
| 
 | 
   253  | 
  is 76, but user interfaces might adapt the margin automatically when
  | 
| 
 | 
   254  | 
  resizing windows.
  | 
| 
 | 
   255  | 
  | 
| 
 | 
   256  | 
  \item \verb|print_depth|~\isa{n} limits the printing depth of the
 | 
| 
 | 
   257  | 
  ML toplevel pretty printer; the precise effect depends on the ML
  | 
| 
 | 
   258  | 
  compiler and run-time system.  Typically \isa{n} should be less
 | 
| 
 | 
   259  | 
  than 10.  Bigger values such as 100--1000 are useful for debugging.
  | 
| 
 | 
   260  | 
  | 
| 
 | 
   261  | 
  \end{description}%
 | 
| 
 | 
   262  | 
\end{isamarkuptext}%
 | 
| 
 | 
   263  | 
\isamarkuptrue%
  | 
| 
 | 
   264  | 
%
  | 
| 
 | 
   265  | 
\isamarkupsection{Mixfix annotations%
 | 
| 
 | 
   266  | 
}
  | 
| 
 | 
   267  | 
\isamarkuptrue%
  | 
| 
 | 
   268  | 
%
  | 
| 
 | 
   269  | 
\begin{isamarkuptext}%
 | 
| 
 | 
   270  | 
Mixfix annotations specify concrete \emph{inner syntax} of
 | 
| 
 | 
   271  | 
  Isabelle types and terms.  Some commands such as \hyperlink{command.typedecl}{\mbox{\isa{\isacommand{typedecl}}}} admit infixes only, while \hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}} etc.\
 | 
| 
 | 
   272  | 
  support the full range of general mixfixes and binders.  Fixed
  | 
| 
 | 
   273  | 
  parameters in toplevel theorem statements, locale specifications
  | 
| 
 | 
   274  | 
  also admit mixfix annotations.
  | 
| 
 | 
   275  | 
  | 
| 
 | 
   276  | 
  \indexouternonterm{infix}\indexouternonterm{mixfix}\indexouternonterm{structmixfix}
 | 
| 
 | 
   277  | 
  \begin{rail}
 | 
| 
 | 
   278  | 
    infix: '(' ('infix' | 'infixl' | 'infixr') string nat ')'
 | 
| 
 | 
   279  | 
    ;
  | 
| 
 | 
   280  | 
    mixfix: infix | '(' string prios? nat? ')' | '(' 'binder' string prios? nat ')'
 | 
| 
 | 
   281  | 
    ;
  | 
| 
 | 
   282  | 
    structmixfix: mixfix | '(' 'structure' ')'
 | 
| 
 | 
   283  | 
    ;
  | 
| 
 | 
   284  | 
  | 
| 
 | 
   285  | 
    prios: '[' (nat + ',') ']'
  | 
| 
 | 
   286  | 
    ;
  | 
| 
 | 
   287  | 
  \end{rail}
 | 
| 
 | 
   288  | 
  | 
| 
 | 
   289  | 
  Here the \railtok{string} specifications refer to the actual mixfix
 | 
| 
 | 
   290  | 
  template, which may include literal text, spacing, blocks, and
  | 
| 
 | 
   291  | 
  arguments (denoted by ``\isa{{\isacharunderscore}}''); the special symbol
 | 
| 
 | 
   292  | 
  ``\verb|\<index>|'' (printed as ``\isa{{\isachardoublequote}{\isasymindex}{\isachardoublequote}}'') represents an index
 | 
| 
 | 
   293  | 
  argument that specifies an implicit structure reference (see also
  | 
| 
 | 
   294  | 
  \secref{sec:locale}).  Infix and binder declarations provide common
 | 
| 
 | 
   295  | 
  abbreviations for particular mixfix declarations.  So in practice,
  | 
| 
 | 
   296  | 
  mixfix templates mostly degenerate to literal text for concrete
  | 
| 
 | 
   297  | 
  syntax, such as ``\verb|++|'' for an infix symbol.
  | 
| 
 | 
   298  | 
  | 
| 
 | 
   299  | 
  \medskip In full generality, mixfix declarations work as follows.
  | 
| 
 | 
   300  | 
  Suppose a constant \isa{{\isachardoublequote}c\ {\isacharcolon}{\isacharcolon}\ {\isasymtau}\isactrlsub {\isadigit{1}}\ {\isasymRightarrow}\ {\isasymdots}\ {\isasymtau}\isactrlsub n\ {\isasymRightarrow}\ {\isasymtau}{\isachardoublequote}} is
 | 
| 
 | 
   301  | 
  annotated by \isa{{\isachardoublequote}{\isacharparenleft}mixfix\ {\isacharbrackleft}p\isactrlsub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ p\isactrlsub n{\isacharbrackright}\ p{\isacharparenright}{\isachardoublequote}}, where \isa{{\isachardoublequote}mixfix{\isachardoublequote}} is a string \isa{{\isachardoublequote}d\isactrlsub {\isadigit{0}}\ {\isacharunderscore}\ d\isactrlsub {\isadigit{1}}\ {\isacharunderscore}\ {\isasymdots}\ {\isacharunderscore}\ d\isactrlsub n{\isachardoublequote}} consisting of
 | 
| 
 | 
   302  | 
  delimiters that surround argument positions as indicated by
  | 
| 
 | 
   303  | 
  underscores.
  | 
| 
 | 
   304  | 
  | 
| 
 | 
   305  | 
  Altogether this determines a production for a context-free priority
  | 
| 
 | 
   306  | 
  grammar, where for each argument \isa{{\isachardoublequote}i{\isachardoublequote}} the syntactic category
 | 
| 
 | 
   307  | 
  is determined by \isa{{\isachardoublequote}{\isasymtau}\isactrlsub i{\isachardoublequote}} (with priority \isa{{\isachardoublequote}p\isactrlsub i{\isachardoublequote}}), and
 | 
| 
 | 
   308  | 
  the result category is determined from \isa{{\isachardoublequote}{\isasymtau}{\isachardoublequote}} (with
 | 
| 
 | 
   309  | 
  priority \isa{{\isachardoublequote}p{\isachardoublequote}}).  Priority specifications are optional, with
 | 
| 
 | 
   310  | 
  default 0 for arguments and 1000 for the result.
  | 
| 
 | 
   311  | 
  | 
| 
 | 
   312  | 
  Since \isa{{\isachardoublequote}{\isasymtau}{\isachardoublequote}} may be again a function type, the constant
 | 
| 
 | 
   313  | 
  type scheme may have more argument positions than the mixfix
  | 
| 
 | 
   314  | 
  pattern.  Printing a nested application \isa{{\isachardoublequote}c\ t\isactrlsub {\isadigit{1}}\ {\isasymdots}\ t\isactrlsub m{\isachardoublequote}} for
 | 
| 
 | 
   315  | 
  \isa{{\isachardoublequote}m\ {\isachargreater}\ n{\isachardoublequote}} works by attaching concrete notation only to the
 | 
| 
 | 
   316  | 
  innermost part, essentially by printing \isa{{\isachardoublequote}{\isacharparenleft}c\ t\isactrlsub {\isadigit{1}}\ {\isasymdots}\ t\isactrlsub n{\isacharparenright}\ {\isasymdots}\ t\isactrlsub m{\isachardoublequote}}
 | 
| 
 | 
   317  | 
  instead.  If a term has fewer arguments than specified in the mixfix
  | 
| 
 | 
   318  | 
  template, the concrete syntax is ignored.
  | 
| 
 | 
   319  | 
  | 
| 
 | 
   320  | 
  \medskip A mixfix template may also contain additional directives
  | 
| 
 | 
   321  | 
  for pretty printing, notably spaces, blocks, and breaks.  The
  | 
| 
 | 
   322  | 
  general template format is a sequence over any of the following
  | 
| 
 | 
   323  | 
  entities.
  | 
| 
 | 
   324  | 
  | 
| 
 | 
   325  | 
  \begin{description}
 | 
| 
 | 
   326  | 
  | 
| 
 | 
   327  | 
  \item \isa{{\isachardoublequote}d{\isachardoublequote}} is a delimiter, namely a non-empty sequence of
 | 
| 
 | 
   328  | 
  characters other than the following special characters:
  | 
| 
 | 
   329  | 
  | 
| 
 | 
   330  | 
  \smallskip
  | 
| 
 | 
   331  | 
  \begin{tabular}{ll}
 | 
| 
 | 
   332  | 
    \verb|'| & single quote \\
  | 
| 
 | 
   333  | 
    \verb|_| & underscore \\
  | 
| 
 | 
   334  | 
    \isa{{\isachardoublequote}{\isasymindex}{\isachardoublequote}} & index symbol \\
 | 
| 
 | 
   335  | 
    \verb|(| & open parenthesis \\
  | 
| 
 | 
   336  | 
    \verb|)| & close parenthesis \\
  | 
| 
 | 
   337  | 
    \verb|/| & slash \\
  | 
| 
 | 
   338  | 
  \end{tabular}
 | 
| 
 | 
   339  | 
  \medskip
  | 
| 
 | 
   340  | 
  | 
| 
 | 
   341  | 
  \item \verb|'| escapes the special meaning of these
  | 
| 
 | 
   342  | 
  meta-characters, producing a literal version of the following
  | 
| 
 | 
   343  | 
  character, unless that is a blank.
  | 
| 
 | 
   344  | 
  | 
| 
 | 
   345  | 
  A single quote followed by a blank separates delimiters, without
  | 
| 
 | 
   346  | 
  affecting printing, but input tokens may have additional white space
  | 
| 
 | 
   347  | 
  here.
  | 
| 
 | 
   348  | 
  | 
| 
 | 
   349  | 
  \item \verb|_| is an argument position, which stands for a
  | 
| 
 | 
   350  | 
  certain syntactic category in the underlying grammar.
  | 
| 
 | 
   351  | 
  | 
| 
 | 
   352  | 
  \item \isa{{\isachardoublequote}{\isasymindex}{\isachardoublequote}} is an indexed argument position; this is the place
 | 
| 
 | 
   353  | 
  where implicit structure arguments can be attached.
  | 
| 
 | 
   354  | 
  | 
| 
 | 
   355  | 
  \item \isa{{\isachardoublequote}s{\isachardoublequote}} is a non-empty sequence of spaces for printing.
 | 
| 
 | 
   356  | 
  This and the following specifications do not affect parsing at all.
  | 
| 
 | 
   357  | 
  | 
| 
 | 
   358  | 
  \item \verb|(|\isa{n} opens a pretty printing block.  The
 | 
| 
 | 
   359  | 
  optional number specifies how much indentation to add when a line
  | 
| 
 | 
   360  | 
  break occurs within the block.  If the parenthesis is not followed
  | 
| 
 | 
   361  | 
  by digits, the indentation defaults to 0.  A block specified via
  | 
| 
 | 
   362  | 
  \verb|(00| is unbreakable.
  | 
| 
 | 
   363  | 
  | 
| 
 | 
   364  | 
  \item \verb|)| closes a pretty printing block.
  | 
| 
 | 
   365  | 
  | 
| 
 | 
   366  | 
  \item \verb|//| forces a line break.
  | 
| 
 | 
   367  | 
  | 
| 
 | 
   368  | 
  \item \verb|/|\isa{s} allows a line break.  Here \isa{s}
 | 
| 
 | 
   369  | 
  stands for the string of spaces (zero or more) right after the
  | 
| 
 | 
   370  | 
  slash.  These spaces are printed if the break is \emph{not} taken.
 | 
| 
 | 
   371  | 
  | 
| 
 | 
   372  | 
  \end{description}
 | 
| 
 | 
   373  | 
  | 
| 
 | 
   374  | 
  For example, the template \verb|(_ +/ _)| specifies an infix
  | 
| 
 | 
   375  | 
  operator.  There are two argument positions; the delimiter
  | 
| 
 | 
   376  | 
  \verb|+| is preceded by a space and followed by a space or
  | 
| 
 | 
   377  | 
  line break; the entire phrase is a pretty printing block.
  | 
| 
 | 
   378  | 
  | 
| 
 | 
   379  | 
  The general idea of pretty printing with blocks and breaks is also
  | 
| 
 | 
   380  | 
  described in \cite{paulson-ml2}.%
 | 
| 
 | 
   381  | 
\end{isamarkuptext}%
 | 
| 
 | 
   382  | 
\isamarkuptrue%
  | 
| 
 | 
   383  | 
%
  | 
| 
 | 
   384  | 
\isamarkupsection{Explicit term notation%
 | 
| 
 | 
   385  | 
}
  | 
| 
 | 
   386  | 
\isamarkuptrue%
  | 
| 
 | 
   387  | 
%
  | 
| 
 | 
   388  | 
\begin{isamarkuptext}%
 | 
| 
 | 
   389  | 
\begin{matharray}{rcll}
 | 
| 
 | 
   390  | 
    \indexdef{}{command}{notation}\hypertarget{command.notation}{\hyperlink{command.notation}{\mbox{\isa{\isacommand{notation}}}}} & : & \isa{{\isachardoublequote}local{\isacharunderscore}theory\ {\isasymrightarrow}\ local{\isacharunderscore}theory{\isachardoublequote}} \\
 | 
| 
 | 
   391  | 
    \indexdef{}{command}{no\_notation}\hypertarget{command.no-notation}{\hyperlink{command.no-notation}{\mbox{\isa{\isacommand{no{\isacharunderscore}notation}}}}} & : & \isa{{\isachardoublequote}local{\isacharunderscore}theory\ {\isasymrightarrow}\ local{\isacharunderscore}theory{\isachardoublequote}} \\
 | 
| 
 | 
   392  | 
  \end{matharray}
 | 
| 
 | 
   393  | 
  | 
| 
 | 
   394  | 
  \begin{rail}
 | 
| 
 | 
   395  | 
    ('notation' | 'no\_notation') target? mode? (nameref structmixfix + 'and')
 | 
| 
 | 
   396  | 
    ;
  | 
| 
 | 
   397  | 
  \end{rail}
 | 
| 
 | 
   398  | 
  | 
| 
 | 
   399  | 
  \begin{description}
 | 
| 
 | 
   400  | 
  | 
| 
 | 
   401  | 
  \item \hyperlink{command.notation}{\mbox{\isa{\isacommand{notation}}}}~\isa{{\isachardoublequote}c\ {\isacharparenleft}mx{\isacharparenright}{\isachardoublequote}} associates mixfix
 | 
| 
 | 
   402  | 
  syntax with an existing constant or fixed variable.  This is a
  | 
| 
 | 
   403  | 
  robust interface to the underlying \hyperlink{command.syntax}{\mbox{\isa{\isacommand{syntax}}}} primitive
 | 
| 
 | 
   404  | 
  (\secref{sec:syn-trans}).  Type declaration and internal syntactic
 | 
| 
 | 
   405  | 
  representation of the given entity is retrieved from the context.
  | 
| 
 | 
   406  | 
  
  | 
| 
 | 
   407  | 
  \item \hyperlink{command.no-notation}{\mbox{\isa{\isacommand{no{\isacharunderscore}notation}}}} is similar to \hyperlink{command.notation}{\mbox{\isa{\isacommand{notation}}}},
 | 
| 
 | 
   408  | 
  but removes the specified syntax annotation from the present
  | 
| 
 | 
   409  | 
  context.
  | 
| 
 | 
   410  | 
  | 
| 
 | 
   411  | 
  \end{description}%
 | 
| 
 | 
   412  | 
\end{isamarkuptext}%
 | 
| 
 | 
   413  | 
\isamarkuptrue%
  | 
| 
 | 
   414  | 
%
  | 
| 
 | 
   415  | 
\isamarkupsection{The Pure syntax \label{sec:pure-syntax}%
 | 
| 
 | 
   416  | 
}
  | 
| 
 | 
   417  | 
\isamarkuptrue%
  | 
| 
 | 
   418  | 
%
  | 
| 
 | 
   419  | 
\isamarkupsubsection{Priority grammars \label{sec:priority-grammar}%
 | 
| 
 | 
   420  | 
}
  | 
| 
 | 
   421  | 
\isamarkuptrue%
  | 
| 
 | 
   422  | 
%
  | 
| 
 | 
   423  | 
\begin{isamarkuptext}%
 | 
| 
 | 
   424  | 
A context-free grammar consists of a set of \emph{terminal
 | 
| 
 | 
   425  | 
  symbols}, a set of \emph{nonterminal symbols} and a set of
 | 
| 
 | 
   426  | 
  \emph{productions}.  Productions have the form \isa{{\isachardoublequote}A\ {\isacharequal}\ {\isasymgamma}{\isachardoublequote}},
 | 
| 
 | 
   427  | 
  where \isa{A} is a nonterminal and \isa{{\isasymgamma}} is a string of
 | 
| 
 | 
   428  | 
  terminals and nonterminals.  One designated nonterminal is called
  | 
| 
 | 
   429  | 
  the \emph{root symbol}.  The language defined by the grammar
 | 
| 
 | 
   430  | 
  consists of all strings of terminals that can be derived from the
  | 
| 
 | 
   431  | 
  root symbol by applying productions as rewrite rules.
  | 
| 
 | 
   432  | 
  | 
| 
 | 
   433  | 
  The standard Isabelle parser for inner syntax uses a \emph{priority
 | 
| 
 | 
   434  | 
  grammar}.  Each nonterminal is decorated by an integer priority:
  | 
| 
 | 
   435  | 
  \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup p\isactrlsup {\isacharparenright}{\isachardoublequote}}.  In a derivation, \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup p\isactrlsup {\isacharparenright}{\isachardoublequote}} may be rewritten
 | 
| 
 | 
   436  | 
  using a production \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup q\isactrlsup {\isacharparenright}\ {\isacharequal}\ {\isasymgamma}{\isachardoublequote}} only if \isa{{\isachardoublequote}p\ {\isasymle}\ q{\isachardoublequote}}.  Any
 | 
| 
 | 
   437  | 
  priority grammar can be translated into a normal context-free
  | 
| 
 | 
   438  | 
  grammar by introducing new nonterminals and productions.
  | 
| 
 | 
   439  | 
  | 
| 
 | 
   440  | 
  \medskip Formally, a set of context free productions \isa{G}
 | 
| 
 | 
   441  | 
  induces a derivation relation \isa{{\isachardoublequote}{\isasymlongrightarrow}\isactrlsub G{\isachardoublequote}} as follows.  Let \isa{{\isasymalpha}} and \isa{{\isasymbeta}} denote strings of terminal or nonterminal symbols.
 | 
| 
 | 
   442  | 
  Then \isa{{\isachardoublequote}{\isasymalpha}\ A\isactrlsup {\isacharparenleft}\isactrlsup p\isactrlsup {\isacharparenright}\ {\isasymbeta}\ {\isasymlongrightarrow}\isactrlsub G\ {\isasymalpha}\ {\isasymgamma}\ {\isasymbeta}{\isachardoublequote}} holds if and only if \isa{G}
 | 
| 
 | 
   443  | 
  contains some production \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup q\isactrlsup {\isacharparenright}\ {\isacharequal}\ {\isasymgamma}{\isachardoublequote}} for \isa{{\isachardoublequote}p\ {\isasymle}\ q{\isachardoublequote}}.
 | 
| 
 | 
   444  | 
  | 
| 
 | 
   445  | 
  \medskip The following grammar for arithmetic expressions
  | 
| 
 | 
   446  | 
  demonstrates how binding power and associativity of operators can be
  | 
| 
 | 
   447  | 
  enforced by priorities.
  | 
| 
 | 
   448  | 
  | 
| 
 | 
   449  | 
  \begin{center}
 | 
| 
 | 
   450  | 
  \begin{tabular}{rclr}
 | 
| 
 | 
   451  | 
  \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharequal}{\isachardoublequote}} & \verb|(| \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{0}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \verb|)| \\
 | 
| 
 | 
   452  | 
  \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharequal}{\isachardoublequote}} & \verb|0| \\
 | 
| 
 | 
   453  | 
  \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{0}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharequal}{\isachardoublequote}} & \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{0}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \verb|+| \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \\
 | 
| 
 | 
   454  | 
  \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{2}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharequal}{\isachardoublequote}} & \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{3}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \verb|*| \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{2}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \\
 | 
| 
 | 
   455  | 
  \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{3}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharequal}{\isachardoublequote}} & \verb|-| \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{3}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \\
 | 
| 
 | 
   456  | 
  \end{tabular}
 | 
| 
 | 
   457  | 
  \end{center}
 | 
| 
 | 
   458  | 
  The choice of priorities determines that \verb|-| binds
  | 
| 
 | 
   459  | 
  tighter than \verb|*|, which binds tighter than \verb|+|.  Furthermore \verb|+| associates to the left and
  | 
| 
 | 
   460  | 
  \verb|*| to the right.
  | 
| 
 | 
   461  | 
  | 
| 
 | 
   462  | 
  \medskip For clarity, grammars obey these conventions:
  | 
| 
 | 
   463  | 
  \begin{itemize}
 | 
| 
 | 
   464  | 
  | 
| 
 | 
   465  | 
  \item All priorities must lie between 0 and 1000.
  | 
| 
 | 
   466  | 
  | 
| 
 | 
   467  | 
  \item Priority 0 on the right-hand side and priority 1000 on the
  | 
| 
 | 
   468  | 
  left-hand side may be omitted.
  | 
| 
 | 
   469  | 
  | 
| 
 | 
   470  | 
  \item The production \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup p\isactrlsup {\isacharparenright}\ {\isacharequal}\ {\isasymalpha}{\isachardoublequote}} is written as \isa{{\isachardoublequote}A\ {\isacharequal}\ {\isasymalpha}\ {\isacharparenleft}p{\isacharparenright}{\isachardoublequote}}, i.e.\ the priority of the left-hand side actually appears in
 | 
| 
 | 
   471  | 
  a column on the far right.
  | 
| 
 | 
   472  | 
  | 
| 
 | 
   473  | 
  \item Alternatives are separated by \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}}.
 | 
| 
 | 
   474  | 
  | 
| 
 | 
   475  | 
  \item Repetition is indicated by dots \isa{{\isachardoublequote}{\isacharparenleft}{\isasymdots}{\isacharparenright}{\isachardoublequote}} in an informal
 | 
| 
 | 
   476  | 
  but obvious way.
  | 
| 
 | 
   477  | 
  | 
| 
 | 
   478  | 
  \end{itemize}
 | 
| 
 | 
   479  | 
  | 
| 
 | 
   480  | 
  Using these conventions, the example grammar specification above
  | 
| 
 | 
   481  | 
  takes the form:
  | 
| 
 | 
   482  | 
  \begin{center}
 | 
| 
 | 
   483  | 
  \begin{tabular}{rclc}
 | 
| 
 | 
   484  | 
    \isa{A} & \isa{{\isachardoublequote}{\isacharequal}{\isachardoublequote}} & \verb|(| \isa{A} \verb|)| \\
 | 
| 
 | 
   485  | 
              & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|0| & \qquad\qquad \\
 | 
| 
 | 
   486  | 
              & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{A} \verb|+| \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{0}}{\isacharparenright}{\isachardoublequote}} \\
 | 
| 
 | 
   487  | 
              & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{3}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \verb|*| \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{2}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{2}}{\isacharparenright}{\isachardoublequote}} \\
 | 
| 
 | 
   488  | 
              & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|-| \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{3}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{3}}{\isacharparenright}{\isachardoublequote}} \\
 | 
| 
 | 
   489  | 
  \end{tabular}
 | 
| 
 | 
   490  | 
  \end{center}%
 | 
| 
 | 
   491  | 
\end{isamarkuptext}%
 | 
| 
 | 
   492  | 
\isamarkuptrue%
  | 
| 
 | 
   493  | 
%
  | 
| 
 | 
   494  | 
\isamarkupsubsection{The Pure grammar%
 | 
| 
 | 
   495  | 
}
  | 
| 
 | 
   496  | 
\isamarkuptrue%
  | 
| 
 | 
   497  | 
%
  | 
| 
 | 
   498  | 
\begin{isamarkuptext}%
 | 
| 
 | 
   499  | 
The priority grammar of the \isa{{\isachardoublequote}Pure{\isachardoublequote}} theory is defined as follows:
 | 
| 
 | 
   500  | 
  | 
| 
 | 
   501  | 
  %FIXME syntax for "index" (?)
  | 
| 
 | 
   502  | 
  %FIXME "op" versions of ==> etc. (?)
  | 
| 
 | 
   503  | 
  | 
| 
 | 
   504  | 
  \begin{center}
 | 
| 
 | 
   505  | 
  \begin{supertabular}{rclr}
 | 
| 
 | 
   506  | 
  | 
| 
 | 
   507  | 
  \indexdef{inner}{syntax}{any}\hypertarget{syntax.inner.any}{\hyperlink{syntax.inner.any}{\mbox{\isa{any}}}} & = & \isa{{\isachardoublequote}prop\ \ {\isacharbar}\ \ logic{\isachardoublequote}} \\\\
 | 
| 
 | 
   508  | 
  | 
| 
 | 
   509  | 
  \indexdef{inner}{syntax}{prop}\hypertarget{syntax.inner.prop}{\hyperlink{syntax.inner.prop}{\mbox{\isa{prop}}}} & = & \verb|(| \isa{prop} \verb|)| \\
 | 
| 
 | 
   510  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}prop\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{4}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \verb|::| \isa{type} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{3}}{\isacharparenright}{\isachardoublequote}} \\
 | 
| 
 | 
   511  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}any\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{3}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \verb|=?=| \isa{{\isachardoublequote}any\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{2}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{2}}{\isacharparenright}{\isachardoublequote}} \\
 | 
| 
 | 
   512  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}any\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{3}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \verb|==| \isa{{\isachardoublequote}any\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{2}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{2}}{\isacharparenright}{\isachardoublequote}} \\
 | 
| 
 | 
   513  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}any\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{3}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \isa{{\isachardoublequote}{\isasymequiv}{\isachardoublequote}} \isa{{\isachardoublequote}any\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{2}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{2}}{\isacharparenright}{\isachardoublequote}} \\
 | 
| 
28857
 | 
   514  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}prop\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{3}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \verb|&&&| \isa{{\isachardoublequote}prop\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{2}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{2}}{\isacharparenright}{\isachardoublequote}} \\
 | 
| 
28762
 | 
   515  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}prop\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{2}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \verb|==>| \isa{{\isachardoublequote}prop\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{1}}{\isacharparenright}{\isachardoublequote}} \\
 | 
| 
 | 
   516  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}prop\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{2}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \isa{{\isachardoublequote}{\isasymLongrightarrow}{\isachardoublequote}} \isa{{\isachardoublequote}prop\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{1}}{\isacharparenright}{\isachardoublequote}} \\
 | 
| 
 | 
   517  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|[|\verb,|,\verb|| \isa{prop} \verb|;| \isa{{\isachardoublequote}{\isasymdots}{\isachardoublequote}} \verb|;| \isa{prop} \verb||\verb,|,\verb|]| \verb|==>| \isa{{\isachardoublequote}prop\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{1}}{\isacharparenright}{\isachardoublequote}} \\
 | 
| 
 | 
   518  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isasymlbrakk}{\isachardoublequote}} \isa{prop} \verb|;| \isa{{\isachardoublequote}{\isasymdots}{\isachardoublequote}} \verb|;| \isa{prop} \isa{{\isachardoublequote}{\isasymrbrakk}{\isachardoublequote}} \isa{{\isachardoublequote}{\isasymLongrightarrow}{\isachardoublequote}} \isa{{\isachardoublequote}prop\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{1}}{\isacharparenright}{\isachardoublequote}} \\
 | 
| 
 | 
   519  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|!!| \isa{idts} \verb|.| \isa{prop} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{0}}{\isacharparenright}{\isachardoublequote}} \\
 | 
| 
 | 
   520  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isasymAnd}{\isachardoublequote}} \isa{idts} \verb|.| \isa{prop} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{0}}{\isacharparenright}{\isachardoublequote}} \\
 | 
| 
 | 
   521  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|OFCLASS| \verb|(| \isa{type} \verb|,| \isa{logic} \verb|)| \\
 | 
| 
 | 
   522  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|SORT_CONSTRAINT| \verb|(| \isa{type} \verb|)| \\
 | 
| 
28857
 | 
   523  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|TERM| \isa{logic} \\
 | 
| 
28762
 | 
   524  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|PROP| \isa{aprop} \\\\
 | 
| 
 | 
   525  | 
  | 
| 
28857
 | 
   526  | 
  \indexdef{inner}{syntax}{aprop}\hypertarget{syntax.inner.aprop}{\hyperlink{syntax.inner.aprop}{\mbox{\isa{aprop}}}} & = & \verb|(| \isa{aprop} \verb|)| \\
 | 
| 
 | 
   527  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}id\ \ {\isacharbar}\ \ longid\ \ {\isacharbar}\ \ var\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|_|\isa{{\isachardoublequote}\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|...| \\
 | 
| 
 | 
   528  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|CONST| \isa{{\isachardoublequote}id\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|CONST| \isa{{\isachardoublequote}longid{\isachardoublequote}} \\
 | 
| 
28762
 | 
   529  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}logic\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isacharparenright}\ \ any\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isacharparenright}\ {\isasymdots}\ any\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{9}}{\isadigit{9}}{\isadigit{9}}{\isacharparenright}{\isachardoublequote}} \\\\
 | 
| 
 | 
   530  | 
  | 
| 
 | 
   531  | 
  \indexdef{inner}{syntax}{logic}\hypertarget{syntax.inner.logic}{\hyperlink{syntax.inner.logic}{\mbox{\isa{logic}}}} & = & \verb|(| \isa{logic} \verb|)| \\
 | 
| 
 | 
   532  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}logic\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{4}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \verb|::| \isa{type} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{3}}{\isacharparenright}{\isachardoublequote}} \\
 | 
| 
 | 
   533  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}id\ \ {\isacharbar}\ \ longid\ \ {\isacharbar}\ \ var\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|_|\isa{{\isachardoublequote}\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|...| \\
 | 
| 
28857
 | 
   534  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|CONST| \isa{{\isachardoublequote}id\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|CONST| \isa{{\isachardoublequote}longid{\isachardoublequote}} \\
 | 
| 
28762
 | 
   535  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}logic\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isacharparenright}\ \ any\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isacharparenright}\ {\isasymdots}\ any\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{9}}{\isadigit{9}}{\isadigit{9}}{\isacharparenright}{\isachardoublequote}} \\
 | 
| 
 | 
   536  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|%| \isa{pttrns} \verb|.| \isa{{\isachardoublequote}any\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{3}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{3}}{\isacharparenright}{\isachardoublequote}} \\
 | 
| 
 | 
   537  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isasymlambda}} \isa{pttrns} \verb|.| \isa{{\isachardoublequote}any\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{3}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{3}}{\isacharparenright}{\isachardoublequote}} \\
 | 
| 
 | 
   538  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|TYPE| \verb|(| \isa{type} \verb|)| \\\\
 | 
| 
 | 
   539  | 
  | 
| 
 | 
   540  | 
  \indexdef{inner}{syntax}{idt}\hypertarget{syntax.inner.idt}{\hyperlink{syntax.inner.idt}{\mbox{\isa{idt}}}} & = & \verb|(| \isa{idt} \verb|)|\isa{{\isachardoublequote}\ \ {\isacharbar}\ \ id\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|_| \\
 | 
| 
 | 
   541  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{id} \verb|::| \isa{type} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{0}}{\isacharparenright}{\isachardoublequote}} \\
 | 
| 
 | 
   542  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|_| \verb|::| \isa{type} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{0}}{\isacharparenright}{\isachardoublequote}} \\\\
 | 
| 
 | 
   543  | 
  | 
| 
 | 
   544  | 
  \indexdef{inner}{syntax}{idts}\hypertarget{syntax.inner.idts}{\hyperlink{syntax.inner.idts}{\mbox{\isa{idts}}}} & = & \isa{{\isachardoublequote}idt\ \ {\isacharbar}\ \ idt\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isacharparenright}\ idts{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{0}}{\isacharparenright}{\isachardoublequote}} \\\\
 | 
| 
 | 
   545  | 
  | 
| 
 | 
   546  | 
  \indexdef{inner}{syntax}{pttrn}\hypertarget{syntax.inner.pttrn}{\hyperlink{syntax.inner.pttrn}{\mbox{\isa{pttrn}}}} & = & \isa{idt} \\\\
 | 
| 
 | 
   547  | 
  | 
| 
 | 
   548  | 
  \indexdef{inner}{syntax}{pttrns}\hypertarget{syntax.inner.pttrns}{\hyperlink{syntax.inner.pttrns}{\mbox{\isa{pttrns}}}} & = & \isa{{\isachardoublequote}pttrn\ \ {\isacharbar}\ \ pttrn\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isacharparenright}\ pttrns{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{0}}{\isacharparenright}{\isachardoublequote}} \\\\
 | 
| 
 | 
   549  | 
  | 
| 
 | 
   550  | 
  \indexdef{inner}{syntax}{type}\hypertarget{syntax.inner.type}{\hyperlink{syntax.inner.type}{\mbox{\isa{type}}}} & = & \verb|(| \isa{type} \verb|)| \\
 | 
| 
 | 
   551  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}tid\ \ {\isacharbar}\ \ tvar\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|_| \\
 | 
| 
 | 
   552  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}tid{\isachardoublequote}} \verb|::| \isa{{\isachardoublequote}sort\ \ {\isacharbar}\ \ tvar\ \ {\isachardoublequote}}\verb|::| \isa{{\isachardoublequote}sort\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|_| \verb|::| \isa{{\isachardoublequote}sort{\isachardoublequote}} \\
 | 
| 
 | 
   553  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}id\ \ {\isacharbar}\ \ type\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isacharparenright}\ id\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|(| \isa{type} \verb|,| \isa{{\isachardoublequote}{\isasymdots}{\isachardoublequote}} \verb|,| \isa{type} \verb|)| \isa{id} \\
 | 
| 
 | 
   554  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}longid\ \ {\isacharbar}\ \ type\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isacharparenright}\ longid\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|(| \isa{type} \verb|,| \isa{{\isachardoublequote}{\isasymdots}{\isachardoublequote}} \verb|,| \isa{type} \verb|)| \isa{longid} \\
 | 
| 
 | 
   555  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}type\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \verb|=>| \isa{type} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{0}}{\isacharparenright}{\isachardoublequote}} \\
 | 
| 
 | 
   556  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}type\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \isa{{\isachardoublequote}{\isasymRightarrow}{\isachardoublequote}} \isa{type} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{0}}{\isacharparenright}{\isachardoublequote}} \\
 | 
| 
 | 
   557  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|[| \isa{type} \verb|,| \isa{{\isachardoublequote}{\isasymdots}{\isachardoublequote}} \verb|,| \isa{type} \verb|]| \verb|=>| \isa{type} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{0}}{\isacharparenright}{\isachardoublequote}} \\
 | 
| 
 | 
   558  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|[| \isa{type} \verb|,| \isa{{\isachardoublequote}{\isasymdots}{\isachardoublequote}} \verb|,| \isa{type} \verb|]| \isa{{\isachardoublequote}{\isasymRightarrow}{\isachardoublequote}} \isa{type} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{0}}{\isacharparenright}{\isachardoublequote}} \\\\
 | 
| 
 | 
   559  | 
  | 
| 
 | 
   560  | 
  \indexdef{inner}{syntax}{sort}\hypertarget{syntax.inner.sort}{\hyperlink{syntax.inner.sort}{\mbox{\isa{sort}}}} & = & \isa{{\isachardoublequote}id\ \ {\isacharbar}\ \ longid\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|{}|\isa{{\isachardoublequote}\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|{| \isa{{\isachardoublequote}{\isacharparenleft}id\ {\isacharbar}\ longid{\isacharparenright}{\isachardoublequote}} \verb|,| \isa{{\isachardoublequote}{\isasymdots}{\isachardoublequote}} \verb|,| \isa{{\isachardoublequote}{\isacharparenleft}id\ {\isacharbar}\ longid{\isacharparenright}{\isachardoublequote}} \verb|}| \\
 | 
| 
 | 
   561  | 
  \end{supertabular}
 | 
| 
 | 
   562  | 
  \end{center}
 | 
| 
 | 
   563  | 
  | 
| 
 | 
   564  | 
  \medskip Here literal terminals are printed \verb|verbatim|;
  | 
| 
 | 
   565  | 
  see also \secref{sec:inner-lex} for further token categories of the
 | 
| 
 | 
   566  | 
  inner syntax.  The meaning of the nonterminals defined by the above
  | 
| 
 | 
   567  | 
  grammar is as follows:
  | 
| 
 | 
   568  | 
  | 
| 
 | 
   569  | 
  \begin{description}
 | 
| 
 | 
   570  | 
  | 
| 
 | 
   571  | 
  \item \indexref{inner}{syntax}{any}\hyperlink{syntax.inner.any}{\mbox{\isa{any}}} denotes any term.
 | 
| 
 | 
   572  | 
  | 
| 
 | 
   573  | 
  \item \indexref{inner}{syntax}{prop}\hyperlink{syntax.inner.prop}{\mbox{\isa{prop}}} denotes meta-level propositions,
 | 
| 
 | 
   574  | 
  which are terms of type \isa{prop}.  The syntax of such formulae of
 | 
| 
 | 
   575  | 
  the meta-logic is carefully distinguished from usual conventions for
  | 
| 
 | 
   576  | 
  object-logics.  In particular, plain \isa{{\isachardoublequote}{\isasymlambda}{\isachardoublequote}}-term notation is
 | 
| 
 | 
   577  | 
  \emph{not} recognized as \hyperlink{syntax.inner.prop}{\mbox{\isa{prop}}}.
 | 
| 
 | 
   578  | 
  | 
| 
 | 
   579  | 
  \item \indexref{inner}{syntax}{aprop}\hyperlink{syntax.inner.aprop}{\mbox{\isa{aprop}}} denotes atomic propositions, which
 | 
| 
 | 
   580  | 
  are embedded into regular \hyperlink{syntax.inner.prop}{\mbox{\isa{prop}}} by means of an
 | 
| 
 | 
   581  | 
  explicit \verb|PROP| token.
  | 
| 
 | 
   582  | 
  | 
| 
 | 
   583  | 
  Terms of type \isa{prop} with non-constant head, e.g.\ a plain
 | 
| 
 | 
   584  | 
  variable, are printed in this form.  Constants that yield type \isa{prop} are expected to provide their own concrete syntax; otherwise
 | 
| 
 | 
   585  | 
  the printed version will appear like \hyperlink{syntax.inner.logic}{\mbox{\isa{logic}}} and
 | 
| 
 | 
   586  | 
  cannot be parsed again as \hyperlink{syntax.inner.prop}{\mbox{\isa{prop}}}.
 | 
| 
 | 
   587  | 
  | 
| 
 | 
   588  | 
  \item \indexref{inner}{syntax}{logic}\hyperlink{syntax.inner.logic}{\mbox{\isa{logic}}} denotes arbitrary terms of a
 | 
| 
 | 
   589  | 
  logical type, excluding type \isa{prop}.  This is the main
 | 
| 
 | 
   590  | 
  syntactic category of object-logic entities, covering plain \isa{{\isasymlambda}}-term notation (variables, abstraction, application), plus
 | 
| 
 | 
   591  | 
  anything defined by the user.
  | 
| 
 | 
   592  | 
  | 
| 
 | 
   593  | 
  When specifying notation for logical entities, all logical types
  | 
| 
 | 
   594  | 
  (excluding \isa{prop}) are \emph{collapsed} to this single category
 | 
| 
 | 
   595  | 
  of \hyperlink{syntax.inner.logic}{\mbox{\isa{logic}}}.
 | 
| 
 | 
   596  | 
  | 
| 
 | 
   597  | 
  \item \indexref{inner}{syntax}{idt}\hyperlink{syntax.inner.idt}{\mbox{\isa{idt}}} denotes identifiers, possibly
 | 
| 
 | 
   598  | 
  constrained by types.
  | 
| 
 | 
   599  | 
  | 
| 
 | 
   600  | 
  \item \indexref{inner}{syntax}{idts}\hyperlink{syntax.inner.idts}{\mbox{\isa{idts}}} denotes a sequence of \indexref{inner}{syntax}{idt}\hyperlink{syntax.inner.idt}{\mbox{\isa{idt}}}.  This is the most basic category for variables in
 | 
| 
 | 
   601  | 
  iterated binders, such as \isa{{\isachardoublequote}{\isasymlambda}{\isachardoublequote}} or \isa{{\isachardoublequote}{\isasymAnd}{\isachardoublequote}}.
 | 
| 
 | 
   602  | 
  | 
| 
 | 
   603  | 
  \item \indexref{inner}{syntax}{pttrn}\hyperlink{syntax.inner.pttrn}{\mbox{\isa{pttrn}}} and \indexref{inner}{syntax}{pttrns}\hyperlink{syntax.inner.pttrns}{\mbox{\isa{pttrns}}}
 | 
| 
 | 
   604  | 
  denote patterns for abstraction, cases bindings etc.  In Pure, these
  | 
| 
 | 
   605  | 
  categories start as a merely copy of \hyperlink{syntax.inner.idt}{\mbox{\isa{idt}}} and
 | 
| 
 | 
   606  | 
  \hyperlink{syntax.inner.idts}{\mbox{\isa{idts}}}, respectively.  Object-logics may add
 | 
| 
 | 
   607  | 
  additional productions for binding forms.
  | 
| 
 | 
   608  | 
  | 
| 
 | 
   609  | 
  \item \indexref{inner}{syntax}{type}\hyperlink{syntax.inner.type}{\mbox{\isa{type}}} denotes types of the meta-logic.
 | 
| 
 | 
   610  | 
  | 
| 
 | 
   611  | 
  \item \indexref{inner}{syntax}{sort}\hyperlink{syntax.inner.sort}{\mbox{\isa{sort}}} denotes meta-level sorts.
 | 
| 
 | 
   612  | 
  | 
| 
 | 
   613  | 
  \end{description}
 | 
| 
 | 
   614  | 
  | 
| 
 | 
   615  | 
  Here are some further explanations of certain syntax features.
  | 
| 
 | 
   616  | 
  | 
| 
 | 
   617  | 
  \begin{itemize}
 | 
| 
 | 
   618  | 
  | 
| 
 | 
   619  | 
  \item In \hyperlink{syntax.inner.idts}{\mbox{\isa{idts}}}, note that \isa{{\isachardoublequote}x\ {\isacharcolon}{\isacharcolon}\ nat\ y{\isachardoublequote}} is
 | 
| 
 | 
   620  | 
  parsed as \isa{{\isachardoublequote}x\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}nat\ y{\isacharparenright}{\isachardoublequote}}, treating \isa{y} like a type
 | 
| 
 | 
   621  | 
  constructor applied to \isa{nat}.  To avoid this interpretation,
 | 
| 
 | 
   622  | 
  write \isa{{\isachardoublequote}{\isacharparenleft}x\ {\isacharcolon}{\isacharcolon}\ nat{\isacharparenright}\ y{\isachardoublequote}} with explicit parentheses.
 | 
| 
 | 
   623  | 
  | 
| 
 | 
   624  | 
  \item Similarly, \isa{{\isachardoublequote}x\ {\isacharcolon}{\isacharcolon}\ nat\ y\ {\isacharcolon}{\isacharcolon}\ nat{\isachardoublequote}} is parsed as \isa{{\isachardoublequote}x\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}nat\ y\ {\isacharcolon}{\isacharcolon}\ nat{\isacharparenright}{\isachardoublequote}}.  The correct form is \isa{{\isachardoublequote}{\isacharparenleft}x\ {\isacharcolon}{\isacharcolon}\ nat{\isacharparenright}\ {\isacharparenleft}y\ {\isacharcolon}{\isacharcolon}\ nat{\isacharparenright}{\isachardoublequote}}, or \isa{{\isachardoublequote}{\isacharparenleft}x\ {\isacharcolon}{\isacharcolon}\ nat{\isacharparenright}\ y\ {\isacharcolon}{\isacharcolon}\ nat{\isachardoublequote}} if \isa{y} is last in the
 | 
| 
 | 
   625  | 
  sequence of identifiers.
  | 
| 
 | 
   626  | 
  | 
| 
 | 
   627  | 
  \item Type constraints for terms bind very weakly.  For example,
  | 
| 
 | 
   628  | 
  \isa{{\isachardoublequote}x\ {\isacharless}\ y\ {\isacharcolon}{\isacharcolon}\ nat{\isachardoublequote}} is normally parsed as \isa{{\isachardoublequote}{\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}\ {\isacharcolon}{\isacharcolon}\ nat{\isachardoublequote}}, unless \isa{{\isachardoublequote}{\isacharless}{\isachardoublequote}} has a very low priority, in which case the
 | 
| 
 | 
   629  | 
  input is likely to be ambiguous.  The correct form is \isa{{\isachardoublequote}x\ {\isacharless}\ {\isacharparenleft}y\ {\isacharcolon}{\isacharcolon}\ nat{\isacharparenright}{\isachardoublequote}}.
 | 
| 
 | 
   630  | 
  | 
| 
 | 
   631  | 
  \item Constraints may be either written with two literal colons
  | 
| 
 | 
   632  | 
  ``\verb|::|'' or the double-colon symbol \verb|\<Colon>|,
  | 
| 
 | 
   633  | 
  which actually looks exactly the same in some {\LaTeX} styles.
 | 
| 
 | 
   634  | 
  | 
| 
 | 
   635  | 
  \item Dummy variables (written as underscore) may occur in different
  | 
| 
 | 
   636  | 
  roles.
  | 
| 
 | 
   637  | 
  | 
| 
 | 
   638  | 
  \begin{description}
 | 
| 
 | 
   639  | 
  | 
| 
 | 
   640  | 
  \item A type ``\isa{{\isachardoublequote}{\isacharunderscore}{\isachardoublequote}}'' or ``\isa{{\isachardoublequote}{\isacharunderscore}\ {\isacharcolon}{\isacharcolon}\ sort{\isachardoublequote}}'' acts like an
 | 
| 
 | 
   641  | 
  anonymous inference parameter, which is filled-in according to the
  | 
| 
 | 
   642  | 
  most general type produced by the type-checking phase.
  | 
| 
 | 
   643  | 
  | 
| 
 | 
   644  | 
  \item A bound ``\isa{{\isachardoublequote}{\isacharunderscore}{\isachardoublequote}}'' refers to a vacuous abstraction, where
 | 
| 
 | 
   645  | 
  the body does not refer to the binding introduced here.  As in the
  | 
| 
 | 
   646  | 
  term \isa{{\isachardoublequote}{\isasymlambda}x\ {\isacharunderscore}{\isachardot}\ x{\isachardoublequote}}, which is \isa{{\isachardoublequote}{\isasymalpha}{\isachardoublequote}}-equivalent to \isa{{\isachardoublequote}{\isasymlambda}x\ y{\isachardot}\ x{\isachardoublequote}}.
 | 
| 
 | 
   647  | 
  | 
| 
 | 
   648  | 
  \item A free ``\isa{{\isachardoublequote}{\isacharunderscore}{\isachardoublequote}}'' refers to an implicit outer binding.
 | 
| 
 | 
   649  | 
  Higher definitional packages usually allow forms like \isa{{\isachardoublequote}f\ x\ {\isacharunderscore}\ {\isacharequal}\ x{\isachardoublequote}}.
 | 
| 
 | 
   650  | 
  | 
| 
 | 
   651  | 
  \item A schematic ``\isa{{\isachardoublequote}{\isacharunderscore}{\isachardoublequote}}'' (within a term pattern, see
 | 
| 
 | 
   652  | 
  \secref{sec:term-decls}) refers to an anonymous variable that is
 | 
| 
 | 
   653  | 
  implicitly abstracted over its context of locally bound variables.
  | 
| 
 | 
   654  | 
  For example, this allows pattern matching of \isa{{\isachardoublequote}{\isacharbraceleft}x{\isachardot}\ f\ x\ {\isacharequal}\ g\ x{\isacharbraceright}{\isachardoublequote}} against \isa{{\isachardoublequote}{\isacharbraceleft}x{\isachardot}\ {\isacharunderscore}\ {\isacharequal}\ {\isacharunderscore}{\isacharbraceright}{\isachardoublequote}}, or even \isa{{\isachardoublequote}{\isacharbraceleft}{\isacharunderscore}{\isachardot}\ {\isacharunderscore}\ {\isacharequal}\ {\isacharunderscore}{\isacharbraceright}{\isachardoublequote}} by
 | 
| 
 | 
   655  | 
  using both bound and schematic dummies.
  | 
| 
 | 
   656  | 
  | 
| 
 | 
   657  | 
  \end{description}
 | 
| 
 | 
   658  | 
  | 
| 
 | 
   659  | 
  \item The three literal dots ``\verb|...|'' may be also
  | 
| 
 | 
   660  | 
  written as ellipsis symbol \verb|\<dots>|.  In both cases this
  | 
| 
 | 
   661  | 
  refers to a special schematic variable, which is bound in the
  | 
| 
 | 
   662  | 
  context.  This special term abbreviation works nicely with
  | 
| 
 | 
   663  | 
  calculational reasoning (\secref{sec:calculation}).
 | 
| 
 | 
   664  | 
  | 
| 
 | 
   665  | 
  \end{itemize}%
 | 
| 
 | 
   666  | 
\end{isamarkuptext}%
 | 
| 
 | 
   667  | 
\isamarkuptrue%
  | 
| 
 | 
   668  | 
%
  | 
| 
 | 
   669  | 
\isamarkupsection{Lexical matters \label{sec:inner-lex}%
 | 
| 
 | 
   670  | 
}
  | 
| 
 | 
   671  | 
\isamarkuptrue%
  | 
| 
 | 
   672  | 
%
  | 
| 
 | 
   673  | 
\begin{isamarkuptext}%
 | 
| 
 | 
   674  | 
The inner lexical syntax vaguely resembles the outer one
  | 
| 
 | 
   675  | 
  (\secref{sec:outer-lex}), but some details are different.  There are
 | 
| 
 | 
   676  | 
  two main categories of inner syntax tokens:
  | 
| 
 | 
   677  | 
  | 
| 
 | 
   678  | 
  \begin{enumerate}
 | 
| 
 | 
   679  | 
  | 
| 
 | 
   680  | 
  \item \emph{delimiters} --- the literal tokens occurring in
 | 
| 
 | 
   681  | 
  productions of the given priority grammar (cf.\
  | 
| 
 | 
   682  | 
  \secref{sec:priority-grammar});
 | 
| 
 | 
   683  | 
  | 
| 
 | 
   684  | 
  \item \emph{named tokens} --- various categories of identifiers etc.
 | 
| 
 | 
   685  | 
  | 
| 
 | 
   686  | 
  \end{enumerate}
 | 
| 
 | 
   687  | 
  | 
| 
 | 
   688  | 
  Delimiters override named tokens and may thus render certain
  | 
| 
 | 
   689  | 
  identifiers inaccessible.  Sometimes the logical context admits
  | 
| 
 | 
   690  | 
  alternative ways to refer to the same entity, potentially via
  | 
| 
 | 
   691  | 
  qualified names.
  | 
| 
 | 
   692  | 
  | 
| 
 | 
   693  | 
  \medskip The categories for named tokens are defined once and for
  | 
| 
 | 
   694  | 
  all as follows, reusing some categories of the outer token syntax
  | 
| 
 | 
   695  | 
  (\secref{sec:outer-lex}).
 | 
| 
 | 
   696  | 
  | 
| 
 | 
   697  | 
  \begin{center}
 | 
| 
 | 
   698  | 
  \begin{supertabular}{rcl}
 | 
| 
 | 
   699  | 
    \indexdef{inner}{syntax}{id}\hypertarget{syntax.inner.id}{\hyperlink{syntax.inner.id}{\mbox{\isa{id}}}} & = & \indexref{}{syntax}{ident}\hyperlink{syntax.ident}{\mbox{\isa{ident}}} \\
 | 
| 
 | 
   700  | 
    \indexdef{inner}{syntax}{longid}\hypertarget{syntax.inner.longid}{\hyperlink{syntax.inner.longid}{\mbox{\isa{longid}}}} & = & \indexref{}{syntax}{longident}\hyperlink{syntax.longident}{\mbox{\isa{longident}}} \\
 | 
| 
 | 
   701  | 
    \indexdef{inner}{syntax}{var}\hypertarget{syntax.inner.var}{\hyperlink{syntax.inner.var}{\mbox{\isa{var}}}} & = & \indexref{}{syntax}{var}\hyperlink{syntax.var}{\mbox{\isa{var}}} \\
 | 
| 
 | 
   702  | 
    \indexdef{inner}{syntax}{tid}\hypertarget{syntax.inner.tid}{\hyperlink{syntax.inner.tid}{\mbox{\isa{tid}}}} & = & \indexref{}{syntax}{typefree}\hyperlink{syntax.typefree}{\mbox{\isa{typefree}}} \\
 | 
| 
 | 
   703  | 
    \indexdef{inner}{syntax}{tvar}\hypertarget{syntax.inner.tvar}{\hyperlink{syntax.inner.tvar}{\mbox{\isa{tvar}}}} & = & \indexref{}{syntax}{typevar}\hyperlink{syntax.typevar}{\mbox{\isa{typevar}}} \\
 | 
| 
 | 
   704  | 
    \indexdef{inner}{syntax}{num}\hypertarget{syntax.inner.num}{\hyperlink{syntax.inner.num}{\mbox{\isa{num}}}} & = & \indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}}\isa{{\isachardoublequote}\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|-|\indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}} \\
 | 
| 
 | 
   705  | 
    \indexdef{inner}{syntax}{xnum}\hypertarget{syntax.inner.xnum}{\hyperlink{syntax.inner.xnum}{\mbox{\isa{xnum}}}} & = & \verb|#|\indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}}\isa{{\isachardoublequote}\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|#-|\indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}} \\
 | 
| 
 | 
   706  | 
  | 
| 
 | 
   707  | 
    \indexdef{inner}{syntax}{xstr}\hypertarget{syntax.inner.xstr}{\hyperlink{syntax.inner.xstr}{\mbox{\isa{xstr}}}} & = & \verb|''| \isa{{\isachardoublequote}{\isasymdots}{\isachardoublequote}} \verb|''| \\
 | 
| 
 | 
   708  | 
  \end{supertabular}
 | 
| 
 | 
   709  | 
  \end{center}
 | 
| 
 | 
   710  | 
  | 
| 
 | 
   711  | 
  The token categories \indexref{inner}{syntax}{num}\hyperlink{syntax.inner.num}{\mbox{\isa{num}}}, \indexref{inner}{syntax}{xnum}\hyperlink{syntax.inner.xnum}{\mbox{\isa{xnum}}}, and \indexref{inner}{syntax}{xstr}\hyperlink{syntax.inner.xstr}{\mbox{\isa{xstr}}} are not used in Pure.
 | 
| 
 | 
   712  | 
  Object-logics may implement numerals and string constants by adding
  | 
| 
 | 
   713  | 
  appropriate syntax declarations, together with some translation
  | 
| 
 | 
   714  | 
  functions (e.g.\ see Isabelle/HOL).%
  | 
| 
 | 
   715  | 
\end{isamarkuptext}%
 | 
| 
 | 
   716  | 
\isamarkuptrue%
  | 
| 
 | 
   717  | 
%
  | 
| 
 | 
   718  | 
\isamarkupsection{Syntax and translations \label{sec:syn-trans}%
 | 
| 
 | 
   719  | 
}
  | 
| 
 | 
   720  | 
\isamarkuptrue%
  | 
| 
 | 
   721  | 
%
  | 
| 
 | 
   722  | 
\begin{isamarkuptext}%
 | 
| 
 | 
   723  | 
\begin{matharray}{rcl}
 | 
| 
 | 
   724  | 
    \indexdef{}{command}{nonterminals}\hypertarget{command.nonterminals}{\hyperlink{command.nonterminals}{\mbox{\isa{\isacommand{nonterminals}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
 | 
| 
 | 
   725  | 
    \indexdef{}{command}{syntax}\hypertarget{command.syntax}{\hyperlink{command.syntax}{\mbox{\isa{\isacommand{syntax}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
 | 
| 
 | 
   726  | 
    \indexdef{}{command}{no\_syntax}\hypertarget{command.no-syntax}{\hyperlink{command.no-syntax}{\mbox{\isa{\isacommand{no{\isacharunderscore}syntax}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
 | 
| 
 | 
   727  | 
    \indexdef{}{command}{translations}\hypertarget{command.translations}{\hyperlink{command.translations}{\mbox{\isa{\isacommand{translations}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
 | 
| 
 | 
   728  | 
    \indexdef{}{command}{no\_translations}\hypertarget{command.no-translations}{\hyperlink{command.no-translations}{\mbox{\isa{\isacommand{no{\isacharunderscore}translations}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
 | 
| 
 | 
   729  | 
  \end{matharray}
 | 
| 
 | 
   730  | 
  | 
| 
 | 
   731  | 
  \begin{rail}
 | 
| 
 | 
   732  | 
    'nonterminals' (name +)
  | 
| 
 | 
   733  | 
    ;
  | 
| 
 | 
   734  | 
    ('syntax' | 'no\_syntax') mode? (constdecl +)
 | 
| 
 | 
   735  | 
    ;
  | 
| 
 | 
   736  | 
    ('translations' | 'no\_translations') (transpat ('==' | '=>' | '<=' | rightleftharpoons | rightharpoonup | leftharpoondown) transpat +)
 | 
| 
 | 
   737  | 
    ;
  | 
| 
 | 
   738  | 
  | 
| 
 | 
   739  | 
    mode: ('(' ( name | 'output' | name 'output' ) ')')
 | 
| 
 | 
   740  | 
    ;
  | 
| 
 | 
   741  | 
    transpat: ('(' nameref ')')? string
 | 
| 
 | 
   742  | 
    ;
  | 
| 
 | 
   743  | 
  \end{rail}
 | 
| 
 | 
   744  | 
  | 
| 
 | 
   745  | 
  \begin{description}
 | 
| 
 | 
   746  | 
  
  | 
| 
 | 
   747  | 
  \item \hyperlink{command.nonterminals}{\mbox{\isa{\isacommand{nonterminals}}}}~\isa{c} declares a type
 | 
| 
 | 
   748  | 
  constructor \isa{c} (without arguments) to act as purely syntactic
 | 
| 
 | 
   749  | 
  type: a nonterminal symbol of the inner syntax.
  | 
| 
 | 
   750  | 
  | 
| 
 | 
   751  | 
  \item \hyperlink{command.syntax}{\mbox{\isa{\isacommand{syntax}}}}~\isa{{\isachardoublequote}{\isacharparenleft}mode{\isacharparenright}\ decls{\isachardoublequote}} is similar to
 | 
| 
 | 
   752  | 
  \hyperlink{command.consts}{\mbox{\isa{\isacommand{consts}}}}~\isa{decls}, except that the actual logical
 | 
| 
 | 
   753  | 
  signature extension is omitted.  Thus the context free grammar of
  | 
| 
 | 
   754  | 
  Isabelle's inner syntax may be augmented in arbitrary ways,
  | 
| 
 | 
   755  | 
  independently of the logic.  The \isa{mode} argument refers to the
 | 
| 
 | 
   756  | 
  print mode that the grammar rules belong; unless the \indexref{}{keyword}{output}\hyperlink{keyword.output}{\mbox{\isa{\isakeyword{output}}}} indicator is given, all productions are added both to the
 | 
| 
 | 
   757  | 
  input and output grammar.
  | 
| 
 | 
   758  | 
  
  | 
| 
 | 
   759  | 
  \item \hyperlink{command.no-syntax}{\mbox{\isa{\isacommand{no{\isacharunderscore}syntax}}}}~\isa{{\isachardoublequote}{\isacharparenleft}mode{\isacharparenright}\ decls{\isachardoublequote}} removes grammar
 | 
| 
 | 
   760  | 
  declarations (and translations) resulting from \isa{decls}, which
 | 
| 
 | 
   761  | 
  are interpreted in the same manner as for \hyperlink{command.syntax}{\mbox{\isa{\isacommand{syntax}}}} above.
 | 
| 
 | 
   762  | 
  
  | 
| 
 | 
   763  | 
  \item \hyperlink{command.translations}{\mbox{\isa{\isacommand{translations}}}}~\isa{rules} specifies syntactic
 | 
| 
 | 
   764  | 
  translation rules (i.e.\ macros): parse~/ print rules (\isa{{\isachardoublequote}{\isasymrightleftharpoons}{\isachardoublequote}}),
 | 
| 
 | 
   765  | 
  parse rules (\isa{{\isachardoublequote}{\isasymrightharpoonup}{\isachardoublequote}}), or print rules (\isa{{\isachardoublequote}{\isasymleftharpoondown}{\isachardoublequote}}).
 | 
| 
 | 
   766  | 
  Translation patterns may be prefixed by the syntactic category to be
  | 
| 
 | 
   767  | 
  used for parsing; the default is \isa{logic}.
 | 
| 
 | 
   768  | 
  
  | 
| 
 | 
   769  | 
  \item \hyperlink{command.no-translations}{\mbox{\isa{\isacommand{no{\isacharunderscore}translations}}}}~\isa{rules} removes syntactic
 | 
| 
 | 
   770  | 
  translation rules, which are interpreted in the same manner as for
  | 
| 
 | 
   771  | 
  \hyperlink{command.translations}{\mbox{\isa{\isacommand{translations}}}} above.
 | 
| 
 | 
   772  | 
  | 
| 
 | 
   773  | 
  \end{description}%
 | 
| 
 | 
   774  | 
\end{isamarkuptext}%
 | 
| 
 | 
   775  | 
\isamarkuptrue%
  | 
| 
 | 
   776  | 
%
  | 
| 
 | 
   777  | 
\isamarkupsection{Syntax translation functions \label{sec:tr-funs}%
 | 
| 
 | 
   778  | 
}
  | 
| 
 | 
   779  | 
\isamarkuptrue%
  | 
| 
 | 
   780  | 
%
  | 
| 
 | 
   781  | 
\begin{isamarkuptext}%
 | 
| 
 | 
   782  | 
\begin{matharray}{rcl}
 | 
| 
 | 
   783  | 
    \indexdef{}{command}{parse\_ast\_translation}\hypertarget{command.parse-ast-translation}{\hyperlink{command.parse-ast-translation}{\mbox{\isa{\isacommand{parse{\isacharunderscore}ast{\isacharunderscore}translation}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
 | 
| 
 | 
   784  | 
    \indexdef{}{command}{parse\_translation}\hypertarget{command.parse-translation}{\hyperlink{command.parse-translation}{\mbox{\isa{\isacommand{parse{\isacharunderscore}translation}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
 | 
| 
 | 
   785  | 
    \indexdef{}{command}{print\_translation}\hypertarget{command.print-translation}{\hyperlink{command.print-translation}{\mbox{\isa{\isacommand{print{\isacharunderscore}translation}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
 | 
| 
 | 
   786  | 
    \indexdef{}{command}{typed\_print\_translation}\hypertarget{command.typed-print-translation}{\hyperlink{command.typed-print-translation}{\mbox{\isa{\isacommand{typed{\isacharunderscore}print{\isacharunderscore}translation}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
 | 
| 
 | 
   787  | 
    \indexdef{}{command}{print\_ast\_translation}\hypertarget{command.print-ast-translation}{\hyperlink{command.print-ast-translation}{\mbox{\isa{\isacommand{print{\isacharunderscore}ast{\isacharunderscore}translation}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
 | 
| 
 | 
   788  | 
  \end{matharray}
 | 
| 
 | 
   789  | 
  | 
| 
 | 
   790  | 
  \begin{rail}
 | 
| 
 | 
   791  | 
  ( 'parse\_ast\_translation' | 'parse\_translation' | 'print\_translation' |
  | 
| 
 | 
   792  | 
    'typed\_print\_translation' | 'print\_ast\_translation' ) ('(advanced)')? text
 | 
| 
 | 
   793  | 
  ;
  | 
| 
 | 
   794  | 
  \end{rail}
 | 
| 
 | 
   795  | 
  | 
| 
 | 
   796  | 
  Syntax translation functions written in ML admit almost arbitrary
  | 
| 
 | 
   797  | 
  manipulations of Isabelle's inner syntax.  Any of the above commands
  | 
| 
 | 
   798  | 
  have a single \railqtok{text} argument that refers to an ML
 | 
| 
 | 
   799  | 
  expression of appropriate type, which are as follows by default:
  | 
| 
 | 
   800  | 
  | 
| 
 | 
   801  | 
%FIXME proper antiquotations
  | 
| 
 | 
   802  | 
\begin{ttbox}
 | 
| 
 | 
   803  | 
val parse_ast_translation   : (string * (ast list -> ast)) list
  | 
| 
 | 
   804  | 
val parse_translation       : (string * (term list -> term)) list
  | 
| 
 | 
   805  | 
val print_translation       : (string * (term list -> term)) list
  | 
| 
 | 
   806  | 
val typed_print_translation :
  | 
| 
 | 
   807  | 
  (string * (bool -> typ -> term list -> term)) list
  | 
| 
 | 
   808  | 
val print_ast_translation   : (string * (ast list -> ast)) list
  | 
| 
 | 
   809  | 
\end{ttbox}
 | 
| 
 | 
   810  | 
  | 
| 
 | 
   811  | 
  If the \isa{{\isachardoublequote}{\isacharparenleft}advanced{\isacharparenright}{\isachardoublequote}} option is given, the corresponding
 | 
| 
 | 
   812  | 
  translation functions may depend on the current theory or proof
  | 
| 
 | 
   813  | 
  context.  This allows to implement advanced syntax mechanisms, as
  | 
| 
 | 
   814  | 
  translations functions may refer to specific theory declarations or
  | 
| 
 | 
   815  | 
  auxiliary proof data.
  | 
| 
 | 
   816  | 
  | 
| 
 | 
   817  | 
  See also \cite[\S8]{isabelle-ref} for more information on the
 | 
| 
 | 
   818  | 
  general concept of syntax transformations in Isabelle.
  | 
| 
 | 
   819  | 
  | 
| 
 | 
   820  | 
%FIXME proper antiquotations
  | 
| 
 | 
   821  | 
\begin{ttbox}
 | 
| 
 | 
   822  | 
val parse_ast_translation:
  | 
| 
 | 
   823  | 
  (string * (Proof.context -> ast list -> ast)) list
  | 
| 
 | 
   824  | 
val parse_translation:
  | 
| 
 | 
   825  | 
  (string * (Proof.context -> term list -> term)) list
  | 
| 
 | 
   826  | 
val print_translation:
  | 
| 
 | 
   827  | 
  (string * (Proof.context -> term list -> term)) list
  | 
| 
 | 
   828  | 
val typed_print_translation:
  | 
| 
 | 
   829  | 
  (string * (Proof.context -> bool -> typ -> term list -> term)) list
  | 
| 
 | 
   830  | 
val print_ast_translation:
  | 
| 
 | 
   831  | 
  (string * (Proof.context -> ast list -> ast)) list
  | 
| 
 | 
   832  | 
\end{ttbox}%
 | 
| 
 | 
   833  | 
\end{isamarkuptext}%
 | 
| 
 | 
   834  | 
\isamarkuptrue%
  | 
| 
 | 
   835  | 
%
  | 
| 
 | 
   836  | 
\isamarkupsection{Inspecting the syntax%
 | 
| 
 | 
   837  | 
}
  | 
| 
 | 
   838  | 
\isamarkuptrue%
  | 
| 
 | 
   839  | 
%
  | 
| 
 | 
   840  | 
\begin{isamarkuptext}%
 | 
| 
 | 
   841  | 
\begin{matharray}{rcl}
 | 
| 
 | 
   842  | 
    \indexdef{}{command}{print\_syntax}\hypertarget{command.print-syntax}{\hyperlink{command.print-syntax}{\mbox{\isa{\isacommand{print{\isacharunderscore}syntax}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}context\ {\isasymrightarrow}{\isachardoublequote}} \\
 | 
| 
 | 
   843  | 
  \end{matharray}
 | 
| 
 | 
   844  | 
  | 
| 
 | 
   845  | 
  \begin{description}
 | 
| 
 | 
   846  | 
  | 
| 
 | 
   847  | 
  \item \hyperlink{command.print-syntax}{\mbox{\isa{\isacommand{print{\isacharunderscore}syntax}}}} prints the inner syntax of the
 | 
| 
 | 
   848  | 
  current context.  The output can be quite large; the most important
  | 
| 
 | 
   849  | 
  sections are explained below.
  | 
| 
 | 
   850  | 
  | 
| 
 | 
   851  | 
  \begin{description}
 | 
| 
 | 
   852  | 
  | 
| 
 | 
   853  | 
  \item \isa{{\isachardoublequote}lexicon{\isachardoublequote}} lists the delimiters of the inner token
 | 
| 
 | 
   854  | 
  language; see \secref{sec:inner-lex}.
 | 
| 
 | 
   855  | 
  | 
| 
 | 
   856  | 
  \item \isa{{\isachardoublequote}prods{\isachardoublequote}} lists the productions of the underlying
 | 
| 
 | 
   857  | 
  priority grammar; see \secref{sec:priority-grammar}.
 | 
| 
 | 
   858  | 
  | 
| 
 | 
   859  | 
  The nonterminal \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup p\isactrlsup {\isacharparenright}{\isachardoublequote}} is rendered in plain text as \isa{{\isachardoublequote}A{\isacharbrackleft}p{\isacharbrackright}{\isachardoublequote}}; delimiters are quoted.  Many productions have an extra
 | 
| 
 | 
   860  | 
  \isa{{\isachardoublequote}{\isasymdots}\ {\isacharequal}{\isachargreater}\ name{\isachardoublequote}}.  These names later become the heads of parse
 | 
| 
 | 
   861  | 
  trees; they also guide the pretty printer.
  | 
| 
 | 
   862  | 
  | 
| 
 | 
   863  | 
  Productions without such parse tree names are called \emph{copy
 | 
| 
 | 
   864  | 
  productions}.  Their right-hand side must have exactly one
  | 
| 
 | 
   865  | 
  nonterminal symbol (or named token).  The parser does not create a
  | 
| 
 | 
   866  | 
  new parse tree node for copy productions, but simply returns the
  | 
| 
 | 
   867  | 
  parse tree of the right-hand symbol.
  | 
| 
 | 
   868  | 
  | 
| 
 | 
   869  | 
  If the right-hand side of a copy production consists of a single
  | 
| 
 | 
   870  | 
  nonterminal without any delimiters, then it is called a \emph{chain
 | 
| 
 | 
   871  | 
  production}.  Chain productions act as abbreviations: conceptually,
  | 
| 
 | 
   872  | 
  they are removed from the grammar by adding new productions.
  | 
| 
 | 
   873  | 
  Priority information attached to chain productions is ignored; only
  | 
| 
 | 
   874  | 
  the dummy value \isa{{\isachardoublequote}{\isacharminus}{\isadigit{1}}{\isachardoublequote}} is displayed.
 | 
| 
 | 
   875  | 
  | 
| 
28857
 | 
   876  | 
  \item \isa{{\isachardoublequote}print\ modes{\isachardoublequote}} lists the alternative print modes
 | 
| 
28762
 | 
   877  | 
  provided by this grammar; see \secref{sec:print-modes}.
 | 
| 
 | 
   878  | 
  | 
| 
 | 
   879  | 
  \item \isa{{\isachardoublequote}parse{\isacharunderscore}rules{\isachardoublequote}} and \isa{{\isachardoublequote}print{\isacharunderscore}rules{\isachardoublequote}} relate to
 | 
| 
 | 
   880  | 
  syntax translations (macros); see \secref{sec:syn-trans}.
 | 
| 
 | 
   881  | 
  | 
| 
 | 
   882  | 
  \item \isa{{\isachardoublequote}parse{\isacharunderscore}ast{\isacharunderscore}translation{\isachardoublequote}} and \isa{{\isachardoublequote}print{\isacharunderscore}ast{\isacharunderscore}translation{\isachardoublequote}} list sets of constants that invoke
 | 
| 
 | 
   883  | 
  translation functions for abstract syntax trees, which are only
  | 
| 
 | 
   884  | 
  required in very special situations; see \secref{sec:tr-funs}.
 | 
| 
 | 
   885  | 
  | 
| 
 | 
   886  | 
  \item \isa{{\isachardoublequote}parse{\isacharunderscore}translation{\isachardoublequote}} and \isa{{\isachardoublequote}print{\isacharunderscore}translation{\isachardoublequote}}
 | 
| 
 | 
   887  | 
  list the sets of constants that invoke regular translation
  | 
| 
 | 
   888  | 
  functions; see \secref{sec:tr-funs}.
 | 
| 
 | 
   889  | 
  | 
| 
 | 
   890  | 
  \end{description}
 | 
| 
 | 
   891  | 
  
  | 
| 
 | 
   892  | 
  \end{description}%
 | 
| 
 | 
   893  | 
\end{isamarkuptext}%
 | 
| 
 | 
   894  | 
\isamarkuptrue%
  | 
| 
 | 
   895  | 
%
  | 
| 
 | 
   896  | 
\isadelimtheory
  | 
| 
 | 
   897  | 
%
  | 
| 
 | 
   898  | 
\endisadelimtheory
  | 
| 
 | 
   899  | 
%
  | 
| 
 | 
   900  | 
\isatagtheory
  | 
| 
 | 
   901  | 
\isacommand{end}\isamarkupfalse%
 | 
| 
 | 
   902  | 
%
  | 
| 
 | 
   903  | 
\endisatagtheory
  | 
| 
 | 
   904  | 
{\isafoldtheory}%
 | 
| 
 | 
   905  | 
%
  | 
| 
 | 
   906  | 
\isadelimtheory
  | 
| 
 | 
   907  | 
%
  | 
| 
 | 
   908  | 
\endisadelimtheory
  | 
| 
 | 
   909  | 
\isanewline
  | 
| 
 | 
   910  | 
\end{isabellebody}%
 | 
| 
 | 
   911  | 
%%% Local Variables:
  | 
| 
 | 
   912  | 
%%% mode: latex
  | 
| 
 | 
   913  | 
%%% TeX-master: "root"
  | 
| 
 | 
   914  | 
%%% End:
  |