doc-src/IsarRef/Thy/document/syntax.tex
changeset 26760 2de4ba348f06
parent 26756 6634a641b961
child 26767 cc127cc0951b
--- a/doc-src/IsarRef/Thy/document/syntax.tex	Tue Apr 29 13:39:54 2008 +0200
+++ b/doc-src/IsarRef/Thy/document/syntax.tex	Tue Apr 29 13:41:11 2008 +0200
@@ -41,23 +41,23 @@
   syntax is that of Isabelle/Isar theory sources (specifications and
   proofs).  As a general rule, inner syntax entities may occur only as
   \emph{atomic entities} within outer syntax.  For example, the string
-  \texttt{"x + y"} and identifier \texttt{z} are legal term
-  specifications within a theory, while \texttt{x + y} is not.
+  \verb|"x + y"| and identifier \verb|z| are legal term
+  specifications within a theory, while \verb|x + y| without
+  quotes is not.
 
   Printed theory documents usually omit quotes to gain readability
-  (this is a matter of {\LaTeX} macro setup, say via
-  \verb,\isabellestyle,, see also \cite{isabelle-sys}).  Experienced
+  (this is a matter of {\LaTeX} macro setup, say via \verb|\isabellestyle|, see also \cite{isabelle-sys}).  Experienced
   users of Isabelle/Isar may easily reconstruct the lost technical
   information, while mere readers need not care about quotes at all.
 
   \medskip Isabelle/Isar input may contain any number of input
-  termination characters ``\texttt{;}'' (semicolon) to separate
+  termination characters ``\verb|;|'' (semicolon) to separate
   commands explicitly.  This is particularly useful in interactive
   shell sessions to make clear where the current command is intended
   to end.  Otherwise, the interpreter loop will continue to issue a
-  secondary prompt ``\verb,#,'' until an end-of-command is clearly
-  recognized from the input syntax, e.g.\ encounter of the next
-  command keyword.
+  secondary prompt ``\verb|#|'' until an end-of-command is
+  clearly recognized from the input syntax, e.g.\ encounter of the
+  next command keyword.
 
   More advanced interfaces such as Proof~General \cite{proofgeneral}
   do not require explicit semicolons, the amount of input text is
@@ -74,9 +74,10 @@
     syntax).  The standard setup should work correctly with any of the
     ``official'' logic images derived from Isabelle/HOL (including
     HOLCF etc.).  Users of alternative logics may need to tell
-    Proof~General explicitly, e.g.\ by giving an option \verb,-k ZF,
-    (in conjunction with \verb,-l ZF, to specify the default logic
-    image).
+    Proof~General explicitly, e.g.\ by giving an option \verb|-k ZF|
+    (in conjunction with \verb|-l ZF|, to specify the default
+    logic image).  Note that option \verb|-L| does both
+    of this at the same time.
   \end{warn}%
 \end{isamarkuptext}%
 \isamarkuptrue%
@@ -122,29 +123,28 @@
   \end{matharray}
 
   The syntax of \isa{string} admits any characters, including
-  newlines; ``\verb|"|'' (double-quote) and ``\verb|\|'' (backslash)
-  need to be escaped by a backslash; arbitrary character codes may be
-  specified as ``\verb|\|$ddd$'', with 3 decimal digits.  Alternative
-  strings according to \isa{altstring} are analogous, using single
-  back-quotes instead.  The body of \isa{verbatim} may consist of
-  any text not containing ``\verb,*,\verb,},''; this allows convenient
-  inclusion of quotes without further escapes.  The greek letters do
-  \emph{not} include \verb,\<lambda>,, which is already used differently in
-  the meta-logic.
+  newlines; ``\verb|"|'' (double-quote) and ``\verb|\|'' (backslash) need to be escaped by a backslash; arbitrary
+  character codes may be specified as ``\verb|\|\isa{ddd}'',
+  with three decimal digits.  Alternative strings according to
+  \isa{altstring} are analogous, using single back-quotes instead.
+  The body of \isa{verbatim} may consist of any text not
+  containing ``\verb|*|\verb|}|''; this allows
+  convenient inclusion of quotes without further escapes.  The greek
+  letters do \emph{not} include \verb|\<lambda>|, which is already used
+  differently in the meta-logic.
 
   Common mathematical symbols such as \isa{{\isasymforall}} are represented in
