Theory Documents
theory Documents imports Main begin
section ‹Concrete Syntax \label{sec:concrete-syntax}›
text ‹
  The core concept of Isabelle's framework for concrete syntax is that
  of \bfindex{mixfix annotations}.  Associated with any kind of
  constant declaration, mixfixes affect both the grammar productions
  for the parser and output templates for the pretty printer.
  In full generality, parser and pretty printer configuration is a
  subtle affair~\<^cite>‹"isabelle-isar-ref"›.  Your syntax specifications need
  to interact properly with the existing setup of Isabelle/Pure and
  Isabelle/HOL\@.  To avoid creating ambiguities with existing
  elements, it is particularly important to give new syntactic
  constructs the right precedence.
  Below we introduce a few simple syntax declaration
  forms that already cover many common situations fairly well.
›
subsection ‹Infix Annotations›
text ‹
  Syntax annotations may be included wherever constants are declared,
  such as \isacommand{definition} and \isacommand{primrec} --- and also
  \isacommand{datatype}, which declares constructor operations.
  Type-constructors may be annotated as well, although this is less
  frequently encountered in practice (the infix type ‹×› comes
  to mind).
  Infix declarations\index{infix annotations} provide a useful special
  case of mixfixes.  The following example of the exclusive-or
  operation on boolean values illustrates typical infix declarations.
›
definition xor :: "bool ⇒ bool ⇒ bool"    (infixl "[+]" 60)
where "A [+] B ≡ (A ∧ ¬ B) ∨ (¬ A ∧ B)"
text ‹
  \noindent Now ‹xor A B› and ‹A [+] B› refer to the
  same expression internally.  Any curried function with at least two
  arguments may be given infix syntax.  For partial applications with
  fewer than two operands, the operator is enclosed in parentheses.
  For instance, ‹xor› without arguments is represented as
  ‹([+])›; together with ordinary function application, this
  turns ‹xor A› into ‹([+]) A›.
  The keyword \isakeyword{infixl} seen above specifies an
  infix operator that is nested to the \emph{left}: in iterated
  applications the more complex expression appears on the left-hand
  side, and \<^term>‹A [+] B [+] C› stands for ‹(A [+] B) [+]
  C›.  Similarly, \isakeyword{infixr} means nesting to the
  \emph{right}, reading \<^term>‹A [+] B [+] C› as ‹A [+] (B
  [+] C)›.  A \emph{non-oriented} declaration via \isakeyword{infix}
  would render \<^term>‹A [+] B [+] C› illegal, but demand explicit
  parentheses to indicate the intended grouping.
  The string @{text [source] "[+]"} in our annotation refers to the
  concrete syntax to represent the operator (a literal token), while
  the number ‹60› determines the precedence of the construct:
  the syntactic priorities of the arguments and result.  Isabelle/HOL
  already uses up many popular combinations of ASCII symbols for its
  own use, including both ‹+› and ‹++›.  Longer
  character combinations are more likely to be still available for
  user extensions, such as our~‹[+]›.
  Operator precedences have a range of 0--1000.  Very low or high
  priorities are reserved for the meta-logic.  HOL syntax mainly uses
  the range of 10--100: the equality infix ‹=› is centered at
  50; logical connectives (like ‹∨› and ‹∧›) are
  below 50; algebraic ones (like ‹+› and ‹*›) are
  above 50.  User syntax should strive to coexist with common HOL
  forms, or use the mostly unused range 100--900.
