draft release;
authorwenzelm
Tue, 24 Aug 1999 15:38:18 +0200
changeset 7335 abba35b98892
parent 7334 a90fc1e5fb19
child 7336 ff05ab18ac5a
draft release;
doc-src/IsarRef/basics.tex
doc-src/IsarRef/generic.tex
doc-src/IsarRef/hol.tex
doc-src/IsarRef/intro.tex
doc-src/IsarRef/isar-ref.tex
doc-src/IsarRef/pure.tex
doc-src/IsarRef/syntax.tex
--- a/doc-src/IsarRef/basics.tex	Tue Aug 24 11:54:13 1999 +0200
+++ b/doc-src/IsarRef/basics.tex	Tue Aug 24 15:38:18 1999 +0200
@@ -13,16 +13,16 @@
 
 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 levels, and left with the final end of proof (e.g.\ 
+$\LEMMANAME$ at the theory level, and left with the final end of proof (e.g.\ 
 via $\QEDNAME$).  Some theory extension mechanisms require proof as well, such
-as the HOL $\isarkeyword{typedef}$ mechanism that only works for non-empty
-representing sets.
+as the HOL $\isarkeyword{typedef}$ which only works for non-empty representing
+sets.
 
 New-style theory files may still be associated with an ML file consisting of
 plain old tactic scripts.  There is no longer any ML binding generated for the
 theory and theorems, though.  Functions \texttt{theory}, \texttt{thm}, and
-\texttt{thms} may be used to retrieve this information from ML (see also
-\cite{isabelle-ref}).  Nevertheless, migration between classic Isabelle and
+\texttt{thms} may be used to retrieve this information from ML
+\cite{isabelle-ref}.  Nevertheless, migration between classic Isabelle and
 Isabelle/Isar is relatively easy.  Thus users may start to benefit from
 interactive theory development even before they have any idea of the Isar
 proof language.
@@ -30,15 +30,15 @@
 \begin{warn}
   Proof~General does \emph{not} support mixed interactive development of
   classic Isabelle theory files and tactic scripts together with Isar
-  documents at the same time.  The \texttt{isa} and \texttt{isar} versions of
-  Proof~General appear as two different theorem proving systems; only one
-  prover may be active at any time.
+  documents at the same time.  The ``\texttt{isa}'' and ``\texttt{isar}''
+  versions of Proof~General are handled as two different theorem proving
+  systems, only one may be active at the same time.
 \end{warn}
 
 
 \section{The Isar proof language}
 
-This rather important section has not been written yet!  Refer to
+Sorry, this rather important section has not been written yet!  Refer to
 \cite{Wenzel:1999:TPHOL} for the time being.
 
 \subsection{Commands}
--- a/doc-src/IsarRef/generic.tex	Tue Aug 24 11:54:13 1999 +0200
+++ b/doc-src/IsarRef/generic.tex	Tue Aug 24 15:38:18 1999 +0200
@@ -14,8 +14,8 @@
   erule^* & : & \isarmeth \\[0.5ex]
   fold & : & \isarmeth \\
   unfold & : & \isarmeth \\[0.5ex]
+  succeed & : & \isarmeth \\
   fail & : & \isarmeth \\
-  succeed & : & \isarmeth \\
 \end{matharray}
 
 \begin{rail}
@@ -25,12 +25,12 @@
 
 \begin{descr}
 \item [``$-$''] does nothing but insert the forward chaining facts as premises
-  into the goal.  Note that command $\PROOFNAME$ without any method given
-  actually performs a single reduction step using the $rule$ method (see
-  below); thus a plain \emph{do-nothing} proof step would be $\PROOF{-}$
-  rather than $\PROOFNAME$ alone.
-\item [$assumption$] solves some goal by assumption (after inserting the
-  goal's facts).
+  into the goal.  Note that command $\PROOFNAME$ without any method actually
+  performs a single reduction step using the $rule$ method (see below); thus a
+  plain \emph{do-nothing} proof step would be $\PROOF{-}$ rather than
+  $\PROOFNAME$ alone.
+\item [$assumption$] solves some goal by assumption, after inserting the
+  goal's facts.
 \item [$finish$] solves all remaining goals by assumption; this is the default
   terminal proof method for $\QEDNAME$, i.e.\ it usually does not have to be
   spelled out explicitly.
@@ -41,24 +41,20 @@
   
   Note that the classical reasoner introduces another version of $rule$ that
   is able to pick appropriate rules automatically, whenever explicit $thms$
-  are omitted (see \S\ref{sec:classical-basic}) .  That method is the default
-  one for proof steps such as $\PROOFNAME$ and ``$\DDOT$'' (two dots).
-  
+  are omitted (see \S\ref{sec:classical-basic}); that method is the default
+  one for initial proof steps, such as $\PROOFNAME$ and ``$\DDOT$'' (two
+  dots).
 \item [$erule~thms$] is similar to $rule$, but applies rules by
   elim-resolution.  This is an improper method, mainly for experimentation and
-  porting of old script.  Actual elimination proofs are usually done with
-  $rule$ (single step) or $elim$ (multiple steps, see
+  porting of old scripts.  Actual elimination proofs are usually done with
+  $rule$ (single step, involving facts) or $elim$ (multiple steps, see
   \S\ref{sec:classical-basic}).
-  
-\item [$unfold~thms$ and $fold~thms$] expand and fold back again meta-level
-  definitions $thms$ throughout all goals; facts may not be given.
-
+\item [$unfold~thms$ and $fold~thms$] expand and fold back again the given
+  meta-level definitions throughout all goals; facts may not be involved.
+\item [$succeed$] yields a single (unchanged) result; it is the identify of
+  the ``\texttt{,}'' method combinator.
 \item [$fail$] yields an empty result sequence; it is the identify of the
   ``\texttt{|}'' method combinator.
-  
-\item [$succeed$] yields a singleton result, which is unchanged except for the
-  change from $prove$ mode back to $state$; it is the identify of the
-  ``\texttt{,}'' method combinator.
 \end{descr}
 
 
@@ -73,11 +69,11 @@
   OF & : & \isaratt \\
   RS & : & \isaratt \\
   COMP & : & \isaratt \\[0.5ex]
-  where & : & \isaratt \\
-  of & : & \isaratt \\[0.5ex]
+  of & : & \isaratt \\
+  where & : & \isaratt \\[0.5ex]
   standard & : & \isaratt \\
   elimify & : & \isaratt \\
-  export & : & \isaratt \\
+  export^* & : & \isaratt \\
   transfer & : & \isaratt \\
 \end{matharray}
 
@@ -106,24 +102,23 @@
   $thms$ in parallel (cf.\ \texttt{MRS} in \cite[\S5]{isabelle-ref}, but note
   the reversed order).  $RS$ resolves with the $n$-th premise of $thm$; $COMP$
   is a version of $RS$ that does not include the automatic lifting process
-  that is normally desired (see \texttt{RS} and \texttt{COMP} in
+  that is normally intended (see also \texttt{RS} and \texttt{COMP} in
   \cite[\S5]{isabelle-ref}).
   
-\item [$of~ts$ and $where~insts$] perform positional and named instantiation,
-  respectively.  The terms given in $of$ are substituted for any variables
-  occurring in a theorem from left to right; ``\texttt{_}'' (underscore)
-  indicates to skip a position.
+\item [$of~ts$ and $where~\vec x = \vec t$] perform positional and named
+  instantiation, respectively.  The terms given in $of$ are substituted for
+  any schematic variables occurring in a theorem from left to right;
+  ``\texttt{_}'' (underscore) indicates to skip a position.
  
 \item [$standard$] puts a theorem into the standard form of object-rules, just
   as the ML function \texttt{standard} (see \cite[\S5]{isabelle-ref}).
   
-\item [$elimify$] turns an destruction rule (such as projection $conjunct@1$
-  into an elimination.
+\item [$elimify$] turns an destruction rule into an elimination.
   
 \item [$export$] lifts a local result out of the current proof context,
-  generalizing all fixed variables and discharging all assumptions.  Export is
-  usually done automatically behind the scenes.  This attribute is mainly for
-  experimentation.
+  generalizing all fixed variables and discharging all assumptions.  Note that
+  (partial) export is usually done automatically behind the scenes.  This
+  attribute is mainly for experimentation.
   
 \item [$transfer$] promotes a theorem to the current theory context, which has
   to enclose the former one.  Normally, this is done automatically when rules
@@ -175,24 +170,23 @@
 \begin{descr}
 \item [$\ALSO~(thms)$] maintains the auxiliary $calculation$ register as
   follows.  The first occurrence of $\ALSO$ in some calculational thread
-  initialises $calculation$ by $facts$. Any subsequent $\ALSO$ on the
-  \emph{same} level of block-structure updates $calculation$ by some
-  transitivity rule applied to $calculation$ and $facts$ (in that order).
-  Transitivity rules are picked from the current context plus those given as
-  $thms$ (the latter have precedence).
+  initialises $calculation$ by $facts$. Any subsequent $\ALSO$ on the same
+  level of block-structure updates $calculation$ by some transitivity rule
+  applied to $calculation$ and $facts$ (in that order).  Transitivity rules
+  are picked from the current context plus those given as $thms$ (the latter
+  have precedence).
   
 \item [$\FINALLY~(thms)$] maintaining $calculation$ in the same way as
   $\ALSO$, and concludes the current calculational thread.  The final result
   is exhibited as fact for forward chaining towards the next goal. Basically,
   $\FINALLY$ just abbreviates $\ALSO~\FROM{calculation}$.  A typical proof
-  idiom is $\FINALLY~\SHOW~\VVar{thesis}~\DOT$.
+  idiom is ``$\FINALLY~\SHOW{}{\VVar{thesis}}~\DOT$''.
   
-\item [Attribute $trans$] maintains the set of transitivity rules of the
-  theory or proof context, by adding or deleting the theorems provided as
-  arguments.  The default is adding of rules.
+\item [$trans$] maintains the set of transitivity rules of the theory or proof
+  context, by adding or deleting theorems (the default is to add).
 \end{descr}
 
-See theory \texttt{HOL/Isar_examples/Group} for a simple applications of
+See theory \texttt{HOL/Isar_examples/Group} for a simple application of
 calculations for basic equational reasoning.
 \texttt{HOL/Isar_examples/KnasterTarski} involves a few more advanced
 calculational steps in combination with natural deduction.
@@ -213,6 +207,7 @@
 theories.  See \cite{Wenzel:1997:TPHOL} for more information.  There is also a
 tutorial on \emph{Using Axiomatic Type Classes in Isabelle} that is part of
 the standard Isabelle documentation.
+%FIXME cite
 
 \begin{rail}
   'axclass' classdecl (axmdecl prop comment? +)
@@ -222,21 +217,23 @@
 \end{rail}
 
 \begin{descr}
-\item [$\isarkeyword{axclass}~$] defines an axiomatic type class as the
-  intersection of existing classes, with additional axioms holding.  Class
-  axioms may not contain more than one type variable.  The class axioms (with
-  implicit sort constraints added) are bound to the given names.  Furthermore
-  a class introduction rule is generated, which is employed by method
-  $expand_classes$ in support instantiation proofs of this class.
-
-\item [$\isarkeyword{instance}~c@1 < c@2$ and $\isarkeyword{instance}~c@1 <
-  c@2$] setup up a goal stating the class relation or type arity.  The proof
-  would usually proceed by the $expand_classes$ method, and then establish the
-  characteristic theorems of the type classes involved.  After finishing the
-  proof the theory will be augmented by a type signature declaration
-  corresponding to the resulting theorem.
+\item [$\isarkeyword{axclass}~c < \vec c~axms$] defines an axiomatic type
+  class as the intersection of existing classes, with additional axioms
+  holding.  Class axioms may not contain more than one type variable.  The
+  class axioms (with implicit sort constraints added) are bound to the given
+  names.  Furthermore a class introduction rule is generated, which is
+  employed by method $expand_classes$ in support instantiation proofs of this
+  class.
+  
+\item [$\isarkeyword{instance}~c@1 < c@2$ and $\isarkeyword{instance}~t ::
+  (\vec s)c$] setup up a goal stating the class relation or type arity.  The
+  proof would usually proceed by the $expand_classes$ method, and then
+  establish the characteristic theorems of the type classes involved.  After
+  finishing the proof the theory will be augmented by a type signature
+  declaration corresponding to the resulting theorem.
 \item [Method $expand_classes$] iteratively expands the class introduction
   rules
+%FIXME intro classIs etc;
 \end{descr}
 
 See theory \texttt{HOL/Isar_examples/Group} for a simple example of using
@@ -253,8 +250,11 @@
   asm_simp & : & \isarmeth \\
 \end{matharray}
 
+\railalias{asmsimp}{asm\_simp}
+\railterm{asmsimp}
+
 \begin{rail}
-  'simp' (simpmod * )
+  ('simp' | asmsimp) (simpmod * )
   ;
 
   simpmod: ('add' | 'del' | 'only' | 'other') ':' thmrefs
@@ -263,15 +263,16 @@
 
 \begin{descr}
 \item [Methods $simp$ and $asm_simp$] invoke Isabelle's simplifier, after
-  modifying the context as follows adding or deleting given rules.  The
+  modifying the context by adding or deleting given rules.  The
   \railtoken{only} modifier first removes all other rewrite rules and
   congruences, and then is like \railtoken{add}.  In contrast,
   \railtoken{other} ignores its arguments; nevertheless there may be
   side-effects on the context via attributes.  This provides a back door for
-  arbitrary manipulation of the context.
+  arbitrary context manipulation.
   
   Both of these methods are based on \texttt{asm_full_simp_tac}, see
-  \cite[\S10]{isabelle-ref}.
+  \cite[\S10]{isabelle-ref}; $simp$ removes any exisiting premises of the
+  goal, before inserting the goal facts; $asm_simp$ leaves the premises.
 \end{descr}
 
 \subsection{Modifying the context}
@@ -288,7 +289,7 @@
 
 \begin{descr}
 \item [Attribute $simp$] adds or deletes rules from the theory or proof
-  context.  The default is to add rules.
+  context (the default is to add).
 \end{descr}
 
 
@@ -303,13 +304,13 @@
 \end{matharray}
 
 These attributes provide forward rules for simplification, which should be
-used very rarely.  See the ML function of the same name in
+used only very rarely.  See the ML functions of the same name in
 \cite[\S10]{isabelle-ref} for more information.
 
 
 \section{The Classical Reasoner}
 
-\subsection{Basic step methods}\label{sec:classical-basic}
+\subsection{Basic methods}\label{sec:classical-basic}
 
 \indexisarmeth{rule}\indexisarmeth{default}\indexisarmeth{contradiction}
 \begin{matharray}{rcl}
@@ -329,14 +330,16 @@
   over the primitive one (see \S\ref{sec:pure-meth}).  In the case that no
   rules are provided as arguments, it automatically determines elimination and
   introduction rules from the context (see also \S\ref{sec:classical-mod}).
-  In that form it is the default method for basic proof steps.
+  In that form it is the default method for basic proof steps, such as
+  $\PROOFNAME$ and ``$\DDOT$'' (two dots).
   
 \item [Methods $intro$ and $elim$] repeatedly refine some goal by intro- or
   elim-resolution, after having inserted the facts.  Omitting the arguments
   refers to any suitable rules from the context, otherwise only the explicitly
-  given ones may be applied.  The latter form admits better control of what is
-  actually happening, thus it is appropriate as an initial proof method that
-  splits up certain connectives of the goal, before entering the sub-proof.
+  given ones may be applied.  The latter form admits better control of what
+  actually happens, thus it is very appropriate as an initial method for
+  $\PROOFNAME$ that splits up certain connectives of the goal, before entering
+  the sub-proof.
 
 \item [Method $contradiction$] solves some goal by contradiction: both $A$ and
   $\neg A$ have to be present in the assumptions.
@@ -370,11 +373,10 @@
 
 \begin{descr}
 \item [$blast$] refers to the classical tableau prover (see \texttt{blast_tac}
-  in \cite[\S11]{isabelle-ref}).  The optional argument specifies a applies a
+  in \cite[\S11]{isabelle-ref}).  The optional argument specifies a
   user-supplied search bound (default 20).
 \item [$fast$, $best$, $slow$, $slow_best$] refer to the generic classical
-  reasoner (see the corresponding tactics \texttt{fast_tac} etc.\ in
-  \cite[\S11]{isabelle-ref}).
+  reasoner (see \cite[\S11]{isabelle-ref}, tactic \texttt{fast_tac} etc).
 \end{descr}
 
 Any of above methods support additional modifiers of the context of classical
@@ -403,7 +405,8 @@
   simplification and classical reasoning tactics.  See \texttt{force_tac} and
   \texttt{auto_tac} in \cite[\S11]{isabelle-ref} for more information.  The
   modifier arguments correspond to those given in \S\ref{sec:simp} and
-  \S\ref{sec:classical-auto}.
+  \S\ref{sec:classical-auto}.  Note that the ones related to the Simplifier
+  are prefixed by \railtoken{simp} here.
 \end{descr}
 
 \subsection{Modifying the context}\label{sec:classical-mod}
@@ -422,13 +425,13 @@
 \end{rail}
 
 \begin{descr}
-\item [Attributes $intro$, $elim$, and $dest$] add introduction, elimination,
-  and destruct rules, respectively.  By default, rules are considered as
-  \emph{safe}, while a single ``!'' classifies as \emph{unsafe}, and ``!!'' as
-  \emph{extra} (i.e.\ not applied in the search-oriented automatic methods).
-
-\item [Attribute $delrule$] deletes introduction or elimination rules from the
-  context.  Destruction rules would have to be turned into elimination rules
+\item [$intro$, $elim$, $dest$] add introduction, elimination, destruct rules,
+  respectively.  By default, rules are considered as \emph{safe}, while a
+  single ``!'' classifies as \emph{unsafe}, and ``!!'' as \emph{extra} (i.e.\ 
+  not applied in the search-oriented automatic methods).
+  
+\item [$delrule$] deletes introduction or elimination rules from the context.
+  Note that destruction rules would have to be turned into elimination rules
   first, e.g.\ by using the $elimify$ attribute.
 \end{descr}
 
--- a/doc-src/IsarRef/hol.tex	Tue Aug 24 11:54:13 1999 +0200
+++ b/doc-src/IsarRef/hol.tex	Tue Aug 24 15:38:18 1999 +0200
@@ -24,9 +24,9 @@
 \item [$\isarkeyword{typedef}~(\vec\alpha)t = A$] sets up a goal stating
   non-emptiness of the set $A$.  After finishing the proof, the theory will be
   augmented by a Gordon/HOL-style type definition.  See \cite{isabelle-HOL}
-  for more information.  Note that user-level work usually does not directly
-  refer to the HOL $\isarkeyword{typedef}$ primitive, but uses more advanced
-  packages such as $\isarkeyword{record}$ (see \S\ref{sec:record}) or
+  for more information.  Note that user-level theories usually do not directly
+  refer to the HOL $\isarkeyword{typedef}$ primitive, but use more advanced
+  packages such as $\isarkeyword{record}$ (see \S\ref{sec:record}) and
   $\isarkeyword{datatype}$ (see \S\ref{sec:datatype}).
 \end{descr}
 
@@ -51,11 +51,10 @@
 \item [$\isarkeyword{record}~(\vec\alpha)t = \tau + \vec c :: \vec\sigma$]
   defines extensible record type $(\vec\alpha)t$, derived from the optional
   parent record $\tau$ by adding new field components $\vec c :: \vec\sigma$.
+  See \cite{isabelle-HOL,NaraschewskiW-TPHOLs98} for more information only
+  simply-typed extensible records.
 \end{descr}
 
-See \cite{isabelle-HOL,NaraschewskiW-TPHOLs98} for more information only
-simply-typed extensible records.
-
 
 \section{Datatypes}\label{sec:datatype}
 
@@ -86,7 +85,8 @@
 \end{descr}
 
 See \cite{isabelle-HOL} for more details on datatypes.  Note that the theory
-syntax above has been slightly simplified over the old version.
+syntax above has been slightly simplified over the old version, usually
+requiring more quotes and less parentheses.
 
 
 \section{Recursive functions}
@@ -113,7 +113,7 @@
   functions (using the TFL package).
 \end{descr}
 
-See \cite{isabelle-HOL} for more information.
+See \cite{isabelle-HOL} for more information on both mechanisms.
 
 
 \section{(Co)Inductive sets}
@@ -178,7 +178,7 @@
   by $insts$.  The latter either consists of pairs $P$ $x$ (induction
   predicate and variable), where $P$ is optional.  If $kind$ is omitted, the
   default is to pick a datatype induction rule according to the type of some
-  induction variable (at least one has to be given in that case).
+  induction variable, which may not be omitted that case.
 \end{descr}
 
 
--- a/doc-src/IsarRef/intro.tex	Tue Aug 24 11:54:13 1999 +0200
+++ b/doc-src/IsarRef/intro.tex	Tue Aug 24 15:38:18 1999 +0200
@@ -38,9 +38,9 @@
 PROOFGENERAL_OPTIONS=""
 \end{ttbox}
 You may have to change \texttt{\$ISABELLE_HOME/contrib/ProofGeneral} to the
-actual installation directory of Proof~General.  Now the capital
+actual installation directory of Proof~General.  From now on, the capital
 \texttt{Isabelle} executable refers to the \texttt{ProofGeneral/isar}
-interface.  It's usage is as follows:
+interface.  Its usage is as follows:
 \begin{ttbox}
 Usage: interface [OPTIONS] [FILES ...]
 
@@ -53,12 +53,11 @@
 Starts Proof General for Isabelle/Isar with proof documents FILES
 (default Scratch.thy).
 \end{ttbox}
-Note that the defaults for these options may be overridden via the
-\texttt{PROOFGENERAL_OPTIONS} setting.  For example, GNU
-Emacs\footnote{Version 20.x required.} with loading of the startup file
-enabled may be configured as follows:\footnote{The interface disables
+Apart from the command line, the defaults for these options may be overridden
+via the \texttt{PROOFGENERAL_OPTIONS} setting as well.  For example, plain GNU
+Emacs with loading of the startup file enabled\footnote{The interface disables
   \texttt{.emacs} by default to ensure successful startup despite any strange
-  commands that tend to occur there.}
+  commands that tend to occur there.} may be configured as follows:
 \begin{ttbox}
 PROOFGENERAL_OPTIONS="-p emacs -u true"
 \end{ttbox}
--- a/doc-src/IsarRef/isar-ref.tex	Tue Aug 24 11:54:13 1999 +0200
+++ b/doc-src/IsarRef/isar-ref.tex	Tue Aug 24 15:38:18 1999 +0200
@@ -51,7 +51,7 @@
   Isabelle/Isar input may consist either of \emph{proper document
     constructors}, or \emph{improper auxiliary commands} (for diagnostics,
   exploration etc.).  Proof texts consisting of proper document constructors
-  only admit a purely static reading, thus being intelligible later without
+  only, admit a purely static reading, thus being intelligible later without
   requiring dynamic replay that is so typical for traditional proof scripts.
   Any of the Isabelle/Isar commands may be executed in single-steps, so
   basically the interpreter has a proof text debugger already built-in.
@@ -63,12 +63,11 @@
   including forward and backward tracing of partial documents; intermediate
   states may be inspected by diagnostic commands.
   
-  The Isar subsystem of Isabelle is tightly integrated into the Isabelle/Pure
-  meta-logic implementation.  Theories, theorems, proof procedures etc.\ may
-  be used interchangeably between Isabelle-classic proof scripts and
-  Isabelle/Isar documents.  Isar is as generic as Isabelle, able to support a
-  wide range of object-logics.  The current end-user setup is mainly for
-  Isabelle/HOL.
+  The Isar subsystem is tightly integrated into the Isabelle/Pure meta-logic
+  implementation.  Theories, theorems, proof procedures etc.\ may be used
+  interchangeably between Isabelle-classic proof scripts and Isabelle/Isar
+  documents.  Isar is as generic as Isabelle, able to support a wide range of
+  object-logics.  The current end-user setup is mainly for Isabelle/HOL.
 \end{abstract}
 
 \pagenumbering{roman} \tableofcontents \clearfirst
--- a/doc-src/IsarRef/pure.tex	Tue Aug 24 11:54:13 1999 +0200
+++ b/doc-src/IsarRef/pure.tex	Tue Aug 24 15:38:18 1999 +0200
@@ -1,19 +1,19 @@
 
-\chapter{Basic Isar elements}\label{ch:pure-syntax}
+\chapter{Basic Isar Elements}\label{ch:pure-syntax}
 
 Subsequently, we introduce the main part of the basic Isar theory and proof
 commands as provided by Isabelle/Pure.  Chapter~\ref{ch:gen-tools} describes
-further Isar elements as provided by generic tools and packages that are
-either part of Pure Isabelle, or pre-loaded by most object logics (such as the
-Simplifier).  See chapter~\ref{ch:hol-tools} for actual object-logic specific
-elements (for Isabelle/HOL).
+further Isar elements as provided by generic tools and packages (such as the
+Simplifier) that are either part of Pure Isabelle, or pre-loaded by most
+object logics.  See chapter~\ref{ch:hol-tools} for actual object-logic
+specific elements (for Isabelle/HOL).
 
 \medskip
 
 Isar commands may be either \emph{proper} document constructors, or
 \emph{improper commands} (indicated by $^*$).  Some proof methods and
-attributes introduced later may be classified as improper as well.  Improper
-Isar language elements might be helpful when developing proof documents, while
+attributes introduced later are classified as improper as well.  Improper Isar
+language elements might be helpful when developing proof documents, while
 their use is strongly discouraged for the final version.  Typical examples are
 diagnostic commands that print terms or theorems according to the current
 context; other commands even emulate old-style tactical theorem proving, which
@@ -33,9 +33,9 @@
 
 Isabelle/Isar ``new-style'' theories are either defined via theory files or
 interactively.  Both actual theory specifications and proofs are handled
-uniformly --- occasionally definitional mechanisms even require some manual
-proof.  In contrast, ``old-style'' Isabelle theories support batch processing
-only, with the proof scripts collected in separate ML files.
+uniformly --- occasionally definitional mechanisms even require some explicit
+proof as well.  In contrast, ``old-style'' Isabelle theories support batch
+processing only, with the proof scripts collected in separate ML files.
 
 The first command of any theory has to be $\THEORY$, starting a new theory
 based on the merge of existing ones.  The theory context may be also changed
@@ -61,7 +61,9 @@
   $\THEORY$ is executed interactively).  The optional $\isarkeyword{files}$
   specification declares additional dependencies on ML files.  Unless put in
   parentheses, any file will be loaded immediately via $\isarcmd{use}$ (see
-  also \S\ref{sec:ML}).
+  also \S\ref{sec:ML}).  The optional ML file \texttt{$A$.ML} that may be
+  associated with any theory should \emph{not} be included in
+  $\isarkeyword{files}$.
   
 \item [$\CONTEXT~B$] enters an existing theory context $B$, basically in
   read-only mode, so only a limited set of commands may be performed.  Just as
@@ -112,14 +114,13 @@
 \begin{descr}
 \item [$\isarkeyword{title}~title~author~date$] specifies the document title
   just as in typical {\LaTeX} documents.
-\item [$\isarkeyword{chapter}~text$, $\isarkeyword{section}~text$,
-  $\isarkeyword{subsection}~text$, and $\isarkeyword{subsubsection}~text$]
-  mark chapter and section headings.
-\item [$\TEXT~text$] specifies an actual body of prose text, including
-  references to formal entities.\footnote{The latter feature is not yet
-    exploited.  Nevertheless, any text of the form
-    \texttt{\at\ttlbrace\dots\ttrbrace} should be considered as reserved for
-    future use.}
+\item [$\isarkeyword{chapter}$, $\isarkeyword{section}$,
+  $\isarkeyword{subsection}$, and $\isarkeyword{subsubsection}$] mark chapter
+  and section headings.
+\item [$\TEXT$] specifies an actual body of prose text, including references
+  to formal entities.\footnote{The latter feature is not yet exploited.
+    Nevertheless, any text of the form \texttt{\at\ttlbrace\dots\ttrbrace}
+    should be considered as reserved for future use.}
 \end{descr}
 
 
@@ -142,9 +143,8 @@
 \end{rail}
 
 \begin{descr}
-\item [$\isarkeyword{classes}~c<\vec c ~\dots$] declares class $c$ to be a
-  subclass of existing classes $\vec c$.  Cyclic class structures are ruled
-  out.
+\item [$\isarkeyword{classes}~c<\vec c$] declares class $c$ to be a subclass
+  of existing classes $\vec c$.  Cyclic class structures are ruled out.
 \item [$\isarkeyword{classrel}~c@1<c@2$] states a subclass relation between
   existing classes $c@1$ and $c@2$.  This is done axiomatically!  The
   $\isarkeyword{instance}$ command (see \S\ref{sec:axclass}) provides a way
@@ -177,7 +177,7 @@
 \end{rail}
 
 \begin{descr}
-\item [$\TYPES~(\vec\alpha)t = \tau~\dots$] introduces \emph{type synonym}
+\item [$\TYPES~(\vec\alpha)t = \tau$] introduces \emph{type synonym}
   $(\vec\alpha)t$ for existing type $\tau$.  Unlike actual type definitions,
   as are available in Isabelle/HOL for example, type synonyms are just purely
   syntactic abbreviations, without any logical significance.  Internally, type
@@ -189,9 +189,9 @@
 \item [$\isarkeyword{nonterminals}~\vec c$] declares $0$-ary type constructors
   $\vec c$ to act as purely syntactic types, i.e.\ nonterminal symbols of
   Isabelle's inner syntax of terms or types.
-\item [$\isarkeyword{arities}~t::(\vec s)s~\dots$] augments Isabelle's
-  order-sorted signature of types by new type constructor arities.  This is
-  done axiomatically!  The $\isarkeyword{instance}$ command (see
+\item [$\isarkeyword{arities}~t::(\vec s)s$] augments Isabelle's order-sorted
+  signature of types by new type constructor arities.  This is done
+  axiomatically!  The $\isarkeyword{instance}$ command (see
   \S\ref{sec:axclass}) provides a way introduce proven type arities.
 \end{descr}
 
@@ -218,15 +218,15 @@
 \end{rail}
 
 \begin{descr}
-\item [$\CONSTS~c::\sigma~\dots$] declares constant $c$ to have any instance
-  of type scheme $\sigma$.  The optional mixfix annotations may attach
-  concrete syntax constants.
-\item [$\DEFS~name: eqn~\dots$] introduces $eqn$ as a definitional axiom for
-  some existing constant.  See \cite[\S6]{isabelle-ref} for more details on
-  the form of equations admitted as constant definitions.
-\item [$\isarkeyword{constdefs}~c::\sigma~eqn~\dots$] combines constant
-  declarations and definitions, using canonical name $c_def$ for the
-  definitional axiom.
+\item [$\CONSTS~c::\sigma$] declares constant $c$ to have any instance of type
+  scheme $\sigma$.  The optional mixfix annotations may attach concrete syntax
+  constants.
+\item [$\DEFS~name: eqn$] introduces $eqn$ as a definitional axiom for some
+  existing constant.  See \cite[\S6]{isabelle-ref} for more details on the
+  form of equations admitted as constant definitions.
+\item [$\isarkeyword{constdefs}~c::\sigma~eqn$] combines declarations and
+  definitions of constants, using canonical name $c_def$ for the definitional
+  axiom.
 \end{descr}
 
 
@@ -251,13 +251,14 @@
 \item [$\isarkeyword{syntax}~(mode)~decls$] is similar to $\CONSTS~decls$,
   except that the actual logical signature extension is omitted.  Thus the
   context free grammar of Isabelle's inner syntax may be augmented in
-  arbitrary ways.  The $mode$ argument refers to the print mode that the
-  grammar rules belong; unless there is the \texttt{output} flag given, all
-  productions are added both to the input and output grammar.
+  arbitrary ways, independently of the logic.  The $mode$ argument refers to
+  the print mode that the grammar rules belong; unless there is the
+  \texttt{output} flag given, all productions are added both to the input and
+  output grammar.
 \item [$\isarkeyword{translations}~rules$] specifies syntactic translation
   rules (also known as \emph{macros}): parse/print rules (\texttt{==}), parse
-  rules (\texttt{=>}), print rules (\texttt{<=}).  Translation patterns may be
-  prefixed by the syntactic category to be used for parsing; the default is
+  rules (\texttt{=>}), or print rules (\texttt{<=}).  Translation patterns may
+  be prefixed by the syntactic category to be used for parsing; the default is
   \texttt{logic}.
 \end{descr}
 
@@ -279,16 +280,16 @@
 \end{rail}
 
 \begin{descr}
-\item [$\isarkeyword{axioms}~name: \phi~\dots$] introduces arbitrary
-  statements as logical axioms.  In fact, axioms are ``axiomatic theorems'',
-  and may be referred just as any other theorem later.
+\item [$\isarkeyword{axioms}~a: \phi$] introduces arbitrary statements as
+  logical axioms.  In fact, axioms are ``axiomatic theorems'', and may be
+  referred later just as any other theorem.
   
   Axioms are usually only introduced when declaring new logical systems.
   Everyday work is typically done the hard way, with proper definitions and
   actual theorems.
-\item [$\isarkeyword{theorems}~name = thms$] stores lists of existing theorems
-  as $name$.  Typical applications would also involve attributes (to augment
-  the default simpset, for example).
+\item [$\isarkeyword{theorems}~a = \vec b$] stores lists of existing theorems.
+  Typical applications would also involve attributes, to augment the default
+  Simplifier context, for example.
 \item [$\isarkeyword{lemmas}$] is similar to $\isarkeyword{theorems}$, but
   tags the results as ``lemma''.
 \end{descr}
@@ -302,9 +303,9 @@
   \isarcmd{local} & : & \isartrans{theory}{theory} \\
 \end{matharray}
 
-Isabelle organises any kind of names (of types, constants, theorems etc.)  by
+Isabelle organizes any kind of names (of types, constants, theorems etc.)  by
 hierarchically structured name spaces.  Normally the user never has to control
-the behaviour of name space entry by hand, yet the following commands provide
+the behavior of name space entry by hand, yet the following commands provide
 some way to do so.
 
 \begin{descr}
@@ -342,10 +343,11 @@
   \S\ref{sec:begin-thy}).
 \item [$\isarkeyword{ML}~text$] reads and executes ML commands from $text$.
   The theory context is passed just as for $\isarkeyword{use}$.
+%FIXME picked up again!?
 \item [$\isarkeyword{setup}~text$] changes the current theory context by
   applying setup functions $text$, which has to be an ML expression of type
   $(theory \to theory)~list$.  The $\isarkeyword{setup}$ command is the usual
-  way to initialise object-logic specific tools and packages written in ML.
+  way to initialize object-logic specific tools and packages written in ML.
 \end{descr}
 
 
@@ -389,7 +391,7 @@
 \begin{descr}
 \item [$\isarkeyword{oracle}~name=text$] declares oracle $name$ to be ML
   function $text$, which has to be of type $Sign\mathord.sg \times
-  Object\mathord.T \to term)$.
+  Object\mathord.T \to term$.
 \end{descr}
 
 
@@ -397,9 +399,9 @@
 
 Proof commands provide transitions of Isar/VM machine configurations, which
 are block-structured, consisting of a stack of nodes with three main
-components: logical \emph{proof context}, local \emph{facts}, and open
-\emph{goals}.  Isar/VM transitions are \emph{typed} according to the following
-three three different modes of operation:
+components: logical proof context, current facts, and open goals.  Isar/VM
+transitions are \emph{typed} according to the following three three different
+modes of operation:
 \begin{descr}
 \item [$proof(prove)$] means that a new goal has just been stated that is now
   to be \emph{proven}; the next command may refine it by some proof method
@@ -408,7 +410,7 @@
   augmented by \emph{stating} additional assumptions, intermediate result etc.
 \item [$proof(chain)$] is an intermediate mode between $proof(state)$ and
   $proof(prove)$: existing facts have been just picked up in order to use them
-  when refining the goal to be claimed next.
+  when refining the goal claimed next.
 \end{descr}
 
 
@@ -454,7 +456,7 @@
 hand, a local theorem is created that may be used as a fact in subsequent
 proof steps.  On the other hand, any result $\phi$ exported from the context
 becomes conditional wrt.\ the assumption.  Thus, solving an enclosing goal
-using this result would basically introduce a new subgoal stemming from the
+using such a result would basically introduce a new subgoal stemming from the
 assumption.  How this situation is handled depends on the actual version of
 assumption command used: while $\ASSUMENAME$ solves the subgoal by unifying
 with some premise of the goal, $\PRESUMENAME$ leaves the subgoal unchanged to
@@ -482,20 +484,22 @@
 \begin{descr}
 \item [$\FIX{x}$] introduces a local \emph{arbitrary, but fixed} variable $x$.
 \item [$\ASSUME{a}{\Phi}$ and $\PRESUME{a}{\Phi}$] introduce local theorems
-  $\Phi$.  Subsequent results applied to some enclosing goal (e.g.\ via
-  $\SHOWNAME$) are handled as follows: $\ASSUMENAME$ expects to be able to
-  unify with existing premises in the goal, while $\PRESUMENAME$ leaves $\Phi$
-  as new subgoals.  Note that several lists of assumptions may be given
-  (separated by $\isarkeyword{and}$); the resulting list of facts consists of
-  all of these concatenated.
+  $\Phi$ by assumption.  Subsequent results applied to an enclosing goal
+  (e.g.\ via $\SHOWNAME$) are handled as follows: $\ASSUMENAME$ expects to be
+  able to unify with existing premises in the goal, while $\PRESUMENAME$
+  leaves $\Phi$ as new subgoals.
+  
+  Several lists of assumptions may be given (separated by
+  $\isarkeyword{and}$); the resulting list of facts consists of all of these
+  concatenated.
 \item [$\DEF{a}{x \equiv t}$] introduces a local (non-polymorphic) definition.
   In results exported from the context, $x$ is replaced by $t$.  Basically,
-  $\DEF{}{x \equiv t}$ abbreviates $\FIX{x}~\PRESUME{}{x \equiv t}$ (the
-  resulting hypothetical equation is solved by reflexivity, though).
+  $\DEF{}{x \equiv t}$ abbreviates $\FIX{x}~\PRESUME{}{x \equiv t}$, with the
+  resulting hypothetical equation solved by reflexivity.
 \end{descr}
 
-The internal register $prems$\indexisarreg{prems} refers to the current list
-of assumptions.
+The special theorem name $prems$\indexisarreg{prems} refers to all current
+assumptions.
 
 
 \subsection{Facts and forward chaining}
@@ -509,12 +513,10 @@
 \end{matharray}
 
 New facts are established either by assumption or proof of local statements.
-Any facts will usually be involved in proofs of further results: either as
-explicit arguments of proof methods or when forward chaining towards the next
-goal via $\THEN$ (and variants).  Note that the internal register of
-\emph{current facts} may be referred as theorem list
-$facts$.\indexisarreg{facts}
-
+Any fact will usually be involved in further proofs, either as explicit
+arguments of proof methods or when forward chaining towards the next goal via
+$\THEN$ (and variants).  Note that the special theorem name
+$facts$.\indexisarreg{facts} refers to the most recently established facts.
 \begin{rail}
   'note' thmdef? thmrefs comment?
   ;
@@ -529,12 +531,12 @@
   as $a$.  Note that attributes may be involved as well, both on the left and
   right hand sides.
 \item [$\THEN$] indicates forward chaining by the current facts in order to
-  establish the goal to be claimed next.  The initial proof method invoked to
-  solve that will be offered these facts to do ``anything appropriate'' (see
-  also \S\ref{sec:proof-steps}).  For example, method $rule$ (see
+  establish the goal claimed next.  The initial proof method invoked to refine
+  that will be offered these facts to do ``anything appropriate'' (see also
+  \S\ref{sec:proof-steps}).  For example, method $rule$ (see
   \S\ref{sec:pure-meth}) would do an elimination rather than an introduction.
-\item [$\FROM{\vec b}$] abbreviates $\NOTE{}{\vec b}~\THEN$; also note that
-  $\THEN$ is equivalent to $\FROM{facts}$.
+\item [$\FROM{\vec b}$] abbreviates $\NOTE{}{\vec b}~\THEN$; thus $\THEN$ is
+  equivalent to $\FROM{facts}$.
 \item [$\WITH{\vec b}$] abbreviates $\FROM{\vec b~facts}$; thus the forward
   chaining is from earlier facts together with the current ones.
 \end{descr}
@@ -569,18 +571,18 @@
 \end{rail}
 
 \begin{descr}
-\item [$\THEOREM{name}{\phi}$] enters proof mode with $\phi$ as main goal,
+\item [$\THEOREM{a}{\phi}$] enters proof mode with $\phi$ as main goal,
   eventually resulting in some theorem $\turn \phi$, which will be stored in
   the theory.
 \item [$\LEMMANAME$] is similar to $\THEOREMNAME$, but tags the result as
   ``lemma''.
-\item [$\HAVE{name}{\phi}$] claims a local goal, eventually resulting in a
+\item [$\HAVE{a}{\phi}$] claims a local goal, eventually resulting in a
   theorem with the current assumption context as hypotheses.
-\item [$\SHOW{name}{\phi}$] is similar to $\HAVE{name}{\phi}$, but solves some
+\item [$\SHOW{a}{\phi}$] is similar to $\HAVE{a}{\phi}$, but solves some
   pending goal with the result \emph{exported} into the corresponding context.
-\item [$\HENCE{name}{\phi}$] abbreviates $\THEN~\HAVE{name}{\phi}$, i.e.\ 
-  claims a local goal to be proven by forward chaining the current facts.
-\item [$\THUS{name}{\phi}$] abbreviates $\THEN~\SHOW{name}{\phi}$.
+\item [$\HENCE{a}{\phi}$] abbreviates $\THEN~\HAVE{a}{\phi}$, i.e.\ claims a
+  local goal to be proven by forward chaining the current facts.
+\item [$\THUS{a}{\phi}$] abbreviates $\THEN~\SHOW{a}{\phi}$.
 \end{descr}
 
 
@@ -597,20 +599,19 @@
   \isarcmd{sorry} & : & \isartrans{proof(prove)}{proof(state) ~|~ theory} \\
 \end{matharray}
 
-Arbitrary goal refinements via tactics is considered harmful.  Consequently
-the Isar framework admits proof methods to be invoked in two places only.
+Arbitrary goal refinement via tactics is considered harmful.  Consequently the
+Isar framework admits proof methods to be invoked in two places only.
 \begin{enumerate}
 \item An \emph{initial} refinement step $\PROOF{m@1}$ reduces a newly stated
-  intermediate goal to a number of sub-goals that are to be solved later.
-  Facts are passed to $m@1$ for forward chaining if so indicated by
-  $proof(chain)$ mode.
+  goal to a number of sub-goals that are to be solved later.  Facts are passed
+  to $m@1$ for forward chaining if so indicated by $proof(chain)$ mode.
   
-\item A \emph{terminal} conclusion step $\QED{m@2}$ solves any remaining
-  pending goals completely.  No facts are passed to $m@2$.
+\item A \emph{terminal} conclusion step $\QED{m@2}$ solves any remaining goals
+  completely.  No facts are passed to $m@2$.
 \end{enumerate}
 
-The only other proper way to affect pending goals is by $\SHOWNAME$, which
-involves an explicit statement of what is solved.
+The only other proper way to affect pending goals is by $\SHOWNAME$ (or
+$\THUSNAME$), which involves an explicit statement of what is to be solved.
 
 \medskip
 
@@ -644,27 +645,26 @@
 \end{rail}
 
 \begin{descr}
-\item [$\PROOF{m@1}$] refines the pending goal by proof method $m@1$; facts
-  for forward chaining are passed if so indicated by $proof(chain)$.
-\item [$\QED{m@2}$] refines any remaining goals by proof method $m@1$ and
-  concludes the sub-proof.  If the goal had been $\SHOWNAME$, some pending
-  sub-goal is solved as well by the rule resulting from the result exported to
-  the enclosing goal context.  Thus $\QEDNAME$ may fail for two reasons:
-  either $m@2$ fails to solve all remaining goals completely, or the resulting
-  rule does not resolve with any enclosing goal.  Debugging such a situation
-  might involve temporarily changing $\SHOWNAME$ into $\HAVENAME$, or
-  weakening the local context by replacing $\ASSUMENAME$ by $\PRESUMENAME$.
+\item [$\PROOF{m@1}$] refines the goal by proof method $m@1$; facts for
+  forward chaining are passed if so indicated by $proof(chain)$ mode.
+\item [$\QED{m@2}$] refines any remaining goals by proof method $m@2$ and
+  concludes the sub-proof.  If the goal had been $\SHOWNAME$ (or $\THUSNAME$),
+  some pending sub-goal is solved as well by the rule resulting from the
+  result exported to the enclosing goal context.  Thus $\QEDNAME$ may fail for
+  two reasons: either $m@2$ fails to solve all remaining goals completely, or
+  the resulting rule does not resolve with any enclosing goal.  Debugging such
+  a situation might involve temporarily changing $\SHOWNAME$ into $\HAVENAME$,
+  or softening the local context by replacing $\ASSUMENAME$ by $\PRESUMENAME$.
 \item [$\BYY{m@1}{m@2}$] is a \emph{terminal proof}; it abbreviates
   $\PROOF{m@1}~\QED{m@2}$, with automatic backtracking across both methods.
   Debugging an unsuccessful $\BYY{m@1}{m@2}$ commands might be done by simply
   expanding the abbreviation by hand; note that $\PROOF{m@1}$ is usually
   sufficient to see what is going wrong.
 \item [``$\DDOT$''] is a \emph{default proof}; it abbreviates $\BY{default}$.
-\item [``$\DOT$''] is a \emph{trivial proof}, it abbreviates $\BY{-}$, where
-  method ``$-$'' does nothing except inserting any facts into the proof state.
+\item [``$\DOT$''] is a \emph{trivial proof}; it abbreviates $\BY{-}$.
 \item [$\isarkeyword{sorry}$] is a \emph{fake proof}; provided that
   \texttt{quick_and_dirty} is enabled, $\isarkeyword{sorry}$ pretends to solve
-  the goal without much ado.  Of course, the result is a fake theorem only,
+  the goal without further ado.  Of course, the result is a fake theorem only,
   involving some oracle in its internal derivation object (this is indicated
   as $[!]$ in the printed result).  The main application of
   $\isarkeyword{sorry}$ is to support top-down proof development.
@@ -697,10 +697,10 @@
 \end{rail}
 
 \begin{descr}
-\item [$\isarkeyword{apply}~m$] applies proof method $m$ in the
+\item [$\isarkeyword{apply}~(m)$] applies proof method $m$ in the
   plain-old-tactic sense.  Facts for forward chaining are ignored.
-\item [$\isarkeyword{then_apply}~m$] is similar to $\isarkeyword{apply}$, but
-  observes the goal's facts.
+\item [$\isarkeyword{then_apply}~(m)$] is similar to $\isarkeyword{apply}$,
+  but observes the goal's facts.
 \item [$\isarkeyword{back}$] does back-tracking over the result sequence of
   the last proof command.  Basically, any proof command may return multiple
   results.
@@ -747,15 +747,16 @@
 \end{descr}
 
 A few \emph{automatic} term abbreviations\index{automatic abbreviation} for
-goals and facts are available as well.  For any open goal, $\VVar{thesis}$
-refers to its object-logical statement, $\VVar{thesis_prop}$ to the full
-proposition (which may be a rule), and $\VVar{thesis_concl}$ to its (atomic)
-conclusion.
+goals and facts are available as well.  For any open goal,
+$\VVar{thesis_prop}$ refers to the full proposition (which may be a rule),
+$\VVar{thesis_concl}$ to its (atomic) conclusion, and $\VVar{thesis}$ to its
+object-logical statement.  The latter two abstract over any meta-level
+parameters.
 
 Facts (i.e.\ assumptions and finished goals) that have an application $f(x)$
 as object-logic statement get $x$ bound to the special text variable
 ``$\dots$'' (three dots).  The canonical application of this feature are
-calculational proofs, see \S\ref{sec:calculation}.
+calculational proofs (see \S\ref{sec:calculation}).
 
 
 \subsection{Block structure}
@@ -786,18 +787,14 @@
 \item [$\isarkeyword{\{\{}$ and $\isarkeyword{\}\}}$] explicitly open and
   close blocks.  Any current facts pass through $\isarkeyword{\{\{}$
   unchanged, while $\isarkeyword{\}\}}$ causes them to be \emph{exported} into
-  the enclosing context.  Thus fixed variables are generalised, assumptions
-  discharged, and local definitions eliminated.
+  the enclosing context.  Thus fixed variables are generalized, assumptions
+  discharged, and local definitions eliminated.  There is no difference of
+  $\ASSUMENAME$ and $\PRESUMENAME$ here.
 \end{descr}
 
 
 \section{Other commands}
 
-The following commands are not part of the actual proper or improper
-Isabelle/Isar syntax, but assist interactive development, for example.  Also
-note that $undo$ does not apply here, since the theory or proof configuration
-is not changed.
-
 \subsection{Diagnostics}
 
 \indexisarcmd{typ}\indexisarcmd{term}\indexisarcmd{prop}\indexisarcmd{thm}
@@ -808,6 +805,10 @@
   \isarcmd{thm} & : & \isarkeep{theory~|~proof} \\
 \end{matharray}
 
+These commands are not part of the actual Isabelle/Isar syntax, but assist
+interactive development.  Also note that $undo$ does not apply here, since the
+theory or proof configuration is not changed.
+
 \begin{rail}
   'typ' type
   ;
@@ -826,7 +827,8 @@
 \item [$\isarkeyword{thm}~thms$] retrieves lists of theorems from the current
   theory or proof context.  Note that any attributes included in the theorem
   specifications are applied to a temporary context derived from the current
-  theory or proof; the result is discarded.
+  theory or proof; the result is discarded, i.e.\ attributes involved in
+  $thms$ only have a temporary effect.
 \end{descr}
 
 
@@ -850,11 +852,14 @@
 \item [$\isarkeyword{use_thy}$, $\isarkeyword{use_thy_only}$,
   $\isarkeyword{update_thy}$, and $\isarkeyword{update_thy_only}$] load some
   theory given as $name$ argument.  These commands are exactly the same as the
-  corresponding ML functions (see also \cite[\S1 and \S6]{isabelle-ref}).
-  Note that both the ML and Isar versions of these commands may load new- and
+  corresponding ML functions (see also \cite[\S1,\S6]{isabelle-ref}).  Note
+  that both the ML and Isar versions of these commands may load new- and
   old-style theories alike.
 \end{descr}
 
+Note that these system commands are scarcely used when working with
+Proof~General, since loading of theories is done fully automatic.
+
 
 %%% Local Variables: 
 %%% mode: latex
--- a/doc-src/IsarRef/syntax.tex	Tue Aug 24 11:54:13 1999 +0200
+++ b/doc-src/IsarRef/syntax.tex	Tue Aug 24 15:38:18 1999 +0200
@@ -1,10 +1,9 @@
 
-\chapter{Isar document syntax}
+\chapter{Isar Document Syntax}
 
 We give a complete reference of all basic syntactic entities underlying the
-the Isabelle/Isar document syntax.  This chapter will not introduce any actual
-theory and proof commands, though (cf.\ chapter~\ref{ch:pure-syntax} and
-later).
+Isabelle/Isar document syntax.  Actual theory and proof commands will be
+introduced later on.
 
 \medskip
 
@@ -17,10 +16,10 @@
 are legal term specifications, while \texttt{x + y} is not.
 
 \begin{warn}
-  Note that old-style Isabelle theories used to fake parts of the inner type
-  syntax, with complicated rules when quotes may be omitted.  Despite the
-  minor drawback of requiring quotes more often, Isabelle/Isar is simpler and
-  more robust in that respect.
+  Note that Isabelle theories used to fake parts of the inner type syntax,
+  with complicated rules when quotes may be omitted.  Despite the minor
+  drawback of requiring quotes more often, Isabelle/Isar is simpler and more
+  robust in that respect.
 \end{warn}
 
 
@@ -31,6 +30,7 @@
 as given in \cite{isabelle-ref}.  These different levels of syntax should not
 be confused, though.
 
+%FIXME keyword, command
 \begin{matharray}{rcl}
   ident & = & letter~quasiletter^* \\
   longident & = & ident\verb,.,ident~\dots~ident \\
@@ -58,22 +58,22 @@
 ML-style control character notation is not supported.  The body of
 \texttt{verbatim} may consist of any text not containing \verb|*}|.
 
-Comments take the form \texttt{(*~\dots~*)} and may be nested, just as in ML.
-Note that these are \emph{source} comments only, which are stripped after
-lexical analysis of the input.  The Isar document syntax also provides several
+Comments take the form \texttt{(*~\dots~*)} and may be nested, just as in
+ML.\footnote{Proof~General may get confused by nested comments, though.} Note
+that these are \emph{source} comments only, which are stripped after lexical
+analysis of the input.  The Isar document syntax also provides several
 elements of \emph{formal comments} that are actually part of the text (see
 \S\ref{sec:comments}, \S\ref{sec:formal-cmt-thy}, \S\ref{sec:formal-cmt-prf}).
 
 
 \section{Common syntax entities}
 
-The Isar proof and theory language syntax has been carefully designed with
-orthogonality in mind.  Subsequently, we introduce several basic syntactic
-entities, such as names, terms, theorem specifications, which have been
-factored out of the actual Isar language elements described later.
+Subsequently, we introduce several basic syntactic entities, such as names,
+terms, theorem specifications, which have been factored out of the actual Isar
+language elements to be described later.
 
 Note that some of the basic syntactic entities introduced below act much like
-tokens rather than nonterminals, in particular for the sake of error messages.
+tokens rather than nonterminals, especially for the sake of error messages.
 E.g.\ syntax elements such as $\CONSTS$ referring to \railqtoken{name} or
 \railqtoken{type} would really report a missing name or type rather than any
 of the constituent primitive tokens such as \railtoken{ident} or
@@ -108,12 +108,12 @@
 are admitted as well.  Almost any of the Isar commands may be annotated by
 some marginal \railnonterm{comment} of the form \texttt{--} \railqtoken{text}.
 Note that this kind of comment is actually part of the language, while source
-level comments \verb|(*|\dots\verb|*)| are already stripped at the lexical
-level.  A few commands such as $\PROOFNAME$ admit additional markup with a
-``level of interest'': \texttt{\%} followed by an optional number $n$ (default
-$n = 1$) indicates that the respective part of the document becomes $n$ levels
-more obscure; \texttt{\%\%} means that interest drops by $\infty$ ---
-\emph{abandon every hope, who enter here}.
+level comments \verb|(*|\dots\verb|*)| are stripped at the lexical level.  A
+few commands such as $\PROOFNAME$ admit additional markup with a ``level of
+interest'': \texttt{\%} followed by an optional number $n$ (default $n = 1$)
+indicates that the respective part of the document becomes $n$ levels more
+obscure; \texttt{\%\%} means that interest drops by $\infty$ --- abandon every
+hope, who enter here.
 
 \indexoutertoken{text}\indexouternonterm{comment}\indexouternonterm{interest}
 \begin{rail}
@@ -126,10 +126,10 @@
 \end{rail}
 
 
-\subsection{Type classes, Sorts and arities}
+\subsection{Type classes, sorts and arities}
 
 The syntax of sorts and arities is given directly at the outer level.  Note
-that this is in contrast to that types and terms (see \ref{sec:types-terms}).
+that this is in contrast to types and terms (see \ref{sec:types-terms}).
 
 \indexouternonterm{sort}\indexouternonterm{arity}\indexouternonterm{simplearity}
 \indexouternonterm{classdecl}
@@ -148,12 +148,12 @@
 \subsection{Types and terms}\label{sec:types-terms}
 
 The actual inner Isabelle syntax, that of types and terms of the logic, is far
-too flexible in order to be modelled explicitly at the outer theory level.
-Basically, any such entity has to be quoted at the outer level to turn it into
-a single token (the parsing and type-checking is performed later).  For
-convenience, a slightly more liberal convention is adopted: quotes may be
-omitted for any type or term that is already \emph{atomic at the outer level}.
-E.g.\ one may write just \texttt{x} instead of \texttt{"x"}.
+too advanced in order to be modelled explicitly at the outer theory level.
+Basically, any such entity has to be quoted here to turn it into a single
+token (the parsing and type-checking is performed later).  For convenience, a
+slightly more liberal convention is adopted: quotes may be omitted for any
+type or term that is already \emph{atomic} at the outer level.  E.g.\ one may
+write just \texttt{x} instead of \texttt{"x"}.
 
 \indexoutertoken{type}\indexoutertoken{term}\indexoutertoken{prop}
 \begin{rail}
@@ -215,12 +215,12 @@
 \subsection{Attributes and theorems}\label{sec:syn-att}
 
 Attributes (and proof methods, see \S\ref{sec:syn-meth}) have their own
-``semi-inner'' syntax, in the sense that input conforming \railnonterm{args}
-below are parsed by the attribute a second time.  The attribute argument
-specifications may be any sequence of atomic entities (identifiers, strings
-etc.), or properly bracketed argument lists.  Below \railqtoken{atom} refers
-to any atomic entity, including \railtoken{keyword}s conforming to
-\railtoken{symident}.
+``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 entities
+(identifiers, strings etc.), or properly bracketed argument lists.  Below
+\railqtoken{atom} refers to any atomic entity, including \railtoken{keyword}s
+conforming to \railtoken{symident}.
 
 \indexoutertoken{atom}\indexouternonterm{args}\indexouternonterm{attributes}
 \begin{rail}
@@ -240,7 +240,7 @@
 Existing theorems are given by \railnonterm{thmref} and \railnonterm{thmrefs}
 (the former requires an actual singleton result).  Any of these theorem
 specifications may include lists of attributes both on the left and right hand
-sides; attributes are applied to the any immediately preceding theorem.
+sides; attributes are applied to any immediately preceding theorem.
 
 \indexouternonterm{thmdecl}\indexouternonterm{axmdecl}
 \indexouternonterm{thmdef}\indexouternonterm{thmrefs}
@@ -267,10 +267,10 @@
 ``\texttt{,}'' (sequential composition), ``\texttt{|}'' (alternative choices),
 ``\texttt{?}'' (try), ``\texttt{*}'' (repeat ${} \ge 0$ times), ``\texttt{+}''
 (repeat ${} > 0$ times).  In practice, proof methods are usually just a comma
-separated list of (\railqtoken{nameref}~\railnonterm{args}) specifications.
+separated list of \railqtoken{nameref}~\railnonterm{args} specifications.
 Thus the syntax is similar to that of attributes, with plain parentheses
-instead of square brackets (see also \S\ref{sec:syn-att}).  Note that
-parentheses may be dropped for single method specifications without arguments.
+instead of square brackets.  Note that parentheses may be dropped for single
+method specifications without arguments.
 
 \indexouternonterm{method}
 \begin{rail}