-  Isabelle as \verb,\<forall>,.  There are infinitely many Isabelle symbols
-  like this, although proper presentation is left to front-end tools
-  such as {\LaTeX} or Proof~General with the X-Symbol package.  A list
-  of standard Isabelle symbols that work well with these tools is
-  given in \cite[appendix~A]{isabelle-sys}.
+  Isabelle as \verb|\<forall>|.  There are infinitely many Isabelle
+  symbols like this, although proper presentation is left to front-end
+  tools such as {\LaTeX} or Proof~General with the X-Symbol package.
+  A list of standard Isabelle symbols that work well with these tools
+  is given in \cite[appendix~A]{isabelle-sys}.
   
-  Source comments take the form \texttt{(*~\dots~*)} and may be
-  nested, although user-interface tools might prevent this.  Note that
-  \texttt{(*~\dots~*)} indicate source comments only, which are
-  stripped after lexical analysis of the input.  The Isar document
-  syntax also provides formal comments that are considered as part of
-  the text (see \S\ref{sec:comments}).%
+  Source comments take the form \verb|(*|~\isa{{\isasymdots}}~\verb|*)| and may be nested, although user-interface
+  tools might prevent this.  Note that this form indicates source
+  comments only, which are stripped after lexical analysis of the
+  input.  The Isar document syntax also provides formal comments that
+  are considered as part of the text (see \secref{sec:comments}).%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -168,8 +168,9 @@
   constants, theorems etc.\ that are to be \emph{declared} or
   \emph{defined} (so qualified identifiers are excluded here).  Quoted
   strings provide an escape for non-identifier names or those ruled
-  out by outer syntax keywords (e.g.\ \verb|"let"|).  Already existing
-  objects are usually referenced by \railqtok{nameref}.
+  out by outer syntax keywords (e.g.\ quoted \verb|"let"|).
+  Already existing objects are usually referenced by
+  \railqtok{nameref}.
 
   \indexoutertoken{name}\indexoutertoken{parname}\indexoutertoken{nameref}
   \indexoutertoken{int}
@@ -192,12 +193,11 @@
 %
 \begin{isamarkuptext}%
 Large chunks of plain \railqtok{text} are usually given
-  \railtok{verbatim}, i.e.\ enclosed in
-  \verb,{,\verb,*,~\dots~\verb,*,\verb,},.  For convenience, any of
-  the smaller text units conforming to \railqtok{nameref} are admitted
-  as well.  A marginal \railnonterm{comment} is of the form
-  \texttt{--} \railqtok{text}.  Any number of these may occur within
-  Isabelle/Isar commands.
+  \railtok{verbatim}, i.e.\ enclosed in \verb|{|\verb|*|~\isa{{\isasymdots}}~\verb|*|\verb|}|.  For convenience,
+  any of the smaller text units conforming to \railqtok{nameref} are
+  admitted as well.  A marginal \railnonterm{comment} is of the form
+  \verb|--| \railqtok{text}.  Any number of these may occur
+  within Isabelle/Isar commands.
 
   \indexoutertoken{text}\indexouternonterm{comment}
   \begin{rail}
@@ -248,10 +248,10 @@
   is performed internally later).  For convenience, a slightly more
   liberal convention is adopted: quotes may be omitted for any type or
   term that is already atomic at the outer level.  For example, one
