--- a/doc-src/IsarRef/Thy/document/intro.tex Tue Apr 29 13:39:54 2008 +0200
+++ b/doc-src/IsarRef/Thy/document/intro.tex Tue Apr 29 13:41:11 2008 +0200
@@ -88,11 +88,11 @@
\isamarkuptrue%
%
\begin{isamarkuptext}%
-Isar is already part of Isabelle. The low-level \texttt{isabelle} binary
- provides option \texttt{-I} to run the Isabelle/Isar interaction loop at
- startup, rather than the raw ML top-level. So the most basic way to do
- anything with Isabelle/Isar is as follows:
-\begin{ttbox}
+Isar is already part of Isabelle. The low-level \verb|isabelle| binary provides option \verb|-I| to run the
+ Isabelle/Isar interaction loop at startup, rather than the raw ML
+ top-level. So the most basic way to do anything with Isabelle/Isar
+ is as follows:
+\begin{ttbox} % FIXME update
isabelle -I HOL\medskip
\out{> Welcome to Isabelle/HOL (Isabelle2005)}\medskip
theory Foo imports Main begin;
@@ -101,10 +101,9 @@
end;
\end{ttbox}
- Note that any Isabelle/Isar command may be retracted by
- \texttt{undo}. See the Isabelle/Isar Quick Reference
- (appendix~\ref{ap:refcard}) for a comprehensive overview of
- available commands and other language elements.%
+ Note that any Isabelle/Isar command may be retracted by \isa{\isacommand{undo}}. See the Isabelle/Isar Quick Reference
+ (\appref{ap:refcard}) for a comprehensive overview of available
+ commands and other language elements.%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -137,11 +136,10 @@
Proof~General (including XEmacs or GNU Emacs). The default
configuration of Isabelle is smart enough to detect the
Proof~General distribution in several canonical places (e.g.\
- \texttt{\$ISABELLE_HOME/contrib/ProofGeneral}). Thus the capital
- \texttt{Isabelle} executable would already refer to the
- \texttt{ProofGeneral/isar} interface without further ado. The
- Isabelle interface script provides several options; pass \verb,-?,
- to see its usage.
+ \verb|$ISABELLE_HOME/contrib/ProofGeneral|). Thus the
+ capital \verb|Isabelle| executable would already refer to the
+ \verb|ProofGeneral/isar| interface without further ado. The
+ Isabelle interface script provides several options; pass \verb|-?| to see its usage.
With the proper Isabelle interface setup, Isar documents may now be edited by
visiting appropriate theory files, e.g.\
@@ -151,8 +149,8 @@
Beginners may note the tool bar for navigating forward and backward
through the text (this depends on the local Emacs installation).
Consult the Proof~General documentation \cite{proofgeneral} for
- further basic command sequences, in particular ``\texttt{C-c
- C-return}'' and ``\texttt{C-c u}''.
+ further basic command sequences, in particular ``\verb|C-c C-return|''
+ and ``\verb|C-c u|''.
\medskip Proof~General may be also configured manually by giving
Isabelle settings like this (see also \cite{isabelle-sys}):
@@ -161,8 +159,8 @@
ISABELLE_INTERFACE=\$ISABELLE_HOME/contrib/ProofGeneral/isar/interface
PROOFGENERAL_OPTIONS=""
\end{ttbox}
- You may have to change \texttt{\$ISABELLE_HOME/contrib/ProofGeneral}
- to the actual installation directory of Proof~General.
+ You may have to change \verb|$ISABELLE_HOME/contrib/ProofGeneral| to the actual installation
+ directory of Proof~General.
\medskip Apart from the Isabelle command line, defaults for
interface options may be given by the \texttt{PROOFGENERAL_OPTIONS}
@@ -172,13 +170,12 @@
PROOFGENERAL_OPTIONS="-p xemacs-mule"
\end{ttbox}
- Occasionally, a user's \verb,~/.emacs, file contains code that is
- incompatible with the (X)Emacs version used by Proof~General,
- causing the interface startup to fail prematurely. Here the
- \texttt{-u false} option helps to get the interface process up and
- running. Note that additional Lisp customization code may reside in
- \texttt{proofgeneral-settings.el} of \texttt{\$ISABELLE_HOME/etc} or
- \texttt{\$ISABELLE_HOME_USER/etc}.%
+ Occasionally, a user's \verb|~/.emacs| file contains code
+ that is incompatible with the (X)Emacs version used by
+ Proof~General, causing the interface startup to fail prematurely.
+ Here the \texttt{-u false} option helps to get the interface process
+ up and running. Note that additional Lisp customization code may
+ reside in \texttt{proofgeneral-settings.el} of \verb|$ISABELLE_HOME/etc| or \verb|$ISABELLE_HOME_USER/etc|.%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -197,9 +194,9 @@
\medskip Using proper mathematical symbols in Isabelle theories can
be very convenient for readability of large formulas. On the other
hand, the plain ASCII sources easily become somewhat unintelligible.
- For example, $\Longrightarrow$ would appear as \verb,\<Longrightarrow>, according
+ For example, \isa{{\isasymLongrightarrow}} would appear as \verb|\<Longrightarrow>| according
the default set of Isabelle symbols. Nevertheless, the Isabelle
- document preparation system (see \S\ref{sec:document-prep}) will be
+ document preparation system (see \secref{sec:document-prep}) will be
happy to print non-ASCII symbols properly. It is even possible to
invent additional notation beyond the display capabilities of Emacs
and X-Symbol.%
@@ -236,28 +233,11 @@
The Isar proof language is embedded into the new theory format as a
proper sub-language. Proof mode is entered by stating some
- $\THEOREMNAME$ or $\LEMMANAME$ at the theory level, and left again
- with the final conclusion (e.g.\ via $\QEDNAME$). A few theory
- specification mechanisms also require some proof, such as HOL's
- $\isarkeyword{typedef}$ which demands non-emptiness of the
- representing sets.
-
- New-style theory files may still be associated with separate ML
- files consisting of plain old tactic scripts. There is no longer
- any ML binding generated for the theory and theorems, though. ML
- functions \texttt{theory}, \texttt{thm}, and \texttt{thms} retrieve
- this information from the context \cite{isabelle-ref}.
- Nevertheless, migration between classic Isabelle and Isabelle/Isar
- is relatively easy. Thus users may start to benefit from
- interactive theory development and document preparation, even before
- they have any idea of the Isar proof language at all.
-
- Manual conversion of existing tactic scripts may be done by running
- two separate Proof~General sessions, one for replaying the old
- script and the other for the emerging Isabelle/Isar document. Also
- note that Isar supports emulation commands and methods that support
- traditional tactic scripts within new-style theories, see
- appendix~\ref{ap:conv} for more information.%
+ \isa{\isacommand{theorem}} or \isa{\isacommand{lemma}} at the theory level, and
+ left again with the final conclusion (e.g.\ via \isa{\isacommand{qed}}).
+ A few theory specification mechanisms also require some proof, such
+ as HOL's \isa{\isacommand{typedef}} which demands non-emptiness of the
+ representing sets.%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -277,35 +257,37 @@
Isabelle generates {\LaTeX} output as part of the run of a
\emph{logic session} (see also \cite{isabelle-sys}). Getting
started with a working configuration for common situations is quite
- easy by using the Isabelle \texttt{mkdir} and \texttt{make} tools.
- First invoke
+ easy by using the Isabelle \verb|mkdir| and \verb|make|
+ tools. First invoke
\begin{ttbox}
isatool mkdir Foo
\end{ttbox}
- to initialize a separate directory for session \texttt{Foo} --- it
- is safe to experiment, since \texttt{isatool mkdir} never overwrites
- existing files. Ensure that \texttt{Foo/ROOT.ML} holds ML commands
- to load all theories required for this session; furthermore
- \texttt{Foo/document/root.tex} should include any special {\LaTeX}
- macro packages required for your document (the default is usually
- sufficient as a start).
+ to initialize a separate directory for session \verb|Foo| ---
+ it is safe to experiment, since \texttt{isatool mkdir} never
+ overwrites existing files. Ensure that \verb|Foo/ROOT.ML|
+ holds ML commands to load all theories required for this session;
+ furthermore \verb|Foo/document/root.tex| should include any
+ special {\LaTeX} macro packages required for your document (the
+ default is usually sufficient as a start).
- The session is controlled by a separate \texttt{IsaMakefile} (with
- crude source dependencies by default). This file is located one
- level up from the \texttt{Foo} directory location. Now invoke
+ The session is controlled by a separate \verb|IsaMakefile|
+ (with crude source dependencies by default). This file is located
+ one level up from the \verb|Foo| directory location. Now
+ invoke
\begin{ttbox}
isatool make Foo
\end{ttbox}
- to run the \texttt{Foo} session, with browser information and
+ to run the \verb|Foo| session, with browser information and
document preparation enabled. Unless any errors are reported by
Isabelle or {\LaTeX}, the output will appear inside the directory
- \texttt{ISABELLE_BROWSER_INFO}, as reported by the batch job in
+ \verb|ISABELLE_BROWSER_INFO|, as reported by the batch job in
verbose mode.
- \medskip You may also consider to tune the \texttt{usedir} options
- in \texttt{IsaMakefile}, for example to change the output format
- from \texttt{pdf} to \texttt{dvi}, or activate the \texttt{-D}
- option to retain a second copy of the generated {\LaTeX} sources.
+ \medskip You may also consider to tune the \verb|usedir|
+ options in \verb|IsaMakefile|, for example to change the output
+ format from \verb|pdf| to \verb|dvi|, or activate the
+ \verb|-D| option to retain a second copy of the generated
+ {\LaTeX} sources.
\medskip See \emph{The Isabelle System Manual} \cite{isabelle-sys}
for further details on Isabelle logic sessions and theory
@@ -335,11 +317,10 @@
\medskip The present text really is only a reference manual on
Isabelle/Isar, not a tutorial. Nevertheless, we will attempt to
give some clues of how the concepts introduced here may be put into
- practice. Appendix~\ref{ap:refcard} provides a quick reference card
- of the most common Isabelle/Isar language elements.
- Appendix~\ref{ap:conv} offers some practical hints on converting
- existing Isabelle theories and proof scripts to the new format
- (without restructuring proofs).
+ practice. \Appref{ap:refcard} provides a quick reference card of
+ the most common Isabelle/Isar language elements. \Appref{ap:conv}
+ offers some practical hints on converting existing Isabelle theories
+ and proof scripts to the new format (without restructuring proofs).
Further issues concerning the Isar concepts are covered in the
literature
--- 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}
--- a/doc-src/IsarRef/Thy/intro.thy Tue Apr 29 13:39:54 2008 +0200
+++ b/doc-src/IsarRef/Thy/intro.thy Tue Apr 29 13:41:11 2008 +0200
@@ -64,11 +64,12 @@
subsection {* Terminal sessions *}
text {*
- Isar is already part of Isabelle. The low-level \texttt{isabelle} binary
- provides option \texttt{-I} to run the Isabelle/Isar interaction loop at
- startup, rather than the raw ML top-level. So the most basic way to do
- anything with Isabelle/Isar is as follows:
-\begin{ttbox}
+ Isar is already part of Isabelle. The low-level @{verbatim
+ isabelle} binary provides option @{verbatim "-I"} to run the
+ Isabelle/Isar interaction loop at startup, rather than the raw ML
+ top-level. So the most basic way to do anything with Isabelle/Isar
+ is as follows:
+\begin{ttbox} % FIXME update
isabelle -I HOL\medskip
\out{> Welcome to Isabelle/HOL (Isabelle2005)}\medskip
theory Foo imports Main begin;
@@ -77,10 +78,10 @@
end;
\end{ttbox}
- Note that any Isabelle/Isar command may be retracted by
- \texttt{undo}. See the Isabelle/Isar Quick Reference
- (appendix~\ref{ap:refcard}) for a comprehensive overview of
- available commands and other language elements.
+ Note that any Isabelle/Isar command may be retracted by @{command
+ "undo"}. See the Isabelle/Isar Quick Reference
+ (\appref{ap:refcard}) for a comprehensive overview of available
+ commands and other language elements.
*}
@@ -109,11 +110,11 @@
Proof~General (including XEmacs or GNU Emacs). The default
configuration of Isabelle is smart enough to detect the
Proof~General distribution in several canonical places (e.g.\
- \texttt{\$ISABELLE_HOME/contrib/ProofGeneral}). Thus the capital
- \texttt{Isabelle} executable would already refer to the
- \texttt{ProofGeneral/isar} interface without further ado. The
- Isabelle interface script provides several options; pass \verb,-?,
- to see its usage.
+ @{verbatim "$ISABELLE_HOME/contrib/ProofGeneral"}). Thus the
+ capital @{verbatim Isabelle} executable would already refer to the
+ @{verbatim "ProofGeneral/isar"} interface without further ado. The
+ Isabelle interface script provides several options; pass @{verbatim
+ "-?"} to see its usage.
With the proper Isabelle interface setup, Isar documents may now be edited by
visiting appropriate theory files, e.g.\
@@ -123,8 +124,8 @@
Beginners may note the tool bar for navigating forward and backward
through the text (this depends on the local Emacs installation).
Consult the Proof~General documentation \cite{proofgeneral} for
- further basic command sequences, in particular ``\texttt{C-c
- C-return}'' and ``\texttt{C-c u}''.
+ further basic command sequences, in particular ``@{verbatim "C-c C-return"}''
+ and ``@{verbatim "C-c u"}''.
\medskip Proof~General may be also configured manually by giving
Isabelle settings like this (see also \cite{isabelle-sys}):
@@ -133,8 +134,9 @@
ISABELLE_INTERFACE=\$ISABELLE_HOME/contrib/ProofGeneral/isar/interface
PROOFGENERAL_OPTIONS=""
\end{ttbox}
- You may have to change \texttt{\$ISABELLE_HOME/contrib/ProofGeneral}
- to the actual installation directory of Proof~General.
+ You may have to change @{verbatim
+ "$ISABELLE_HOME/contrib/ProofGeneral"} to the actual installation
+ directory of Proof~General.
\medskip Apart from the Isabelle command line, defaults for
interface options may be given by the \texttt{PROOFGENERAL_OPTIONS}
@@ -144,13 +146,13 @@
PROOFGENERAL_OPTIONS="-p xemacs-mule"
\end{ttbox}
- Occasionally, a user's \verb,~/.emacs, file contains code that is
- incompatible with the (X)Emacs version used by Proof~General,
- causing the interface startup to fail prematurely. Here the
- \texttt{-u false} option helps to get the interface process up and
- running. Note that additional Lisp customization code may reside in
- \texttt{proofgeneral-settings.el} of \texttt{\$ISABELLE_HOME/etc} or
- \texttt{\$ISABELLE_HOME_USER/etc}.
+ Occasionally, a user's @{verbatim "~/.emacs"} file contains code
+ that is incompatible with the (X)Emacs version used by
+ Proof~General, causing the interface startup to fail prematurely.
+ Here the \texttt{-u false} option helps to get the interface process
+ up and running. Note that additional Lisp customization code may
+ reside in \texttt{proofgeneral-settings.el} of @{verbatim
+ "$ISABELLE_HOME/etc"} or @{verbatim "$ISABELLE_HOME_USER/etc"}.
*}
@@ -167,9 +169,9 @@
\medskip Using proper mathematical symbols in Isabelle theories can
be very convenient for readability of large formulas. On the other
hand, the plain ASCII sources easily become somewhat unintelligible.
- For example, $\Longrightarrow$ would appear as \verb,\<Longrightarrow>, according
+ For example, @{text "\<Longrightarrow>"} would appear as @{verbatim "\<Longrightarrow>"} according
the default set of Isabelle symbols. Nevertheless, the Isabelle
- document preparation system (see \S\ref{sec:document-prep}) will be
+ document preparation system (see \secref{sec:document-prep}) will be
happy to print non-ASCII symbols properly. It is even possible to
invent additional notation beyond the display capabilities of Emacs
and X-Symbol.
@@ -204,28 +206,11 @@
The Isar proof language is embedded into the new theory format as a
proper sub-language. Proof mode is entered by stating some
- $\THEOREMNAME$ or $\LEMMANAME$ at the theory level, and left again
- with the final conclusion (e.g.\ via $\QEDNAME$). A few theory
- specification mechanisms also require some proof, such as HOL's
- $\isarkeyword{typedef}$ which demands non-emptiness of the
+ @{command "theorem"} or @{command "lemma"} at the theory level, and
+ left again with the final conclusion (e.g.\ via @{command "qed"}).
+ A few theory specification mechanisms also require some proof, such
+ as HOL's @{command "typedef"} which demands non-emptiness of the
representing sets.
-
- New-style theory files may still be associated with separate ML
- files consisting of plain old tactic scripts. There is no longer
- any ML binding generated for the theory and theorems, though. ML
- functions \texttt{theory}, \texttt{thm}, and \texttt{thms} retrieve
- this information from the context \cite{isabelle-ref}.
- Nevertheless, migration between classic Isabelle and Isabelle/Isar
- is relatively easy. Thus users may start to benefit from
- interactive theory development and document preparation, even before
- they have any idea of the Isar proof language at all.
-
- Manual conversion of existing tactic scripts may be done by running
- two separate Proof~General sessions, one for replaying the old
- script and the other for the emerging Isabelle/Isar document. Also
- note that Isar supports emulation commands and methods that support
- traditional tactic scripts within new-style theories, see
- appendix~\ref{ap:conv} for more information.
*}
@@ -243,35 +228,37 @@
Isabelle generates {\LaTeX} output as part of the run of a
\emph{logic session} (see also \cite{isabelle-sys}). Getting
started with a working configuration for common situations is quite
- easy by using the Isabelle \texttt{mkdir} and \texttt{make} tools.
- First invoke
+ easy by using the Isabelle @{verbatim mkdir} and @{verbatim make}
+ tools. First invoke
\begin{ttbox}
isatool mkdir Foo
\end{ttbox}
- to initialize a separate directory for session \texttt{Foo} --- it
- is safe to experiment, since \texttt{isatool mkdir} never overwrites
- existing files. Ensure that \texttt{Foo/ROOT.ML} holds ML commands
- to load all theories required for this session; furthermore
- \texttt{Foo/document/root.tex} should include any special {\LaTeX}
- macro packages required for your document (the default is usually
- sufficient as a start).
+ to initialize a separate directory for session @{verbatim Foo} ---
+ it is safe to experiment, since \texttt{isatool mkdir} never
+ overwrites existing files. Ensure that @{verbatim "Foo/ROOT.ML"}
+ holds ML commands to load all theories required for this session;
+ furthermore @{verbatim "Foo/document/root.tex"} should include any
+ special {\LaTeX} macro packages required for your document (the
+ default is usually sufficient as a start).
- The session is controlled by a separate \texttt{IsaMakefile} (with
- crude source dependencies by default). This file is located one
- level up from the \texttt{Foo} directory location. Now invoke
+ The session is controlled by a separate @{verbatim IsaMakefile}
+ (with crude source dependencies by default). This file is located
+ one level up from the @{verbatim Foo} directory location. Now
+ invoke
\begin{ttbox}
isatool make Foo
\end{ttbox}
- to run the \texttt{Foo} session, with browser information and
+ to run the @{verbatim Foo} session, with browser information and
document preparation enabled. Unless any errors are reported by
Isabelle or {\LaTeX}, the output will appear inside the directory
- \texttt{ISABELLE_BROWSER_INFO}, as reported by the batch job in
+ @{verbatim ISABELLE_BROWSER_INFO}, as reported by the batch job in
verbose mode.
- \medskip You may also consider to tune the \texttt{usedir} options
- in \texttt{IsaMakefile}, for example to change the output format
- from \texttt{pdf} to \texttt{dvi}, or activate the \texttt{-D}
- option to retain a second copy of the generated {\LaTeX} sources.
+ \medskip You may also consider to tune the @{verbatim usedir}
+ options in @{verbatim IsaMakefile}, for example to change the output
+ format from @{verbatim pdf} to @{verbatim dvi}, or activate the
+ @{verbatim "-D"} option to retain a second copy of the generated
+ {\LaTeX} sources.
\medskip See \emph{The Isabelle System Manual} \cite{isabelle-sys}
for further details on Isabelle logic sessions and theory
@@ -299,11 +286,10 @@
\medskip The present text really is only a reference manual on
Isabelle/Isar, not a tutorial. Nevertheless, we will attempt to
give some clues of how the concepts introduced here may be put into
- practice. Appendix~\ref{ap:refcard} provides a quick reference card
- of the most common Isabelle/Isar language elements.
- Appendix~\ref{ap:conv} offers some practical hints on converting
- existing Isabelle theories and proof scripts to the new format
- (without restructuring proofs).
+ practice. \Appref{ap:refcard} provides a quick reference card of
+ the most common Isabelle/Isar language elements. \Appref{ap:conv}
+ offers some practical hints on converting existing Isabelle theories
+ and proof scripts to the new format (without restructuring proofs).
Further issues concerning the Isar concepts are covered in the
literature
--- a/doc-src/IsarRef/Thy/syntax.thy Tue Apr 29 13:39:54 2008 +0200
+++ b/doc-src/IsarRef/Thy/syntax.thy Tue Apr 29 13:41:11 2008 +0200
@@ -23,23 +23,24 @@
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.
+ @{verbatim "\"x + y\""} and identifier @{verbatim z} are legal term
+ specifications within a theory, while @{verbatim "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 @{verbatim
+ "\\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 ``@{verbatim ";"}'' (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 ``@{verbatim "#"}'' 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
@@ -56,9 +57,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 @{verbatim "-k ZF"}
+ (in conjunction with @{verbatim "-l ZF"}, to specify the default
+ logic image). Note that option @{verbatim "-L"} does both
+ of this at the same time.
\end{warn}
*}
@@ -102,29 +104,30 @@
\end{matharray}
The syntax of @{syntax 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 @{syntax altstring} are analogous, using single
- back-quotes instead. The body of @{syntax 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; ``@{verbatim "\""}'' (double-quote) and ``@{verbatim
+ "\\"}'' (backslash) need to be escaped by a backslash; arbitrary
+ character codes may be specified as ``@{verbatim "\\"}@{text ddd}'',
+ with three decimal digits. Alternative strings according to
+ @{syntax altstring} are analogous, using single back-quotes instead.
+ The body of @{syntax verbatim} may consist of any text not
+ containing ``@{verbatim "*"}@{verbatim "}"}''; this allows
+ convenient inclusion of quotes without further escapes. The greek
+ letters do \emph{not} include @{verbatim "\<lambda>"}, which is already used
+ differently in the meta-logic.
Common mathematical symbols such as @{text \<forall>} 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 @{verbatim \<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 @{verbatim "(*"}~@{text
+ "\<dots>"}~@{verbatim "*)"} 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}).
*}
@@ -144,8 +147,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 @{verbatim "\"let\""}).
+ Already existing objects are usually referenced by
+ \railqtok{nameref}.
\indexoutertoken{name}\indexoutertoken{parname}\indexoutertoken{nameref}
\indexoutertoken{int}
@@ -166,12 +170,12 @@
text {*
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 @{verbatim "{"}@{verbatim
+ "*"}~@{text "\<dots>"}~@{verbatim "*"}@{verbatim "}"}. For convenience,
+ any of the smaller text units conforming to \railqtok{nameref} are
+ admitted as well. A marginal \railnonterm{comment} is of the form
+ @{verbatim "--"} \railqtok{text}. Any number of these may occur
+ within Isabelle/Isar commands.
\indexoutertoken{text}\indexouternonterm{comment}
\begin{rail}
@@ -218,10 +222,11 @@
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 @{text "\<forall>"} are available
- as well, provided these have not been superseded by commands or
- other keywords already (e.g.\ \texttt{=} or \texttt{+}).
+ may just write @{verbatim x} instead of quoted @{verbatim "\"x\""}.
+ Note that symbolic identifiers (e.g.\ @{verbatim "++"} or @{text
+ "\<forall>"} are available as well, provided these have not been superseded
+ by commands or other keywords already (such as @{verbatim "="} or
+ @{verbatim "+"}).
\indexoutertoken{type}\indexoutertoken{term}\indexoutertoken{prop}
\begin{rail}
@@ -234,8 +239,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 ``@{verbatim _}'' (underscore), which
+ means to skip a position.
\indexoutertoken{inst}\indexoutertoken{insts}
\begin{rail}
@@ -263,9 +268,9 @@
text {*
Mixfix annotations specify concrete \emph{inner} syntax of Isabelle
types and terms. Some commands such as @{command "types"} (see
- \S\ref{sec:types-pure}) admit infixes only, while @{command
- "consts"} (see \S\ref{sec:consts}) and @{command "syntax"} (see
- \S\ref{sec:syn-trans}) support the full range of general mixfixes
+ \secref{sec:types-pure}) admit infixes only, while @{command
+ "consts"} (see \secref{sec:consts}) and @{command "syntax"} (see
+ \secref{sec:syn-trans}) support the full range of general mixfixes
and binders.
\indexouternonterm{infix}\indexouternonterm{mixfix}\indexouternonterm{structmixfix}
@@ -283,14 +288,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 ``@{verbatim _}'');
+ the special symbol ``@{verbatim "\<index>"}'' (printed as ``@{text "\<index>"}'')
+ 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 ``@{verbatim "++"}'' for
+ an infix symbol, or ``@{verbatim "++"}@{text "\<index>"}'' for an infix of
+ an implicit structure.
*}
@@ -298,13 +304,15 @@
text {*
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 @{text 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 ``@{verbatim ","}'' (sequential composition),
+ ``@{verbatim "|"}'' (alternative choices), ``@{verbatim "?"}''
+ (try), ``@{verbatim "+"}'' (repeat at least once), ``@{verbatim
+ "["}@{text n}@{verbatim "]"}'' (restriction to first @{text n}
+ sub-goals, with default @{text "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).
\indexouternonterm{method}
\begin{rail}
@@ -316,17 +324,19 @@
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 ``@{text "[n]"}''
evaluates a method expression within a sandbox consisting of the
- first @{text n} sub-goals (which need to exist). For example,
- @{text "simp_all[3]"} simplifies the first three sub-goals, while
- @{text "(rule foo, simp_all)[]"} simplifies all new goals that
- emerge from applying rule @{text "foo"} to the originally first one.
+ first @{text n} sub-goals (which need to exist). For example, the
+ method ``@{text "simp_all[3]"}'' simplifies the first three
+ sub-goals, while ``@{text "(rule foo, simp_all)[]"}'' simplifies all
+ new goals that emerge from applying rule @{text "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 @{text "[!]"} refers to all
- goals, and @{text "[n-]"} to all goals starting from @{text "n"},
+ individual tactic being involved. Here ``@{text "[!]"}'' refers to
+ all goals, and ``@{text "[n-]"}'' to all goals starting from @{text
+ "n"}.
\indexouternonterm{goalspec}
\begin{rail}
@@ -339,7 +349,7 @@
subsection {* Attributes and theorems \label{sec:syn-att} *}
text {*
- 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
@@ -370,13 +380,13 @@
There are three forms of theorem references:
\begin{enumerate}
- \item named facts @{text "a"}
+ \item named facts @{text "a"},
- \item selections from named facts @{text "a(i, j - k)"}
+ \item selections from named facts @{text "a(i)"} or @{text "a(j - k)"},
\item literal fact propositions using @{syntax_ref altstring} syntax
- $\backquote\phi\backquote$, (see also method @{method_ref fact} in
- \S\ref{sec:pure-meth-att}).
+ @{verbatim "`"}@{text "\<phi>"}@{verbatim "`"} (see also method
+ @{method_ref fact} in \secref{sec:pure-meth-att}).
\end{enumerate}
@@ -386,12 +396,12 @@
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
- ``@{text "[[simproc a]]"}'' --- 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 @{command "declare"} and @{command "using"}.
+ An extra pair of brackets around attributes (like ``@{text
+ "[[simproc a]]"}'') 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 @{command "declare"} and @{command "using"}.
\indexouternonterm{axmdecl}\indexouternonterm{thmdecl}
\indexouternonterm{thmdef}\indexouternonterm{thmref}
@@ -457,7 +467,7 @@
\railnonterm{props} typically admit separate typings or namings via
another level of iteration, with explicit @{keyword_ref "and"}
separators; e.g.\ see @{command "fix"} and @{command "assume"} in
- \S\ref{sec:proof-context}.
+ \secref{sec:proof-context}.
*}
@@ -485,19 +495,19 @@
@{antiquotation_def ML_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 ``@{text "@{term [show_types] \"f x = a + x\"}"}''
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.
+ @{text "@{thm sym [no_vars]}"} 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
@@ -530,8 +540,9 @@
\end{rail}
Note that the syntax of antiquotations may \emph{not} include source
- comments \texttt{(*~\dots~*)} or verbatim text
- \verb|{*|~\dots~\verb|*|\verb|}|.
+ comments @{verbatim "(*"}~@{text "\<dots>"}~@{verbatim "*)"} or verbatim
+ text @{verbatim "{"}@{verbatim "*"}~@{text "\<dots>"}~@{verbatim
+ "*"}@{verbatim "}"}.
\begin{descr}
@@ -539,11 +550,11 @@
guaranteed to refer to a valid ancestor theory in the current
context.
- \item [@{text "@{thm a\<^sub>1 \<dots> a\<^sub>n}"}] prints theorems @{text
- "a\<^sub>1 \<dots> a\<^sub>n"}. Note that attribute specifications may be
- included as well (see also \S\ref{sec:syn-att}); the @{attribute_ref
- no_vars} rule (see \S\ref{sec:misc-meth-att}) would be particularly
- useful to suppress printing of schematic variables.
+ \item [@{text "@{thm a\<^sub>1 \<dots> a\<^sub>n}"}] prints theorems
+ @{text "a\<^sub>1 \<dots> a\<^sub>n"}. Note that attribute specifications
+ may be included as well (see also \secref{sec:syn-att}); the
+ @{attribute_ref no_vars} rule (see \secref{sec:misc-meth-att}) would
+ be particularly useful to suppress printing of schematic variables.
\item [@{text "@{prop \<phi>}"}] prints a well-typed proposition @{text
"\<phi>"}.
@@ -594,7 +605,7 @@
\item [@{text "@{full_prf a\<^sub>1 \<dots> a\<^sub>n}"}] is like @{text
"@{prf a\<^sub>1 \<dots> a\<^sub>n}"}, 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 ``@{verbatim _}'' placeholders there.
\item [@{text "@{ML s}"}, @{text "@{ML_type s}"}, and @{text
"@{ML_struct s}"}] check text @{text s} as ML value, type, and
@@ -618,7 +629,7 @@
in Horn-clause normal form @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C"}.
\item [@{text "prem1"}, \dots, @{text "prem9"}] extract premise
- number $1$, \dots, $9$, respectively, from from a rule in
+ number @{text "1, \<dots>, 9"}, respectively, from from a rule in
Horn-clause normal form @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C"}
\end{descr}