›
subsection ‹Mathematical Symbols \label{sec:syntax-symbols}›
text ‹
  Concrete syntax based on ASCII characters has inherent limitations.
  Mathematical notation demands a larger repertoire of glyphs.
  Several standards of extended character sets have been proposed over
  decades, but none has become universally available so far.  Isabelle
  has its own notion of \bfindex{symbols} as the smallest entities of
  source text, without referring to internal encodings.  There are
  three kinds of such ``generalized characters'':
  \begin{enumerate}
  \item 7-bit ASCII characters
  \item named symbols: \verb,\,\verb,<,$ident$\verb,>,
  \item named control symbols: \verb,\,\verb,<^,$ident$\verb,>,
  \end{enumerate}
  Here $ident$ is any sequence of letters. 
  This results in an infinite store of symbols, whose
  interpretation is left to further front-end tools.  For example, the
  Isabelle document processor (see \S\ref{sec:document-preparation})
  display the \verb,\,\verb,<forall>, symbol as~‹∀›.
  A list of standard Isabelle symbols is given in
  \<^cite>‹"isabelle-isar-ref"›.  You may introduce your own
  interpretation of further symbols by configuring the appropriate
  front-end tool accordingly, e.g.\ by defining certain {\LaTeX}
  macros (see also \S\ref{sec:doc-prep-symbols}).  There are also a
  few predefined control symbols, such as \verb,\,\verb,<^sub>, and
  \verb,\,\verb,<^sup>, for sub- and superscript of the subsequent
  printable symbol, respectively.  For example, ▩‹A⇧⋆›, is
  output as ‹A⇧⋆›.
  A number of symbols are considered letters by the Isabelle lexer and
  can be used as part of identifiers. These are the greek letters
  ‹α› (\verb+\+\verb+<alpha>+), ‹β›
  (\verb+\+\verb+<beta>+), etc. (excluding ‹λ›),
  special letters like ‹𝒜› (\verb+\+\verb+<A>+) and ‹𝔄› (\verb+\+\verb+<AA>+).  Moreover the control symbol
  \verb+\+\verb+<^sub>+ may be used to subscript a single letter or digit
  in the trailing part of an identifier. This means that the input
  \medskip
  {\small\noindent ▩‹∀α⇩1. α⇩1 = Π⇩𝒜›}
  \medskip
  \noindent is recognized as the term \<^term>‹∀α⇩1. α⇩1 = Π⇩𝒜› 
  by Isabelle.
  Replacing our previous definition of ‹xor› by the
  following specifies an Isabelle symbol for the new operator:
›
hide_const xor
setup ‹Sign.add_path "version1"›
definition xor :: "bool ⇒ bool ⇒ bool"    (infixl "⊕" 60)
where "A ⊕ B ≡ (A ∧ ¬ B) ∨ (¬ A ∧ B)"
setup ‹Sign.local_path›
text ‹
  It is possible to provide alternative syntax forms
  through the \bfindex{print mode} concept~\<^cite>‹"isabelle-isar-ref"›.  By
  convention, the mode of ``$xsymbols$'' is enabled whenever
  Proof~General's X-Symbol mode or {\LaTeX} output is active.  Now
  consider the following hybrid declaration of ‹xor›:
›
hide_const xor
setup ‹Sign.add_path "version2"›
definition xor :: "bool ⇒ bool ⇒ bool"    (infixl "[+]\<ignore>" 60)
where "A [+]\<ignore> B ≡ (A ∧ ¬ B) ∨ (¬ A ∧ B)"
notation (xsymbols) xor (infixl "⊕\<ignore>" 60)
setup ‹Sign.local_path›
text ‹\noindent
The \commdx{notation} command associates a mixfix
annotation with a known constant.  The print mode specification,
here ‹(xsymbols)›, is optional.
We may now write ‹A [+] B› or ‹A ⊕ B› in input, while
output uses the nicer syntax of $xsymbols$ whenever that print mode is
active.  Such an arrangement is particularly useful for interactive
development, where users may type ASCII text and see mathematical
symbols displayed during proofs.›
subsection ‹Prefix Annotations›
text ‹
  Prefix syntax annotations\index{prefix annotation} are another form
  of mixfixes \<^cite>‹"isabelle-isar-ref"›, without any template arguments or
  priorities --- just some literal syntax.  The following example
  associates common symbols with the constructors of a datatype.
›