-  may just write \texttt{x} instead of \texttt{"x"}.  Note that
-  symbolic identifiers (e.g.\ \texttt{++} or \isa{{\isasymforall}} are available
-  as well, provided these have not been superseded by commands or
-  other keywords already (e.g.\ \texttt{=} or \texttt{+}).
+  may just write \verb|x| instead of quoted \verb|"x"|.
+  Note that symbolic identifiers (e.g.\ \verb|++| or \isa{{\isasymforall}} are available as well, provided these have not been superseded
+  by commands or other keywords already (such as \verb|=| or
+  \verb|+|).
 
   \indexoutertoken{type}\indexoutertoken{term}\indexoutertoken{prop}
   \begin{rail}
@@ -264,8 +264,8 @@
   \end{rail}
 
   Positional instantiations are indicated by giving a sequence of
-  terms, or the placeholder ``$\_$'' (underscore), which means to skip
-  a position.
+  terms, or the placeholder ``\verb|_|'' (underscore), which
+  means to skip a position.
 
   \indexoutertoken{inst}\indexoutertoken{insts}
   \begin{rail}
@@ -295,8 +295,8 @@
 \begin{isamarkuptext}%
 Mixfix annotations specify concrete \emph{inner} syntax of Isabelle
   types and terms.  Some commands such as \isa{\isacommand{types}} (see
-  \S\ref{sec:types-pure}) admit infixes only, while \isa{\isacommand{consts}} (see \S\ref{sec:consts}) and \isa{\isacommand{syntax}} (see
-  \S\ref{sec:syn-trans}) support the full range of general mixfixes
+  \secref{sec:types-pure}) admit infixes only, while \isa{\isacommand{consts}} (see \secref{sec:consts}) and \isa{\isacommand{syntax}} (see
+  \secref{sec:syn-trans}) support the full range of general mixfixes
   and binders.
 
   \indexouternonterm{infix}\indexouternonterm{mixfix}\indexouternonterm{structmixfix}
@@ -314,14 +314,15 @@
 
   Here the \railtok{string} specifications refer to the actual mixfix
   template (see also \cite{isabelle-ref}), which may include literal
-  text, spacing, blocks, and arguments (denoted by ``$_$''); the
-  special symbol \verb,\<index>, (printed as ``\i'') represents an index
-  argument that specifies an implicit structure reference (see also
-  \S\ref{sec:locale}).  Infix and binder declarations provide common
-  abbreviations for particular mixfix declarations.  So in practice,
-  mixfix templates mostly degenerate to literal text for concrete
-  syntax, such as ``\verb,++,'' for an infix symbol, or
-  ``\verb,++,\i'' for an infix of an implicit structure.%
+  text, spacing, blocks, and arguments (denoted by ``\verb|_|'');
+  the special symbol ``\verb|\<index>|'' (printed as ``\isa{{\isasymindex}}'')
+  represents an index argument that specifies an implicit structure
+  reference (see also \secref{sec:locale}).  Infix and binder
+  declarations provide common abbreviations for particular mixfix
+  declarations.  So in practice, mixfix templates mostly degenerate to
+  literal text for concrete syntax, such as ``\verb|++|'' for
+  an infix symbol, or ``\verb|++|\isa{{\isasymindex}}'' for an infix of
+  an implicit structure.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -331,13 +332,14 @@
 %
 \begin{isamarkuptext}%
 Proof methods are either basic ones, or expressions composed of
-  methods via ``\texttt{,}'' (sequential composition), ``\texttt{|}''
-  (alternative choices), ``\texttt{?}'' (try), ``\texttt{+}'' (repeat
-  at least once), ``\texttt{[$n$]}'' (restriction to first \isa{n}
-  sub-goals, default $n = 1$).  In practice, proof methods are usually
-  just a comma separated list of \railqtok{nameref}~\railnonterm{args}
-  specifications.  Note that parentheses may be dropped for single
-  method specifications (with no arguments).
+  methods via ``\verb|,|'' (sequential composition),
+  ``\verb||\verb,|,\verb||'' (alternative choices), ``\verb|?|'' 
+  (try), ``\verb|+|'' (repeat at least once), ``\verb|[|\isa{n}\verb|]|'' (restriction to first \isa{n}
+  sub-goals, with default \isa{n\ {\isacharequal}\ {\isadigit{1}}}).  In practice, proof
+  methods are usually just a comma separated list of
+  \railqtok{nameref}~\railnonterm{args} specifications.  Note that
+  parentheses may be dropped for single method specifications (with no
+  arguments).
 
   \indexouternonterm{method}
   \begin{rail}
@@ -349,17 +351,18 @@
 
   Proper Isar proof methods do \emph{not} admit arbitrary goal
   addressing, but refer either to the first sub-goal or all sub-goals
-  uniformly.  The goal restriction operator ``\texttt{[$n$]}''
+  uniformly.  The goal restriction operator ``\isa{{\isacharbrackleft}n{\isacharbrackright}}''
   evaluates a method expression within a sandbox consisting of the
-  first \isa{n} sub-goals (which need to exist).  For example,
-  \isa{simp{\isacharunderscore}all{\isacharbrackleft}{\isadigit{3}}{\isacharbrackright}} simplifies the first three sub-goals, while
-  \isa{{\isacharparenleft}rule\ foo{\isacharcomma}\ simp{\isacharunderscore}all{\isacharparenright}{\isacharbrackleft}{\isacharbrackright}} simplifies all new goals that
-  emerge from applying rule \isa{foo} to the originally first one.
+  first \isa{n} sub-goals (which need to exist).  For example, the
+  method ``\isa{simp{\isacharunderscore}all{\isacharbrackleft}{\isadigit{3}}{\isacharbrackright}}'' simplifies the first three
+  sub-goals, while ``\isa{{\isacharparenleft}rule\ foo{\isacharcomma}\ simp{\isacharunderscore}all{\isacharparenright}{\isacharbrackleft}{\isacharbrackright}}'' simplifies all
+  new goals that emerge from applying rule \isa{foo} to the
+  originally first one.
 
   Improper methods, notably tactic emulations, offer a separate
   low-level goal addressing scheme as explicit argument to the
-  individual tactic being involved.  Here \isa{{\isacharbrackleft}{\isacharbang}{\isacharbrackright}} refers to all
-  goals, and \isa{{\isacharbrackleft}n{\isacharminus}{\isacharbrackright}} to all goals starting from \isa{n},
+  individual tactic being involved.  Here ``\isa{{\isacharbrackleft}{\isacharbang}{\isacharbrackright}}'' refers to
+  all goals, and ``\isa{{\isacharbrackleft}n{\isacharminus}{\isacharbrackright}}'' to all goals starting from \isa{n}.
 
   \indexouternonterm{goalspec}
   \begin{rail}
@@ -374,7 +377,7 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-Attributes (and proof methods, see \S\ref{sec:syn-meth}) have their
+Attributes (and proof methods, see \secref{sec:syn-meth}) have their
   own ``semi-inner'' syntax, in the sense that input conforming to
   \railnonterm{args} below is parsed by the attribute a second time.
   The attribute argument specifications may be any sequence of atomic
@@ -405,13 +408,13 @@
   There are three forms of theorem references:
   \begin{enumerate}
   
-  \item named facts \isa{a}
+  \item named facts \isa{a},
 
-  \item selections from named facts \isa{a{\isacharparenleft}i{\isacharcomma}\ j\ {\isacharminus}\ k{\isacharparenright}}
+  \item selections from named facts \isa{a{\isacharparenleft}i{\isacharparenright}} or \isa{a{\isacharparenleft}j\ {\isacharminus}\ k{\isacharparenright}},
 
   \item literal fact propositions using \indexref{}{syntax}{altstring}\isa{altstring} syntax
-  $\backquote\phi\backquote$, (see also method \indexref{}{method}{fact}\isa{fact} in
-  \S\ref{sec:pure-meth-att}).
+  \verb|`|\isa{{\isasymphi}}\verb|`| (see also method
+  \indexref{}{method}{fact}\isa{fact} in \secref{sec:pure-meth-att}).
 
   \end{enumerate}
 
@@ -421,12 +424,11 @@
   not stored within the theorem database of the theory or proof
   context, but any given attributes are applied nonetheless.
 
-  An extra pair of brackets around attribute declarations --- such as
-  ``\isa{{\isacharbrackleft}{\isacharbrackleft}simproc\ a{\isacharbrackright}{\isacharbrackright}}'' --- abbreviates a theorem reference
-  involving an internal dummy fact, which will be ignored later on.
-  So only the effect of the attribute on the background context will
-  persist.  This form of in-place declarations is particularly useful
-  with commands like \isa{\isacommand{declare}} and \isa{\isacommand{using}}.
+  An extra pair of brackets around attributes (like ``\isa{{\isacharbrackleft}{\isacharbrackleft}simproc\ a{\isacharbrackright}{\isacharbrackright}}'') abbreviates a theorem reference involving an
+  internal dummy fact, which will be ignored later on.  So only the
+  effect of the attribute on the background context will persist.
+  This form of in-place declarations is particularly useful with
+  commands like \isa{\isacommand{declare}} and \isa{\isacommand{using}}.
 
   \indexouternonterm{axmdecl}\indexouternonterm{thmdecl}
   \indexouternonterm{thmdef}\indexouternonterm{thmref}
@@ -492,7 +494,7 @@
   \railnonterm{props} typically admit separate typings or namings via
   another level of iteration, with explicit \indexref{}{keyword}{and}\isa{\isakeyword{and}}
   separators; e.g.\ see \isa{\isacommand{fix}} and \isa{\isacommand{assume}} in
-  \S\ref{sec:proof-context}.%
+  \secref{sec:proof-context}.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -522,19 +524,19 @@
     \indexdef{}{antiquotation}{ML-struct}\isa{ML{\isacharunderscore}struct} & : & \isarantiq \\
   \end{matharray}
 
-  The text body of formal comments (see also \S\ref{sec:comments}) may
-  contain antiquotations of logical entities, such as theorems, terms
-  and types, which are to be presented in the final output produced by
-  the Isabelle document preparation system (see also
-  \S\ref{sec:document-prep}).
+  The text body of formal comments (see also \secref{sec:comments})
+  may contain antiquotations of logical entities, such as theorems,
+  terms and types, which are to be presented in the final output
+  produced by the Isabelle document preparation system (see also
+  \secref{sec:document-prep}).
 
   Thus embedding of ``\isa{{\isacharat}{\isacharbraceleft}term\ {\isacharbrackleft}show{\isacharunderscore}types{\isacharbrackright}\ {\isachardoublequote}f\ x\ {\isacharequal}\ a\ {\isacharplus}\ x{\isachardoublequote}{\isacharbraceright}}''
   within a text block would cause
   \isa{{\isacharparenleft}f{\isasymColon}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}\ {\isacharparenleft}x{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}a{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isacharplus}\ x} to appear in the final {\LaTeX} document.  Also note that theorem
   antiquotations may involve attributes as well.  For example,
-  \texttt{{\at}{\ttlbrace}thm~sym~[no_vars]{\ttrbrace}} would print
-  the statement where all schematic variables have been replaced by
-  fixed ones, which are easier to read.
+  \isa{{\isacharat}{\isacharbraceleft}thm\ sym\ {\isacharbrackleft}no{\isacharunderscore}vars{\isacharbrackright}{\isacharbraceright}} would print the theorem's
+  statement where all schematic variables have been replaced by fixed
+  ones, which are easier to read.
 
   \begin{rail}
     atsign lbrace antiquotation rbrace
@@ -567,8 +569,8 @@
   \end{rail}
 
   Note that the syntax of antiquotations may \emph{not} include source
-  comments \texttt{(*~\dots~*)} or verbatim text
-  \verb|{*|~\dots~\verb|*|\verb|}|.
+  comments \verb|(*|~\isa{{\isasymdots}}~\verb|*)| or verbatim
+  text \verb|{|\verb|*|~\isa{{\isasymdots}}~\verb|*|\verb|}|.
 
   \begin{descr}
   
@@ -576,9 +578,11 @@
   guaranteed to refer to a valid ancestor theory in the current
   context.
 
-  \item [\isa{{\isacharat}{\isacharbraceleft}thm\ a\isactrlsub {\isadigit{1}}\ {\isasymdots}\ a\isactrlsub n{\isacharbraceright}}] prints theorems \isa{a\isactrlsub {\isadigit{1}}\ {\isasymdots}\ a\isactrlsub n}.  Note that attribute specifications may be
-  included as well (see also \S\ref{sec:syn-att}); the \indexref{}{attribute}{no-vars}\isa{no{\isacharunderscore}vars} rule (see \S\ref{sec:misc-meth-att}) would be particularly
-  useful to suppress printing of schematic variables.
+  \item [\isa{{\isacharat}{\isacharbraceleft}thm\ a\isactrlsub {\isadigit{1}}\ {\isasymdots}\ a\isactrlsub n{\isacharbraceright}}] prints theorems
+  \isa{a\isactrlsub {\isadigit{1}}\ {\isasymdots}\ a\isactrlsub n}.  Note that attribute specifications
+  may be included as well (see also \secref{sec:syn-att}); the
+  \indexref{}{attribute}{no-vars}\isa{no{\isacharunderscore}vars} rule (see \secref{sec:misc-meth-att}) would
+  be particularly useful to suppress printing of schematic variables.
 
   \item [\isa{{\isacharat}{\isacharbraceleft}prop\ {\isasymphi}{\isacharbraceright}}] prints a well-typed proposition \isa{{\isasymphi}}.
 
@@ -624,7 +628,7 @@
   
   \item [\isa{{\isacharat}{\isacharbraceleft}full{\isacharunderscore}prf\ a\isactrlsub {\isadigit{1}}\ {\isasymdots}\ a\isactrlsub n{\isacharbraceright}}] is like \isa{{\isacharat}{\isacharbraceleft}prf\ a\isactrlsub {\isadigit{1}}\ {\isasymdots}\ a\isactrlsub n{\isacharbraceright}}, but displays the full proof terms,
   i.e.\ also displays information omitted in the compact proof term,
-  which is denoted by ``$_$'' placeholders there.
+  which is denoted by ``\verb|_|'' placeholders there.
   
   \item [\isa{{\isacharat}{\isacharbraceleft}ML\ s{\isacharbraceright}}, \isa{{\isacharat}{\isacharbraceleft}ML{\isacharunderscore}type\ s{\isacharbraceright}}, and \isa{{\isacharat}{\isacharbraceleft}ML{\isacharunderscore}struct\ s{\isacharbraceright}}] check text \isa{s} as ML value, type, and
   structure, respectively.  The source is displayed verbatim.
@@ -646,7 +650,7 @@
   in Horn-clause normal form \isa{A\isactrlsub {\isadigit{1}}\ {\isasymLongrightarrow}\ {\isasymdots}\ A\isactrlsub n\ {\isasymLongrightarrow}\ C}.
   
   \item [\isa{prem{\isadigit{1}}}, \dots, \isa{prem{\isadigit{9}}}] extract premise
-  number $1$, \dots, $9$, respectively, from from a rule in
+  number \isa{{\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isadigit{9}}}, respectively, from from a rule in
   Horn-clause normal form \isa{A\isactrlsub {\isadigit{1}}\ {\isasymLongrightarrow}\ {\isasymdots}\ A\isactrlsub n\ {\isasymLongrightarrow}\ C}
 
   \end{descr}