Merged, overriding earlier fix.
authorballarin
Mon, 19 Jan 2009 21:20:18 +0100
changeset 29567 286c01be90cb
parent 29565 3f8b24fcfbd6 (diff)
parent 29566 937baa077df2 (current diff)
child 29568 ba144750086d
Merged, overriding earlier fix.
doc-src/Locales/Locales/Examples.thy
doc-src/Locales/Locales/Examples3.thy
doc-src/Locales/Locales/document/Examples.tex
doc-src/Locales/Locales/document/Examples3.tex
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Admin/churn	Mon Jan 19 21:20:18 2009 +0100
@@ -0,0 +1,5 @@
+#!/bin/bash
+
+ADMIN="$(cd "$(dirname "$0")"; cd "$(pwd -P)"; pwd)"
+cd "$ADMIN"
+hg churn --aliases user-aliases --progress
--- a/Admin/isatest/isatest-statistics	Mon Jan 19 20:37:08 2009 +0100
+++ b/Admin/isatest/isatest-statistics	Mon Jan 19 21:20:18 2009 +0100
@@ -1,6 +1,5 @@
 #!/usr/bin/env bash
 #
-# $Id$
 # Author: Makarius
 #
 # DESCRIPTION: Produce statistics from isatest session logs.
@@ -93,7 +92,7 @@
 
 for SESSION in $SESSIONS
 do
-  fgrep "$SESSION " "$ALL_DATA" > "$SESSION_DATA"
+  grep "^${SESSION} " "$ALL_DATA" > "$SESSION_DATA"
   PLOT="plot [] [0:] \"$SESSION_DATA\" using 2:3 smooth sbezier title \"interpolated cpu time\", \"$SESSION_DATA\" using 2:3 smooth csplines title \"cpu time\""
   if [ "$PARALLEL" = true ]; then
     PLOT="${PLOT}, \"$SESSION_DATA\" using 2:4 smooth sbezier title \"interpolated elapsed time\", \"$SESSION_DATA\" using 2:4 smooth csplines title \"elapsed time\""
--- a/Admin/isatest/isatest-stats	Mon Jan 19 20:37:08 2009 +0100
+++ b/Admin/isatest/isatest-stats	Mon Jan 19 21:20:18 2009 +0100
@@ -31,28 +31,33 @@
   HOL-Word \
   HOL-ex \
   ZF \
-  ZF-Constructible\
+  ZF-Constructible \
   ZF-UNITY"
 
 AFP_SESSIONS="\
-  CoreC++\
-  LinearQuantifierElim\
-  HOL-DiskPaxos\
-  HOL-Fermat3_4\
-  HOL-Flyspeck-Tame\
-  HOL-Group-Ring-Module\
-  HOL-JinjaThreads\
-  HOL-Jinja\
-  HOL-JiveDataStoreModel\
-  HOL-POPLmark-deBruijn\
-  HOL-Program-Conflict-Analysis\
-  HOL-RSAPSS\
-  HOL-Recursion-Theory-I\
-  HOL-SumSquares\
-  HOL-Topology\
-  HOL-Valuation\
-  Simpl-BDD\
-  Simpl"
+  CoreC++ \
+  HOL-BytecodeLogicJmlTypes \
+  HOL-DiskPaxos \
+  HOL-Fermat3_4 \
+  HOL-Flyspeck-Tame \
+  HOL-Group-Ring-Module \
+  HOL-Jinja \
+  HOL-JinjaThreads \
+  HOL-JiveDataStoreModel \
+  HOL-POPLmark-deBruijn \
+  HOL-Program-Conflict-Analysis \
+  HOL-RSAPSS \
+  HOL-Recursion-Theory-I \
+  HOL-SATSolverVerification \
+  HOL-SIFPL \
+  HOL-SenSocialChoice \
+  HOL-Slicing \
+  HOL-SumSquares \
+  HOL-Topology \
+  HOL-Valuation \
+  LinearQuantifierElim \
+  Simpl \
+  Simpl-BDD"
 
 for PLATFORM in $PLATFORMS
 do
--- a/Admin/isatest/settings/at-mac-poly-5.1-para	Mon Jan 19 20:37:08 2009 +0100
+++ b/Admin/isatest/settings/at-mac-poly-5.1-para	Mon Jan 19 21:20:18 2009 +0100
@@ -4,7 +4,7 @@
   ML_SYSTEM="polyml-5.2.1"
   ML_PLATFORM="x86-darwin"
   ML_HOME="$POLYML_HOME/$ML_PLATFORM"
-  ML_OPTIONS="--immutable 800 --mutable 1200"
+  ML_OPTIONS="--mutable 200 --immutable 800"
 
 
 ISABELLE_HOME_USER=~/isabelle-at-mac-poly-e
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Admin/user-aliases	Mon Jan 19 21:20:18 2009 +0100
@@ -0,0 +1,4 @@
+lcp paulson
+norbert.schirmer@web.de schirmer
+urbanc@in.tum.de urbanc
+nipkow@lapbroy100.local nipkow
\ No newline at end of file
--- a/CONTRIBUTORS	Mon Jan 19 20:37:08 2009 +0100
+++ b/CONTRIBUTORS	Mon Jan 19 21:20:18 2009 +0100
@@ -7,12 +7,19 @@
 Contributions to this Isabelle version
 --------------------------------------
 
+* December 2008: Clemens Ballarin, TUM
+  New locale implementation.
+
 * December 2008: Armin Heller, TUM and Alexander Krauss, TUM
   Method "sizechange" for advanced termination proofs.
 
 * November 2008: Timothy Bourke, NICTA
   Performance improvement (factor 50) for find_theorems.
 
+* 2008: Florian Haftmann, TUM
+  Various extensions and restructurings in HOL, improvements
+  in evaluation mechanisms, new module binding.ML for name bindings.
+
 * October 2008: Fabian Immler, TUM
   ATP manager for Sledgehammer, based on ML threads instead of Posix
   processes.  Additional ATP wrappers, including remote SystemOnTPTP
--- a/NEWS	Mon Jan 19 20:37:08 2009 +0100
+++ b/NEWS	Mon Jan 19 21:20:18 2009 +0100
@@ -236,6 +236,13 @@
     src/HOL/Real/rat_arith.ML ~> src/HOL/Tools
     src/HOL/Real/real_arith.ML ~> src/HOL/Tools
 
+    src/HOL/Library/Array.thy ~> src/HOL/Imperative_HOL
+    src/HOL/Library/Heap_Monad.thy ~> src/HOL/Imperative_HOL
+    src/HOL/Library/Heap.thy ~> src/HOL/Imperative_HOL
+    src/HOL/Library/Imperative_HOL.thy ~> src/HOL/Imperative_HOL
+    src/HOL/Library/Ref.thy ~> src/HOL/Imperative_HOL
+    src/HOL/Library/Relational.thy ~> src/HOL/Imperative_HOL
+
 * If methods "eval" and "evaluation" encounter a structured proof state
 with !!/==>, only the conclusion is evaluated to True (if possible),
 avoiding strange error messages.
--- a/README_REPOSITORY	Mon Jan 19 20:37:08 2009 +0100
+++ b/README_REPOSITORY	Mon Jan 19 21:20:18 2009 +0100
@@ -30,7 +30,7 @@
 
 
 Initial configuration
-=====================
+---------------------
 
 Always use Mercurial version 1.0 or later, such as 1.0.1 or 1.0.2.
 The old 0.9.x versions do not work in a multi-user environment with
@@ -84,7 +84,7 @@
 
 
 Shared pull/push access
-=======================
+-----------------------
 
 The entry point http://isabelle.in.tum.de/repos/isabelle is world
 readable, both via plain web browsing and the hg client as described
@@ -136,7 +136,7 @@
 
 
 Content discipline
-==================
+------------------
 
 Old-style centralized version control is occasionally compared to "a
 library where everybody scribbles into the books".  Or seen the other
@@ -183,7 +183,7 @@
 
 
 Building Isabelle from the repository version
-=============================================
+---------------------------------------------
 
 Compared to a proper distribution or development snapshot, a
 repository version of Isabelle lacks textual version identifiers in
--- a/doc-src/IsarAdvanced/Classes/Thy/Classes.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/doc-src/IsarAdvanced/Classes/Thy/Classes.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -368,15 +368,15 @@
 text {*
   \noindent The connection to the type system is done by means
   of a primitive axclass
-*}
+*} setup %invisible {* Sign.add_path "foo" *}
 
 axclass %quote idem < type
-  idem: "f (f x) = f x"
+  idem: "f (f x) = f x" setup %invisible {* Sign.parent_path *}
 
 text {* \noindent together with a corresponding interpretation: *}
 
 interpretation %quote idem_class:
-  idem ["f \<Colon> (\<alpha>\<Colon>idem) \<Rightarrow> \<alpha>"]
+  idem "f \<Colon> (\<alpha>\<Colon>idem) \<Rightarrow> \<alpha>"
 proof qed (rule idem)
 
 text {*
@@ -459,7 +459,7 @@
   of monoids for lists:
 *}
 
-interpretation %quote list_monoid: monoid [append "[]"]
+interpretation %quote list_monoid!: monoid append "[]"
   proof qed auto
 
 text {*
@@ -474,10 +474,10 @@
   "replicate 0 _ = []"
   | "replicate (Suc n) xs = xs @ replicate n xs"
 
-interpretation %quote list_monoid: monoid [append "[]"] where
+interpretation %quote list_monoid!: monoid append "[]" where
   "monoid.pow_nat append [] = replicate"
 proof -
-  interpret monoid [append "[]"] ..
+  interpret monoid append "[]" ..
   show "monoid.pow_nat append [] = replicate"
   proof
     fix n
--- a/doc-src/IsarAdvanced/Classes/Thy/document/Classes.tex	Mon Jan 19 20:37:08 2009 +0100
+++ b/doc-src/IsarAdvanced/Classes/Thy/document/Classes.tex	Mon Jan 19 21:20:18 2009 +0100
@@ -655,7 +655,23 @@
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
+\isadeliminvisible
+\ %
+\endisadeliminvisible
+%
+\isataginvisible
+\isacommand{setup}\isamarkupfalse%
+\ {\isacharverbatimopen}\ Sign{\isachardot}add{\isacharunderscore}path\ {\isachardoublequote}foo{\isachardoublequote}\ {\isacharverbatimclose}%
+\endisataginvisible
+{\isafoldinvisible}%
+%
+\isadeliminvisible
+%
+\endisadeliminvisible
+\isanewline
+%
 \isadelimquote
+\isanewline
 %
 \endisadelimquote
 %
@@ -670,6 +686,20 @@
 %
 \endisadelimquote
 %
+\isadeliminvisible
+\ %
+\endisadeliminvisible
+%
+\isataginvisible
+\isacommand{setup}\isamarkupfalse%
+\ {\isacharverbatimopen}\ Sign{\isachardot}parent{\isacharunderscore}path\ {\isacharverbatimclose}%
+\endisataginvisible
+{\isafoldinvisible}%
+%
+\isadeliminvisible
+%
+\endisadeliminvisible
+%
 \begin{isamarkuptext}%
 \noindent together with a corresponding interpretation:%
 \end{isamarkuptext}%
@@ -682,7 +712,7 @@
 \isatagquote
 \isacommand{interpretation}\isamarkupfalse%
 \ idem{\isacharunderscore}class{\isacharcolon}\isanewline
-\ \ idem\ {\isacharbrackleft}{\isachardoublequoteopen}f\ {\isasymColon}\ {\isacharparenleft}{\isasymalpha}{\isasymColon}idem{\isacharparenright}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}{\isacharbrackright}\isanewline
+\ \ idem\ {\isachardoublequoteopen}f\ {\isasymColon}\ {\isacharparenleft}{\isasymalpha}{\isasymColon}idem{\isacharparenright}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\isanewline
 \isacommand{proof}\isamarkupfalse%
 \ \isacommand{qed}\isamarkupfalse%
 \ {\isacharparenleft}rule\ idem{\isacharparenright}%
@@ -844,7 +874,7 @@
 %
 \isatagquote
 \isacommand{interpretation}\isamarkupfalse%
-\ list{\isacharunderscore}monoid{\isacharcolon}\ monoid\ {\isacharbrackleft}append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}{\isacharbrackright}\isanewline
+\ list{\isacharunderscore}monoid{\isacharbang}{\isacharcolon}\ monoid\ append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline
 \ \ \isacommand{proof}\isamarkupfalse%
 \ \isacommand{qed}\isamarkupfalse%
 \ auto%
@@ -875,12 +905,12 @@
 \ \ {\isacharbar}\ {\isachardoublequoteopen}replicate\ {\isacharparenleft}Suc\ n{\isacharparenright}\ xs\ {\isacharequal}\ xs\ {\isacharat}\ replicate\ n\ xs{\isachardoublequoteclose}\isanewline
 \isanewline
 \isacommand{interpretation}\isamarkupfalse%
-\ list{\isacharunderscore}monoid{\isacharcolon}\ monoid\ {\isacharbrackleft}append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}{\isacharbrackright}\ \isakeyword{where}\isanewline
+\ list{\isacharunderscore}monoid{\isacharbang}{\isacharcolon}\ monoid\ append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
 \ \ {\isachardoublequoteopen}monoid{\isachardot}pow{\isacharunderscore}nat\ append\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ replicate{\isachardoublequoteclose}\isanewline
 \isacommand{proof}\isamarkupfalse%
 \ {\isacharminus}\isanewline
 \ \ \isacommand{interpret}\isamarkupfalse%
-\ monoid\ {\isacharbrackleft}append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}{\isacharbrackright}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
+\ monoid\ append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{show}\isamarkupfalse%
 \ {\isachardoublequoteopen}monoid{\isachardot}pow{\isacharunderscore}nat\ append\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ replicate{\isachardoublequoteclose}\isanewline
@@ -1231,7 +1261,7 @@
 \hspace*{0pt}\\
 \hspace*{0pt}fun inverse{\char95}int i = IntInf.{\char126}~i;\\
 \hspace*{0pt}\\
-\hspace*{0pt}val neutral{\char95}int :~IntInf.int = (0 :~IntInf.int);\\
+\hspace*{0pt}val neutral{\char95}int :~IntInf.int = (0 :~IntInf.int)\\
 \hspace*{0pt}\\
 \hspace*{0pt}fun mult{\char95}int i j = IntInf.+ (i,~j);\\
 \hspace*{0pt}\\
@@ -1258,7 +1288,7 @@
 \hspace*{0pt} ~~~else inverse A{\char95}~(pow{\char95}nat (monoid{\char95}group A{\char95}) (nat (IntInf.{\char126}~k)) x));\\
 \hspace*{0pt}\\
 \hspace*{0pt}val example :~IntInf.int =\\
-\hspace*{0pt} ~pow{\char95}int group{\char95}int (10 :~IntInf.int) ({\char126}2 :~IntInf.int);\\
+\hspace*{0pt} ~pow{\char95}int group{\char95}int (10 :~IntInf.int) ({\char126}2 :~IntInf.int)\\
 \hspace*{0pt}\\
 \hspace*{0pt}end;~(*struct Example*)%
 \end{isamarkuptext}%
--- a/doc-src/IsarAdvanced/Codegen/Thy/Introduction.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/doc-src/IsarAdvanced/Codegen/Thy/Introduction.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -128,7 +128,7 @@
       \tikzstyle process=[ellipse, draw, thick, color = green, fill = white];
       \tikzstyle process_arrow=[->, semithick, color = green];
       \node (HOL) at (0, 4) [style=entity] {@{text "Isabelle/HOL"} theory};
-      \node (eqn) at (2, 2) [style=entity] {defining equations};
+      \node (eqn) at (2, 2) [style=entity] {code equations};
       \node (iml) at (2, 0) [style=entity] {intermediate language};
       \node (seri) at (1, 0) [style=process] {serialisation};
       \node (SML) at (0, 3) [style=entity] {@{text SML}};
@@ -153,12 +153,12 @@
   The code generator employs a notion of executability
   for three foundational executable ingredients known
   from functional programming:
-  \emph{defining equations}, \emph{datatypes}, and
-  \emph{type classes}.  A defining equation as a first approximation
+  \emph{code equations}, \emph{datatypes}, and
+  \emph{type classes}.  A code equation as a first approximation
   is a theorem of the form @{text "f t\<^isub>1 t\<^isub>2 \<dots> t\<^isub>n \<equiv> t"}
   (an equation headed by a constant @{text f} with arguments
   @{text "t\<^isub>1 t\<^isub>2 \<dots> t\<^isub>n"} and right hand side @{text t}).
-  Code generation aims to turn defining equations
+  Code generation aims to turn code equations
   into a functional program.  This is achieved by three major
   components which operate sequentially, i.e. the result of one is
   the input
@@ -168,7 +168,7 @@
 
     \item Out of the vast collection of theorems proven in a
       \qn{theory}, a reasonable subset modelling
-      defining equations is \qn{selected}.
+      code equations is \qn{selected}.
 
     \item On those selected theorems, certain
       transformations are carried out
@@ -177,7 +177,7 @@
       specifications into equivalent but executable counterparts.
       The result is a structured collection of \qn{code theorems}.
 
-    \item Before the selected defining equations are continued with,
+    \item Before the selected code equations are continued with,
       they can be \qn{preprocessed}, i.e. subjected to theorem
       transformations.  This \qn{preprocessor} is an interface which
       allows to apply
@@ -185,12 +185,12 @@
       to code generation;  motivating examples are shown below, see
       \secref{sec:preproc}.
       The result of the preprocessing step is a structured collection
-      of defining equations.
+      of code equations.
 
-    \item These defining equations are \qn{translated} to a program
+    \item These code equations are \qn{translated} to a program
       in an abstract intermediate language.  Think of it as a kind
       of \qt{Mini-Haskell} with four \qn{statements}: @{text data}
-      (for datatypes), @{text fun} (stemming from defining equations),
+      (for datatypes), @{text fun} (stemming from code equations),
       also @{text class} and @{text inst} (for type classes).
 
     \item Finally, the abstract program is \qn{serialised} into concrete
--- a/doc-src/IsarAdvanced/Codegen/Thy/ML.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/doc-src/IsarAdvanced/Codegen/Thy/ML.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -24,9 +24,9 @@
   \begin{mldecls}
   @{index_ML Code.add_eqn: "thm -> theory -> theory"} \\
   @{index_ML Code.del_eqn: "thm -> theory -> theory"} \\
-  @{index_ML Code.add_eqnl: "string * (thm * bool) list Lazy.T -> theory -> theory"} \\
-  @{index_ML Code.map_pre: "(MetaSimplifier.simpset -> MetaSimplifier.simpset) -> theory -> theory"} \\
-  @{index_ML Code.map_post: "(MetaSimplifier.simpset -> MetaSimplifier.simpset) -> theory -> theory"} \\
+  @{index_ML Code.add_eqnl: "string * (thm * bool) list lazy -> theory -> theory"} \\
+  @{index_ML Code.map_pre: "(simpset -> simpset) -> theory -> theory"} \\
+  @{index_ML Code.map_post: "(simpset -> simpset) -> theory -> theory"} \\
   @{index_ML Code.add_functrans: "string * (theory -> (thm * bool) list -> (thm * bool) list option)
     -> theory -> theory"} \\
   @{index_ML Code.del_functrans: "string -> theory -> theory"} \\
@@ -45,7 +45,7 @@
      theorem @{text "thm"} from executable content, if present.
 
   \item @{ML Code.add_eqnl}~@{text "(const, lthms)"}~@{text "thy"} adds
-     suspended defining equations @{text lthms} for constant
+     suspended code equations @{text lthms} for constant
      @{text const} to executable content.
 
   \item @{ML Code.map_pre}~@{text "f"}~@{text "thy"} changes
@@ -53,11 +53,11 @@
 
   \item @{ML Code.add_functrans}~@{text "(name, f)"}~@{text "thy"} adds
      function transformer @{text f} (named @{text name}) to executable content;
-     @{text f} is a transformer of the defining equations belonging
+     @{text f} is a transformer of the code equations belonging
      to a certain function definition, depending on the
      current theory context.  Returning @{text NONE} indicates that no
      transformation took place;  otherwise, the whole process will be iterated
-     with the new defining equations.
+     with the new code equations.
 
   \item @{ML Code.del_functrans}~@{text "name"}~@{text "thy"} removes
      function transformer named @{text name} from executable content.
@@ -80,7 +80,7 @@
   \begin{mldecls}
   @{index_ML Code_Unit.read_const: "theory -> string -> string"} \\
   @{index_ML Code_Unit.head_eqn: "theory -> thm -> string * ((string * sort) list * typ)"} \\
-  @{index_ML Code_Unit.rewrite_eqn: "MetaSimplifier.simpset -> thm -> thm"} \\
+  @{index_ML Code_Unit.rewrite_eqn: "simpset -> thm -> thm"} \\
   \end{mldecls}
 
   \begin{description}
@@ -89,12 +89,12 @@
      reads a constant as a concrete term expression @{text s}.
 
   \item @{ML Code_Unit.head_eqn}~@{text thy}~@{text thm}
-     extracts the constant and its type from a defining equation @{text thm}.
+     extracts the constant and its type from a code equation @{text thm}.
 
   \item @{ML Code_Unit.rewrite_eqn}~@{text ss}~@{text thm}
-     rewrites a defining equation @{text thm} with a simpset @{text ss};
+     rewrites a code equation @{text thm} with a simpset @{text ss};
      only arguments and right hand side are rewritten,
-     not the head of the defining equation.
+     not the head of the code equation.
 
   \end{description}
 
--- a/doc-src/IsarAdvanced/Codegen/Thy/Program.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/doc-src/IsarAdvanced/Codegen/Thy/Program.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -10,7 +10,7 @@
   We have already seen how by default equations stemming from
   @{command definition}/@{command primrec}/@{command fun}
   statements are used for code generation.  This default behaviour
-  can be changed, e.g. by providing different defining equations.
+  can be changed, e.g. by providing different code equations.
   All kinds of customisation shown in this section is \emph{safe}
   in the sense that the user does not have to worry about
   correctness -- all programs generatable that way are partially
@@ -21,7 +21,7 @@
 
 text {*
   Coming back to our introductory example, we
-  could provide an alternative defining equations for @{const dequeue}
+  could provide an alternative code equations for @{const dequeue}
   explicitly:
 *}
 
@@ -36,7 +36,7 @@
 text {*
   \noindent The annotation @{text "[code]"} is an @{text Isar}
   @{text attribute} which states that the given theorems should be
-  considered as defining equations for a @{text fun} statement --
+  considered as code equations for a @{text fun} statement --
   the corresponding constant is determined syntactically.  The resulting code:
 *}
 
@@ -59,13 +59,13 @@
 code_thms %quote dequeue
 
 text {*
-  \noindent prints a table with \emph{all} defining equations
+  \noindent prints a table with \emph{all} code equations
   for @{const dequeue}, including
-  \emph{all} defining equations those equations depend
+  \emph{all} code equations those equations depend
   on recursively.
   
   Similarly, the @{command code_deps} command shows a graph
-  visualising dependencies between defining equations.
+  visualising dependencies between code equations.
 *}
 
 subsection {* @{text class} and @{text instantiation} *}
@@ -155,7 +155,7 @@
 
   The \emph{simpset} allows to employ the full generality of the Isabelle
   simplifier.  Due to the interpretation of theorems
-  as defining equations, rewrites are applied to the right
+  as code equations, rewrites are applied to the right
   hand side and the arguments of the left hand side of an
   equation, but never to the constant heading the left hand side.
   An important special case are \emph{inline theorems} which may be
@@ -207,7 +207,7 @@
   the @{command print_codesetup} command.
   @{command code_thms} provides a convenient
   mechanism to inspect the impact of a preprocessor setup
-  on defining equations.
+  on code equations.
 
   \begin{warn}
     The attribute \emph{code unfold}
@@ -351,7 +351,7 @@
   an explicit class @{class eq} with a corresponding operation
   @{const eq_class.eq} such that @{thm eq [no_vars]}.
   The preprocessing framework does the rest by propagating the
-  @{class eq} constraints through all dependent defining equations.
+  @{class eq} constraints through all dependent code equations.
   For datatypes, instances of @{class eq} are implicitly derived
   when possible.  For other types, you may instantiate @{text eq}
   manually like any other type class.
@@ -410,7 +410,7 @@
 text %quote {*@{code_stmts "op \<le> \<Colon> _ \<times> _ \<Rightarrow> _ \<times> _ \<Rightarrow> bool" (SML)}*}
 
 text {*
-  In some cases, the automatically derived defining equations
+  In some cases, the automatically derived code equations
   for equality on a particular type may not be appropriate.
   As example, watch the following datatype representing
   monomorphic parametric types (where type constructors
@@ -493,7 +493,7 @@
   on the right hand side of its first equation the constant
   @{const empty_queue} occurs which is unspecified.
 
-  Normally, if constants without any defining equations occur in
+  Normally, if constants without any code equations occur in
   a program, the code generator complains (since in most cases
   this is not what the user expects).  But such constants can also
   be thought of as function definitions with no equations which
--- a/doc-src/IsarAdvanced/Codegen/Thy/Setup.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/doc-src/IsarAdvanced/Codegen/Thy/Setup.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -4,8 +4,8 @@
 begin
 
 ML {* no_document use_thys
-  ["Efficient_Nat", "Code_Char_chr", "Product_ord", "Imperative_HOL",
-   "~~/src/HOL/Complex/ex/ReflectedFerrack"] *}
+  ["Efficient_Nat", "Code_Char_chr", "Product_ord", "~~/src/HOL/Imperative_HOL/Imperative_HOL",
+   "~~/src/HOL/ex/ReflectedFerrack"] *}
 
 ML_val {* Code_Target.code_width := 74 *}
 
--- a/doc-src/IsarAdvanced/Codegen/Thy/document/Adaption.tex	Mon Jan 19 20:37:08 2009 +0100
+++ b/doc-src/IsarAdvanced/Codegen/Thy/document/Adaption.tex	Mon Jan 19 21:20:18 2009 +0100
@@ -281,9 +281,9 @@
 \hspace*{0pt}and less{\char95}eq{\char95}nat (Suc m) n = less{\char95}nat m n\\
 \hspace*{0pt} ~| less{\char95}eq{\char95}nat Zero{\char95}nat n = True;\\
 \hspace*{0pt}\\
-\hspace*{0pt}fun in{\char95}interval (k, l) n = anda (less{\char95}eq{\char95}nat k n) (less{\char95}eq{\char95}nat n l);\\
+\hspace*{0pt}fun in{\char95}interval (k,~l) n = anda (less{\char95}eq{\char95}nat k n) (less{\char95}eq{\char95}nat n l);\\
 \hspace*{0pt}\\
-\hspace*{0pt}end; (*struct Example*)%
+\hspace*{0pt}end;~(*struct Example*)%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -357,9 +357,9 @@
 \hspace*{0pt}and less{\char95}eq{\char95}nat (Suc m) n = less{\char95}nat m n\\
 \hspace*{0pt} ~| less{\char95}eq{\char95}nat Zero{\char95}nat n = true;\\
 \hspace*{0pt}\\
-\hspace*{0pt}fun in{\char95}interval (k, l) n = (less{\char95}eq{\char95}nat k n) andalso (less{\char95}eq{\char95}nat n l);\\
+\hspace*{0pt}fun in{\char95}interval (k,~l) n = (less{\char95}eq{\char95}nat k n) andalso (less{\char95}eq{\char95}nat n l);\\
 \hspace*{0pt}\\
-\hspace*{0pt}end; (*struct Example*)%
+\hspace*{0pt}end;~(*struct Example*)%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -414,9 +414,9 @@
 \hspace*{0pt}and less{\char95}eq{\char95}nat (Suc m) n = less{\char95}nat m n\\
 \hspace*{0pt} ~| less{\char95}eq{\char95}nat Zero{\char95}nat n = true;\\
 \hspace*{0pt}\\
-\hspace*{0pt}fun in{\char95}interval (k, l) n = less{\char95}eq{\char95}nat k n andalso less{\char95}eq{\char95}nat n l;\\
+\hspace*{0pt}fun in{\char95}interval (k,~l) n = less{\char95}eq{\char95}nat k n andalso less{\char95}eq{\char95}nat n l;\\
 \hspace*{0pt}\\
-\hspace*{0pt}end; (*struct Example*)%
+\hspace*{0pt}end;~(*struct Example*)%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
--- a/doc-src/IsarAdvanced/Codegen/Thy/document/Introduction.tex	Mon Jan 19 20:37:08 2009 +0100
+++ b/doc-src/IsarAdvanced/Codegen/Thy/document/Introduction.tex	Mon Jan 19 21:20:18 2009 +0100
@@ -164,20 +164,20 @@
 \hspace*{0pt}\\
 \hspace*{0pt}datatype 'a queue = Queue of 'a list * 'a list;\\
 \hspace*{0pt}\\
-\hspace*{0pt}val empty :~'a queue = Queue ([], []);\\
+\hspace*{0pt}val empty :~'a queue = Queue ([],~[])\\
 \hspace*{0pt}\\
-\hspace*{0pt}fun dequeue (Queue ([], [])) = (NONE, Queue ([], []))\\
-\hspace*{0pt} ~| dequeue (Queue (xs, y ::~ys)) = (SOME y, Queue (xs, ys))\\
-\hspace*{0pt} ~| dequeue (Queue (v ::~va, [])) =\\
+\hspace*{0pt}fun dequeue (Queue ([],~[])) = (NONE,~Queue ([],~[]))\\
+\hspace*{0pt} ~| dequeue (Queue (xs,~y ::~ys)) = (SOME y,~Queue (xs,~ys))\\
+\hspace*{0pt} ~| dequeue (Queue (v ::~va,~[])) =\\
 \hspace*{0pt} ~~~let\\
 \hspace*{0pt} ~~~~~val y ::~ys = rev (v ::~va);\\
 \hspace*{0pt} ~~~in\\
-\hspace*{0pt} ~~~~~(SOME y, Queue ([], ys))\\
+\hspace*{0pt} ~~~~~(SOME y,~Queue ([],~ys))\\
 \hspace*{0pt} ~~~end;\\
 \hspace*{0pt}\\
-\hspace*{0pt}fun enqueue x (Queue (xs, ys)) = Queue (x ::~xs, ys);\\
+\hspace*{0pt}fun enqueue x (Queue (xs,~ys)) = Queue (x ::~xs,~ys);\\
 \hspace*{0pt}\\
-\hspace*{0pt}end; (*struct Example*)%
+\hspace*{0pt}end;~(*struct Example*)%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -233,31 +233,31 @@
 \hspace*{0pt}module Example where {\char123}\\
 \hspace*{0pt}\\
 \hspace*{0pt}\\
-\hspace*{0pt}foldla ::~forall a b. (a -> b -> a) -> a -> [b] -> a;\\
+\hspace*{0pt}foldla ::~forall a b.~(a -> b -> a) -> a -> [b] -> a;\\
 \hspace*{0pt}foldla f a [] = a;\\
 \hspace*{0pt}foldla f a (x :~xs) = foldla f (f a x) xs;\\
 \hspace*{0pt}\\
-\hspace*{0pt}rev ::~forall a. [a] -> [a];\\
+\hspace*{0pt}rev ::~forall a.~[a] -> [a];\\
 \hspace*{0pt}rev xs = foldla ({\char92}~xsa x -> x :~xsa) [] xs;\\
 \hspace*{0pt}\\
-\hspace*{0pt}list{\char95}case ::~forall t a. t -> (a -> [a] -> t) -> [a] -> t;\\
+\hspace*{0pt}list{\char95}case ::~forall t a.~t -> (a -> [a] -> t) -> [a] -> t;\\
 \hspace*{0pt}list{\char95}case f1 f2 (a :~list) = f2 a list;\\
 \hspace*{0pt}list{\char95}case f1 f2 [] = f1;\\
 \hspace*{0pt}\\
 \hspace*{0pt}data Queue a = Queue [a] [a];\\
 \hspace*{0pt}\\
-\hspace*{0pt}empty ::~forall a. Queue a;\\
+\hspace*{0pt}empty ::~forall a.~Queue a;\\
 \hspace*{0pt}empty = Queue [] [];\\
 \hspace*{0pt}\\
-\hspace*{0pt}dequeue ::~forall a. Queue a -> (Maybe a, Queue a);\\
-\hspace*{0pt}dequeue (Queue [] []) = (Nothing, Queue [] []);\\
-\hspace*{0pt}dequeue (Queue xs (y :~ys)) = (Just y, Queue xs ys);\\
+\hspace*{0pt}dequeue ::~forall a.~Queue a -> (Maybe a,~Queue a);\\
+\hspace*{0pt}dequeue (Queue [] []) = (Nothing,~Queue [] []);\\
+\hspace*{0pt}dequeue (Queue xs (y :~ys)) = (Just y,~Queue xs ys);\\
 \hspace*{0pt}dequeue (Queue (v :~va) []) =\\
 \hspace*{0pt} ~let {\char123}\\
 \hspace*{0pt} ~~~(y :~ys) = rev (v :~va);\\
-\hspace*{0pt} ~{\char125}~in (Just y, Queue [] ys);\\
+\hspace*{0pt} ~{\char125}~in (Just y,~Queue [] ys);\\
 \hspace*{0pt}\\
-\hspace*{0pt}enqueue ::~forall a. a -> Queue a -> Queue a;\\
+\hspace*{0pt}enqueue ::~forall a.~a -> Queue a -> Queue a;\\
 \hspace*{0pt}enqueue x (Queue xs ys) = Queue (x :~xs) ys;\\
 \hspace*{0pt}\\
 \hspace*{0pt}{\char125}%
@@ -293,7 +293,7 @@
       \tikzstyle process=[ellipse, draw, thick, color = green, fill = white];
       \tikzstyle process_arrow=[->, semithick, color = green];
       \node (HOL) at (0, 4) [style=entity] {\isa{Isabelle{\isacharslash}HOL} theory};
-      \node (eqn) at (2, 2) [style=entity] {defining equations};
+      \node (eqn) at (2, 2) [style=entity] {code equations};
       \node (iml) at (2, 0) [style=entity] {intermediate language};
       \node (seri) at (1, 0) [style=process] {serialisation};
       \node (SML) at (0, 3) [style=entity] {\isa{SML}};
@@ -318,12 +318,12 @@
   The code generator employs a notion of executability
   for three foundational executable ingredients known
   from functional programming:
-  \emph{defining equations}, \emph{datatypes}, and
-  \emph{type classes}.  A defining equation as a first approximation
+  \emph{code equations}, \emph{datatypes}, and
+  \emph{type classes}.  A code equation as a first approximation
   is a theorem of the form \isa{f\ t\isactrlisub {\isadigit{1}}\ t\isactrlisub {\isadigit{2}}\ {\isasymdots}\ t\isactrlisub n\ {\isasymequiv}\ t}
   (an equation headed by a constant \isa{f} with arguments
   \isa{t\isactrlisub {\isadigit{1}}\ t\isactrlisub {\isadigit{2}}\ {\isasymdots}\ t\isactrlisub n} and right hand side \isa{t}).
-  Code generation aims to turn defining equations
+  Code generation aims to turn code equations
   into a functional program.  This is achieved by three major
   components which operate sequentially, i.e. the result of one is
   the input
@@ -333,7 +333,7 @@
 
     \item Out of the vast collection of theorems proven in a
       \qn{theory}, a reasonable subset modelling
-      defining equations is \qn{selected}.
+      code equations is \qn{selected}.
 
     \item On those selected theorems, certain
       transformations are carried out
@@ -342,7 +342,7 @@
       specifications into equivalent but executable counterparts.
       The result is a structured collection of \qn{code theorems}.
 
-    \item Before the selected defining equations are continued with,
+    \item Before the selected code equations are continued with,
       they can be \qn{preprocessed}, i.e. subjected to theorem
       transformations.  This \qn{preprocessor} is an interface which
       allows to apply
@@ -350,12 +350,12 @@
       to code generation;  motivating examples are shown below, see
       \secref{sec:preproc}.
       The result of the preprocessing step is a structured collection
-      of defining equations.
+      of code equations.
 
-    \item These defining equations are \qn{translated} to a program
+    \item These code equations are \qn{translated} to a program
       in an abstract intermediate language.  Think of it as a kind
       of \qt{Mini-Haskell} with four \qn{statements}: \isa{data}
-      (for datatypes), \isa{fun} (stemming from defining equations),
+      (for datatypes), \isa{fun} (stemming from code equations),
       also \isa{class} and \isa{inst} (for type classes).
 
     \item Finally, the abstract program is \qn{serialised} into concrete
--- a/doc-src/IsarAdvanced/Codegen/Thy/document/ML.tex	Mon Jan 19 20:37:08 2009 +0100
+++ b/doc-src/IsarAdvanced/Codegen/Thy/document/ML.tex	Mon Jan 19 21:20:18 2009 +0100
@@ -54,9 +54,9 @@
 \begin{mldecls}
   \indexml{Code.add\_eqn}\verb|Code.add_eqn: thm -> theory -> theory| \\
   \indexml{Code.del\_eqn}\verb|Code.del_eqn: thm -> theory -> theory| \\
-  \indexml{Code.add\_eqnl}\verb|Code.add_eqnl: string * (thm * bool) list Lazy.T -> theory -> theory| \\
-  \indexml{Code.map\_pre}\verb|Code.map_pre: (MetaSimplifier.simpset -> MetaSimplifier.simpset) -> theory -> theory| \\
-  \indexml{Code.map\_post}\verb|Code.map_post: (MetaSimplifier.simpset -> MetaSimplifier.simpset) -> theory -> theory| \\
+  \indexml{Code.add\_eqnl}\verb|Code.add_eqnl: string * (thm * bool) list lazy -> theory -> theory| \\
+  \indexml{Code.map\_pre}\verb|Code.map_pre: (simpset -> simpset) -> theory -> theory| \\
+  \indexml{Code.map\_post}\verb|Code.map_post: (simpset -> simpset) -> theory -> theory| \\
   \indexml{Code.add\_functrans}\verb|Code.add_functrans: string * (theory -> (thm * bool) list -> (thm * bool) list option)|\isasep\isanewline%
 \verb|    -> theory -> theory| \\
   \indexml{Code.del\_functrans}\verb|Code.del_functrans: string -> theory -> theory| \\
@@ -75,7 +75,7 @@
      theorem \isa{thm} from executable content, if present.
 
   \item \verb|Code.add_eqnl|~\isa{{\isacharparenleft}const{\isacharcomma}\ lthms{\isacharparenright}}~\isa{thy} adds
-     suspended defining equations \isa{lthms} for constant
+     suspended code equations \isa{lthms} for constant
      \isa{const} to executable content.
 
   \item \verb|Code.map_pre|~\isa{f}~\isa{thy} changes
@@ -83,11 +83,11 @@
 
   \item \verb|Code.add_functrans|~\isa{{\isacharparenleft}name{\isacharcomma}\ f{\isacharparenright}}~\isa{thy} adds
      function transformer \isa{f} (named \isa{name}) to executable content;
-     \isa{f} is a transformer of the defining equations belonging
+     \isa{f} is a transformer of the code equations belonging
      to a certain function definition, depending on the
      current theory context.  Returning \isa{NONE} indicates that no
      transformation took place;  otherwise, the whole process will be iterated
-     with the new defining equations.
+     with the new code equations.
 
   \item \verb|Code.del_functrans|~\isa{name}~\isa{thy} removes
      function transformer named \isa{name} from executable content.
@@ -126,7 +126,7 @@
 \begin{mldecls}
   \indexml{Code\_Unit.read\_const}\verb|Code_Unit.read_const: theory -> string -> string| \\
   \indexml{Code\_Unit.head\_eqn}\verb|Code_Unit.head_eqn: theory -> thm -> string * ((string * sort) list * typ)| \\
-  \indexml{Code\_Unit.rewrite\_eqn}\verb|Code_Unit.rewrite_eqn: MetaSimplifier.simpset -> thm -> thm| \\
+  \indexml{Code\_Unit.rewrite\_eqn}\verb|Code_Unit.rewrite_eqn: simpset -> thm -> thm| \\
   \end{mldecls}
 
   \begin{description}
@@ -135,12 +135,12 @@
      reads a constant as a concrete term expression \isa{s}.
 
   \item \verb|Code_Unit.head_eqn|~\isa{thy}~\isa{thm}
-     extracts the constant and its type from a defining equation \isa{thm}.
+     extracts the constant and its type from a code equation \isa{thm}.
 
   \item \verb|Code_Unit.rewrite_eqn|~\isa{ss}~\isa{thm}
-     rewrites a defining equation \isa{thm} with a simpset \isa{ss};
+     rewrites a code equation \isa{thm} with a simpset \isa{ss};
      only arguments and right hand side are rewritten,
-     not the head of the defining equation.
+     not the head of the code equation.
 
   \end{description}%
 \end{isamarkuptext}%
--- a/doc-src/IsarAdvanced/Codegen/Thy/document/Program.tex	Mon Jan 19 20:37:08 2009 +0100
+++ b/doc-src/IsarAdvanced/Codegen/Thy/document/Program.tex	Mon Jan 19 21:20:18 2009 +0100
@@ -30,7 +30,7 @@
 We have already seen how by default equations stemming from
   \hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}/\hyperlink{command.primrec}{\mbox{\isa{\isacommand{primrec}}}}/\hyperlink{command.fun}{\mbox{\isa{\isacommand{fun}}}}
   statements are used for code generation.  This default behaviour
-  can be changed, e.g. by providing different defining equations.
+  can be changed, e.g. by providing different code equations.
   All kinds of customisation shown in this section is \emph{safe}
   in the sense that the user does not have to worry about
   correctness -- all programs generatable that way are partially
@@ -44,7 +44,7 @@
 %
 \begin{isamarkuptext}%
 Coming back to our introductory example, we
-  could provide an alternative defining equations for \isa{dequeue}
+  could provide an alternative code equations for \isa{dequeue}
   explicitly:%
 \end{isamarkuptext}%
 \isamarkuptrue%
@@ -73,7 +73,7 @@
 \begin{isamarkuptext}%
 \noindent The annotation \isa{{\isacharbrackleft}code{\isacharbrackright}} is an \isa{Isar}
   \isa{attribute} which states that the given theorems should be
-  considered as defining equations for a \isa{fun} statement --
+  considered as code equations for a \isa{fun} statement --
   the corresponding constant is determined syntactically.  The resulting code:%
 \end{isamarkuptext}%
 \isamarkuptrue%
@@ -87,10 +87,10 @@
 \begin{isamarkuptext}%
 \isatypewriter%
 \noindent%
-\hspace*{0pt}dequeue ::~forall a. Queue a -> (Maybe a, Queue a);\\
-\hspace*{0pt}dequeue (Queue xs (y :~ys)) = (Just y, Queue xs ys);\\
+\hspace*{0pt}dequeue ::~forall a.~Queue a -> (Maybe a,~Queue a);\\
+\hspace*{0pt}dequeue (Queue xs (y :~ys)) = (Just y,~Queue xs ys);\\
 \hspace*{0pt}dequeue (Queue xs []) =\\
-\hspace*{0pt} ~(if nulla xs then (Nothing, Queue [] [])\\
+\hspace*{0pt} ~(if nulla xs then (Nothing,~Queue [] [])\\
 \hspace*{0pt} ~~~else dequeue (Queue [] (rev xs)));%
 \end{isamarkuptext}%
 \isamarkuptrue%
@@ -132,13 +132,13 @@
 \endisadelimquote
 %
 \begin{isamarkuptext}%
-\noindent prints a table with \emph{all} defining equations
+\noindent prints a table with \emph{all} code equations
   for \isa{dequeue}, including
-  \emph{all} defining equations those equations depend
+  \emph{all} code equations those equations depend
   on recursively.
   
   Similarly, the \hyperlink{command.code-deps}{\mbox{\isa{\isacommand{code{\isacharunderscore}deps}}}} command shows a graph
-  visualising dependencies between defining equations.%
+  visualising dependencies between code equations.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -286,7 +286,7 @@
 \hspace*{0pt} ~neutral ::~a;\\
 \hspace*{0pt}{\char125};\\
 \hspace*{0pt}\\
-\hspace*{0pt}pow ::~forall a. (Monoid a) => Nat -> a -> a;\\
+\hspace*{0pt}pow ::~forall a.~(Monoid a) => Nat -> a -> a;\\
 \hspace*{0pt}pow Zero{\char95}nat a = neutral;\\
 \hspace*{0pt}pow (Suc n) a = mult a (pow n a);\\
 \hspace*{0pt}\\
@@ -346,7 +346,7 @@
 \hspace*{0pt}type 'a semigroup = {\char123}mult :~'a -> 'a -> 'a{\char125};\\
 \hspace*{0pt}fun mult (A{\char95}:'a semigroup) = {\char35}mult A{\char95};\\
 \hspace*{0pt}\\
-\hspace*{0pt}type 'a monoid = {\char123}Program{\char95}{\char95}semigroup{\char95}monoid :~'a semigroup, neutral :~'a{\char125};\\
+\hspace*{0pt}type 'a monoid = {\char123}Program{\char95}{\char95}semigroup{\char95}monoid :~'a semigroup,~neutral :~'a{\char125};\\
 \hspace*{0pt}fun semigroup{\char95}monoid (A{\char95}:'a monoid) = {\char35}Program{\char95}{\char95}semigroup{\char95}monoid A{\char95};\\
 \hspace*{0pt}fun neutral (A{\char95}:'a monoid) = {\char35}neutral A{\char95};\\
 \hspace*{0pt}\\
@@ -356,7 +356,7 @@
 \hspace*{0pt}fun plus{\char95}nat (Suc m) n = plus{\char95}nat m (Suc n)\\
 \hspace*{0pt} ~| plus{\char95}nat Zero{\char95}nat n = n;\\
 \hspace*{0pt}\\
-\hspace*{0pt}val neutral{\char95}nat :~nat = Suc Zero{\char95}nat;\\
+\hspace*{0pt}val neutral{\char95}nat :~nat = Suc Zero{\char95}nat\\
 \hspace*{0pt}\\
 \hspace*{0pt}fun mult{\char95}nat Zero{\char95}nat n = Zero{\char95}nat\\
 \hspace*{0pt} ~| mult{\char95}nat (Suc m) n = plus{\char95}nat n (mult{\char95}nat m n);\\
@@ -364,12 +364,12 @@
 \hspace*{0pt}val semigroup{\char95}nat = {\char123}mult = mult{\char95}nat{\char125}~:~nat semigroup;\\
 \hspace*{0pt}\\
 \hspace*{0pt}val monoid{\char95}nat =\\
-\hspace*{0pt} ~{\char123}Program{\char95}{\char95}semigroup{\char95}monoid = semigroup{\char95}nat, neutral = neutral{\char95}nat{\char125}~:\\
+\hspace*{0pt} ~{\char123}Program{\char95}{\char95}semigroup{\char95}monoid = semigroup{\char95}nat,~neutral = neutral{\char95}nat{\char125}~:\\
 \hspace*{0pt} ~nat monoid;\\
 \hspace*{0pt}\\
 \hspace*{0pt}fun bexp n = pow monoid{\char95}nat n (Suc (Suc Zero{\char95}nat));\\
 \hspace*{0pt}\\
-\hspace*{0pt}end; (*struct Example*)%
+\hspace*{0pt}end;~(*struct Example*)%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -398,7 +398,7 @@
 
   The \emph{simpset} allows to employ the full generality of the Isabelle
   simplifier.  Due to the interpretation of theorems
-  as defining equations, rewrites are applied to the right
+  as code equations, rewrites are applied to the right
   hand side and the arguments of the left hand side of an
   equation, but never to the constant heading the left hand side.
   An important special case are \emph{inline theorems} which may be
@@ -489,7 +489,7 @@
   the \hyperlink{command.print-codesetup}{\mbox{\isa{\isacommand{print{\isacharunderscore}codesetup}}}} command.
   \hyperlink{command.code-thms}{\mbox{\isa{\isacommand{code{\isacharunderscore}thms}}}} provides a convenient
   mechanism to inspect the impact of a preprocessor setup
-  on defining equations.
+  on code equations.
 
   \begin{warn}
     The attribute \emph{code unfold}
@@ -675,7 +675,7 @@
 \hspace*{0pt} ~| plus{\char95}nat m Zero{\char95}nat = m\\
 \hspace*{0pt} ~| plus{\char95}nat Zero{\char95}nat n = n;\\
 \hspace*{0pt}\\
-\hspace*{0pt}end; (*struct Example*)%
+\hspace*{0pt}end;~(*struct Example*)%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -794,7 +794,7 @@
 \hspace*{0pt} ~~~~~~~~~~~~else collect{\char95}duplicates A{\char95}~xs (z ::~ys) zs)\\
 \hspace*{0pt} ~~~~~else collect{\char95}duplicates A{\char95}~(z ::~xs) (z ::~ys) zs);\\
 \hspace*{0pt}\\
-\hspace*{0pt}end; (*struct Example*)%
+\hspace*{0pt}end;~(*struct Example*)%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -811,7 +811,7 @@
   an explicit class \isa{eq} with a corresponding operation
   \isa{eq{\isacharunderscore}class{\isachardot}eq} such that \isa{eq{\isacharunderscore}class{\isachardot}eq\ {\isacharequal}\ op\ {\isacharequal}}.
   The preprocessing framework does the rest by propagating the
-  \isa{eq} constraints through all dependent defining equations.
+  \isa{eq} constraints through all dependent code equations.
   For datatypes, instances of \isa{eq} are implicitly derived
   when possible.  For other types, you may instantiate \isa{eq}
   manually like any other type class.
@@ -918,7 +918,7 @@
 \hspace*{0pt}type 'a eq = {\char123}eq :~'a -> 'a -> bool{\char125};\\
 \hspace*{0pt}fun eq (A{\char95}:'a eq) = {\char35}eq A{\char95};\\
 \hspace*{0pt}\\
-\hspace*{0pt}type 'a ord = {\char123}less{\char95}eq :~'a -> 'a -> bool, less :~'a -> 'a -> bool{\char125};\\
+\hspace*{0pt}type 'a ord = {\char123}less{\char95}eq :~'a -> 'a -> bool,~less :~'a -> 'a -> bool{\char125};\\
 \hspace*{0pt}fun less{\char95}eq (A{\char95}:'a ord) = {\char35}less{\char95}eq A{\char95};\\
 \hspace*{0pt}fun less (A{\char95}:'a ord) = {\char35}less A{\char95};\\
 \hspace*{0pt}\\
@@ -930,16 +930,16 @@
 \hspace*{0pt}type 'a order = {\char123}Orderings{\char95}{\char95}preorder{\char95}order :~'a preorder{\char125};\\
 \hspace*{0pt}fun preorder{\char95}order (A{\char95}:'a order) = {\char35}Orderings{\char95}{\char95}preorder{\char95}order A{\char95};\\
 \hspace*{0pt}\\
-\hspace*{0pt}fun less{\char95}eqa (A1{\char95}, A2{\char95}) B{\char95}~(x1, y1) (x2, y2) =\\
+\hspace*{0pt}fun less{\char95}eqa (A1{\char95},~A2{\char95}) B{\char95}~(x1,~y1) (x2,~y2) =\\
 \hspace*{0pt} ~less ((ord{\char95}preorder o preorder{\char95}order) A2{\char95}) x1 x2 orelse\\
 \hspace*{0pt} ~~~eqop A1{\char95}~x1 x2 andalso\\
 \hspace*{0pt} ~~~~~less{\char95}eq ((ord{\char95}preorder o preorder{\char95}order) B{\char95}) y1 y2\\
-\hspace*{0pt} ~| less{\char95}eqa (A1{\char95}, A2{\char95}) B{\char95}~(x1, y1) (x2, y2) =\\
+\hspace*{0pt} ~| less{\char95}eqa (A1{\char95},~A2{\char95}) B{\char95}~(x1,~y1) (x2,~y2) =\\
 \hspace*{0pt} ~~~less ((ord{\char95}preorder o preorder{\char95}order) A2{\char95}) x1 x2 orelse\\
 \hspace*{0pt} ~~~~~eqop A1{\char95}~x1 x2 andalso\\
 \hspace*{0pt} ~~~~~~~less{\char95}eq ((ord{\char95}preorder o preorder{\char95}order) B{\char95}) y1 y2;\\
 \hspace*{0pt}\\
-\hspace*{0pt}end; (*struct Example*)%
+\hspace*{0pt}end;~(*struct Example*)%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -951,7 +951,7 @@
 \endisadelimquote
 %
 \begin{isamarkuptext}%
-In some cases, the automatically derived defining equations
+In some cases, the automatically derived code equations
   for equality on a particular type may not be appropriate.
   As example, watch the following datatype representing
   monomorphic parametric types (where type constructors
@@ -1052,10 +1052,10 @@
 \hspace*{0pt} ~| list{\char95}all2 p xs [] = null xs\\
 \hspace*{0pt} ~| list{\char95}all2 p [] ys = null ys;\\
 \hspace*{0pt}\\
-\hspace*{0pt}fun eq{\char95}monotype (Mono (tyco1, typargs1)) (Mono (tyco2, typargs2)) =\\
+\hspace*{0pt}fun eq{\char95}monotype (Mono (tyco1,~typargs1)) (Mono (tyco2,~typargs2)) =\\
 \hspace*{0pt} ~eq{\char95}nat tyco1 tyco2 andalso list{\char95}all2 eq{\char95}monotype typargs1 typargs2;\\
 \hspace*{0pt}\\
-\hspace*{0pt}end; (*struct Example*)%
+\hspace*{0pt}end;~(*struct Example*)%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -1108,12 +1108,12 @@
 \begin{isamarkuptext}%
 \isatypewriter%
 \noindent%
-\hspace*{0pt}strict{\char95}dequeue ::~forall a. Queue a -> (a, Queue a);\\
-\hspace*{0pt}strict{\char95}dequeue (Queue xs (y :~ys)) = (y, Queue xs ys);\\
+\hspace*{0pt}strict{\char95}dequeue ::~forall a.~Queue a -> (a,~Queue a);\\
+\hspace*{0pt}strict{\char95}dequeue (Queue xs (y :~ys)) = (y,~Queue xs ys);\\
 \hspace*{0pt}strict{\char95}dequeue (Queue xs []) =\\
 \hspace*{0pt} ~let {\char123}\\
 \hspace*{0pt} ~~~(y :~ys) = rev xs;\\
-\hspace*{0pt} ~{\char125}~in (y, Queue [] ys);%
+\hspace*{0pt} ~{\char125}~in (y,~Queue [] ys);%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -1165,7 +1165,7 @@
   on the right hand side of its first equation the constant
   \isa{empty{\isacharunderscore}queue} occurs which is unspecified.
 
-  Normally, if constants without any defining equations occur in
+  Normally, if constants without any code equations occur in
   a program, the code generator complains (since in most cases
   this is not what the user expects).  But such constants can also
   be thought of as function definitions with no equations which
@@ -1204,13 +1204,13 @@
 \begin{isamarkuptext}%
 \isatypewriter%
 \noindent%
-\hspace*{0pt}empty{\char95}queue ::~forall a. a;\\
+\hspace*{0pt}empty{\char95}queue ::~forall a.~a;\\
 \hspace*{0pt}empty{\char95}queue = error {\char34}empty{\char95}queue{\char34};\\
 \hspace*{0pt}\\
-\hspace*{0pt}strict{\char95}dequeue' ::~forall a. Queue a -> (a, Queue a);\\
+\hspace*{0pt}strict{\char95}dequeue' ::~forall a.~Queue a -> (a,~Queue a);\\
 \hspace*{0pt}strict{\char95}dequeue' (Queue xs []) =\\
 \hspace*{0pt} ~(if nulla xs then empty{\char95}queue else strict{\char95}dequeue' (Queue [] (rev xs)));\\
-\hspace*{0pt}strict{\char95}dequeue' (Queue xs (y :~ys)) = (y, Queue xs ys);%
+\hspace*{0pt}strict{\char95}dequeue' (Queue xs (y :~ys)) = (y,~Queue xs ys);%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
--- a/doc-src/IsarAdvanced/Codegen/Thy/examples/Example.hs	Mon Jan 19 20:37:08 2009 +0100
+++ b/doc-src/IsarAdvanced/Codegen/Thy/examples/Example.hs	Mon Jan 19 21:20:18 2009 +0100
@@ -1,3 +1,5 @@
+{-# OPTIONS_GHC -fglasgow-exts #-}
+
 module Example where {
 
 
--- a/doc-src/IsarAdvanced/Codegen/Thy/examples/example.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/doc-src/IsarAdvanced/Codegen/Thy/examples/example.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -11,7 +11,7 @@
 
 datatype 'a queue = Queue of 'a list * 'a list;
 
-val empty : 'a queue = Queue ([], []);
+val empty : 'a queue = Queue ([], [])
 
 fun dequeue (Queue ([], [])) = (NONE, Queue ([], []))
   | dequeue (Queue (xs, y :: ys)) = (SOME y, Queue (xs, ys))
--- a/doc-src/IsarAdvanced/Functions/Thy/document/Functions.tex	Mon Jan 19 20:37:08 2009 +0100
+++ b/doc-src/IsarAdvanced/Functions/Thy/document/Functions.tex	Mon Jan 19 21:20:18 2009 +0100
@@ -440,7 +440,7 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-The the key to this error message is the matrix at the bottom. The rows
+The key to this error message is the matrix at the bottom. The rows
   of that matrix correspond to the different recursive calls (In our
   case, there is just one). The columns are the function's arguments 
   (expressed through different measure functions, which map the
@@ -674,7 +674,7 @@
 {\isacharbar}\ {\isachardoublequoteopen}And\ X\ X\ {\isacharequal}\ X{\isachardoublequoteclose}%
 \begin{isamarkuptext}%
 This definition is useful, because the equations can directly be used
-  as simplification rules rules. But the patterns overlap: For example,
+  as simplification rules. But the patterns overlap: For example,
   the expression \isa{And\ T\ T} is matched by both the first and
   the second equation. By default, Isabelle makes the patterns disjoint by
   splitting them up, producing instances:%
@@ -829,13 +829,21 @@
   either \isa{{\isadigit{0}}}, \isa{{\isadigit{1}}} or \isa{n\ {\isacharplus}\ {\isadigit{2}}}:
 
   \begin{isabelle}%
-\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}P\ x{\isachardot}\ {\isasymlbrakk}x\ {\isacharequal}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\ x\ {\isacharequal}\ {\isadigit{1}}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\ {\isasymAnd}n{\isachardot}\ x\ {\isacharequal}\ n\ {\isacharplus}\ {\isadigit{2}}\ {\isasymLongrightarrow}\ P{\isasymrbrakk}\ {\isasymLongrightarrow}\ P%
+\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}P\ x{\isachardot}\ {\isasymlbrakk}x\ {\isacharequal}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\ x\ {\isacharequal}\ {\isadigit{1}}\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\ {\isasymAnd}n{\isachardot}\ x\ {\isacharequal}\ n\ {\isacharplus}\ {\isadigit{2}}\ {\isasymLongrightarrow}\ P{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\isanewline
+\ {\isadigit{2}}{\isachardot}\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ {\isadigit{1}}\ {\isacharequal}\ {\isadigit{1}}\isanewline
+\ {\isadigit{3}}{\isachardot}\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}\ {\isasymLongrightarrow}\ {\isadigit{1}}\ {\isacharequal}\ {\isadigit{1}}\isanewline
+\ {\isadigit{4}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ {\isadigit{0}}\ {\isacharequal}\ n\ {\isacharplus}\ {\isadigit{2}}\ {\isasymLongrightarrow}\ {\isadigit{1}}\ {\isacharequal}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ n\ {\isacharplus}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ {\isacharparenleft}Suc\ n{\isacharparenright}\isanewline
+\ {\isadigit{5}}{\isachardot}\ {\isadigit{1}}\ {\isacharequal}\ {\isadigit{1}}\ {\isasymLongrightarrow}\ {\isadigit{1}}\ {\isacharequal}\ {\isadigit{1}}\isanewline
+\ {\isadigit{6}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ {\isadigit{1}}\ {\isacharequal}\ n\ {\isacharplus}\ {\isadigit{2}}\ {\isasymLongrightarrow}\ {\isadigit{1}}\ {\isacharequal}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ n\ {\isacharplus}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ {\isacharparenleft}Suc\ n{\isacharparenright}\isanewline
+\ {\isadigit{7}}{\isachardot}\ {\isasymAnd}n\ na{\isachardot}\isanewline
+\isaindent{\ {\isadigit{7}}{\isachardot}\ \ \ \ }n\ {\isacharplus}\ {\isadigit{2}}\ {\isacharequal}\ na\ {\isacharplus}\ {\isadigit{2}}\ {\isasymLongrightarrow}\isanewline
+\isaindent{\ {\isadigit{7}}{\isachardot}\ \ \ \ }fib{\isadigit{2}}{\isacharunderscore}sumC\ n\ {\isacharplus}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ na\ {\isacharplus}\ fib{\isadigit{2}}{\isacharunderscore}sumC\ {\isacharparenleft}Suc\ na{\isacharparenright}%
 \end{isabelle}
 
   This is an arithmetic triviality, but unfortunately the
   \isa{arith} method cannot handle this specific form of an
   elimination rule. However, we can use the method \isa{atomize{\isacharunderscore}elim} to do an ad-hoc conversion to a disjunction of
-  existentials, which can then be soved by the arithmetic decision procedure.
+  existentials, which can then be solved by the arithmetic decision procedure.
   Pattern compatibility and termination are automatic as usual.%
 \end{isamarkuptxt}%
 \isamarkuptrue%
--- a/doc-src/IsarOverview/Isar/Logic.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/doc-src/IsarOverview/Isar/Logic.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -34,54 +34,51 @@
 be enclosed in double quotes. However, we will continue to do so for
 uniformity.
 
-Trivial proofs, in particular those by assumption, should be trivial
-to perform. Proof ``.'' does just that (and a bit more). Thus
-naming of assumptions is often superfluous: *}
+Instead of applying fact @{text a} via the @{text rule} method, we can
+also push it directly onto our goal.  The proof is then immediate,
+which is formally written as ``.'' in Isar: *}
 lemma "A \<longrightarrow> A"
 proof
-  assume "A"
-  show "A" .
+  assume a: "A"
+  from a show "A" .
 qed
 
-text{* To hide proofs by assumption further, \isakeyword{by}@{text"(method)"}
-first applies @{text method} and then tries to solve all remaining subgoals
-by assumption: *}
+text{* We can also push several facts towards a goal, and put another
+rule in between to establish some result that is one step further
+removed.  We illustrate this by introducing a trivial conjunction: *}
 lemma "A \<longrightarrow> A \<and> A"
 proof
-  assume "A"
-  show "A \<and> A" by(rule conjI)
+  assume a: "A"
+  from a and a show "A \<and> A" by(rule conjI)
 qed
 text{*\noindent Rule @{thm[source]conjI} is of course @{thm conjI}.
-A drawback of implicit proofs by assumption is that it
-is no longer obvious where an assumption is used.
 
 Proofs of the form \isakeyword{by}@{text"(rule"}~\emph{name}@{text")"}
-can be abbreviated to ``..''  if \emph{name} refers to one of the
+can be abbreviated to ``..'' if \emph{name} refers to one of the
 predefined introduction rules (or elimination rules, see below): *}
 
 lemma "A \<longrightarrow> A \<and> A"
 proof
-  assume "A"
-  show "A \<and> A" ..
+  assume a: "A"
+  from a and a show "A \<and> A" ..
 qed
 text{*\noindent
 This is what happens: first the matching introduction rule @{thm[source]conjI}
-is applied (first ``.''), then the two subgoals are solved by assumption
-(second ``.''). *}
+is applied (first ``.''), the remaining problem is solved immediately (second ``.''). *}
 
 subsubsection{*Elimination rules*}
 
 text{*A typical elimination rule is @{thm[source]conjE}, $\land$-elimination:
 @{thm[display,indent=5]conjE}  In the following proof it is applied
 by hand, after its first (\emph{major}) premise has been eliminated via
-@{text"[OF AB]"}: *}
+@{text"[OF ab]"}: *}
 lemma "A \<and> B \<longrightarrow> B \<and> A"
 proof
-  assume AB: "A \<and> B"
+  assume ab: "A \<and> B"
   show "B \<and> A"
-  proof (rule conjE[OF AB])  -- {*@{text"conjE[OF AB]"}: @{thm conjE[OF AB]} *}
-    assume "A" "B"
-    show ?thesis ..
+  proof (rule conjE[OF ab])  -- {*@{text"conjE[OF ab]"}: @{thm conjE[OF ab]} *}
+    assume a: "A" and b: "B"
+    from b and a show ?thesis ..
   qed
 qed
 text{*\noindent Note that the term @{text"?thesis"} always stands for the
@@ -106,11 +103,11 @@
 
 lemma "A \<and> B \<longrightarrow> B \<and> A"
 proof
-  assume AB: "A \<and> B"
-  from AB show "B \<and> A"
+  assume ab: "A \<and> B"
+  from ab show "B \<and> A"
   proof
-    assume "A" "B"
-    show ?thesis ..
+    assume a: "A" and b: "B"
+    from b and a show ?thesis ..
   qed
 qed
 
@@ -120,15 +117,16 @@
 such that the proof of each proposition builds on the previous proposition.
 \end{quote}
 The previous proposition can be referred to via the fact @{text this}.
-This greatly reduces the need for explicit naming of propositions:
+This greatly reduces the need for explicit naming of propositions.  We also
+rearrange the additional inner assumptions into proper order for immediate use:
 *}
 lemma "A \<and> B \<longrightarrow> B \<and> A"
 proof
   assume "A \<and> B"
   from this show "B \<and> A"
   proof
-    assume "A" "B"
-    show ?thesis ..
+    assume "B" "A"
+    from this show ?thesis ..
   qed
 qed
 
@@ -199,11 +197,11 @@
     assume nn: "\<not> (\<not> A \<or> \<not> B)"
     have "\<not> A"
     proof
-      assume "A"
+      assume a: "A"
       have "\<not> B"
       proof
-        assume "B"
-        have "A \<and> B" ..
+        assume b: "B"
+        from a and b have "A \<and> B" ..
         with n show False ..
       qed
       hence "\<not> A \<or> \<not> B" ..
@@ -282,28 +280,28 @@
 \isakeyword{assumes} and \isakeyword{shows} elements which allow direct
 naming of assumptions: *}
 
-lemma assumes AB: "large_A \<and> large_B"
+lemma assumes ab: "large_A \<and> large_B"
   shows "large_B \<and> large_A" (is "?B \<and> ?A")
 proof
-  from AB show "?B" ..
+  from ab show "?B" ..
 next
-  from AB show "?A" ..
+  from ab show "?A" ..
 qed
 text{*\noindent Note the difference between @{text ?AB}, a term, and
-@{text AB}, a fact.
+@{text ab}, a fact.
 
 Finally we want to start the proof with $\land$-elimination so we
 don't have to perform it twice, as above. Here is a slick way to
 achieve this: *}
 
-lemma assumes AB: "large_A \<and> large_B"
+lemma assumes ab: "large_A \<and> large_B"
   shows "large_B \<and> large_A" (is "?B \<and> ?A")
-using AB
+using ab
 proof
-  assume "?A" "?B" show ?thesis ..
+  assume "?B" "?A" thus ?thesis ..
 qed
 text{*\noindent Command \isakeyword{using} can appear before a proof
-and adds further facts to those piped into the proof. Here @{text AB}
+and adds further facts to those piped into the proof. Here @{text ab}
 is the only such fact and it triggers $\land$-elimination. Another
 frequent idiom is as follows:
 \begin{center}
@@ -319,23 +317,23 @@
 not be what we had in mind.
 A simple ``@{text"-"}'' prevents this \emph{faux pas}: *}
 
-lemma assumes AB: "A \<or> B" shows "B \<or> A"
+lemma assumes ab: "A \<or> B" shows "B \<or> A"
 proof -
-  from AB show ?thesis
+  from ab show ?thesis
   proof
-    assume A show ?thesis ..
+    assume A thus ?thesis ..
   next
-    assume B show ?thesis ..
+    assume B thus ?thesis ..
   qed
 qed
 text{*\noindent Alternatively one can feed @{prop"A \<or> B"} directly
 into the proof, thus triggering the elimination rule: *}
-lemma assumes AB: "A \<or> B" shows "B \<or> A"
-using AB
+lemma assumes ab: "A \<or> B" shows "B \<or> A"
+using ab
 proof
-  assume A show ?thesis ..
+  assume A thus ?thesis ..
 next
-  assume B show ?thesis ..
+  assume B thus ?thesis ..
 qed
 text{* \noindent Remember that eliminations have priority over
 introductions.
@@ -416,7 +414,7 @@
   proof              -- "@{thm[source]exE}: @{thm exE}"
     fix x
     assume "P(f x)"
-    show ?thesis ..  -- "@{thm[source]exI}: @{thm exI}"
+    thus ?thesis ..  -- "@{thm[source]exI}: @{thm exI}"
   qed
 qed
 text{*\noindent Explicit $\exists$-elimination as seen above can become
@@ -499,12 +497,12 @@
       assume "y \<in> ?S"
       hence "y \<notin> f y"   by simp
       hence "y \<notin> ?S"    by(simp add: `?S = f y`)
-      thus False         by contradiction
+      with `y \<in> ?S` show False by contradiction
     next
       assume "y \<notin> ?S"
       hence "y \<in> f y"   by simp
       hence "y \<in> ?S"    by(simp add: `?S = f y`)
-      thus False         by contradiction
+      with `y \<notin> ?S` show False by contradiction
     qed
   qed
 qed
--- a/doc-src/IsarOverview/Isar/document/Logic.tex	Mon Jan 19 20:37:08 2009 +0100
+++ b/doc-src/IsarOverview/Isar/document/Logic.tex	Mon Jan 19 21:20:18 2009 +0100
@@ -97,9 +97,9 @@
 be enclosed in double quotes. However, we will continue to do so for
 uniformity.
 
-Trivial proofs, in particular those by assumption, should be trivial
-to perform. Proof ``.'' does just that (and a bit more). Thus
-naming of assumptions is often superfluous:%
+Instead of applying fact \isa{a} via the \isa{rule} method, we can
+also push it directly onto our goal.  The proof is then immediate,
+which is formally written as ``.'' in Isar:%
 \end{isamarkuptext}%
 \isamarkuptrue%
 \isacommand{lemma}\isamarkupfalse%
@@ -113,8 +113,9 @@
 \isacommand{proof}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{assume}\isamarkupfalse%
-\ {\isachardoublequoteopen}A{\isachardoublequoteclose}\isanewline
-\ \ \isacommand{show}\isamarkupfalse%
+\ a{\isacharcolon}\ {\isachardoublequoteopen}A{\isachardoublequoteclose}\isanewline
+\ \ \isacommand{from}\isamarkupfalse%
+\ a\ \isacommand{show}\isamarkupfalse%
 \ {\isachardoublequoteopen}A{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse%
 \isanewline
 \isacommand{qed}\isamarkupfalse%
@@ -127,9 +128,9 @@
 \endisadelimproof
 %
 \begin{isamarkuptext}%
-To hide proofs by assumption further, \isakeyword{by}\isa{{\isacharparenleft}method{\isacharparenright}}
-first applies \isa{method} and then tries to solve all remaining subgoals
-by assumption:%
+We can also push several facts towards a goal, and put another
+rule in between to establish some result that is one step further
+removed.  We illustrate this by introducing a trivial conjunction:%
 \end{isamarkuptext}%
 \isamarkuptrue%
 \isacommand{lemma}\isamarkupfalse%
@@ -143,8 +144,9 @@
 \isacommand{proof}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{assume}\isamarkupfalse%
-\ {\isachardoublequoteopen}A{\isachardoublequoteclose}\isanewline
-\ \ \isacommand{show}\isamarkupfalse%
+\ a{\isacharcolon}\ {\isachardoublequoteopen}A{\isachardoublequoteclose}\isanewline
+\ \ \isacommand{from}\isamarkupfalse%
+\ a\ \isakeyword{and}\ a\ \isacommand{show}\isamarkupfalse%
 \ {\isachardoublequoteopen}A\ {\isasymand}\ A{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
 {\isacharparenleft}rule\ conjI{\isacharparenright}\isanewline
 \isacommand{qed}\isamarkupfalse%
@@ -158,11 +160,9 @@
 %
 \begin{isamarkuptext}%
 \noindent Rule \isa{conjI} is of course \isa{{\isasymlbrakk}{\isacharquery}P{\isacharsemicolon}\ {\isacharquery}Q{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}P\ {\isasymand}\ {\isacharquery}Q}.
-A drawback of implicit proofs by assumption is that it
-is no longer obvious where an assumption is used.
 
 Proofs of the form \isakeyword{by}\isa{{\isacharparenleft}rule}~\emph{name}\isa{{\isacharparenright}}
-can be abbreviated to ``..''  if \emph{name} refers to one of the
+can be abbreviated to ``..'' if \emph{name} refers to one of the
 predefined introduction rules (or elimination rules, see below):%
 \end{isamarkuptext}%
 \isamarkuptrue%
@@ -177,8 +177,9 @@
 \isacommand{proof}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{assume}\isamarkupfalse%
-\ {\isachardoublequoteopen}A{\isachardoublequoteclose}\isanewline
-\ \ \isacommand{show}\isamarkupfalse%
+\ a{\isacharcolon}\ {\isachardoublequoteopen}A{\isachardoublequoteclose}\isanewline
+\ \ \isacommand{from}\isamarkupfalse%
+\ a\ \isakeyword{and}\ a\ \isacommand{show}\isamarkupfalse%
 \ {\isachardoublequoteopen}A\ {\isasymand}\ A{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
 \isanewline
 \isacommand{qed}\isamarkupfalse%
@@ -193,8 +194,7 @@
 \begin{isamarkuptext}%
 \noindent
 This is what happens: first the matching introduction rule \isa{conjI}
-is applied (first ``.''), then the two subgoals are solved by assumption
-(second ``.'').%
+is applied (first ``.''), the remaining problem is solved immediately (second ``.'').%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -208,7 +208,7 @@
 \ \ \ \ \ {\isasymlbrakk}{\isacharquery}P\ {\isasymand}\ {\isacharquery}Q{\isacharsemicolon}\ {\isasymlbrakk}{\isacharquery}P{\isacharsemicolon}\ {\isacharquery}Q{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}R{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}R%
 \end{isabelle}  In the following proof it is applied
 by hand, after its first (\emph{major}) premise has been eliminated via
-\isa{{\isacharbrackleft}OF\ AB{\isacharbrackright}}:%
+\isa{{\isacharbrackleft}OF\ ab{\isacharbrackright}}:%
 \end{isamarkuptext}%
 \isamarkuptrue%
 \isacommand{lemma}\isamarkupfalse%
@@ -222,17 +222,18 @@
 \isacommand{proof}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{assume}\isamarkupfalse%
-\ AB{\isacharcolon}\ {\isachardoublequoteopen}A\ {\isasymand}\ B{\isachardoublequoteclose}\isanewline
+\ ab{\isacharcolon}\ {\isachardoublequoteopen}A\ {\isasymand}\ B{\isachardoublequoteclose}\isanewline
 \ \ \isacommand{show}\isamarkupfalse%
 \ {\isachardoublequoteopen}B\ {\isasymand}\ A{\isachardoublequoteclose}\isanewline
 \ \ \isacommand{proof}\isamarkupfalse%
-\ {\isacharparenleft}rule\ conjE{\isacharbrackleft}OF\ AB{\isacharbrackright}{\isacharparenright}\ \ %
-\isamarkupcmt{\isa{conjE{\isacharbrackleft}OF\ AB{\isacharbrackright}}: \isa{{\isacharparenleft}{\isasymlbrakk}A{\isacharsemicolon}\ B{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}R{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharquery}R}%
+\ {\isacharparenleft}rule\ conjE{\isacharbrackleft}OF\ ab{\isacharbrackright}{\isacharparenright}\ \ %
+\isamarkupcmt{\isa{conjE{\isacharbrackleft}OF\ ab{\isacharbrackright}}: \isa{{\isacharparenleft}{\isasymlbrakk}A{\isacharsemicolon}\ B{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}R{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharquery}R}%
 }
 \isanewline
 \ \ \ \ \isacommand{assume}\isamarkupfalse%
-\ {\isachardoublequoteopen}A{\isachardoublequoteclose}\ {\isachardoublequoteopen}B{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isacommand{show}\isamarkupfalse%
+\ a{\isacharcolon}\ {\isachardoublequoteopen}A{\isachardoublequoteclose}\ \isakeyword{and}\ b{\isacharcolon}\ {\isachardoublequoteopen}B{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isacommand{from}\isamarkupfalse%
+\ b\ \isakeyword{and}\ a\ \isacommand{show}\isamarkupfalse%
 \ {\isacharquery}thesis\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{qed}\isamarkupfalse%
@@ -279,15 +280,16 @@
 \isacommand{proof}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{assume}\isamarkupfalse%
-\ AB{\isacharcolon}\ {\isachardoublequoteopen}A\ {\isasymand}\ B{\isachardoublequoteclose}\isanewline
+\ ab{\isacharcolon}\ {\isachardoublequoteopen}A\ {\isasymand}\ B{\isachardoublequoteclose}\isanewline
 \ \ \isacommand{from}\isamarkupfalse%
-\ AB\ \isacommand{show}\isamarkupfalse%
+\ ab\ \isacommand{show}\isamarkupfalse%
 \ {\isachardoublequoteopen}B\ {\isasymand}\ A{\isachardoublequoteclose}\isanewline
 \ \ \isacommand{proof}\isamarkupfalse%
 \isanewline
 \ \ \ \ \isacommand{assume}\isamarkupfalse%
-\ {\isachardoublequoteopen}A{\isachardoublequoteclose}\ {\isachardoublequoteopen}B{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isacommand{show}\isamarkupfalse%
+\ a{\isacharcolon}\ {\isachardoublequoteopen}A{\isachardoublequoteclose}\ \isakeyword{and}\ b{\isacharcolon}\ {\isachardoublequoteopen}B{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isacommand{from}\isamarkupfalse%
+\ b\ \isakeyword{and}\ a\ \isacommand{show}\isamarkupfalse%
 \ {\isacharquery}thesis\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{qed}\isamarkupfalse%
@@ -308,7 +310,8 @@
 such that the proof of each proposition builds on the previous proposition.
 \end{quote}
 The previous proposition can be referred to via the fact \isa{this}.
-This greatly reduces the need for explicit naming of propositions:%
+This greatly reduces the need for explicit naming of propositions.  We also
+rearrange the additional inner assumptions into proper order for immediate use:%
 \end{isamarkuptext}%
 \isamarkuptrue%
 \isacommand{lemma}\isamarkupfalse%
@@ -329,8 +332,9 @@
 \ \ \isacommand{proof}\isamarkupfalse%
 \isanewline
 \ \ \ \ \isacommand{assume}\isamarkupfalse%
-\ {\isachardoublequoteopen}A{\isachardoublequoteclose}\ {\isachardoublequoteopen}B{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isacommand{show}\isamarkupfalse%
+\ {\isachardoublequoteopen}B{\isachardoublequoteclose}\ {\isachardoublequoteopen}A{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isacommand{from}\isamarkupfalse%
+\ this\ \isacommand{show}\isamarkupfalse%
 \ {\isacharquery}thesis\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{qed}\isamarkupfalse%
@@ -455,14 +459,15 @@
 \ \ \ \ \isacommand{proof}\isamarkupfalse%
 \isanewline
 \ \ \ \ \ \ \isacommand{assume}\isamarkupfalse%
-\ {\isachardoublequoteopen}A{\isachardoublequoteclose}\isanewline
+\ a{\isacharcolon}\ {\isachardoublequoteopen}A{\isachardoublequoteclose}\isanewline
 \ \ \ \ \ \ \isacommand{have}\isamarkupfalse%
 \ {\isachardoublequoteopen}{\isasymnot}\ B{\isachardoublequoteclose}\isanewline
 \ \ \ \ \ \ \isacommand{proof}\isamarkupfalse%
 \isanewline
 \ \ \ \ \ \ \ \ \isacommand{assume}\isamarkupfalse%
-\ {\isachardoublequoteopen}B{\isachardoublequoteclose}\isanewline
-\ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse%
+\ b{\isacharcolon}\ {\isachardoublequoteopen}B{\isachardoublequoteclose}\isanewline
+\ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse%
+\ a\ \isakeyword{and}\ b\ \isacommand{have}\isamarkupfalse%
 \ {\isachardoublequoteopen}A\ {\isasymand}\ B{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
 \isanewline
 \ \ \ \ \ \ \ \ \isacommand{with}\isamarkupfalse%
@@ -622,7 +627,7 @@
 \end{isamarkuptext}%
 \isamarkuptrue%
 \isacommand{lemma}\isamarkupfalse%
-\ \isakeyword{assumes}\ AB{\isacharcolon}\ {\isachardoublequoteopen}large{\isacharunderscore}A\ {\isasymand}\ large{\isacharunderscore}B{\isachardoublequoteclose}\isanewline
+\ \isakeyword{assumes}\ ab{\isacharcolon}\ {\isachardoublequoteopen}large{\isacharunderscore}A\ {\isasymand}\ large{\isacharunderscore}B{\isachardoublequoteclose}\isanewline
 \ \ \isakeyword{shows}\ {\isachardoublequoteopen}large{\isacharunderscore}B\ {\isasymand}\ large{\isacharunderscore}A{\isachardoublequoteclose}\ {\isacharparenleft}\isakeyword{is}\ {\isachardoublequoteopen}{\isacharquery}B\ {\isasymand}\ {\isacharquery}A{\isachardoublequoteclose}{\isacharparenright}\isanewline
 %
 \isadelimproof
@@ -633,13 +638,13 @@
 \isacommand{proof}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{from}\isamarkupfalse%
-\ AB\ \isacommand{show}\isamarkupfalse%
+\ ab\ \isacommand{show}\isamarkupfalse%
 \ {\isachardoublequoteopen}{\isacharquery}B{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
 \isanewline
 \isacommand{next}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{from}\isamarkupfalse%
-\ AB\ \isacommand{show}\isamarkupfalse%
+\ ab\ \isacommand{show}\isamarkupfalse%
 \ {\isachardoublequoteopen}{\isacharquery}A{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
 \isanewline
 \isacommand{qed}\isamarkupfalse%
@@ -653,7 +658,7 @@
 %
 \begin{isamarkuptext}%
 \noindent Note the difference between \isa{{\isacharquery}AB}, a term, and
-\isa{AB}, a fact.
+\isa{ab}, a fact.
 
 Finally we want to start the proof with $\land$-elimination so we
 don't have to perform it twice, as above. Here is a slick way to
@@ -661,7 +666,7 @@
 \end{isamarkuptext}%
 \isamarkuptrue%
 \isacommand{lemma}\isamarkupfalse%
-\ \isakeyword{assumes}\ AB{\isacharcolon}\ {\isachardoublequoteopen}large{\isacharunderscore}A\ {\isasymand}\ large{\isacharunderscore}B{\isachardoublequoteclose}\isanewline
+\ \isakeyword{assumes}\ ab{\isacharcolon}\ {\isachardoublequoteopen}large{\isacharunderscore}A\ {\isasymand}\ large{\isacharunderscore}B{\isachardoublequoteclose}\isanewline
 \ \ \isakeyword{shows}\ {\isachardoublequoteopen}large{\isacharunderscore}B\ {\isasymand}\ large{\isacharunderscore}A{\isachardoublequoteclose}\ {\isacharparenleft}\isakeyword{is}\ {\isachardoublequoteopen}{\isacharquery}B\ {\isasymand}\ {\isacharquery}A{\isachardoublequoteclose}{\isacharparenright}\isanewline
 %
 \isadelimproof
@@ -670,11 +675,11 @@
 %
 \isatagproof
 \isacommand{using}\isamarkupfalse%
-\ AB\isanewline
+\ ab\isanewline
 \isacommand{proof}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{assume}\isamarkupfalse%
-\ {\isachardoublequoteopen}{\isacharquery}A{\isachardoublequoteclose}\ {\isachardoublequoteopen}{\isacharquery}B{\isachardoublequoteclose}\ \isacommand{show}\isamarkupfalse%
+\ {\isachardoublequoteopen}{\isacharquery}B{\isachardoublequoteclose}\ {\isachardoublequoteopen}{\isacharquery}A{\isachardoublequoteclose}\ \isacommand{thus}\isamarkupfalse%
 \ {\isacharquery}thesis\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
 \isanewline
 \isacommand{qed}\isamarkupfalse%
@@ -688,7 +693,7 @@
 %
 \begin{isamarkuptext}%
 \noindent Command \isakeyword{using} can appear before a proof
-and adds further facts to those piped into the proof. Here \isa{AB}
+and adds further facts to those piped into the proof. Here \isa{ab}
 is the only such fact and it triggers $\land$-elimination. Another
 frequent idiom is as follows:
 \begin{center}
@@ -706,7 +711,7 @@
 \end{isamarkuptext}%
 \isamarkuptrue%
 \isacommand{lemma}\isamarkupfalse%
-\ \isakeyword{assumes}\ AB{\isacharcolon}\ {\isachardoublequoteopen}A\ {\isasymor}\ B{\isachardoublequoteclose}\ \isakeyword{shows}\ {\isachardoublequoteopen}B\ {\isasymor}\ A{\isachardoublequoteclose}\isanewline
+\ \isakeyword{assumes}\ ab{\isacharcolon}\ {\isachardoublequoteopen}A\ {\isasymor}\ B{\isachardoublequoteclose}\ \isakeyword{shows}\ {\isachardoublequoteopen}B\ {\isasymor}\ A{\isachardoublequoteclose}\isanewline
 %
 \isadelimproof
 %
@@ -716,18 +721,18 @@
 \isacommand{proof}\isamarkupfalse%
 \ {\isacharminus}\isanewline
 \ \ \isacommand{from}\isamarkupfalse%
-\ AB\ \isacommand{show}\isamarkupfalse%
+\ ab\ \isacommand{show}\isamarkupfalse%
 \ {\isacharquery}thesis\isanewline
 \ \ \isacommand{proof}\isamarkupfalse%
 \isanewline
 \ \ \ \ \isacommand{assume}\isamarkupfalse%
-\ A\ \isacommand{show}\isamarkupfalse%
+\ A\ \isacommand{thus}\isamarkupfalse%
 \ {\isacharquery}thesis\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{next}\isamarkupfalse%
 \isanewline
 \ \ \ \ \isacommand{assume}\isamarkupfalse%
-\ B\ \isacommand{show}\isamarkupfalse%
+\ B\ \isacommand{thus}\isamarkupfalse%
 \ {\isacharquery}thesis\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{qed}\isamarkupfalse%
@@ -747,7 +752,7 @@
 \end{isamarkuptext}%
 \isamarkuptrue%
 \isacommand{lemma}\isamarkupfalse%
-\ \isakeyword{assumes}\ AB{\isacharcolon}\ {\isachardoublequoteopen}A\ {\isasymor}\ B{\isachardoublequoteclose}\ \isakeyword{shows}\ {\isachardoublequoteopen}B\ {\isasymor}\ A{\isachardoublequoteclose}\isanewline
+\ \isakeyword{assumes}\ ab{\isacharcolon}\ {\isachardoublequoteopen}A\ {\isasymor}\ B{\isachardoublequoteclose}\ \isakeyword{shows}\ {\isachardoublequoteopen}B\ {\isasymor}\ A{\isachardoublequoteclose}\isanewline
 %
 \isadelimproof
 %
@@ -755,17 +760,17 @@
 %
 \isatagproof
 \isacommand{using}\isamarkupfalse%
-\ AB\isanewline
+\ ab\isanewline
 \isacommand{proof}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{assume}\isamarkupfalse%
-\ A\ \isacommand{show}\isamarkupfalse%
+\ A\ \isacommand{thus}\isamarkupfalse%
 \ {\isacharquery}thesis\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
 \isanewline
 \isacommand{next}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{assume}\isamarkupfalse%
-\ B\ \isacommand{show}\isamarkupfalse%
+\ B\ \isacommand{thus}\isamarkupfalse%
 \ {\isacharquery}thesis\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
 \isanewline
 \isacommand{qed}\isamarkupfalse%
@@ -945,7 +950,7 @@
 \ x\isanewline
 \ \ \ \ \isacommand{assume}\isamarkupfalse%
 \ {\isachardoublequoteopen}P{\isacharparenleft}f\ x{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isacommand{show}\isamarkupfalse%
+\ \ \ \ \isacommand{thus}\isamarkupfalse%
 \ {\isacharquery}thesis\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
 \ \ %
 \isamarkupcmt{\isa{exI}: \isa{{\isacharquery}P\ {\isacharquery}x\ {\isasymLongrightarrow}\ {\isasymexists}x{\isachardot}\ {\isacharquery}P\ x}%
@@ -1155,8 +1160,9 @@
 \ \ \ \ \ \ \isacommand{hence}\isamarkupfalse%
 \ {\isachardoublequoteopen}y\ {\isasymnotin}\ {\isacharquery}S{\isachardoublequoteclose}\ \ \ \ \isacommand{by}\isamarkupfalse%
 {\isacharparenleft}simp\ add{\isacharcolon}\ {\isacharbackquoteopen}{\isacharquery}S\ {\isacharequal}\ f\ y{\isacharbackquoteclose}{\isacharparenright}\isanewline
-\ \ \ \ \ \ \isacommand{thus}\isamarkupfalse%
-\ False\ \ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
+\ \ \ \ \ \ \isacommand{with}\isamarkupfalse%
+\ {\isacharbackquoteopen}y\ {\isasymin}\ {\isacharquery}S{\isacharbackquoteclose}\ \isacommand{show}\isamarkupfalse%
+\ False\ \isacommand{by}\isamarkupfalse%
 \ contradiction\isanewline
 \ \ \ \ \isacommand{next}\isamarkupfalse%
 \isanewline
@@ -1168,8 +1174,9 @@
 \ \ \ \ \ \ \isacommand{hence}\isamarkupfalse%
 \ {\isachardoublequoteopen}y\ {\isasymin}\ {\isacharquery}S{\isachardoublequoteclose}\ \ \ \ \isacommand{by}\isamarkupfalse%
 {\isacharparenleft}simp\ add{\isacharcolon}\ {\isacharbackquoteopen}{\isacharquery}S\ {\isacharequal}\ f\ y{\isacharbackquoteclose}{\isacharparenright}\isanewline
-\ \ \ \ \ \ \isacommand{thus}\isamarkupfalse%
-\ False\ \ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
+\ \ \ \ \ \ \isacommand{with}\isamarkupfalse%
+\ {\isacharbackquoteopen}y\ {\isasymnotin}\ {\isacharquery}S{\isacharbackquoteclose}\ \isacommand{show}\isamarkupfalse%
+\ False\ \isacommand{by}\isamarkupfalse%
 \ contradiction\isanewline
 \ \ \ \ \isacommand{qed}\isamarkupfalse%
 \isanewline
--- a/doc-src/IsarRef/Thy/HOL_Specific.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/doc-src/IsarRef/Thy/HOL_Specific.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,3 @@
-(* $Id$ *)
-
 theory HOL_Specific
 imports Main
 begin
@@ -1163,21 +1161,21 @@
   module name onto another.
 
   \item @{command (HOL) "code_abort"} declares constants which are not
-  required to have a definition by means of defining equations; if
+  required to have a definition by means of code equations; if
   needed these are implemented by program abort instead.
 
   \item @{attribute (HOL) code} explicitly selects (or with option
-  ``@{text "del"}'' deselects) a defining equation for code
-  generation.  Usually packages introducing defining equations provide
+  ``@{text "del"}'' deselects) a code equation for code
+  generation.  Usually packages introducing code equations provide
   a reasonable default setup for selection.
 
   \item @{attribute (HOL) code}~@{text inline} declares (or with
   option ``@{text "del"}'' removes) inlining theorems which are
-  applied as rewrite rules to any defining equation during
+  applied as rewrite rules to any code equation during
   preprocessing.
 
   \item @{command (HOL) "print_codesetup"} gives an overview on
-  selected defining equations, code generator datatypes and
+  selected code equations, code generator datatypes and
   preprocessor setup.
 
   \end{description}
--- a/doc-src/IsarRef/Thy/Spec.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/doc-src/IsarRef/Thy/Spec.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,3 @@
-(* $Id$ *)
-
 theory Spec
 imports Main
 begin
--- a/doc-src/IsarRef/Thy/document/HOL_Specific.tex	Mon Jan 19 20:37:08 2009 +0100
+++ b/doc-src/IsarRef/Thy/document/HOL_Specific.tex	Mon Jan 19 21:20:18 2009 +0100
@@ -3,8 +3,6 @@
 \def\isabellecontext{HOL{\isacharunderscore}Specific}%
 %
 \isadelimtheory
-\isanewline
-\isanewline
 %
 \endisadelimtheory
 %
@@ -1166,21 +1164,21 @@
   module name onto another.
 
   \item \hyperlink{command.HOL.code-abort}{\mbox{\isa{\isacommand{code{\isacharunderscore}abort}}}} declares constants which are not
-  required to have a definition by means of defining equations; if
+  required to have a definition by means of code equations; if
   needed these are implemented by program abort instead.
 
   \item \hyperlink{attribute.HOL.code}{\mbox{\isa{code}}} explicitly selects (or with option
-  ``\isa{{\isachardoublequote}del{\isachardoublequote}}'' deselects) a defining equation for code
-  generation.  Usually packages introducing defining equations provide
+  ``\isa{{\isachardoublequote}del{\isachardoublequote}}'' deselects) a code equation for code
+  generation.  Usually packages introducing code equations provide
   a reasonable default setup for selection.
 
   \item \hyperlink{attribute.HOL.code}{\mbox{\isa{code}}}~\isa{inline} declares (or with
   option ``\isa{{\isachardoublequote}del{\isachardoublequote}}'' removes) inlining theorems which are
-  applied as rewrite rules to any defining equation during
+  applied as rewrite rules to any code equation during
   preprocessing.
 
   \item \hyperlink{command.HOL.print-codesetup}{\mbox{\isa{\isacommand{print{\isacharunderscore}codesetup}}}} gives an overview on
-  selected defining equations, code generator datatypes and
+  selected code equations, code generator datatypes and
   preprocessor setup.
 
   \end{description}%
--- a/doc-src/IsarRef/Thy/document/Spec.tex	Mon Jan 19 20:37:08 2009 +0100
+++ b/doc-src/IsarRef/Thy/document/Spec.tex	Mon Jan 19 21:20:18 2009 +0100
@@ -3,8 +3,6 @@
 \def\isabellecontext{Spec}%
 %
 \isadelimtheory
-\isanewline
-\isanewline
 %
 \endisadelimtheory
 %
--- a/doc-src/LaTeXsugar/Sugar/document/Sugar.tex	Mon Jan 19 20:37:08 2009 +0100
+++ b/doc-src/LaTeXsugar/Sugar/document/Sugar.tex	Mon Jan 19 21:20:18 2009 +0100
@@ -148,7 +148,7 @@
 \begin{isamarkuptext}%
 If you print anything, especially theorems, containing
 schematic variables they are prefixed with a question mark:
-\verb!@!\verb!{thm conjI}! results in \isa{{\isasymlbrakk}{\isacharquery}P{\isacharsemicolon}\ {\isacharquery}Q{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}P\ {\isasymand}\ {\isacharquery}Q}. Most of the time
+\verb!@!\verb!{thm conjI}! results in \isa{{\isasymlbrakk}P{\isacharsemicolon}\ Q{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ {\isasymand}\ Q}. Most of the time
 you would rather not see the question marks. There is an attribute
 \verb!no_vars! that you can attach to the theorem that turns its
 schematic into ordinary free variables: \verb!@!\verb!{thm conjI[no_vars]}!
--- a/doc-src/Locales/Locales/Examples3.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/doc-src/Locales/Locales/Examples3.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -535,7 +535,7 @@
 \end{tabular}
 \end{center}
 \hrule
-\caption{Syntax of Locale Commands (abrigded).}
+\caption{Syntax of Locale Commands (abridged).}
 \label{tab:commands}
 \end{table}
   *}
--- a/doc-src/Locales/Locales/document/Examples3.tex	Mon Jan 19 20:37:08 2009 +0100
+++ b/doc-src/Locales/Locales/document/Examples3.tex	Mon Jan 19 21:20:18 2009 +0100
@@ -828,7 +828,7 @@
 \end{tabular}
 \end{center}
 \hrule
-\caption{Syntax of Locale Commands (abrigded).}
+\caption{Syntax of Locale Commands (abridged).}
 \label{tab:commands}
 \end{table}%
 \end{isamarkuptext}%
--- a/doc-src/System/Thy/Presentation.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/doc-src/System/Thy/Presentation.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -442,6 +442,7 @@
     -D PATH      dump generated document sources into PATH
     -M MAX       multithreading: maximum number of worker threads (default 1)
     -P PATH      set path for remote theory browsing information
+    -Q BOOL      check proofs in parallel (default true)
     -T LEVEL     multithreading: trace level (default 0)
     -V VERSION   declare alternative document VERSION
     -b           build mode (output heap image, using current dir)
@@ -461,6 +462,11 @@
 
   ISABELLE_USEDIR_OPTIONS=
   HOL_USEDIR_OPTIONS=
+
+  ML_PLATFORM=x86-linux
+  ML_HOME=/usr/local/polyml-5.2.1/x86-linux
+  ML_SYSTEM=polyml-5.2.1
+  ML_OPTIONS=-H 500
 \end{ttbox}
 
   Note that the value of the @{setting_ref ISABELLE_USEDIR_OPTIONS}
@@ -574,6 +580,13 @@
   bottle-necks, e.g.\ due to excessive wait states when locking
   critical code sections.
 
+  \medskip The @{verbatim "-Q"} option tells whether individual proofs
+  should be checked in parallel; the default is @{verbatim true}.
+  Note that fine-grained proof parallelism requires considerable
+  amounts of extra memory, since full proof context information is
+  maintained for each independent derivation thread, until checking is
+  completed.
+
   \medskip Any @{tool usedir} session is named by some \emph{session
   identifier}. These accumulate, documenting the way sessions depend
   on others. For example, consider @{verbatim "Pure/FOL/ex"}, which
--- a/doc-src/System/Thy/document/Presentation.tex	Mon Jan 19 20:37:08 2009 +0100
+++ b/doc-src/System/Thy/document/Presentation.tex	Mon Jan 19 21:20:18 2009 +0100
@@ -468,6 +468,7 @@
     -D PATH      dump generated document sources into PATH
     -M MAX       multithreading: maximum number of worker threads (default 1)
     -P PATH      set path for remote theory browsing information
+    -Q BOOL      check proofs in parallel (default true)
     -T LEVEL     multithreading: trace level (default 0)
     -V VERSION   declare alternative document VERSION
     -b           build mode (output heap image, using current dir)
@@ -487,6 +488,11 @@
 
   ISABELLE_USEDIR_OPTIONS=
   HOL_USEDIR_OPTIONS=
+
+  ML_PLATFORM=x86-linux
+  ML_HOME=/usr/local/polyml-5.2.1/x86-linux
+  ML_SYSTEM=polyml-5.2.1
+  ML_OPTIONS=-H 500
 \end{ttbox}
 
   Note that the value of the \indexref{}{setting}{ISABELLE\_USEDIR\_OPTIONS}\hyperlink{setting.ISABELLE-USEDIR-OPTIONS}{\mbox{\isa{\isatt{ISABELLE{\isacharunderscore}USEDIR{\isacharunderscore}OPTIONS}}}}
@@ -590,6 +596,13 @@
   bottle-necks, e.g.\ due to excessive wait states when locking
   critical code sections.
 
+  \medskip The \verb|-Q| option tells whether individual proofs
+  should be checked in parallel; the default is \verb|true|.
+  Note that fine-grained proof parallelism requires considerable
+  amounts of extra memory, since full proof context information is
+  maintained for each independent derivation thread, until checking is
+  completed.
+
   \medskip Any \hyperlink{tool.usedir}{\mbox{\isa{\isatt{usedir}}}} session is named by some \emph{session
   identifier}. These accumulate, documenting the way sessions depend
   on others. For example, consider \verb|Pure/FOL/ex|, which
--- a/doc-src/TutorialI/Documents/document/Documents.tex	Mon Jan 19 20:37:08 2009 +0100
+++ b/doc-src/TutorialI/Documents/document/Documents.tex	Mon Jan 19 21:20:18 2009 +0100
@@ -391,7 +391,7 @@
   session is derived from a single parent, usually an object-logic
   image like \texttt{HOL}.  This results in an overall tree structure,
   which is reflected by the output location in the file system
-  (usually rooted at \verb,~/isabelle/browser_info,).
+  (usually rooted at \verb,~/.isabelle/browser_info,).
 
   \medskip The easiest way to manage Isabelle sessions is via
   \texttt{isabelle mkdir} (generates an initial session source setup)
--- a/doc-src/TutorialI/Types/document/Numbers.tex	Mon Jan 19 20:37:08 2009 +0100
+++ b/doc-src/TutorialI/Types/document/Numbers.tex	Mon Jan 19 21:20:18 2009 +0100
@@ -3,7 +3,6 @@
 \def\isabellecontext{Numbers}%
 %
 \isadelimtheory
-\isanewline
 %
 \endisadelimtheory
 %
--- a/doc-src/more_antiquote.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/doc-src/more_antiquote.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -50,7 +50,7 @@
   end
 
 
-(* class antiquotation *)
+(* class and type constructor antiquotation *)
 
 local
 
@@ -74,6 +74,38 @@
 end;
 
 
+(* code theorem antiquotation *)
+
+local
+
+fun pretty_term ctxt t = Syntax.pretty_term (Variable.auto_fixes t ctxt) t;
+
+fun pretty_thm ctxt = pretty_term ctxt o Thm.full_prop_of;
+
+fun no_vars ctxt thm =
+  let
+    val ctxt' = Variable.set_body false ctxt;
+    val ((_, [thm]), _) = Variable.import_thms true [thm] ctxt';
+  in thm end;
+
+fun pretty_code_thm src ctxt raw_const =
+  let
+    val thy = ProofContext.theory_of ctxt;
+    val const = Code_Unit.check_const thy raw_const;
+    val (_, funcgr) = Code_Funcgr.make thy [const];
+    val thms = Code_Funcgr.eqns funcgr const
+      |> map_filter (fn (thm, linear) => if linear then SOME thm else NONE)
+      |> map (no_vars ctxt o AxClass.overload thy);
+  in ThyOutput.output_list pretty_thm src ctxt thms end;
+
+in
+
+val _ = O.add_commands
+  [("code_thms", ThyOutput.args Args.term pretty_code_thm)];
+
+end;
+
+
 (* code antiquotation *)
 
 local
--- a/lib/Tools/usedir	Mon Jan 19 20:37:08 2009 +0100
+++ b/lib/Tools/usedir	Mon Jan 19 21:20:18 2009 +0100
@@ -19,6 +19,7 @@
   echo "    -D PATH      dump generated document sources into PATH"
   echo "    -M MAX       multithreading: maximum number of worker threads (default 1)"
   echo "    -P PATH      set path for remote theory browsing information"
+  echo "    -Q BOOL      check proofs in parallel (default true)"
   echo "    -T LEVEL     multithreading: trace level (default 0)"
   echo "    -V VERSION   declare alternative document VERSION"
   echo "    -b           build mode (output heap image, using current dir)"
@@ -72,6 +73,7 @@
 DUMP=""
 MAXTHREADS="1"
 RPATH=""
+PARALLEL_PROOFS="true"
 TRACETHREADS="0"
 DOCUMENT_VERSIONS=""
 BUILD=""
@@ -89,7 +91,7 @@
 function getoptions()
 {
   OPTIND=1
-  while getopts "C:D:M:P:T:V:bc:d:f:g:i:m:p:rs:v:" OPT
+  while getopts "C:D:M:P:Q:T:V:bc:d:f:g:i:m:p:rs:v:" OPT
   do
     case "$OPT" in
       C)
@@ -101,15 +103,18 @@
         ;;
       M)
         if [ "$OPTARG" = max ]; then
-	  MAXTHREADS=0
-	else
+          MAXTHREADS=0
+        else
           check_number "$OPTARG"
           MAXTHREADS="$OPTARG"
-	fi
+        fi
         ;;
       P)
         RPATH="$OPTARG"
         ;;
+      Q)
+        PARALLEL_PROOFS="$OPTARG"
+        ;;
       T)
         check_number "$OPTARG"
         TRACETHREADS="$OPTARG"
@@ -232,7 +237,7 @@
   [ "$COMPRESS" = true ] && OPT_C="-c"
 
   "$ISABELLE_PROCESS" \
-    -e "Session.use_dir \"$ROOT_FILE\" true [$MODES] $RESET $INFO \"$DOC\" $DOCUMENT_GRAPH [$DOCUMENT_VERSIONS] \"$PARENT\" \"$SESSION\" ($COPY_DUMP, \"$DUMP\") \"$RPATH\" $PROOFS $VERBOSE $MAXTHREADS $TRACETHREADS;" \
+    -e "Session.use_dir \"$ROOT_FILE\" true [$MODES] $RESET $INFO \"$DOC\" $DOCUMENT_GRAPH [$DOCUMENT_VERSIONS] \"$PARENT\" \"$SESSION\" ($COPY_DUMP, \"$DUMP\") \"$RPATH\" $PROOFS $VERBOSE $MAXTHREADS $TRACETHREADS $PARALLEL_PROOFS;" \
     $OPT_C -q -w $LOGIC $NAME > "$LOG"
   RC="$?"
 else
@@ -241,7 +246,7 @@
   LOG="$LOGDIR/$ITEM"
 
   "$ISABELLE_PROCESS" \
-    -e "Session.use_dir \"$ROOT_FILE\" false [$MODES] $RESET $INFO \"$DOC\" $DOCUMENT_GRAPH [$DOCUMENT_VERSIONS] \"$PARENT\" \"$SESSION\" ($COPY_DUMP, \"$DUMP\") \"$RPATH\" $PROOFS $VERBOSE $MAXTHREADS $TRACETHREADS; quit();" \
+    -e "Session.use_dir \"$ROOT_FILE\" false [$MODES] $RESET $INFO \"$DOC\" $DOCUMENT_GRAPH [$DOCUMENT_VERSIONS] \"$PARENT\" \"$SESSION\" ($COPY_DUMP, \"$DUMP\") \"$RPATH\" $PROOFS $VERBOSE $MAXTHREADS $TRACETHREADS $PARALLEL_PROOFS; quit();" \
     -r -q "$LOGIC" > "$LOG"
   RC="$?"
   cd ..
--- a/lib/html/isabelle.css	Mon Jan 19 20:37:08 2009 +0100
+++ b/lib/html/isabelle.css	Mon Jan 19 21:20:18 2009 +0100
@@ -20,20 +20,20 @@
 
 /* inner and outer syntax markup */
 
-.tfree, tfree          { color: purple; }
-.tvar, tvar            { color: purple; }
-.free, free            { color: blue; }
-.skolem, skolem        { color: brown; }
-.bound, bound          { color: green; }
-.var, var              { color: blue; }
-.num, num              { }
-.xnum, xnum            { }
-.xstr, xstr            { color: brown; }
-.literal, literal      { font-weight: bold; }
-                      
+.tfree, tfree                 { color: purple; }
+.tvar, tvar                   { color: purple; }
+.free, free                   { color: blue; }
+.skolem, skolem               { color: brown; }
+.bound, bound                 { color: green; }
+.var, var                     { color: blue; }
+.numeral, numeral             { }
+.literal, literal             { font-weight: bold; }
+.inner_string, inner_string   { color: brown; }
+.inner_comment, inner_comment { color: #8B0000; }
+
 .loc, loc              { color: brown; }
 .tclass, tclass        { color: red; }
-          
+
 .keyword, keyword      { font-weight: bold; }
 .command, command      { font-weight: bold; }
 .ident, ident          { }
--- a/src/CCL/Wfd.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/CCL/Wfd.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      CCL/Wfd.thy
-    ID:         $Id$
     Author:     Martin Coen, Cambridge University Computer Laboratory
     Copyright   1993  University of Cambridge
 *)
@@ -435,9 +434,9 @@
   | get_bno l n (Bound m) = (m-length(l),n)
 
 (* Not a great way of identifying induction hypothesis! *)
-fun could_IH x = could_unify(x,hd (prems_of @{thm rcallT})) orelse
-                 could_unify(x,hd (prems_of @{thm rcall2T})) orelse
-                 could_unify(x,hd (prems_of @{thm rcall3T}))
+fun could_IH x = Term.could_unify(x,hd (prems_of @{thm rcallT})) orelse
+                 Term.could_unify(x,hd (prems_of @{thm rcall2T})) orelse
+                 Term.could_unify(x,hd (prems_of @{thm rcall3T}))
 
 fun IHinst tac rls = SUBGOAL (fn (Bi,i) =>
   let val bvs = bvars Bi []
@@ -451,7 +450,7 @@
 
 fun is_rigid_prog t =
      case (Logic.strip_assums_concl t) of
-        (Const("Trueprop",_) $ (Const("mem",_) $ a $ _)) => (term_vars a = [])
+        (Const("Trueprop",_) $ (Const("mem",_) $ a $ _)) => null (Term.add_vars a [])
        | _ => false
 in
 
--- a/src/FOL/IsaMakefile	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/FOL/IsaMakefile	Mon Jan 19 21:20:18 2009 +0100
@@ -47,7 +47,7 @@
 
 $(LOG)/FOL-ex.gz: $(OUT)/FOL ex/First_Order_Logic.thy ex/If.thy		\
   ex/IffOracle.thy ex/Nat.thy ex/Natural_Numbers.thy	\
-  ex/NewLocaleSetup.thy ex/NewLocaleTest.thy    \
+  ex/LocaleTest.thy    \
   ex/Miniscope.thy ex/Prolog.thy ex/ROOT.ML ex/Classical.thy		\
   ex/document/root.tex ex/Foundation.thy ex/Intuitionistic.thy		\
   ex/Intro.thy ex/Propositional_Int.thy ex/Propositional_Cla.thy	\
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/FOL/ex/LocaleTest.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -0,0 +1,483 @@
+(*  Title:      FOL/ex/NewLocaleTest.thy
+    Author:     Clemens Ballarin, TU Muenchen
+
+Testing environment for locale expressions.
+*)
+
+theory LocaleTest
+imports FOL
+begin
+
+typedecl int arities int :: "term"
+consts plus :: "int => int => int" (infixl "+" 60)
+  zero :: int ("0")
+  minus :: "int => int" ("- _")
+
+axioms
+  int_assoc: "(x + y::int) + z = x + (y + z)"
+  int_zero: "0 + x = x"
+  int_minus: "(-x) + x = 0"
+  int_minus2: "-(-x) = x"
+
+section {* Inference of parameter types *}
+
+locale param1 = fixes p
+print_locale! param1
+
+locale param2 = fixes p :: 'b
+print_locale! param2
+
+(*
+locale param_top = param2 r for r :: "'b :: {}"
+  Fails, cannot generalise parameter.
+*)
+
+locale param3 = fixes p (infix ".." 50)
+print_locale! param3
+
+locale param4 = fixes p :: "'a => 'a => 'a" (infix ".." 50)
+print_locale! param4
+
+
+subsection {* Incremental type constraints *}
+
+locale constraint1 =
+  fixes  prod (infixl "**" 65)
+  assumes l_id: "x ** y = x"
+  assumes assoc: "(x ** y) ** z = x ** (y ** z)"
+print_locale! constraint1
+
+locale constraint2 =
+  fixes p and q
+  assumes "p = q"
+print_locale! constraint2
+
+
+section {* Inheritance *}
+
+locale semi =
+  fixes prod (infixl "**" 65)
+  assumes assoc: "(x ** y) ** z = x ** (y ** z)"
+print_locale! semi thm semi_def
+
+locale lgrp = semi +
+  fixes one and inv
+  assumes lone: "one ** x = x"
+    and linv: "inv(x) ** x = one"
+print_locale! lgrp thm lgrp_def lgrp_axioms_def
+
+locale add_lgrp = semi "op ++" for sum (infixl "++" 60) +
+  fixes zero and neg
+  assumes lzero: "zero ++ x = x"
+    and lneg: "neg(x) ++ x = zero"
+print_locale! add_lgrp thm add_lgrp_def add_lgrp_axioms_def
+
+locale rev_lgrp = semi "%x y. y ++ x" for sum (infixl "++" 60)
+print_locale! rev_lgrp thm rev_lgrp_def
+
+locale hom = f: semi f + g: semi g for f and g
+print_locale! hom thm hom_def
+
+locale perturbation = semi + d: semi "%x y. delta(x) ** delta(y)" for delta
+print_locale! perturbation thm perturbation_def
+
+locale pert_hom = d1: perturbation f d1 + d2: perturbation f d2 for f d1 d2
+print_locale! pert_hom thm pert_hom_def
+
+text {* Alternative expression, obtaining nicer names in @{text "semi f"}. *}
+locale pert_hom' = semi f + d1: perturbation f d1 + d2: perturbation f d2 for f d1 d2
+print_locale! pert_hom' thm pert_hom'_def
+
+
+section {* Syntax declarations *}
+
+locale logic =
+  fixes land (infixl "&&" 55)
+    and lnot ("-- _" [60] 60)
+  assumes assoc: "(x && y) && z = x && (y && z)"
+    and notnot: "-- (-- x) = x"
+begin
+
+definition lor (infixl "||" 50) where
+  "x || y = --(-- x && -- y)"
+
+end
+print_locale! logic
+
+locale use_decl = logic + semi "op ||"
+print_locale! use_decl thm use_decl_def
+
+locale extra_type =
+  fixes a :: 'a
+    and P :: "'a => 'b => o"
+begin
+
+definition test :: "'a => o" where
+  "test(x) <-> (ALL b. P(x, b))"
+
+end
+
+term extra_type.test thm extra_type.test_def
+
+interpretation var: extra_type "0" "%x y. x = 0" .
+
+thm var.test_def
+
+
+section {* Foundational versions of theorems *}
+
+thm logic.assoc
+thm logic.lor_def
+
+
+section {* Defines *}
+
+locale logic_def =
+  fixes land (infixl "&&" 55)
+    and lor (infixl "||" 50)
+    and lnot ("-- _" [60] 60)
+  assumes assoc: "(x && y) && z = x && (y && z)"
+    and notnot: "-- (-- x) = x"
+  defines "x || y == --(-- x && -- y)"
+begin
+
+thm lor_def
+(* Can we get rid the the additional hypothesis, caused by LocalTheory.notes? *)
+
+lemma "x || y = --(-- x && --y)"
+  by (unfold lor_def) (rule refl)
+
+end
+
+(* Inheritance of defines *)
+
+locale logic_def2 = logic_def
+begin
+
+lemma "x || y = --(-- x && --y)"
+  by (unfold lor_def) (rule refl)
+
+end
+
+
+section {* Notes *}
+
+(* A somewhat arcane homomorphism example *)
+
+definition semi_hom where
+  "semi_hom(prod, sum, h) <-> (ALL x y. h(prod(x, y)) = sum(h(x), h(y)))"
+
+lemma semi_hom_mult:
+  "semi_hom(prod, sum, h) ==> h(prod(x, y)) = sum(h(x), h(y))"
+  by (simp add: semi_hom_def)
+
+locale semi_hom_loc = prod: semi prod + sum: semi sum
+  for prod and sum and h +
+  assumes semi_homh: "semi_hom(prod, sum, h)"
+  notes semi_hom_mult = semi_hom_mult [OF semi_homh]
+
+thm semi_hom_loc.semi_hom_mult
+(* unspecified, attribute not applied in backgroud theory !!! *)
+
+lemma (in semi_hom_loc) "h(prod(x, y)) = sum(h(x), h(y))"
+  by (rule semi_hom_mult)
+
+(* Referring to facts from within a context specification *)
+
+lemma
+  assumes x: "P <-> P"
+  notes y = x
+  shows True ..
+
+
+section {* Theorem statements *}
+
+lemma (in lgrp) lcancel:
+  "x ** y = x ** z <-> y = z"
+proof
+  assume "x ** y = x ** z"
+  then have "inv(x) ** x ** y = inv(x) ** x ** z" by (simp add: assoc)
+  then show "y = z" by (simp add: lone linv)
+qed simp
+print_locale! lgrp
+
+
+locale rgrp = semi +
+  fixes one and inv
+  assumes rone: "x ** one = x"
+    and rinv: "x ** inv(x) = one"
+begin
+
+lemma rcancel:
+  "y ** x = z ** x <-> y = z"
+proof
+  assume "y ** x = z ** x"
+  then have "y ** (x ** inv(x)) = z ** (x ** inv(x))"
+    by (simp add: assoc [symmetric])
+  then show "y = z" by (simp add: rone rinv)
+qed simp
+
+end
+print_locale! rgrp
+
+
+subsection {* Patterns *}
+
+lemma (in rgrp)
+  assumes "y ** x = z ** x" (is ?a)
+  shows "y = z" (is ?t)
+proof -
+  txt {* Weird proof involving patterns from context element and conclusion. *}
+  {
+    assume ?a
+    then have "y ** (x ** inv(x)) = z ** (x ** inv(x))"
+      by (simp add: assoc [symmetric])
+    then have ?t by (simp add: rone rinv)
+  }
+  note x = this
+  show ?t by (rule x [OF `?a`])
+qed
+
+
+section {* Interpretation between locales: sublocales *}
+
+sublocale lgrp < right: rgrp
+print_facts
+proof unfold_locales
+  {
+    fix x
+    have "inv(x) ** x ** one = inv(x) ** x" by (simp add: linv lone)
+    then show "x ** one = x" by (simp add: assoc lcancel)
+  }
+  note rone = this
+  {
+    fix x
+    have "inv(x) ** x ** inv(x) = inv(x) ** one"
+      by (simp add: linv lone rone)
+    then show "x ** inv(x) = one" by (simp add: assoc lcancel)
+  }
+qed
+
+(* effect on printed locale *)
+
+print_locale! lgrp
+
+(* use of derived theorem *)
+
+lemma (in lgrp)
+  "y ** x = z ** x <-> y = z"
+  apply (rule rcancel)
+  done
+
+(* circular interpretation *)
+
+sublocale rgrp < left: lgrp
+proof unfold_locales
+  {
+    fix x
+    have "one ** (x ** inv(x)) = x ** inv(x)" by (simp add: rinv rone)
+    then show "one ** x = x" by (simp add: assoc [symmetric] rcancel)
+  }
+  note lone = this
+  {
+    fix x
+    have "inv(x) ** (x ** inv(x)) = one ** inv(x)"
+      by (simp add: rinv lone rone)
+    then show "inv(x) ** x = one" by (simp add: assoc [symmetric] rcancel)
+  }
+qed
+
+(* effect on printed locale *)
+
+print_locale! rgrp
+print_locale! lgrp
+
+
+(* Duality *)
+
+locale order =
+  fixes less :: "'a => 'a => o" (infix "<<" 50)
+  assumes refl: "x << x"
+    and trans: "[| x << y; y << z |] ==> x << z"
+
+sublocale order < dual: order "%x y. y << x"
+  apply unfold_locales apply (rule refl) apply (blast intro: trans)
+  done
+
+print_locale! order  (* Only two instances of order. *)
+
+locale order' =
+  fixes less :: "'a => 'a => o" (infix "<<" 50)
+  assumes refl: "x << x"
+    and trans: "[| x << y; y << z |] ==> x << z"
+
+locale order_with_def = order'
+begin
+
+definition greater :: "'a => 'a => o" (infix ">>" 50) where
+  "x >> y <-> y << x"
+
+end
+
+sublocale order_with_def < dual: order' "op >>"
+  apply unfold_locales
+  unfolding greater_def
+  apply (rule refl) apply (blast intro: trans)
+  done
+
+print_locale! order_with_def
+(* Note that decls come after theorems that make use of them. *)
+
+
+(* locale with many parameters ---
+   interpretations generate alternating group A5 *)
+
+
+locale A5 =
+  fixes A and B and C and D and E
+  assumes eq: "A <-> B <-> C <-> D <-> E"
+
+sublocale A5 < 1: A5 _ _ D E C
+print_facts
+  using eq apply (blast intro: A5.intro) done
+
+sublocale A5 < 2: A5 C _ E _ A
+print_facts
+  using eq apply (blast intro: A5.intro) done
+
+sublocale A5 < 3: A5 B C A _ _
+print_facts
+  using eq apply (blast intro: A5.intro) done
+
+(* Any even permutation of parameters is subsumed by the above. *)
+
+print_locale! A5
+
+
+(* Free arguments of instance *)
+
+locale trivial =
+  fixes P and Q :: o
+  assumes Q: "P <-> P <-> Q"
+begin
+
+lemma Q_triv: "Q" using Q by fast
+
+end
+
+sublocale trivial < x: trivial x _
+  apply unfold_locales using Q by fast
+
+print_locale! trivial
+
+context trivial begin thm x.Q [where ?x = True] end
+
+sublocale trivial < y: trivial Q Q
+  by unfold_locales
+  (* Succeeds since previous interpretation is more general. *)
+
+print_locale! trivial  (* No instance for y created (subsumed). *)
+
+
+subsection {* Sublocale, then interpretation in theory *}
+
+interpretation int: lgrp "op +" "0" "minus"
+proof unfold_locales
+qed (rule int_assoc int_zero int_minus)+
+
+thm int.assoc int.semi_axioms
+
+interpretation int2: semi "op +"
+  by unfold_locales  (* subsumed, thm int2.assoc not generated *)
+
+thm int.lone int.right.rone
+  (* the latter comes through the sublocale relation *)
+
+
+subsection {* Interpretation in theory, then sublocale *}
+
+interpretation (* fol: *) logic "op +" "minus"
+(* FIXME declaration of qualified names *)
+  by unfold_locales (rule int_assoc int_minus2)+
+
+locale logic2 =
+  fixes land (infixl "&&" 55)
+    and lnot ("-- _" [60] 60)
+  assumes assoc: "(x && y) && z = x && (y && z)"
+    and notnot: "-- (-- x) = x"
+begin
+
+(* FIXME interpretation below fails
+definition lor (infixl "||" 50) where
+  "x || y = --(-- x && -- y)"
+*)
+
+end
+
+sublocale logic < two: logic2
+  by unfold_locales (rule assoc notnot)+
+
+thm two.assoc
+
+
+subsection {* Declarations and sublocale *}
+
+locale logic_a = logic
+locale logic_b = logic
+
+sublocale logic_a < logic_b
+  by unfold_locales
+
+
+subsection {* Equations *}
+
+locale logic_o =
+  fixes land (infixl "&&" 55)
+    and lnot ("-- _" [60] 60)
+  assumes assoc_o: "(x && y) && z <-> x && (y && z)"
+    and notnot_o: "-- (-- x) <-> x"
+begin
+
+definition lor_o (infixl "||" 50) where
+  "x || y <-> --(-- x && -- y)"
+
+end
+
+interpretation x!: logic_o "op &" "Not"
+  where bool_logic_o: "logic_o.lor_o(op &, Not, x, y) <-> x | y"
+proof -
+  show bool_logic_o: "PROP logic_o(op &, Not)" by unfold_locales fast+
+  show "logic_o.lor_o(op &, Not, x, y) <-> x | y"
+    by (unfold logic_o.lor_o_def [OF bool_logic_o]) fast
+qed
+
+thm x.lor_o_def bool_logic_o
+
+lemma lor_triv: "z <-> z" ..
+
+lemma (in logic_o) lor_triv: "x || y <-> x || y" by fast
+
+thm lor_triv [where z = True] (* Check strict prefix. *)
+  x.lor_triv
+
+
+subsection {* Interpretation in proofs *}
+
+lemma True
+proof
+  interpret "local": lgrp "op +" "0" "minus"
+    by unfold_locales  (* subsumed *)
+  {
+    fix zero :: int
+    assume "!!x. zero + x = x" "!!x. (-x) + x = zero"
+    then interpret local_fixed: lgrp "op +" zero "minus"
+      by unfold_locales
+    thm local_fixed.lone
+  }
+  assume "!!x zero. zero + x = x" "!!x zero. (-x) + x = zero"
+  then interpret local_free: lgrp "op +" zero "minus" for zero
+    by unfold_locales
+  thm local_free.lone [where ?zero = 0]
+qed
+
+end
--- a/src/FOL/ex/NewLocaleSetup.thy	Mon Jan 19 20:37:08 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,64 +0,0 @@
-(*  Title:      FOL/ex/NewLocaleSetup.thy
-    Author:     Clemens Ballarin, TU Muenchen
-
-Testing environment for locale expressions --- experimental.
-*)
-
-theory NewLocaleSetup
-imports FOL
-begin
-
-ML {*
-
-local
-
-structure P = OuterParse and K = OuterKeyword;
-val opt_bang = Scan.optional (P.$$$ "!" >> K true) false;
-
-val locale_val =
-  SpecParse.locale_expression --
-    Scan.optional (P.$$$ "+" |-- P.!!! (Scan.repeat1 SpecParse.context_element)) [] ||
-  Scan.repeat1 SpecParse.context_element >> pair ([], []);
-
-in
-
-val _ =
-  OuterSyntax.command "locale" "define named proof context" K.thy_decl
-    (P.name -- Scan.optional (P.$$$ "=" |-- P.!!! locale_val) (([], []), []) -- P.opt_begin
-      >> (fn ((name, (expr, elems)), begin) =>
-          (begin ? Toplevel.print) o Toplevel.begin_local_theory begin
-            (Expression.add_locale_cmd name name expr elems #>
-              (fn ((target, notes), ctxt) => TheoryTarget.begin target ctxt |>
-                fold (fn (kind, facts) => LocalTheory.notes kind facts #> snd) notes))));
-
-val _ =
-  OuterSyntax.improper_command "print_locales" "print locales of this theory" K.diag
-    (Scan.succeed (Toplevel.no_timing o (Toplevel.unknown_theory o
-  Toplevel.keep (NewLocale.print_locales o Toplevel.theory_of))));
-
-val _ = OuterSyntax.improper_command "print_locale" "print named locale in this theory" K.diag
-  (opt_bang -- P.xname >> (Toplevel.no_timing oo (fn (show_facts, name) =>
-   Toplevel.unknown_theory o Toplevel.keep (fn state =>
-     NewLocale.print_locale (Toplevel.theory_of state) show_facts name))));
-
-val _ =
-  OuterSyntax.command "interpretation"
-    "prove interpretation of locale expression in theory" K.thy_goal
-    (P.!!! SpecParse.locale_expression --
-      Scan.optional (P.$$$ "where" |-- P.and_list1 (SpecParse.opt_thm_name ":" -- P.prop)) []
-        >> (fn (expr, mixin) => Toplevel.print o
-            Toplevel.theory_to_proof (Expression.interpretation_cmd expr mixin)));
-
-val _ =
-  OuterSyntax.command "interpret"
-    "prove interpretation of locale expression in proof context"
-    (K.tag_proof K.prf_goal)
-    (P.!!! SpecParse.locale_expression
-        >> (fn expr => Toplevel.print o
-            Toplevel.proof' (fn int => Expression.interpret_cmd expr int)));
-
-end
-
-*}
-
-end
--- a/src/FOL/ex/NewLocaleTest.thy	Mon Jan 19 20:37:08 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,486 +0,0 @@
-(*  Title:      FOL/ex/NewLocaleTest.thy
-    Author:     Clemens Ballarin, TU Muenchen
-
-Testing environment for locale expressions --- experimental.
-*)
-
-theory NewLocaleTest
-imports NewLocaleSetup
-begin
-
-ML_val {* set Toplevel.debug *}
-
-
-typedecl int arities int :: "term"
-consts plus :: "int => int => int" (infixl "+" 60)
-  zero :: int ("0")
-  minus :: "int => int" ("- _")
-
-axioms
-  int_assoc: "(x + y::int) + z = x + (y + z)"
-  int_zero: "0 + x = x"
-  int_minus: "(-x) + x = 0"
-  int_minus2: "-(-x) = x"
-
-section {* Inference of parameter types *}
-
-locale param1 = fixes p
-print_locale! param1
-
-locale param2 = fixes p :: 'b
-print_locale! param2
-
-(*
-locale param_top = param2 r for r :: "'b :: {}"
-  Fails, cannot generalise parameter.
-*)
-
-locale param3 = fixes p (infix ".." 50)
-print_locale! param3
-
-locale param4 = fixes p :: "'a => 'a => 'a" (infix ".." 50)
-print_locale! param4
-
-
-subsection {* Incremental type constraints *}
-
-locale constraint1 =
-  fixes  prod (infixl "**" 65)
-  assumes l_id: "x ** y = x"
-  assumes assoc: "(x ** y) ** z = x ** (y ** z)"
-print_locale! constraint1
-
-locale constraint2 =
-  fixes p and q
-  assumes "p = q"
-print_locale! constraint2
-
-
-section {* Inheritance *}
-
-locale semi =
-  fixes prod (infixl "**" 65)
-  assumes assoc: "(x ** y) ** z = x ** (y ** z)"
-print_locale! semi thm semi_def
-
-locale lgrp = semi +
-  fixes one and inv
-  assumes lone: "one ** x = x"
-    and linv: "inv(x) ** x = one"
-print_locale! lgrp thm lgrp_def lgrp_axioms_def
-
-locale add_lgrp = semi "op ++" for sum (infixl "++" 60) +
-  fixes zero and neg
-  assumes lzero: "zero ++ x = x"
-    and lneg: "neg(x) ++ x = zero"
-print_locale! add_lgrp thm add_lgrp_def add_lgrp_axioms_def
-
-locale rev_lgrp = semi "%x y. y ++ x" for sum (infixl "++" 60)
-print_locale! rev_lgrp thm rev_lgrp_def
-
-locale hom = f: semi f + g: semi g for f and g
-print_locale! hom thm hom_def
-
-locale perturbation = semi + d: semi "%x y. delta(x) ** delta(y)" for delta
-print_locale! perturbation thm perturbation_def
-
-locale pert_hom = d1: perturbation f d1 + d2: perturbation f d2 for f d1 d2
-print_locale! pert_hom thm pert_hom_def
-
-text {* Alternative expression, obtaining nicer names in @{text "semi f"}. *}
-locale pert_hom' = semi f + d1: perturbation f d1 + d2: perturbation f d2 for f d1 d2
-print_locale! pert_hom' thm pert_hom'_def
-
-
-section {* Syntax declarations *}
-
-locale logic =
-  fixes land (infixl "&&" 55)
-    and lnot ("-- _" [60] 60)
-  assumes assoc: "(x && y) && z = x && (y && z)"
-    and notnot: "-- (-- x) = x"
-begin
-
-definition lor (infixl "||" 50) where
-  "x || y = --(-- x && -- y)"
-
-end
-print_locale! logic
-
-locale use_decl = logic + semi "op ||"
-print_locale! use_decl thm use_decl_def
-
-locale extra_type =
-  fixes a :: 'a
-    and P :: "'a => 'b => o"
-begin
-
-definition test :: "'a => o" where
-  "test(x) <-> (ALL b. P(x, b))"
-
-end
-
-term extra_type.test thm extra_type.test_def
-
-interpretation var: extra_type "0" "%x y. x = 0" .
-
-thm var.test_def
-
-
-section {* Foundational versions of theorems *}
-
-thm logic.assoc
-thm logic.lor_def
-
-
-section {* Defines *}
-
-locale logic_def =
-  fixes land (infixl "&&" 55)
-    and lor (infixl "||" 50)
-    and lnot ("-- _" [60] 60)
-  assumes assoc: "(x && y) && z = x && (y && z)"
-    and notnot: "-- (-- x) = x"
-  defines "x || y == --(-- x && -- y)"
-begin
-
-thm lor_def
-(* Can we get rid the the additional hypothesis, caused by LocalTheory.notes? *)
-
-lemma "x || y = --(-- x && --y)"
-  by (unfold lor_def) (rule refl)
-
-end
-
-(* Inheritance of defines *)
-
-locale logic_def2 = logic_def
-begin
-
-lemma "x || y = --(-- x && --y)"
-  by (unfold lor_def) (rule refl)
-
-end
-
-
-section {* Notes *}
-
-(* A somewhat arcane homomorphism example *)
-
-definition semi_hom where
-  "semi_hom(prod, sum, h) <-> (ALL x y. h(prod(x, y)) = sum(h(x), h(y)))"
-
-lemma semi_hom_mult:
-  "semi_hom(prod, sum, h) ==> h(prod(x, y)) = sum(h(x), h(y))"
-  by (simp add: semi_hom_def)
-
-locale semi_hom_loc = prod: semi prod + sum: semi sum
-  for prod and sum and h +
-  assumes semi_homh: "semi_hom(prod, sum, h)"
-  notes semi_hom_mult = semi_hom_mult [OF semi_homh]
-
-thm semi_hom_loc.semi_hom_mult
-(* unspecified, attribute not applied in backgroud theory !!! *)
-
-lemma (in semi_hom_loc) "h(prod(x, y)) = sum(h(x), h(y))"
-  by (rule semi_hom_mult)
-
-(* Referring to facts from within a context specification *)
-
-lemma
-  assumes x: "P <-> P"
-  notes y = x
-  shows True ..
-
-
-section {* Theorem statements *}
-
-lemma (in lgrp) lcancel:
-  "x ** y = x ** z <-> y = z"
-proof
-  assume "x ** y = x ** z"
-  then have "inv(x) ** x ** y = inv(x) ** x ** z" by (simp add: assoc)
-  then show "y = z" by (simp add: lone linv)
-qed simp
-print_locale! lgrp
-
-
-locale rgrp = semi +
-  fixes one and inv
-  assumes rone: "x ** one = x"
-    and rinv: "x ** inv(x) = one"
-begin
-
-lemma rcancel:
-  "y ** x = z ** x <-> y = z"
-proof
-  assume "y ** x = z ** x"
-  then have "y ** (x ** inv(x)) = z ** (x ** inv(x))"
-    by (simp add: assoc [symmetric])
-  then show "y = z" by (simp add: rone rinv)
-qed simp
-
-end
-print_locale! rgrp
-
-
-subsection {* Patterns *}
-
-lemma (in rgrp)
-  assumes "y ** x = z ** x" (is ?a)
-  shows "y = z" (is ?t)
-proof -
-  txt {* Weird proof involving patterns from context element and conclusion. *}
-  {
-    assume ?a
-    then have "y ** (x ** inv(x)) = z ** (x ** inv(x))"
-      by (simp add: assoc [symmetric])
-    then have ?t by (simp add: rone rinv)
-  }
-  note x = this
-  show ?t by (rule x [OF `?a`])
-qed
-
-
-section {* Interpretation between locales: sublocales *}
-
-sublocale lgrp < right: rgrp
-print_facts
-proof unfold_locales
-  {
-    fix x
-    have "inv(x) ** x ** one = inv(x) ** x" by (simp add: linv lone)
-    then show "x ** one = x" by (simp add: assoc lcancel)
-  }
-  note rone = this
-  {
-    fix x
-    have "inv(x) ** x ** inv(x) = inv(x) ** one"
-      by (simp add: linv lone rone)
-    then show "x ** inv(x) = one" by (simp add: assoc lcancel)
-  }
-qed
-
-(* effect on printed locale *)
-
-print_locale! lgrp
-
-(* use of derived theorem *)
-
-lemma (in lgrp)
-  "y ** x = z ** x <-> y = z"
-  apply (rule rcancel)
-  done
-
-(* circular interpretation *)
-
-sublocale rgrp < left: lgrp
-proof unfold_locales
-  {
-    fix x
-    have "one ** (x ** inv(x)) = x ** inv(x)" by (simp add: rinv rone)
-    then show "one ** x = x" by (simp add: assoc [symmetric] rcancel)
-  }
-  note lone = this
-  {
-    fix x
-    have "inv(x) ** (x ** inv(x)) = one ** inv(x)"
-      by (simp add: rinv lone rone)
-    then show "inv(x) ** x = one" by (simp add: assoc [symmetric] rcancel)
-  }
-qed
-
-(* effect on printed locale *)
-
-print_locale! rgrp
-print_locale! lgrp
-
-
-(* Duality *)
-
-locale order =
-  fixes less :: "'a => 'a => o" (infix "<<" 50)
-  assumes refl: "x << x"
-    and trans: "[| x << y; y << z |] ==> x << z"
-
-sublocale order < dual: order "%x y. y << x"
-  apply unfold_locales apply (rule refl) apply (blast intro: trans)
-  done
-
-print_locale! order  (* Only two instances of order. *)
-
-locale order' =
-  fixes less :: "'a => 'a => o" (infix "<<" 50)
-  assumes refl: "x << x"
-    and trans: "[| x << y; y << z |] ==> x << z"
-
-locale order_with_def = order'
-begin
-
-definition greater :: "'a => 'a => o" (infix ">>" 50) where
-  "x >> y <-> y << x"
-
-end
-
-sublocale order_with_def < dual: order' "op >>"
-  apply unfold_locales
-  unfolding greater_def
-  apply (rule refl) apply (blast intro: trans)
-  done
-
-print_locale! order_with_def
-(* Note that decls come after theorems that make use of them. *)
-
-
-(* locale with many parameters ---
-   interpretations generate alternating group A5 *)
-
-
-locale A5 =
-  fixes A and B and C and D and E
-  assumes eq: "A <-> B <-> C <-> D <-> E"
-
-sublocale A5 < 1: A5 _ _ D E C
-print_facts
-  using eq apply (blast intro: A5.intro) done
-
-sublocale A5 < 2: A5 C _ E _ A
-print_facts
-  using eq apply (blast intro: A5.intro) done
-
-sublocale A5 < 3: A5 B C A _ _
-print_facts
-  using eq apply (blast intro: A5.intro) done
-
-(* Any even permutation of parameters is subsumed by the above. *)
-
-print_locale! A5
-
-
-(* Free arguments of instance *)
-
-locale trivial =
-  fixes P and Q :: o
-  assumes Q: "P <-> P <-> Q"
-begin
-
-lemma Q_triv: "Q" using Q by fast
-
-end
-
-sublocale trivial < x: trivial x _
-  apply unfold_locales using Q by fast
-
-print_locale! trivial
-
-context trivial begin thm x.Q [where ?x = True] end
-
-sublocale trivial < y: trivial Q Q
-  by unfold_locales
-  (* Succeeds since previous interpretation is more general. *)
-
-print_locale! trivial  (* No instance for y created (subsumed). *)
-
-
-subsection {* Sublocale, then interpretation in theory *}
-
-interpretation int: lgrp "op +" "0" "minus"
-proof unfold_locales
-qed (rule int_assoc int_zero int_minus)+
-
-thm int.assoc int.semi_axioms
-
-interpretation int2: semi "op +"
-  by unfold_locales  (* subsumed, thm int2.assoc not generated *)
-
-thm int.lone int.right.rone
-  (* the latter comes through the sublocale relation *)
-
-
-subsection {* Interpretation in theory, then sublocale *}
-
-interpretation (* fol: *) logic "op +" "minus"
-(* FIXME declaration of qualified names *)
-  by unfold_locales (rule int_assoc int_minus2)+
-
-locale logic2 =
-  fixes land (infixl "&&" 55)
-    and lnot ("-- _" [60] 60)
-  assumes assoc: "(x && y) && z = x && (y && z)"
-    and notnot: "-- (-- x) = x"
-begin
-
-(* FIXME interpretation below fails
-definition lor (infixl "||" 50) where
-  "x || y = --(-- x && -- y)"
-*)
-
-end
-
-sublocale logic < two: logic2
-  by unfold_locales (rule assoc notnot)+
-
-thm two.assoc
-
-
-subsection {* Declarations and sublocale *}
-
-locale logic_a = logic
-locale logic_b = logic
-
-sublocale logic_a < logic_b
-  by unfold_locales
-
-
-subsection {* Equations *}
-
-locale logic_o =
-  fixes land (infixl "&&" 55)
-    and lnot ("-- _" [60] 60)
-  assumes assoc_o: "(x && y) && z <-> x && (y && z)"
-    and notnot_o: "-- (-- x) <-> x"
-begin
-
-definition lor_o (infixl "||" 50) where
-  "x || y <-> --(-- x && -- y)"
-
-end
-
-interpretation x!: logic_o "op &" "Not"
-  where bool_logic_o: "logic_o.lor_o(op &, Not, x, y) <-> x | y"
-proof -
-  show bool_logic_o: "PROP logic_o(op &, Not)" by unfold_locales fast+
-  show "logic_o.lor_o(op &, Not, x, y) <-> x | y"
-    by (unfold logic_o.lor_o_def [OF bool_logic_o]) fast
-qed
-
-thm x.lor_o_def bool_logic_o
-
-lemma lor_triv: "z <-> z" ..
-
-lemma (in logic_o) lor_triv: "x || y <-> x || y" by fast
-
-thm lor_triv [where z = True] (* Check strict prefix. *)
-  x.lor_triv
-
-
-subsection {* Interpretation in proofs *}
-
-lemma True
-proof
-  interpret "local": lgrp "op +" "0" "minus"
-    by unfold_locales  (* subsumed *)
-  {
-    fix zero :: int
-    assume "!!x. zero + x = x" "!!x. (-x) + x = zero"
-    then interpret local_fixed: lgrp "op +" zero "minus"
-      by unfold_locales
-    thm local_fixed.lone
-  }
-  assume "!!x zero. zero + x = x" "!!x zero. (-x) + x = zero"
-  then interpret local_free: lgrp "op +" zero "minus" for zero
-    by unfold_locales
-  thm local_free.lone [where ?zero = 0]
-qed
-
-end
--- a/src/FOL/ex/ROOT.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/FOL/ex/ROOT.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -29,5 +29,5 @@
 ];
 
 (*regression test for locales -- sets several global flags!*)
-no_document use_thy "NewLocaleTest";
+no_document use_thy "LocaleTest";
 
--- a/src/FOLP/IFOLP.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/FOLP/IFOLP.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      FOLP/IFOLP.thy
-    ID:         $Id$
     Author:     Martin D Coen, Cambridge University Computer Laboratory
     Copyright   1992  University of Cambridge
 *)
@@ -247,7 +246,7 @@
           and concl = discard_proof (Logic.strip_assums_concl prem)
       in
           if exists (fn hyp => hyp aconv concl) hyps
-          then case distinct (op =) (filter (fn hyp => could_unify (hyp, concl)) hyps) of
+          then case distinct (op =) (filter (fn hyp => Term.could_unify (hyp, concl)) hyps) of
                    [_] => assume_tac i
                  |  _  => no_tac
           else no_tac
@@ -340,6 +339,7 @@
   shows "?a : R"
   apply (insert assms(1) [unfolded ex1_def])
   apply (erule exE conjE | assumption | rule assms(1))+
+  apply (erule assms(2), assumption)
   done
 
 
--- a/src/FOLP/simp.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/FOLP/simp.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
-(*  Title:      FOLP/simp
-    ID:         $Id$
+(*  Title:      FOLP/simp.ML
     Author:     Tobias Nipkow
     Copyright   1993  University of Cambridge
 
@@ -156,21 +155,21 @@
 (*ccs contains the names of the constants possessing congruence rules*)
 fun add_hidden_vars ccs =
   let fun add_hvars tm hvars = case tm of
-              Abs(_,_,body) => add_term_vars(body,hvars)
+              Abs(_,_,body) => OldTerm.add_term_vars(body,hvars)
             | _$_ => let val (f,args) = strip_comb tm 
                      in case f of
                             Const(c,T) => 
                                 if member (op =) ccs c
                                 then fold_rev add_hvars args hvars
-                                else add_term_vars (tm, hvars)
-                          | _ => add_term_vars (tm, hvars)
+                                else OldTerm.add_term_vars (tm, hvars)
+                          | _ => OldTerm.add_term_vars (tm, hvars)
                      end
             | _ => hvars;
   in add_hvars end;
 
 fun add_new_asm_vars new_asms =
     let fun itf (tm, at) vars =
-                if at then vars else add_term_vars(tm,vars)
+                if at then vars else OldTerm.add_term_vars(tm,vars)
         fun add_list(tm,al,vars) = let val (_,tml) = strip_comb tm
                 in if length(tml)=length(al)
                    then fold_rev itf (tml ~~ al) vars
@@ -198,7 +197,7 @@
     val hvars = add_new_asm_vars new_asms (rhs,hvars)
     fun it_asms asm hvars =
         if atomic asm then add_new_asm_vars new_asms (asm,hvars)
-        else add_term_frees(asm,hvars)
+        else OldTerm.add_term_frees(asm,hvars)
     val hvars = fold_rev it_asms asms hvars
     val hvs = map (#1 o dest_Var) hvars
     val refl1_tac = refl_tac 1
@@ -542,7 +541,7 @@
 fun find_subst sg T =
 let fun find (thm::thms) =
         let val (Const(_,cT), va, vb) = dest_red(hd(prems_of thm));
-            val [P] = add_term_vars(concl_of thm,[]) \\ [va,vb]
+            val [P] = OldTerm.add_term_vars(concl_of thm,[]) \\ [va,vb]
             val eqT::_ = binder_types cT
         in if Sign.typ_instance sg (T,eqT) then SOME(thm,va,vb,P)
            else find thms
--- a/src/HOL/Algebra/IntRing.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Algebra/IntRing.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -348,8 +348,8 @@
   "(Idl\<^bsub>\<Z>\<^esub> {k} \<subseteq> Idl\<^bsub>\<Z>\<^esub> {l}) = (l dvd k)"
 apply (subst int_Idl_subset_ideal, subst int_Idl, simp)
 apply (rule, clarify)
-apply (simp add: dvd_def, clarify)
-apply (simp add: int.m_comm)
+apply (simp add: dvd_def)
+apply (simp add: dvd_def mult_ac)
 done
 
 lemma dvds_eq_Idl:
--- a/src/HOL/Algebra/abstract/Ring2.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Algebra/abstract/Ring2.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -1,8 +1,7 @@
-(*
-  Title:     The algebraic hierarchy of rings as axiomatic classes
-  Id:        $Id$
-  Author:    Clemens Ballarin, started 9 December 1996
-  Copyright: Clemens Ballarin
+(*  Title:     HOL/Algebra/abstract/Ring2.thy
+    Author:    Clemens Ballarin
+
+The algebraic hierarchy of rings as axiomatic classes.
 *)
 
 header {* The algebraic hierarchy of rings as type classes *}
@@ -211,7 +210,7 @@
         @{const_name HOL.minus}, @{const_name HOL.one}, @{const_name HOL.times}]
   | ring_ord _ = ~1;
 
-fun termless_ring (a, b) = (Term.term_lpo ring_ord (a, b) = LESS);
+fun termless_ring (a, b) = (TermOrd.term_lpo ring_ord (a, b) = LESS);
 
 val ring_ss = HOL_basic_ss settermless termless_ring addsimps
   [thm "a_assoc", thm "l_zero", thm "l_neg", thm "a_comm", thm "m_assoc",
--- a/src/HOL/Algebra/ringsimp.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Algebra/ringsimp.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,6 +1,4 @@
-(*
-  Id:        $Id$
-  Author:    Clemens Ballarin
+(*  Author:    Clemens Ballarin
 
 Normalisation method for locales ring and cring.
 *)
@@ -48,7 +46,7 @@
     val ops = map (fst o Term.strip_comb) ts;
     fun ord (Const (a, _)) = find_index (fn (Const (b, _)) => a=b | _ => false) ops
       | ord (Free (a, _)) = find_index (fn (Free (b, _)) => a=b | _ => false) ops;
-    fun less (a, b) = (Term.term_lpo ord (a, b) = LESS);
+    fun less (a, b) = (TermOrd.term_lpo ord (a, b) = LESS);
   in asm_full_simp_tac (HOL_ss settermless less addsimps simps) end;
 
 fun algebra_tac ctxt =
--- a/src/HOL/Bali/AxExample.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Bali/AxExample.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -41,7 +41,7 @@
 
 ML {*
 fun inst1_tac ctxt s t st =
-  case AList.lookup (op =) (rev (Term.add_varnames (Thm.prop_of st) [])) s of
+  case AList.lookup (op =) (rev (Term.add_var_names (Thm.prop_of st) [])) s of
   SOME i => instantiate_tac ctxt [((s, i), t)] st | NONE => Seq.empty;
 
 val ax_tac =
--- a/src/HOL/Complex_Main.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Complex_Main.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -1,9 +1,4 @@
-(*  Title:      HOL/Complex_Main.thy
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   2003  University of Cambridge
-*)
-
-header{*Comprehensive Complex Theory*}
+header {* Comprehensive Complex Theory *}
 
 theory Complex_Main
 imports
--- a/src/HOL/Dense_Linear_Order.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Dense_Linear_Order.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -301,7 +301,7 @@
 
 text {* Linear order without upper bounds *}
 
-class_locale linorder_stupid_syntax = linorder
+locale linorder_stupid_syntax = linorder
 begin
 notation
   less_eq  ("op \<sqsubseteq>") and
@@ -311,7 +311,7 @@
 
 end
 
-class_locale linorder_no_ub = linorder_stupid_syntax +
+locale linorder_no_ub = linorder_stupid_syntax +
   assumes gt_ex: "\<exists>y. less x y"
 begin
 lemma ge_ex: "\<exists>y. x \<sqsubseteq> y" using gt_ex by auto
@@ -360,7 +360,7 @@
 
 text {* Linear order without upper bounds *}
 
-class_locale linorder_no_lb = linorder_stupid_syntax +
+locale linorder_no_lb = linorder_stupid_syntax +
   assumes lt_ex: "\<exists>y. less y x"
 begin
 lemma le_ex: "\<exists>y. y \<sqsubseteq> x" using lt_ex by auto
@@ -407,12 +407,12 @@
 end
 
 
-class_locale constr_dense_linear_order = linorder_no_lb + linorder_no_ub +
+locale constr_dense_linear_order = linorder_no_lb + linorder_no_ub +
   fixes between
   assumes between_less: "less x y \<Longrightarrow> less x (between x y) \<and> less (between x y) y"
      and  between_same: "between x x = x"
 
-class_interpretation  constr_dense_linear_order < dense_linear_order 
+sublocale  constr_dense_linear_order < dense_linear_order 
   apply unfold_locales
   using gt_ex lt_ex between_less
     by (auto, rule_tac x="between x y" in exI, simp)
@@ -635,9 +635,9 @@
   using eq_diff_eq[where a= x and b=t and c=0] by simp
 
 
-class_interpretation class_ordered_field_dense_linear_order: constr_dense_linear_order
- ["op <=" "op <"
-   "\<lambda> x y. 1/2 * ((x::'a::{ordered_field,recpower,number_ring}) + y)"]
+interpretation class_ordered_field_dense_linear_order!: constr_dense_linear_order
+ "op <=" "op <"
+   "\<lambda> x y. 1/2 * ((x::'a::{ordered_field,recpower,number_ring}) + y)"
 proof (unfold_locales, dlo, dlo, auto)
   fix x y::'a assume lt: "x < y"
   from  less_half_sum[OF lt] show "x < (x + y) /2" by simp
--- a/src/HOL/Deriv.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Deriv.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -9,7 +9,7 @@
 header{* Differentiation *}
 
 theory Deriv
-imports Lim Univ_Poly
+imports Lim Polynomial
 begin
 
 text{*Standard Definitions*}
@@ -1412,34 +1412,71 @@
 lemma lemma_DERIV_subst: "[| DERIV f x :> D; D = E |] ==> DERIV f x :> E"
 by auto
 
+
 subsection {* Derivatives of univariate polynomials *}
 
-
-  
-primrec pderiv_aux :: "nat => real list => real list" where
-   pderiv_aux_Nil:  "pderiv_aux n [] = []"
-|  pderiv_aux_Cons: "pderiv_aux n (h#t) =
-                     (real n * h)#(pderiv_aux (Suc n) t)"
-
 definition
-  pderiv :: "real list => real list" where
-  "pderiv p = (if p = [] then [] else pderiv_aux 1 (tl p))"
+  pderiv :: "'a::real_normed_field poly \<Rightarrow> 'a poly" where
+  "pderiv = poly_rec 0 (\<lambda>a p p'. p + pCons 0 p')"
+
+lemma pderiv_0 [simp]: "pderiv 0 = 0"
+  unfolding pderiv_def by (simp add: poly_rec_0)
 
+lemma pderiv_pCons: "pderiv (pCons a p) = p + pCons 0 (pderiv p)"
+  unfolding pderiv_def by (simp add: poly_rec_pCons)
+
+lemma coeff_pderiv: "coeff (pderiv p) n = of_nat (Suc n) * coeff p (Suc n)"
+  apply (induct p arbitrary: n, simp)
+  apply (simp add: pderiv_pCons coeff_pCons ring_simps split: nat.split)
+  done
 
-text{*The derivative*}
+lemma pderiv_eq_0_iff: "pderiv p = 0 \<longleftrightarrow> degree p = 0"
+  apply (rule iffI)
+  apply (cases p, simp)
+  apply (simp add: expand_poly_eq coeff_pderiv del: of_nat_Suc)
+  apply (simp add: expand_poly_eq coeff_pderiv coeff_eq_0)
+  done
 
-lemma pderiv_Nil: "pderiv [] = []"
+lemma degree_pderiv: "degree (pderiv p) = degree p - 1"
+  apply (rule order_antisym [OF degree_le])
+  apply (simp add: coeff_pderiv coeff_eq_0)
+  apply (cases "degree p", simp)
+  apply (rule le_degree)
+  apply (simp add: coeff_pderiv del: of_nat_Suc)
+  apply (rule subst, assumption)
+  apply (rule leading_coeff_neq_0, clarsimp)
+  done
 
-apply (simp add: pderiv_def)
-done
-declare pderiv_Nil [simp]
+lemma pderiv_singleton [simp]: "pderiv [:a:] = 0"
+by (simp add: pderiv_pCons)
+
+lemma pderiv_add: "pderiv (p + q) = pderiv p + pderiv q"
+by (rule poly_ext, simp add: coeff_pderiv ring_simps)
+
+lemma pderiv_minus: "pderiv (- p) = - pderiv p"
+by (rule poly_ext, simp add: coeff_pderiv)
+
+lemma pderiv_diff: "pderiv (p - q) = pderiv p - pderiv q"
+by (rule poly_ext, simp add: coeff_pderiv ring_simps)
+
+lemma pderiv_smult: "pderiv (smult a p) = smult a (pderiv p)"
+by (rule poly_ext, simp add: coeff_pderiv ring_simps)
 
-lemma pderiv_singleton: "pderiv [c] = []"
-by (simp add: pderiv_def)
-declare pderiv_singleton [simp]
+lemma pderiv_mult: "pderiv (p * q) = p * pderiv q + q * pderiv p"
+apply (induct p)
+apply simp
+apply (simp add: pderiv_add pderiv_smult pderiv_pCons ring_simps)
+done
 
-lemma pderiv_Cons: "pderiv (h#t) = pderiv_aux 1 t"
-by (simp add: pderiv_def)
+lemma pderiv_power_Suc:
+  "pderiv (p ^ Suc n) = smult (of_nat (Suc n)) (p ^ n) * pderiv p"
+apply (induct n)
+apply simp
+apply (subst power_Suc)
+apply (subst pderiv_mult)
+apply (erule ssubst)
+apply (simp add: smult_add_left ring_simps)
+done
 
 lemma DERIV_cmult2: "DERIV f x :> D ==> DERIV (%x. (f x) * c :: real) x :> D * c"
 by (simp add: DERIV_cmult mult_commute [of _ c])
@@ -1448,33 +1485,18 @@
 by (rule lemma_DERIV_subst, rule DERIV_pow, simp)
 declare DERIV_pow2 [simp] DERIV_pow [simp]
 
-lemma lemma_DERIV_poly1: "\<forall>n. DERIV (%x. (x ^ (Suc n) * poly p x)) x :>
-        x ^ n * poly (pderiv_aux (Suc n) p) x "
-apply (induct "p")
-apply (auto intro!: DERIV_add DERIV_cmult2 
-            simp add: pderiv_def right_distrib real_mult_assoc [symmetric] 
-            simp del: realpow_Suc)
-apply (subst mult_commute) 
-apply (simp del: realpow_Suc) 
-apply (simp add: mult_commute realpow_Suc [symmetric] del: realpow_Suc)
-done
-
-lemma lemma_DERIV_poly: "DERIV (%x. (x ^ (Suc n) * poly p x)) x :>
-        x ^ n * poly (pderiv_aux (Suc n) p) x "
-by (simp add: lemma_DERIV_poly1 del: realpow_Suc)
-
-lemma DERIV_add_const: "DERIV f x :> D ==>  DERIV (%x. a + f x :: real) x :> D"
+lemma DERIV_add_const: "DERIV f x :> D ==>  DERIV (%x. a + f x :: 'a::real_normed_field) x :> D"
 by (rule lemma_DERIV_subst, rule DERIV_add, auto)
 
 lemma poly_DERIV[simp]: "DERIV (%x. poly p x) x :> poly (pderiv p) x"
-apply (induct "p")
-apply (auto simp add: pderiv_Cons)
-apply (rule DERIV_add_const)
+apply (induct p)
+apply simp
+apply (simp add: pderiv_pCons)
 apply (rule lemma_DERIV_subst)
-apply (rule lemma_DERIV_poly [where n=0, simplified], simp) 
+apply (rule DERIV_add DERIV_mult DERIV_const DERIV_ident | assumption)+
+apply simp
 done
 
-
 text{* Consequences of the derivative theorem above*}
 
 lemma poly_differentiable[simp]: "(%x. poly p x) differentiable (x::real)"
@@ -1493,11 +1515,9 @@
 
 lemma poly_IVT_neg: "[| (a::real) < b; 0 < poly p a; poly p b < 0 |]
       ==> \<exists>x. a < x & x < b & (poly p x = 0)"
-apply (insert poly_IVT_pos [where p = "-- p" ]) 
-apply (simp add: poly_minus neg_less_0_iff_less) 
-done
+by (insert poly_IVT_pos [where p = "- p" ]) simp
 
-lemma poly_MVT: "a < b ==>
+lemma poly_MVT: "(a::real) < b ==>
      \<exists>x. a < x & x < b & (poly p b - poly p a = (b - a) * poly (pderiv p) x)"
 apply (drule_tac f = "poly p" in MVT, auto)
 apply (rule_tac x = z in exI)
@@ -1506,136 +1526,7 @@
 
 text{*Lemmas for Derivatives*}
 
-lemma lemma_poly_pderiv_aux_add: "\<forall>p2 n. poly (pderiv_aux n (p1 +++ p2)) x =
-                poly (pderiv_aux n p1 +++ pderiv_aux n p2) x"
-apply (induct "p1", simp, clarify) 
-apply (case_tac "p2")
-apply (auto simp add: right_distrib)
-done
-
-lemma poly_pderiv_aux_add: "poly (pderiv_aux n (p1 +++ p2)) x =
-      poly (pderiv_aux n p1 +++ pderiv_aux n p2) x"
-apply (simp add: lemma_poly_pderiv_aux_add)
-done
-
-lemma lemma_poly_pderiv_aux_cmult: "\<forall>n. poly (pderiv_aux n (c %* p)) x = poly (c %* pderiv_aux n p) x"
-apply (induct "p")
-apply (auto simp add: poly_cmult mult_ac)
-done
-
-lemma poly_pderiv_aux_cmult: "poly (pderiv_aux n (c %* p)) x = poly (c %* pderiv_aux n p) x"
-by (simp add: lemma_poly_pderiv_aux_cmult)
-
-lemma poly_pderiv_aux_minus:
-   "poly (pderiv_aux n (-- p)) x = poly (-- pderiv_aux n p) x"
-apply (simp add: poly_minus_def poly_pderiv_aux_cmult)
-done
-
-lemma lemma_poly_pderiv_aux_mult1: "\<forall>n. poly (pderiv_aux (Suc n) p) x = poly ((pderiv_aux n p) +++ p) x"
-apply (induct "p")
-apply (auto simp add: real_of_nat_Suc left_distrib)
-done
-
-lemma lemma_poly_pderiv_aux_mult: "poly (pderiv_aux (Suc n) p) x = poly ((pderiv_aux n p) +++ p) x"
-by (simp add: lemma_poly_pderiv_aux_mult1)
-
-lemma lemma_poly_pderiv_add: "\<forall>q. poly (pderiv (p +++ q)) x = poly (pderiv p +++ pderiv q) x"
-apply (induct "p", simp, clarify) 
-apply (case_tac "q")
-apply (auto simp add: poly_pderiv_aux_add poly_add pderiv_def)
-done
-
-lemma poly_pderiv_add: "poly (pderiv (p +++ q)) x = poly (pderiv p +++ pderiv q) x"
-by (simp add: lemma_poly_pderiv_add)
-
-lemma poly_pderiv_cmult: "poly (pderiv (c %* p)) x = poly (c %* (pderiv p)) x"
-apply (induct "p")
-apply (auto simp add: poly_pderiv_aux_cmult poly_cmult pderiv_def)
-done
-
-lemma poly_pderiv_minus: "poly (pderiv (--p)) x = poly (--(pderiv p)) x"
-by (simp add: poly_minus_def poly_pderiv_cmult)
-
-lemma lemma_poly_mult_pderiv:
-   "poly (pderiv (h#t)) x = poly ((0 # (pderiv t)) +++ t) x"
-apply (simp add: pderiv_def)
-apply (induct "t")
-apply (auto simp add: poly_add lemma_poly_pderiv_aux_mult)
-done
-
-lemma poly_pderiv_mult: "\<forall>q. poly (pderiv (p *** q)) x =
-      poly (p *** (pderiv q) +++ q *** (pderiv p)) x"
-apply (induct "p")
-apply (auto simp add: poly_add poly_cmult poly_pderiv_cmult poly_pderiv_add poly_mult)
-apply (rule lemma_poly_mult_pderiv [THEN ssubst])
-apply (rule lemma_poly_mult_pderiv [THEN ssubst])
-apply (rule poly_add [THEN ssubst])
-apply (rule poly_add [THEN ssubst])
-apply (simp (no_asm_simp) add: poly_mult right_distrib add_ac mult_ac)
-done
-
-lemma poly_pderiv_exp: "poly (pderiv (p %^ (Suc n))) x =
-         poly ((real (Suc n)) %* (p %^ n) *** pderiv p) x"
-apply (induct "n")
-apply (auto simp add: poly_add poly_pderiv_cmult poly_cmult poly_pderiv_mult
-                      real_of_nat_zero poly_mult real_of_nat_Suc 
-                      right_distrib left_distrib mult_ac)
-done
-
-lemma poly_pderiv_exp_prime: "poly (pderiv ([-a, 1] %^ (Suc n))) x =
-      poly (real (Suc n) %* ([-a, 1] %^ n)) x"
-apply (simp add: poly_pderiv_exp poly_mult del: pexp_Suc)
-apply (simp add: poly_cmult pderiv_def)
-done
-
-
-lemma real_mult_zero_disj_iff[simp]: "(x * y = 0) = (x = (0::real) | y = 0)"
-by simp
-
-lemma pderiv_aux_iszero [rule_format, simp]:
-    "\<forall>n. list_all (%c. c = 0) (pderiv_aux (Suc n) p) = list_all (%c. c = 0) p"
-by (induct "p", auto)
-
-lemma pderiv_aux_iszero_num: "(number_of n :: nat) \<noteq> 0
-      ==> (list_all (%c. c = 0) (pderiv_aux (number_of n) p) =
-      list_all (%c. c = 0) p)"
-unfolding neq0_conv
-apply (rule_tac n1 = "number_of n" and m1 = 0 in less_imp_Suc_add [THEN exE], force)
-apply (rule_tac n1 = "0 + x" in pderiv_aux_iszero [THEN subst])
-apply (simp (no_asm_simp) del: pderiv_aux_iszero)
-done
-
-instance real:: idom_char_0
-apply (intro_classes)
-done
-
-instance real:: recpower_idom_char_0
-apply (intro_classes)
-done
-
-lemma pderiv_iszero [rule_format]:
-     "poly (pderiv p) = poly [] --> (\<exists>h. poly p = poly [h])"
-apply (simp add: poly_zero)
-apply (induct "p", force)
-apply (simp add: pderiv_Cons pderiv_aux_iszero_num del: poly_Cons)
-apply (auto simp add: poly_zero [symmetric])
-done
-
-lemma pderiv_zero_obj: "poly p = poly [] --> (poly (pderiv p) = poly [])"
-apply (simp add: poly_zero)
-apply (induct "p", force)
-apply (simp add: pderiv_Cons pderiv_aux_iszero_num del: poly_Cons)
-done
-
-lemma pderiv_zero: "poly p = poly [] ==> (poly (pderiv p) = poly [])"
-by (blast elim: pderiv_zero_obj [THEN impE])
-declare pderiv_zero [simp]
-
-lemma poly_pderiv_welldef: "poly p = poly q ==> (poly (pderiv p) = poly (pderiv q))"
-apply (cut_tac p = "p +++ --q" in pderiv_zero_obj)
-apply (simp add: fun_eq poly_add poly_minus poly_pderiv_add poly_pderiv_minus del: pderiv_zero)
-done
-
+(* FIXME
 lemma lemma_order_pderiv [rule_format]:
      "\<forall>p q a. 0 < n &
        poly (pderiv p) \<noteq> poly [] &
@@ -1756,7 +1647,7 @@
 apply (simp (no_asm) add: rsquarefree_def order_root del: pmult_Cons)
 apply (simp add: poly_entire del: pmult_Cons)
 done
-
+*)
 
 subsection {* Theorems about Limits *}
 
--- a/src/HOL/Divides.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Divides.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -20,7 +20,7 @@
 
 subsection {* Abstract division in commutative semirings. *}
 
-class semiring_div = comm_semiring_1_cancel + div + 
+class semiring_div = comm_semiring_1_cancel + div +
   assumes mod_div_equality: "a div b * b + a mod b = a"
     and div_by_0 [simp]: "a div 0 = 0"
     and div_0 [simp]: "0 div a = 0"
@@ -33,6 +33,10 @@
   unfolding mult_commute [of b]
   by (rule mod_div_equality)
 
+lemma mod_div_equality': "a mod b + a div b * b = a"
+  using mod_div_equality [of a b]
+  by (simp only: add_ac)
+
 lemma div_mod_equality: "((a div b) * b + a mod b) + c = a + c"
   by (simp add: mod_div_equality)
 
@@ -143,6 +147,173 @@
   then show "b mod a = 0" by simp
 qed
 
+lemma mod_div_trivial [simp]: "a mod b div b = 0"
+proof (cases "b = 0")
+  assume "b = 0"
+  thus ?thesis by simp
+next
+  assume "b \<noteq> 0"
+  hence "a div b + a mod b div b = (a mod b + a div b * b) div b"
+    by (rule div_mult_self1 [symmetric])
+  also have "\<dots> = a div b"
+    by (simp only: mod_div_equality')
+  also have "\<dots> = a div b + 0"
+    by simp
+  finally show ?thesis
+    by (rule add_left_imp_eq)
+qed
+
+lemma mod_mod_trivial [simp]: "a mod b mod b = a mod b"
+proof -
+  have "a mod b mod b = (a mod b + a div b * b) mod b"
+    by (simp only: mod_mult_self1)
+  also have "\<dots> = a mod b"
+    by (simp only: mod_div_equality')
+  finally show ?thesis .
+qed
+
+text {* Addition respects modular equivalence. *}
+
+lemma mod_add_left_eq: "(a + b) mod c = (a mod c + b) mod c"
+proof -
+  have "(a + b) mod c = (a div c * c + a mod c + b) mod c"
+    by (simp only: mod_div_equality)
+  also have "\<dots> = (a mod c + b + a div c * c) mod c"
+    by (simp only: add_ac)
+  also have "\<dots> = (a mod c + b) mod c"
+    by (rule mod_mult_self1)
+  finally show ?thesis .
+qed
+
+lemma mod_add_right_eq: "(a + b) mod c = (a + b mod c) mod c"
+proof -
+  have "(a + b) mod c = (a + (b div c * c + b mod c)) mod c"
+    by (simp only: mod_div_equality)
+  also have "\<dots> = (a + b mod c + b div c * c) mod c"
+    by (simp only: add_ac)
+  also have "\<dots> = (a + b mod c) mod c"
+    by (rule mod_mult_self1)
+  finally show ?thesis .
+qed
+
+lemma mod_add_eq: "(a + b) mod c = (a mod c + b mod c) mod c"
+by (rule trans [OF mod_add_left_eq mod_add_right_eq])
+
+lemma mod_add_cong:
+  assumes "a mod c = a' mod c"
+  assumes "b mod c = b' mod c"
+  shows "(a + b) mod c = (a' + b') mod c"
+proof -
+  have "(a mod c + b mod c) mod c = (a' mod c + b' mod c) mod c"
+    unfolding assms ..
+  thus ?thesis
+    by (simp only: mod_add_eq [symmetric])
+qed
+
+text {* Multiplication respects modular equivalence. *}
+
+lemma mod_mult_left_eq: "(a * b) mod c = ((a mod c) * b) mod c"
+proof -
+  have "(a * b) mod c = ((a div c * c + a mod c) * b) mod c"
+    by (simp only: mod_div_equality)
+  also have "\<dots> = (a mod c * b + a div c * b * c) mod c"
+    by (simp only: left_distrib right_distrib add_ac mult_ac)
+  also have "\<dots> = (a mod c * b) mod c"
+    by (rule mod_mult_self1)
+  finally show ?thesis .
+qed
+
+lemma mod_mult_right_eq: "(a * b) mod c = (a * (b mod c)) mod c"
+proof -
+  have "(a * b) mod c = (a * (b div c * c + b mod c)) mod c"
+    by (simp only: mod_div_equality)
+  also have "\<dots> = (a * (b mod c) + a * (b div c) * c) mod c"
+    by (simp only: left_distrib right_distrib add_ac mult_ac)
+  also have "\<dots> = (a * (b mod c)) mod c"
+    by (rule mod_mult_self1)
+  finally show ?thesis .
+qed
+
+lemma mod_mult_eq: "(a * b) mod c = ((a mod c) * (b mod c)) mod c"
+by (rule trans [OF mod_mult_left_eq mod_mult_right_eq])
+
+lemma mod_mult_cong:
+  assumes "a mod c = a' mod c"
+  assumes "b mod c = b' mod c"
+  shows "(a * b) mod c = (a' * b') mod c"
+proof -
+  have "(a mod c * (b mod c)) mod c = (a' mod c * (b' mod c)) mod c"
+    unfolding assms ..
+  thus ?thesis
+    by (simp only: mod_mult_eq [symmetric])
+qed
+
+lemma mod_mod_cancel:
+  assumes "c dvd b"
+  shows "a mod b mod c = a mod c"
+proof -
+  from `c dvd b` obtain k where "b = c * k"
+    by (rule dvdE)
+  have "a mod b mod c = a mod (c * k) mod c"
+    by (simp only: `b = c * k`)
+  also have "\<dots> = (a mod (c * k) + a div (c * k) * k * c) mod c"
+    by (simp only: mod_mult_self1)
+  also have "\<dots> = (a div (c * k) * (c * k) + a mod (c * k)) mod c"
+    by (simp only: add_ac mult_ac)
+  also have "\<dots> = a mod c"
+    by (simp only: mod_div_equality)
+  finally show ?thesis .
+qed
+
+end
+
+class ring_div = semiring_div + comm_ring_1
+begin
+
+text {* Negation respects modular equivalence. *}
+
+lemma mod_minus_eq: "(- a) mod b = (- (a mod b)) mod b"
+proof -
+  have "(- a) mod b = (- (a div b * b + a mod b)) mod b"
+    by (simp only: mod_div_equality)
+  also have "\<dots> = (- (a mod b) + - (a div b) * b) mod b"
+    by (simp only: minus_add_distrib minus_mult_left add_ac)
+  also have "\<dots> = (- (a mod b)) mod b"
+    by (rule mod_mult_self1)
+  finally show ?thesis .
+qed
+
+lemma mod_minus_cong:
+  assumes "a mod b = a' mod b"
+  shows "(- a) mod b = (- a') mod b"
+proof -
+  have "(- (a mod b)) mod b = (- (a' mod b)) mod b"
+    unfolding assms ..
+  thus ?thesis
+    by (simp only: mod_minus_eq [symmetric])
+qed
+
+text {* Subtraction respects modular equivalence. *}
+
+lemma mod_diff_left_eq: "(a - b) mod c = (a mod c - b) mod c"
+  unfolding diff_minus
+  by (intro mod_add_cong mod_minus_cong) simp_all
+
+lemma mod_diff_right_eq: "(a - b) mod c = (a - b mod c) mod c"
+  unfolding diff_minus
+  by (intro mod_add_cong mod_minus_cong) simp_all
+
+lemma mod_diff_eq: "(a - b) mod c = (a mod c - b mod c) mod c"
+  unfolding diff_minus
+  by (intro mod_add_cong mod_minus_cong) simp_all
+
+lemma mod_diff_cong:
+  assumes "a mod c = a' mod c"
+  assumes "b mod c = b' mod c"
+  shows "(a - b) mod c = (a' - b') mod c"
+  unfolding diff_minus using assms
+  by (intro mod_add_cong mod_minus_cong)
+
 end
 
 
@@ -467,22 +638,14 @@
 done
 
 lemma mod_mult1_eq: "(a*b) mod c = a*(b mod c) mod (c::nat)"
-apply (cases "c = 0", simp)
-apply (blast intro: divmod_rel [THEN divmod_rel_mult1_eq, THEN mod_eq])
-done
+  by (rule mod_mult_right_eq)
 
 lemma mod_mult1_eq': "(a*b) mod (c::nat) = ((a mod c) * b) mod c"
-  apply (rule trans)
-   apply (rule_tac s = "b*a mod c" in trans)
-    apply (rule_tac [2] mod_mult1_eq)
-   apply (simp_all add: mult_commute)
-  done
+  by (rule mod_mult_left_eq)
 
 lemma mod_mult_distrib_mod:
   "(a*b) mod (c::nat) = ((a mod c) * (b mod c)) mod c"
-apply (rule mod_mult1_eq' [THEN trans])
-apply (rule mod_mult1_eq)
-done
+  by (rule mod_mult_eq)
 
 lemma divmod_rel_add1_eq:
   "[| divmod_rel a c aq ar; divmod_rel b c bq br;  c > 0 |]
@@ -497,9 +660,7 @@
 done
 
 lemma mod_add1_eq: "(a+b) mod (c::nat) = (a mod c + b mod c) mod c"
-apply (cases "c = 0", simp)
-apply (blast intro: divmod_rel_add1_eq [THEN mod_eq] divmod_rel)
-done
+  by (rule mod_add_eq)
 
 lemma mod_lemma: "[| (0::nat) < c; r < b |] ==> b * (q mod c) + r < b * c"
   apply (cut_tac m = q and n = c in mod_less_divisor)
@@ -618,11 +779,11 @@
 apply (auto simp add: Suc_diff_le le_mod_geq)
 done
 
-lemma nat_mod_div_trivial [simp]: "m mod n div n = (0 :: nat)"
-  by (cases "n = 0") auto
+lemma nat_mod_div_trivial: "m mod n div n = (0 :: nat)"
+  by simp
 
-lemma nat_mod_mod_trivial [simp]: "m mod n mod n = (m mod n :: nat)"
-  by (cases "n = 0") auto
+lemma nat_mod_mod_trivial: "m mod n mod n = (m mod n :: nat)"
+  by simp
 
 
 subsubsection {* The Divides Relation *}
@@ -639,7 +800,7 @@
 
 text {* @{term "op dvd"} is a partial order *}
 
-class_interpretation dvd: order ["op dvd" "\<lambda>n m \<Colon> nat. n dvd m \<and> \<not> m dvd n"]
+interpretation dvd!: order "op dvd" "\<lambda>n m \<Colon> nat. n dvd m \<and> \<not> m dvd n"
   proof qed (auto intro: dvd_refl dvd_trans dvd_anti_sym)
 
 lemma dvd_diff: "[| k dvd m; k dvd n |] ==> k dvd (m-n :: nat)"
@@ -718,18 +879,6 @@
   apply (simp add: power_add)
   done
 
-lemma mod_add_left_eq: "((a::nat) + b) mod c = (a mod c + b) mod c"
-  apply (rule trans [symmetric])
-   apply (rule mod_add1_eq, simp)
-  apply (rule mod_add1_eq [symmetric])
-  done
-
-lemma mod_add_right_eq: "(a+b) mod (c::nat) = (a + (b mod c)) mod c"
-  apply (rule trans [symmetric])
-   apply (rule mod_add1_eq, simp)
-  apply (rule mod_add1_eq [symmetric])
-  done
-
 lemma nat_zero_less_power_iff [simp]: "(x^n > 0) = (x > (0::nat) | n=0)"
   by (induct n) auto
 
--- a/src/HOL/Finite_Set.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Finite_Set.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -873,7 +873,7 @@
 
 subsection {* Generalized summation over a set *}
 
-class_interpretation comm_monoid_add: comm_monoid_mult ["0::'a::comm_monoid_add" "op +"]
+interpretation comm_monoid_add!: comm_monoid_mult "0::'a::comm_monoid_add" "op +"
   proof qed (auto intro: add_assoc add_commute)
 
 definition setsum :: "('a => 'b) => 'a set => 'b::comm_monoid_add"
@@ -1760,7 +1760,7 @@
 proof (induct rule: finite_induct)
   case empty then show ?case by simp
 next
-  class_interpret ab_semigroup_mult ["op Un"]
+  interpret ab_semigroup_mult "op Un"
     proof qed auto
   case insert 
   then show ?case by simp
@@ -2198,7 +2198,7 @@
   assumes "finite A" "A \<noteq> {}"
   shows "x \<le> fold1 inf A \<longleftrightarrow> (\<forall>a\<in>A. x \<le> a)"
 proof -
-  class_interpret ab_semigroup_idem_mult [inf]
+  interpret ab_semigroup_idem_mult inf
     by (rule ab_semigroup_idem_mult_inf)
   show ?thesis using assms by (induct rule: finite_ne_induct) simp_all
 qed
@@ -2213,7 +2213,7 @@
   proof (induct rule: finite_ne_induct)
     case singleton thus ?case by simp
   next
-    class_interpret ab_semigroup_idem_mult [inf]
+    interpret ab_semigroup_idem_mult inf
       by (rule ab_semigroup_idem_mult_inf)
     case (insert x F)
     from insert(5) have "a = x \<or> a \<in> F" by simp
@@ -2288,7 +2288,7 @@
     and "A \<noteq> {}"
   shows "sup x (\<Sqinter>\<^bsub>fin\<^esub>A) = \<Sqinter>\<^bsub>fin\<^esub>{sup x a|a. a \<in> A}"
 proof -
-  class_interpret ab_semigroup_idem_mult [inf]
+  interpret ab_semigroup_idem_mult inf
     by (rule ab_semigroup_idem_mult_inf)
   from assms show ?thesis
     by (simp add: Inf_fin_def image_def
@@ -2303,7 +2303,7 @@
   case singleton thus ?case
     by (simp add: sup_Inf1_distrib [OF B] fold1_singleton_def [OF Inf_fin_def])
 next
-  class_interpret ab_semigroup_idem_mult [inf]
+  interpret ab_semigroup_idem_mult inf
     by (rule ab_semigroup_idem_mult_inf)
   case (insert x A)
   have finB: "finite {sup x b |b. b \<in> B}"
@@ -2333,7 +2333,7 @@
   assumes "finite A" and "A \<noteq> {}"
   shows "inf x (\<Squnion>\<^bsub>fin\<^esub>A) = \<Squnion>\<^bsub>fin\<^esub>{inf x a|a. a \<in> A}"
 proof -
-  class_interpret ab_semigroup_idem_mult [sup]
+  interpret ab_semigroup_idem_mult sup
     by (rule ab_semigroup_idem_mult_sup)
   from assms show ?thesis
     by (simp add: Sup_fin_def image_def hom_fold1_commute [where h="inf x", OF inf_sup_distrib1])
@@ -2357,7 +2357,7 @@
     thus ?thesis by(simp add: insert(1) B(1))
   qed
   have ne: "{inf a b |a b. a \<in> A \<and> b \<in> B} \<noteq> {}" using insert B by blast
-  class_interpret ab_semigroup_idem_mult [sup]
+  interpret ab_semigroup_idem_mult sup
     by (rule ab_semigroup_idem_mult_sup)
   have "inf (\<Squnion>\<^bsub>fin\<^esub>(insert x A)) (\<Squnion>\<^bsub>fin\<^esub>B) = inf (sup x (\<Squnion>\<^bsub>fin\<^esub>A)) (\<Squnion>\<^bsub>fin\<^esub>B)"
     using insert by (simp add: fold1_insert_idem_def [OF Sup_fin_def])
@@ -2386,7 +2386,7 @@
   assumes "finite A" and "A \<noteq> {}"
   shows "\<Sqinter>\<^bsub>fin\<^esub>A = Inf A"
 proof -
-  class_interpret ab_semigroup_idem_mult [inf]
+    interpret ab_semigroup_idem_mult inf
     by (rule ab_semigroup_idem_mult_inf)
   from assms show ?thesis
   unfolding Inf_fin_def by (induct A set: finite)
@@ -2397,7 +2397,7 @@
   assumes "finite A" and "A \<noteq> {}"
   shows "\<Squnion>\<^bsub>fin\<^esub>A = Sup A"
 proof -
-  class_interpret ab_semigroup_idem_mult [sup]
+  interpret ab_semigroup_idem_mult sup
     by (rule ab_semigroup_idem_mult_sup)
   from assms show ?thesis
   unfolding Sup_fin_def by (induct A set: finite)
@@ -2446,7 +2446,7 @@
   assumes "finite A" and "A \<noteq> {}"
   shows "x < fold1 min A \<longleftrightarrow> (\<forall>a\<in>A. x < a)"
 proof -
-  class_interpret ab_semigroup_idem_mult [min]
+  interpret ab_semigroup_idem_mult min
     by (rule ab_semigroup_idem_mult_min)
   from assms show ?thesis
   by (induct rule: finite_ne_induct)
@@ -2457,7 +2457,7 @@
   assumes "finite A" and "A \<noteq> {}"
   shows "fold1 min A \<le> x \<longleftrightarrow> (\<exists>a\<in>A. a \<le> x)"
 proof -
-  class_interpret ab_semigroup_idem_mult [min]
+  interpret ab_semigroup_idem_mult min
     by (rule ab_semigroup_idem_mult_min)
   from assms show ?thesis
   by (induct rule: finite_ne_induct)
@@ -2468,7 +2468,7 @@
   assumes "finite A" and "A \<noteq> {}"
   shows "fold1 min A < x \<longleftrightarrow> (\<exists>a\<in>A. a < x)"
 proof -
-  class_interpret ab_semigroup_idem_mult [min]
+  interpret ab_semigroup_idem_mult min
     by (rule ab_semigroup_idem_mult_min)
   from assms show ?thesis
   by (induct rule: finite_ne_induct)
@@ -2481,7 +2481,7 @@
 proof cases
   assume "A = B" thus ?thesis by simp
 next
-  class_interpret ab_semigroup_idem_mult [min]
+  interpret ab_semigroup_idem_mult min
     by (rule ab_semigroup_idem_mult_min)
   assume "A \<noteq> B"
   have B: "B = A \<union> (B-A)" using `A \<subseteq> B` by blast
@@ -2515,7 +2515,7 @@
   assumes "finite A" and "A \<noteq> {}"
   shows "Min (insert x A) = min x (Min A)"
 proof -
-  class_interpret ab_semigroup_idem_mult [min]
+  interpret ab_semigroup_idem_mult min
     by (rule ab_semigroup_idem_mult_min)
   from assms show ?thesis by (rule fold1_insert_idem_def [OF Min_def])
 qed
@@ -2524,7 +2524,7 @@
   assumes "finite A" and "A \<noteq> {}"
   shows "Max (insert x A) = max x (Max A)"
 proof -
-  class_interpret ab_semigroup_idem_mult [max]
+  interpret ab_semigroup_idem_mult max
     by (rule ab_semigroup_idem_mult_max)
   from assms show ?thesis by (rule fold1_insert_idem_def [OF Max_def])
 qed
@@ -2533,7 +2533,7 @@
   assumes "finite A" and "A \<noteq> {}"
   shows "Min A \<in> A"
 proof -
-  class_interpret ab_semigroup_idem_mult [min]
+  interpret ab_semigroup_idem_mult min
     by (rule ab_semigroup_idem_mult_min)
   from assms fold1_in show ?thesis by (fastsimp simp: Min_def min_def)
 qed
@@ -2542,7 +2542,7 @@
   assumes "finite A" and "A \<noteq> {}"
   shows "Max A \<in> A"
 proof -
-  class_interpret ab_semigroup_idem_mult [max]
+  interpret ab_semigroup_idem_mult max
     by (rule ab_semigroup_idem_mult_max)
   from assms fold1_in [of A] show ?thesis by (fastsimp simp: Max_def max_def)
 qed
@@ -2551,7 +2551,7 @@
   assumes "finite A" and "A \<noteq> {}" and "finite B" and "B \<noteq> {}"
   shows "Min (A \<union> B) = min (Min A) (Min B)"
 proof -
-  class_interpret ab_semigroup_idem_mult [min]
+  interpret ab_semigroup_idem_mult min
     by (rule ab_semigroup_idem_mult_min)
   from assms show ?thesis
     by (simp add: Min_def fold1_Un2)
@@ -2561,7 +2561,7 @@
   assumes "finite A" and "A \<noteq> {}" and "finite B" and "B \<noteq> {}"
   shows "Max (A \<union> B) = max (Max A) (Max B)"
 proof -
-  class_interpret ab_semigroup_idem_mult [max]
+  interpret ab_semigroup_idem_mult max
     by (rule ab_semigroup_idem_mult_max)
   from assms show ?thesis
     by (simp add: Max_def fold1_Un2)
@@ -2572,7 +2572,7 @@
     and "finite N" and "N \<noteq> {}"
   shows "h (Min N) = Min (h ` N)"
 proof -
-  class_interpret ab_semigroup_idem_mult [min]
+  interpret ab_semigroup_idem_mult min
     by (rule ab_semigroup_idem_mult_min)
   from assms show ?thesis
     by (simp add: Min_def hom_fold1_commute)
@@ -2583,7 +2583,7 @@
     and "finite N" and "N \<noteq> {}"
   shows "h (Max N) = Max (h ` N)"
 proof -
-  class_interpret ab_semigroup_idem_mult [max]
+  interpret ab_semigroup_idem_mult max
     by (rule ab_semigroup_idem_mult_max)
   from assms show ?thesis
     by (simp add: Max_def hom_fold1_commute [of h])
@@ -2593,7 +2593,7 @@
   assumes "finite A" and "x \<in> A"
   shows "Min A \<le> x"
 proof -
-  class_interpret lower_semilattice ["op \<le>" "op <" min]
+  interpret lower_semilattice "op \<le>" "op <" min
     by (rule min_lattice)
   from assms show ?thesis by (simp add: Min_def fold1_belowI)
 qed
@@ -2602,7 +2602,7 @@
   assumes "finite A" and "x \<in> A"
   shows "x \<le> Max A"
 proof -
-  invoke lower_semilattice ["op \<ge>" "op >" max]
+  interpret lower_semilattice "op \<ge>" "op >" max
     by (rule max_lattice)
   from assms show ?thesis by (simp add: Max_def fold1_belowI)
 qed
@@ -2611,7 +2611,7 @@
   assumes "finite A" and "A \<noteq> {}"
   shows "x \<le> Min A \<longleftrightarrow> (\<forall>a\<in>A. x \<le> a)"
 proof -
-  class_interpret lower_semilattice ["op \<le>" "op <" min]
+  interpret lower_semilattice "op \<le>" "op <" min
     by (rule min_lattice)
   from assms show ?thesis by (simp add: Min_def below_fold1_iff)
 qed
@@ -2620,7 +2620,7 @@
   assumes "finite A" and "A \<noteq> {}"
   shows "Max A \<le> x \<longleftrightarrow> (\<forall>a\<in>A. a \<le> x)"
 proof -
-  invoke lower_semilattice ["op \<ge>" "op >" max]
+  interpret lower_semilattice "op \<ge>" "op >" max
     by (rule max_lattice)
   from assms show ?thesis by (simp add: Max_def below_fold1_iff)
 qed
@@ -2629,7 +2629,7 @@
   assumes "finite A" and "A \<noteq> {}"
   shows "x < Min A \<longleftrightarrow> (\<forall>a\<in>A. x < a)"
 proof -
-  class_interpret lower_semilattice ["op \<le>" "op <" min]
+  interpret lower_semilattice "op \<le>" "op <" min
     by (rule min_lattice)
   from assms show ?thesis by (simp add: Min_def strict_below_fold1_iff)
 qed
@@ -2639,7 +2639,7 @@
   shows "Max A < x \<longleftrightarrow> (\<forall>a\<in>A. a < x)"
 proof -
   note Max = Max_def
-  class_interpret linorder ["op \<ge>" "op >"]
+  interpret linorder "op \<ge>" "op >"
     by (rule dual_linorder)
   from assms show ?thesis
     by (simp add: Max strict_below_fold1_iff [folded dual_max])
@@ -2649,7 +2649,7 @@
   assumes "finite A" and "A \<noteq> {}"
   shows "Min A \<le> x \<longleftrightarrow> (\<exists>a\<in>A. a \<le> x)"
 proof -
-  class_interpret lower_semilattice ["op \<le>" "op <" min]
+  interpret lower_semilattice "op \<le>" "op <" min
     by (rule min_lattice)
   from assms show ?thesis
     by (simp add: Min_def fold1_below_iff)
@@ -2660,7 +2660,7 @@
   shows "x \<le> Max A \<longleftrightarrow> (\<exists>a\<in>A. x \<le> a)"
 proof -
   note Max = Max_def
-  class_interpret linorder ["op \<ge>" "op >"]
+  interpret linorder "op \<ge>" "op >"
     by (rule dual_linorder)
   from assms show ?thesis
     by (simp add: Max fold1_below_iff [folded dual_max])
@@ -2670,7 +2670,7 @@
   assumes "finite A" and "A \<noteq> {}"
   shows "Min A < x \<longleftrightarrow> (\<exists>a\<in>A. a < x)"
 proof -
-  class_interpret lower_semilattice ["op \<le>" "op <" min]
+  interpret lower_semilattice "op \<le>" "op <" min
     by (rule min_lattice)
   from assms show ?thesis
     by (simp add: Min_def fold1_strict_below_iff)
@@ -2681,7 +2681,7 @@
   shows "x < Max A \<longleftrightarrow> (\<exists>a\<in>A. x < a)"
 proof -
   note Max = Max_def
-  class_interpret linorder ["op \<ge>" "op >"]
+  interpret linorder "op \<ge>" "op >"
     by (rule dual_linorder)
   from assms show ?thesis
     by (simp add: Max fold1_strict_below_iff [folded dual_max])
@@ -2691,7 +2691,7 @@
   assumes "M \<subseteq> N" and "M \<noteq> {}" and "finite N"
   shows "Min N \<le> Min M"
 proof -
-  class_interpret distrib_lattice ["op \<le>" "op <" min max]
+  interpret distrib_lattice "op \<le>" "op <" min max
     by (rule distrib_lattice_min_max)
   from assms show ?thesis by (simp add: Min_def fold1_antimono)
 qed
@@ -2701,7 +2701,7 @@
   shows "Max M \<le> Max N"
 proof -
   note Max = Max_def
-  class_interpret linorder ["op \<ge>" "op >"]
+  interpret linorder "op \<ge>" "op >"
     by (rule dual_linorder)
   from assms show ?thesis
     by (simp add: Max fold1_antimono [folded dual_max])
--- a/src/HOL/Fundamental_Theorem_Algebra.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Fundamental_Theorem_Algebra.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -3,7 +3,7 @@
 header{*Fundamental Theorem of Algebra*}
 
 theory Fundamental_Theorem_Algebra
-imports Univ_Poly Dense_Linear_Order Complex
+imports Polynomial Dense_Linear_Order Complex
 begin
 
 subsection {* Square root of complex numbers *}
@@ -16,7 +16,7 @@
 
 lemma csqrt[algebra]: "csqrt z ^ 2 = z"
 proof-
-  obtain x y where xy: "z = Complex x y" by (cases z, simp_all)
+  obtain x y where xy: "z = Complex x y" by (cases z)
   {assume y0: "y = 0"
     {assume x0: "x \<ge> 0" 
       then have ?thesis using y0 xy real_sqrt_pow2[OF x0]
@@ -70,10 +70,10 @@
 lemma poly_bound_exists:
   shows "\<exists>m. m > 0 \<and> (\<forall>z. cmod z <= r \<longrightarrow> cmod (poly p z) \<le> m)"
 proof(induct p)
-  case Nil thus ?case by (rule exI[where x=1], simp) 
+  case 0 thus ?case by (rule exI[where x=1], simp) 
 next
-  case (Cons c cs)
-  from Cons.hyps obtain m where m: "\<forall>z. cmod z \<le> r \<longrightarrow> cmod (poly cs z) \<le> m"
+  case (pCons c cs)
+  from pCons.hyps obtain m where m: "\<forall>z. cmod z \<le> r \<longrightarrow> cmod (poly cs z) \<le> m"
     by blast
   let ?k = " 1 + cmod c + \<bar>r * m\<bar>"
   have kp: "?k > 0" using abs_ge_zero[of "r*m"] norm_ge_zero[of c] by arith
@@ -81,48 +81,72 @@
     assume H: "cmod z \<le> r"
     from m H have th: "cmod (poly cs z) \<le> m" by blast
     from H have rp: "r \<ge> 0" using norm_ge_zero[of z] by arith
-    have "cmod (poly (c # cs) z) \<le> cmod c + cmod (z* poly cs z)"
+    have "cmod (poly (pCons c cs) z) \<le> cmod c + cmod (z* poly cs z)"
       using norm_triangle_ineq[of c "z* poly cs z"] by simp
     also have "\<dots> \<le> cmod c + r*m" using mult_mono[OF H th rp norm_ge_zero[of "poly cs z"]] by (simp add: norm_mult)
     also have "\<dots> \<le> ?k" by simp
-    finally have "cmod (poly (c # cs) z) \<le> ?k" .}
+    finally have "cmod (poly (pCons c cs) z) \<le> ?k" .}
   with kp show ?case by blast
 qed
 
 
 text{* Offsetting the variable in a polynomial gives another of same degree *}
-  (* FIXME : Lemma holds also in locale --- fix it later *)
-lemma  poly_offset_lemma:
-  shows "\<exists>b q. (length q = length p) \<and> (\<forall>x. poly (b#q) (x::complex) = (a + x) * poly p x)"
-proof(induct p)
-  case Nil thus ?case by simp
-next
-  case (Cons c cs)
-  from Cons.hyps obtain b q where 
-    bq: "length q = length cs" "\<forall>x. poly (b # q) x = (a + x) * poly cs x"
-    by blast
-  let ?b = "a*c"
-  let ?q = "(b+c)#q"
-  have lg: "length ?q = length (c#cs)" using bq(1) by simp
-  {fix x
-    from bq(2)[rule_format, of x]
-    have "x*poly (b # q) x = x*((a + x) * poly cs x)" by simp
-    hence "poly (?b# ?q) x = (a + x) * poly (c # cs) x"
-      by (simp add: ring_simps)}
-  with lg  show ?case by blast 
-qed
+
+definition
+  "offset_poly p h = poly_rec 0 (\<lambda>a p q. smult h q + pCons a q) p"
+
+lemma offset_poly_0: "offset_poly 0 h = 0"
+  unfolding offset_poly_def by (simp add: poly_rec_0)
+
+lemma offset_poly_pCons:
+  "offset_poly (pCons a p) h =
+    smult h (offset_poly p h) + pCons a (offset_poly p h)"
+  unfolding offset_poly_def by (simp add: poly_rec_pCons)
+
+lemma offset_poly_single: "offset_poly [:a:] h = [:a:]"
+by (simp add: offset_poly_pCons offset_poly_0)
+
+lemma poly_offset_poly: "poly (offset_poly p h) x = poly p (h + x)"
+apply (induct p)
+apply (simp add: offset_poly_0)
+apply (simp add: offset_poly_pCons ring_simps)
+done
+
+lemma offset_poly_eq_0_lemma: "smult c p + pCons a p = 0 \<Longrightarrow> p = 0"
+by (induct p arbitrary: a, simp, force)
 
-    (* FIXME : This one too*)
-lemma poly_offset: "\<exists> q. length q = length p \<and> (\<forall>x. poly q (x::complex) = poly p (a + x))"
-proof (induct p)
-  case Nil thus ?case by simp
-next
-  case (Cons c cs)
-  from Cons.hyps obtain q where q: "length q = length cs" "\<forall>x. poly q x = poly cs (a + x)" by blast
-  from poly_offset_lemma[of q a] obtain b p where 
-    bp: "length p = length q" "\<forall>x. poly (b # p) x = (a + x) * poly q x"
-    by blast
-  thus ?case using q bp by - (rule exI[where x="(c + b)#p"], simp)
+lemma offset_poly_eq_0_iff: "offset_poly p h = 0 \<longleftrightarrow> p = 0"
+apply (safe intro!: offset_poly_0)
+apply (induct p, simp)
+apply (simp add: offset_poly_pCons)
+apply (frule offset_poly_eq_0_lemma, simp)
+done
+
+lemma degree_offset_poly: "degree (offset_poly p h) = degree p"
+apply (induct p)
+apply (simp add: offset_poly_0)
+apply (case_tac "p = 0")
+apply (simp add: offset_poly_0 offset_poly_pCons)
+apply (simp add: offset_poly_pCons)
+apply (subst degree_add_eq_right)
+apply (rule le_less_trans [OF degree_smult_le])
+apply (simp add: offset_poly_eq_0_iff)
+apply (simp add: offset_poly_eq_0_iff)
+done
+
+definition
+  "psize p = (if p = 0 then 0 else Suc (degree p))"
+
+lemma psize_eq_0_iff [simp]: "psize p = 0 \<longleftrightarrow> p = 0"
+  unfolding psize_def by simp
+
+lemma poly_offset: "\<exists> q. psize q = psize p \<and> (\<forall>x. poly q (x::complex) = poly p (a + x))"
+proof (intro exI conjI)
+  show "psize (offset_poly p a) = psize p"
+    unfolding psize_def
+    by (simp add: offset_poly_eq_0_iff degree_offset_poly)
+  show "\<forall>x. poly (offset_poly p a) x = poly p (a + x)"
+    by (simp add: poly_offset_poly)
 qed
 
 text{* An alternative useful formulation of completeness of the reals *}
@@ -474,15 +498,21 @@
   assumes ep: "e > 0" 
   shows "\<exists>d >0. \<forall>w. 0 < cmod (w - z) \<and> cmod (w - z) < d \<longrightarrow> cmod (poly p w - poly p z) < e"
 proof-
-  from poly_offset[of p z] obtain q where q: "length q = length p" "\<And>x. poly q x = poly p (z + x)" by blast
+  obtain q where q: "degree q = degree p" "\<And>x. poly q x = poly p (z + x)"
+  proof
+    show "degree (offset_poly p z) = degree p"
+      by (rule degree_offset_poly)
+    show "\<And>x. poly (offset_poly p z) x = poly p (z + x)"
+      by (rule poly_offset_poly)
+  qed
   {fix w
     note q(2)[of "w - z", simplified]}
   note th = this
   show ?thesis unfolding th[symmetric]
   proof(induct q)
-    case Nil thus ?case  using ep by auto
+    case 0 thus ?case  using ep by auto
   next
-    case (Cons c cs)
+    case (pCons c cs)
     from poly_bound_exists[of 1 "cs"] 
     obtain m where m: "m > 0" "\<And>z. cmod z \<le> 1 \<Longrightarrow> cmod (poly cs z) \<le> m" by blast
     from ep m(1) have em0: "e/m > 0" by (simp add: field_simps)
@@ -621,36 +651,32 @@
 
 text {* Nonzero polynomial in z goes to infinity as z does. *}
 
-instance complex::idom_char_0 by (intro_classes)
-instance complex :: recpower_idom_char_0 by intro_classes
-
 lemma poly_infinity:
-  assumes ex: "list_ex (\<lambda>c. c \<noteq> 0) p"
-  shows "\<exists>r. \<forall>z. r \<le> cmod z \<longrightarrow> d \<le> cmod (poly (a#p) z)"
+  assumes ex: "p \<noteq> 0"
+  shows "\<exists>r. \<forall>z. r \<le> cmod z \<longrightarrow> d \<le> cmod (poly (pCons a p) z)"
 using ex
 proof(induct p arbitrary: a d)
-  case (Cons c cs a d) 
-  {assume H: "list_ex (\<lambda>c. c\<noteq>0) cs"
-    with Cons.hyps obtain r where r: "\<forall>z. r \<le> cmod z \<longrightarrow> d + cmod a \<le> cmod (poly (c # cs) z)" by blast
+  case (pCons c cs a d) 
+  {assume H: "cs \<noteq> 0"
+    with pCons.hyps obtain r where r: "\<forall>z. r \<le> cmod z \<longrightarrow> d + cmod a \<le> cmod (poly (pCons c cs) z)" by blast
     let ?r = "1 + \<bar>r\<bar>"
     {fix z assume h: "1 + \<bar>r\<bar> \<le> cmod z"
       have r0: "r \<le> cmod z" using h by arith
       from r[rule_format, OF r0]
-      have th0: "d + cmod a \<le> 1 * cmod(poly (c#cs) z)" by arith
+      have th0: "d + cmod a \<le> 1 * cmod(poly (pCons c cs) z)" by arith
       from h have z1: "cmod z \<ge> 1" by arith
-      from order_trans[OF th0 mult_right_mono[OF z1 norm_ge_zero[of "poly (c#cs) z"]]]
-      have th1: "d \<le> cmod(z * poly (c#cs) z) - cmod a"
+      from order_trans[OF th0 mult_right_mono[OF z1 norm_ge_zero[of "poly (pCons c cs) z"]]]
+      have th1: "d \<le> cmod(z * poly (pCons c cs) z) - cmod a"
 	unfolding norm_mult by (simp add: ring_simps)
-      from complex_mod_triangle_sub[of "z * poly (c#cs) z" a]
-      have th2: "cmod(z * poly (c#cs) z) - cmod a \<le> cmod (poly (a#c#cs) z)" 
+      from complex_mod_triangle_sub[of "z * poly (pCons c cs) z" a]
+      have th2: "cmod(z * poly (pCons c cs) z) - cmod a \<le> cmod (poly (pCons a (pCons c cs)) z)" 
 	by (simp add: diff_le_eq ring_simps) 
-      from th1 th2 have "d \<le> cmod (poly (a#c#cs) z)"  by arith}
+      from th1 th2 have "d \<le> cmod (poly (pCons a (pCons c cs)) z)"  by arith}
     hence ?case by blast}
   moreover
-  {assume cs0: "\<not> (list_ex (\<lambda>c. c \<noteq> 0) cs)"
-    with Cons.prems have c0: "c \<noteq> 0" by simp
-    from cs0 have cs0': "list_all (\<lambda>c. c = 0) cs" 
-      by (auto simp add: list_all_iff list_ex_iff)
+  {assume cs0: "\<not> (cs \<noteq> 0)"
+    with pCons.prems have c0: "c \<noteq> 0" by simp
+    from cs0 have cs0': "cs = 0" by simp
     {fix z
       assume h: "(\<bar>d\<bar> + cmod a) / cmod c \<le> cmod z"
       from c0 have "cmod c > 0" by simp
@@ -660,8 +686,8 @@
       from complex_mod_triangle_sub[of "z*c" a ]
       have th1: "cmod (z * c) \<le> cmod (a + z * c) + cmod a"
 	by (simp add: ring_simps)
-      from ath[OF th1 th0] have "d \<le> cmod (poly (a # c # cs) z)" 
-	using poly_0[OF cs0'] by simp}
+      from ath[OF th1 th0] have "d \<le> cmod (poly (pCons a (pCons c cs)) z)" 
+        using cs0' by simp}
     then have ?case  by blast}
   ultimately show ?case by blast
 qed simp
@@ -670,57 +696,53 @@
 lemma poly_minimum_modulus:
   "\<exists>z.\<forall>w. cmod (poly p z) \<le> cmod (poly p w)"
 proof(induct p)
-  case (Cons c cs) 
-  {assume cs0: "list_ex (\<lambda>c. c \<noteq> 0) cs"
-    from poly_infinity[OF cs0, of "cmod (poly (c#cs) 0)" c]
-    obtain r where r: "\<And>z. r \<le> cmod z \<Longrightarrow> cmod (poly (c # cs) 0) \<le> cmod (poly (c # cs) z)" by blast
+  case (pCons c cs) 
+  {assume cs0: "cs \<noteq> 0"
+    from poly_infinity[OF cs0, of "cmod (poly (pCons c cs) 0)" c]
+    obtain r where r: "\<And>z. r \<le> cmod z \<Longrightarrow> cmod (poly (pCons c cs) 0) \<le> cmod (poly (pCons c cs) z)" by blast
     have ath: "\<And>z r. r \<le> cmod z \<or> cmod z \<le> \<bar>r\<bar>" by arith
-    from poly_minimum_modulus_disc[of "\<bar>r\<bar>" "c#cs"] 
-    obtain v where v: "\<And>w. cmod w \<le> \<bar>r\<bar> \<Longrightarrow> cmod (poly (c # cs) v) \<le> cmod (poly (c # cs) w)" by blast
+    from poly_minimum_modulus_disc[of "\<bar>r\<bar>" "pCons c cs"] 
+    obtain v where v: "\<And>w. cmod w \<le> \<bar>r\<bar> \<Longrightarrow> cmod (poly (pCons c cs) v) \<le> cmod (poly (pCons c cs) w)" by blast
     {fix z assume z: "r \<le> cmod z"
       from v[of 0] r[OF z] 
-      have "cmod (poly (c # cs) v) \<le> cmod (poly (c # cs) z)"
+      have "cmod (poly (pCons c cs) v) \<le> cmod (poly (pCons c cs) z)"
 	by simp }
     note v0 = this
     from v0 v ath[of r] have ?case by blast}
   moreover
-  {assume cs0: "\<not> (list_ex (\<lambda>c. c\<noteq>0) cs)"
-    hence th:"list_all (\<lambda>c. c = 0) cs" by (simp add: list_all_iff list_ex_iff)
-    from poly_0[OF th] Cons.hyps have ?case by simp}
+  {assume cs0: "\<not> (cs \<noteq> 0)"
+    hence th:"cs = 0" by simp
+    from th pCons.hyps have ?case by simp}
   ultimately show ?case by blast
 qed simp
 
 text{* Constant function (non-syntactic characterization). *}
 definition "constant f = (\<forall>x y. f x = f y)"
 
-lemma nonconstant_length: "\<not> (constant (poly p)) \<Longrightarrow> length p \<ge> 2"
-  unfolding constant_def
+lemma nonconstant_length: "\<not> (constant (poly p)) \<Longrightarrow> psize p \<ge> 2"
+  unfolding constant_def psize_def
   apply (induct p, auto)
-  apply (unfold not_less[symmetric])
-  apply simp
-  apply (rule ccontr)
-  apply auto
   done
  
 lemma poly_replicate_append:
-  "poly ((replicate n 0)@p) (x::'a::{recpower, comm_ring}) = x^n * poly p x"
-  by(induct n, auto simp add: power_Suc ring_simps)
+  "poly (monom 1 n * p) (x::'a::{recpower, comm_ring_1}) = x^n * poly p x"
+  by (simp add: poly_monom)
 
 text {* Decomposition of polynomial, skipping zero coefficients 
   after the first.  *}
 
 lemma poly_decompose_lemma:
  assumes nz: "\<not>(\<forall>z. z\<noteq>0 \<longrightarrow> poly p z = (0::'a::{recpower,idom}))"
-  shows "\<exists>k a q. a\<noteq>0 \<and> Suc (length q + k) = length p \<and> 
-                 (\<forall>z. poly p z = z^k * poly (a#q) z)"
+  shows "\<exists>k a q. a\<noteq>0 \<and> Suc (psize q + k) = psize p \<and> 
+                 (\<forall>z. poly p z = z^k * poly (pCons a q) z)"
+unfolding psize_def
 using nz
 proof(induct p)
-  case Nil thus ?case by simp
+  case 0 thus ?case by simp
 next
-  case (Cons c cs)
+  case (pCons c cs)
   {assume c0: "c = 0"
-    
-    from Cons.hyps Cons.prems c0 have ?case apply auto
+    from pCons.hyps pCons.prems c0 have ?case apply auto
       apply (rule_tac x="k+1" in exI)
       apply (rule_tac x="a" in exI, clarsimp)
       apply (rule_tac x="q" in exI)
@@ -739,26 +761,27 @@
 lemma poly_decompose:
   assumes nc: "~constant(poly p)"
   shows "\<exists>k a q. a\<noteq>(0::'a::{recpower,idom}) \<and> k\<noteq>0 \<and>
-               length q + k + 1 = length p \<and> 
-              (\<forall>z. poly p z = poly p 0 + z^k * poly (a#q) z)"
+               psize q + k + 1 = psize p \<and> 
+              (\<forall>z. poly p z = poly p 0 + z^k * poly (pCons a q) z)"
 using nc 
 proof(induct p)
-  case Nil thus ?case by (simp add: constant_def)
+  case 0 thus ?case by (simp add: constant_def)
 next
-  case (Cons c cs)
+  case (pCons c cs)
   {assume C:"\<forall>z. z \<noteq> 0 \<longrightarrow> poly cs z = 0"
     {fix x y
-      from C have "poly (c#cs) x = poly (c#cs) y" by (cases "x=0", auto)}
-    with Cons.prems have False by (auto simp add: constant_def)}
+      from C have "poly (pCons c cs) x = poly (pCons c cs) y" by (cases "x=0", auto)}
+    with pCons.prems have False by (auto simp add: constant_def)}
   hence th: "\<not> (\<forall>z. z \<noteq> 0 \<longrightarrow> poly cs z = 0)" ..
   from poly_decompose_lemma[OF th] 
   show ?case 
-    apply clarsimp    
+    apply clarsimp
     apply (rule_tac x="k+1" in exI)
     apply (rule_tac x="a" in exI)
     apply simp
     apply (rule_tac x="q" in exI)
     apply (auto simp add: power_Suc)
+    apply (auto simp add: psize_def split: if_splits)
     done
 qed
 
@@ -768,10 +791,10 @@
   assumes nc: "~constant(poly p)"
   shows "\<exists>z::complex. poly p z = 0"
 using nc
-proof(induct n\<equiv> "length p" arbitrary: p rule: nat_less_induct)
-  fix n fix p :: "complex list"
+proof(induct n\<equiv> "psize p" arbitrary: p rule: nat_less_induct)
+  fix n fix p :: "complex poly"
   let ?p = "poly p"
-  assume H: "\<forall>m<n. \<forall>p. \<not> constant (poly p) \<longrightarrow> m = length p \<longrightarrow> (\<exists>(z::complex). poly p z = 0)" and nc: "\<not> constant ?p" and n: "n = length p"
+  assume H: "\<forall>m<n. \<forall>p. \<not> constant (poly p) \<longrightarrow> m = psize p \<longrightarrow> (\<exists>(z::complex). poly p z = 0)" and nc: "\<not> constant ?p" and n: "n = psize p"
   let ?ths = "\<exists>z. ?p z = 0"
 
   from nonconstant_length[OF nc] have n2: "n\<ge> 2" by (simp add: n)
@@ -781,7 +804,7 @@
   moreover
   {assume pc0: "?p c \<noteq> 0"
     from poly_offset[of p c] obtain q where
-      q: "length q = length p" "\<forall>x. poly q x = ?p (c+x)" by blast
+      q: "psize q = psize p" "\<forall>x. poly q x = ?p (c+x)" by blast
     {assume h: "constant (poly q)"
       from q(2) have th: "\<forall>x. poly q (x - c) = ?p x" by auto
       {fix x y
@@ -797,10 +820,12 @@
     let ?a0 = "poly q 0"
     from pc0 pqc0 have a00: "?a0 \<noteq> 0" by simp 
     from a00 
-    have qr: "\<forall>z. poly q z = poly (map (op * (inverse ?a0)) q) z * ?a0"
-      by (simp add: poly_cmult_map)
-    let ?r = "map (op * (inverse ?a0)) q"
-    have lgqr: "length q = length ?r" by simp 
+    have qr: "\<forall>z. poly q z = poly (smult (inverse ?a0) q) z * ?a0"
+      by simp
+    let ?r = "smult (inverse ?a0) q"
+    have lgqr: "psize q = psize ?r"
+      using a00 unfolding psize_def degree_def
+      by (simp add: expand_poly_eq)
     {assume h: "\<And>x y. poly ?r x = poly ?r y"
       {fix x y
 	from qr[rule_format, of x] 
@@ -813,16 +838,16 @@
     from qr[rule_format, of 0] a00  have r01: "poly ?r 0 = 1" by auto
     {fix w 
       have "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w / ?a0) < 1"
-	using qr[rule_format, of w] a00 by simp
+	using qr[rule_format, of w] a00 by (simp add: divide_inverse mult_ac)
       also have "\<dots> \<longleftrightarrow> cmod (poly q w) < cmod ?a0"
 	using a00 unfolding norm_divide by (simp add: field_simps)
       finally have "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w) < cmod ?a0" .}
     note mrmq_eq = this
     from poly_decompose[OF rnc] obtain k a s where 
-      kas: "a\<noteq>0" "k\<noteq>0" "length s + k + 1 = length ?r" 
-      "\<forall>z. poly ?r z = poly ?r 0 + z^k* poly (a#s) z" by blast
+      kas: "a\<noteq>0" "k\<noteq>0" "psize s + k + 1 = psize ?r" 
+      "\<forall>z. poly ?r z = poly ?r 0 + z^k* poly (pCons a s) z" by blast
     {assume "k + 1 = n"
-      with kas(3) lgqr[symmetric] q(1) n[symmetric] have s0:"s=[]" by auto
+      with kas(3) lgqr[symmetric] q(1) n[symmetric] have s0:"s=0" by auto
       {fix w
 	have "cmod (poly ?r w) = cmod (1 + a * w ^ k)" 
 	  using kas(4)[rule_format, of w] s0 r01 by (simp add: ring_simps)}
@@ -831,18 +856,17 @@
       have "\<exists>w. cmod (poly ?r w) < 1" unfolding hth by blast}
     moreover
     {assume kn: "k+1 \<noteq> n"
-      from kn kas(3) q(1) n[symmetric] have k1n: "k + 1 < n" by simp
-      have th01: "\<not> constant (poly (1#((replicate (k - 1) 0)@[a])))" 
-	unfolding constant_def poly_Nil poly_Cons poly_replicate_append
+      from kn kas(3) q(1) n[symmetric] lgqr have k1n: "k + 1 < n" by simp
+      have th01: "\<not> constant (poly (pCons 1 (monom a (k - 1))))" 
+	unfolding constant_def poly_pCons poly_monom
 	using kas(1) apply simp 
 	by (rule exI[where x=0], rule exI[where x=1], simp)
-      from kas(2) have th02: "k+1 = length (1#((replicate (k - 1) 0)@[a]))" 
-	by simp
+      from kas(1) kas(2) have th02: "k+1 = psize (pCons 1 (monom a (k - 1)))"
+	by (simp add: psize_def degree_monom_eq)
       from H[rule_format, OF k1n th01 th02]
       obtain w where w: "1 + w^k * a = 0"
-	unfolding poly_Nil poly_Cons poly_replicate_append
-	using kas(2) by (auto simp add: power_Suc[symmetric, of _ "k - Suc 0"] 
-	  mult_assoc[of _ _ a, symmetric])
+	unfolding poly_pCons poly_monom
+	using kas(2) by (cases k, auto simp add: ring_simps)
       from poly_bound_exists[of "cmod w" s] obtain m where 
 	m: "m > 0" "\<forall>z. cmod z \<le> cmod w \<longrightarrow> cmod (poly s z) \<le> m" by blast
       have w0: "w\<noteq>0" using kas(2) w by (auto simp add: power_0_left)
@@ -901,21 +925,21 @@
 text {* Alternative version with a syntactic notion of constant polynomial. *}
 
 lemma fundamental_theorem_of_algebra_alt:
-  assumes nc: "~(\<exists>a l. a\<noteq> 0 \<and> list_all(\<lambda>b. b = 0) l \<and> p = a#l)"
+  assumes nc: "~(\<exists>a l. a\<noteq> 0 \<and> l = 0 \<and> p = pCons a l)"
   shows "\<exists>z. poly p z = (0::complex)"
 using nc
 proof(induct p)
-  case (Cons c cs)
+  case (pCons c cs)
   {assume "c=0" hence ?case by auto}
   moreover
   {assume c0: "c\<noteq>0"
-    {assume nc: "constant (poly (c#cs))"
+    {assume nc: "constant (poly (pCons c cs))"
       from nc[unfolded constant_def, rule_format, of 0] 
       have "\<forall>w. w \<noteq> 0 \<longrightarrow> poly cs w = 0" by auto 
-      hence "list_all (\<lambda>c. c=0) cs"
+      hence "cs = 0"
 	proof(induct cs)
-	  case (Cons d ds)
-	  {assume "d=0" hence ?case using Cons.prems Cons.hyps by simp}
+	  case (pCons d ds)
+	  {assume "d=0" hence ?case using pCons.prems pCons.hyps by simp}
 	  moreover
 	  {assume d0: "d\<noteq>0"
 	    from poly_bound_exists[of 1 ds] obtain m where 
@@ -925,7 +949,7 @@
 	      x: "x > 0" "x < cmod d / m" "x < 1" by blast
 	    let ?x = "complex_of_real x"
 	    from x have cx: "?x \<noteq> 0"  "cmod ?x \<le> 1" by simp_all
-	    from Cons.prems[rule_format, OF cx(1)]
+	    from pCons.prems[rule_format, OF cx(1)]
 	    have cth: "cmod (?x*poly ds ?x) = cmod d" by (simp add: eq_diff_eq[symmetric])
 	    from m(2)[rule_format, OF cx(2)] x(1)
 	    have th0: "cmod (?x*poly ds ?x) \<le> x*m"
@@ -935,154 +959,252 @@
 	    with cth  have ?case by blast}
 	  ultimately show ?case by blast 
 	qed simp}
-      then have nc: "\<not> constant (poly (c#cs))" using Cons.prems c0 
+      then have nc: "\<not> constant (poly (pCons c cs))" using pCons.prems c0 
 	by blast
       from fundamental_theorem_of_algebra[OF nc] have ?case .}
   ultimately show ?case by blast  
 qed simp
 
+subsection {* Order of polynomial roots *}
+
+definition
+  order :: "'a::{idom,recpower} \<Rightarrow> 'a poly \<Rightarrow> nat"
+where
+  [code del]:
+  "order a p = (LEAST n. \<not> [:-a, 1:] ^ Suc n dvd p)"
+
+lemma degree_power_le: "degree (p ^ n) \<le> degree p * n"
+by (induct n, simp, auto intro: order_trans degree_mult_le)
+
+lemma coeff_linear_power:
+  fixes a :: "'a::{comm_semiring_1,recpower}"
+  shows "coeff ([:a, 1:] ^ n) n = 1"
+apply (induct n, simp_all)
+apply (subst coeff_eq_0)
+apply (auto intro: le_less_trans degree_power_le)
+done
+
+lemma degree_linear_power:
+  fixes a :: "'a::{comm_semiring_1,recpower}"
+  shows "degree ([:a, 1:] ^ n) = n"
+apply (rule order_antisym)
+apply (rule ord_le_eq_trans [OF degree_power_le], simp)
+apply (rule le_degree, simp add: coeff_linear_power)
+done
+
+lemma order_1: "[:-a, 1:] ^ order a p dvd p"
+apply (cases "p = 0", simp)
+apply (cases "order a p", simp)
+apply (subgoal_tac "nat < (LEAST n. \<not> [:-a, 1:] ^ Suc n dvd p)")
+apply (drule not_less_Least, simp)
+apply (fold order_def, simp)
+done
+
+lemma order_2: "p \<noteq> 0 \<Longrightarrow> \<not> [:-a, 1:] ^ Suc (order a p) dvd p"
+unfolding order_def
+apply (rule LeastI_ex)
+apply (rule_tac x="degree p" in exI)
+apply (rule notI)
+apply (drule (1) dvd_imp_degree_le)
+apply (simp only: degree_linear_power)
+done
+
+lemma order:
+  "p \<noteq> 0 \<Longrightarrow> [:-a, 1:] ^ order a p dvd p \<and> \<not> [:-a, 1:] ^ Suc (order a p) dvd p"
+by (rule conjI [OF order_1 order_2])
+
+lemma order_degree:
+  assumes p: "p \<noteq> 0"
+  shows "order a p \<le> degree p"
+proof -
+  have "order a p = degree ([:-a, 1:] ^ order a p)"
+    by (simp only: degree_linear_power)
+  also have "\<dots> \<le> degree p"
+    using order_1 p by (rule dvd_imp_degree_le)
+  finally show ?thesis .
+qed
+
+lemma order_root: "poly p a = 0 \<longleftrightarrow> p = 0 \<or> order a p \<noteq> 0"
+apply (cases "p = 0", simp_all)
+apply (rule iffI)
+apply (rule ccontr, simp)
+apply (frule order_2 [where a=a], simp)
+apply (simp add: poly_eq_0_iff_dvd)
+apply (simp add: poly_eq_0_iff_dvd)
+apply (simp only: order_def)
+apply (drule not_less_Least, simp)
+done
+
+lemma UNIV_nat_infinite:
+  "\<not> finite (UNIV :: nat set)" (is "\<not> finite ?U")
+proof
+  assume "finite ?U"
+  moreover have "Suc (Max ?U) \<in> ?U" ..
+  ultimately have "Suc (Max ?U) \<le> Max ?U" by (rule Max_ge)
+  then show "False" by simp
+qed
+
+lemma UNIV_char_0_infinite:
+  "\<not> finite (UNIV::'a::semiring_char_0 set)"
+proof
+  assume "finite (UNIV::'a set)"
+  with subset_UNIV have "finite (range of_nat::'a set)"
+    by (rule finite_subset)
+  moreover have "inj (of_nat::nat \<Rightarrow> 'a)"
+    by (simp add: inj_on_def)
+  ultimately have "finite (UNIV::nat set)"
+    by (rule finite_imageD)
+  then show "False"
+    by (simp add: UNIV_nat_infinite)
+qed
+
+lemma poly_zero:
+  fixes p :: "'a::{idom,ring_char_0} poly"
+  shows "poly p = poly 0 \<longleftrightarrow> p = 0"
+apply (cases "p = 0", simp_all)
+apply (drule poly_roots_finite)
+apply (auto simp add: UNIV_char_0_infinite)
+done
+
+lemma poly_eq_iff:
+  fixes p q :: "'a::{idom,ring_char_0} poly"
+  shows "poly p = poly q \<longleftrightarrow> p = q"
+  using poly_zero [of "p - q"]
+  by (simp add: expand_fun_eq)
+
+
 subsection{* Nullstellenstatz, degrees and divisibility of polynomials *}
 
 lemma nullstellensatz_lemma:
-  fixes p :: "complex list"
+  fixes p :: "complex poly"
   assumes "\<forall>x. poly p x = 0 \<longrightarrow> poly q x = 0"
   and "degree p = n" and "n \<noteq> 0"
-  shows "p divides (pexp q n)"
+  shows "p dvd (q ^ n)"
 using prems
 proof(induct n arbitrary: p q rule: nat_less_induct)
-  fix n::nat fix p q :: "complex list"
+  fix n::nat fix p q :: "complex poly"
   assume IH: "\<forall>m<n. \<forall>p q.
                  (\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longrightarrow>
-                 degree p = m \<longrightarrow> m \<noteq> 0 \<longrightarrow> p divides (q %^ m)"
+                 degree p = m \<longrightarrow> m \<noteq> 0 \<longrightarrow> p dvd (q ^ m)"
     and pq0: "\<forall>x. poly p x = 0 \<longrightarrow> poly q x = 0" 
     and dpn: "degree p = n" and n0: "n \<noteq> 0"
-  let ?ths = "p divides (q %^ n)"
+  from dpn n0 have pne: "p \<noteq> 0" by auto
+  let ?ths = "p dvd (q ^ n)"
   {fix a assume a: "poly p a = 0"
-    {assume p0: "poly p = poly []" 
-      hence ?ths unfolding divides_def  using pq0 n0
-	apply - apply (rule exI[where x="[]"], rule ext)
-	by (auto simp add: poly_mult poly_exp)}
-    moreover
-    {assume p0: "poly p \<noteq> poly []" 
-      and oa: "order  a p \<noteq> 0"
-      from p0 have pne: "p \<noteq> []" by auto
+    {assume oa: "order a p \<noteq> 0"
       let ?op = "order a p"
-      from p0 have ap: "([- a, 1] %^ ?op) divides p" 
-	"\<not> pexp [- a, 1] (Suc ?op) divides p" using order by blast+ 
-      note oop = order_degree[OF p0, unfolded dpn]
-      {assume q0: "q = []"
-	hence ?ths using n0 unfolding divides_def 
-	  apply simp
-	  apply (rule exI[where x="[]"], rule ext)
-	  by (simp add: divides_def poly_exp poly_mult)}
+      from pne have ap: "([:- a, 1:] ^ ?op) dvd p" 
+	"\<not> [:- a, 1:] ^ (Suc ?op) dvd p" using order by blast+ 
+      note oop = order_degree[OF pne, unfolded dpn]
+      {assume q0: "q = 0"
+	hence ?ths using n0
+          by (simp add: power_0_left)}
       moreover
-      {assume q0: "q\<noteq>[]"
-	from pq0[rule_format, OF a, unfolded poly_linear_divides] q0
-	obtain r where r: "q = pmult [- a, 1] r" by blast
-	from ap[unfolded divides_def] obtain s where
-	  s: "poly p = poly (pmult (pexp [- a, 1] ?op) s)" by blast
-	have s0: "poly s \<noteq> poly []"
-	  using s p0 by (simp add: poly_entire)
-	hence pns0: "poly (pnormalize s) \<noteq> poly []" and sne: "s\<noteq>[]" by auto
+      {assume q0: "q \<noteq> 0"
+	from pq0[rule_format, OF a, unfolded poly_eq_0_iff_dvd]
+	obtain r where r: "q = [:- a, 1:] * r" by (rule dvdE)
+	from ap(1) obtain s where
+	  s: "p = [:- a, 1:] ^ ?op * s" by (rule dvdE)
+	have sne: "s \<noteq> 0"
+	  using s pne by auto
 	{assume ds0: "degree s = 0"
-	  from ds0 pns0 have "\<exists>k. pnormalize s = [k]" unfolding degree_def 
-	    by (cases "pnormalize s", auto)
-	  then obtain k where kpn: "pnormalize s = [k]" by blast
-	  from pns0[unfolded poly_zero] kpn have k: "k \<noteq>0" "poly s = poly [k]"
-	    using poly_normalize[of s] by simp_all
-	  let ?w = "pmult (pmult [1/k] (pexp [-a,1] (n - ?op))) (pexp r n)"
-	  from k r s oop have "poly (pexp q n) = poly (pmult p ?w)"
-	    by - (rule ext, simp add: poly_mult poly_exp poly_cmult poly_add power_add[symmetric] ring_simps power_mult_distrib[symmetric])
-	  hence ?ths unfolding divides_def by blast}
+	  from ds0 have "\<exists>k. s = [:k:]"
+            by (cases s, simp split: if_splits)
+	  then obtain k where kpn: "s = [:k:]" by blast
+          from sne kpn have k: "k \<noteq> 0" by simp
+	  let ?w = "([:1/k:] * ([:-a,1:] ^ (n - ?op))) * (r ^ n)"
+          from k oop [of a] have "q ^ n = p * ?w"
+            apply -
+            apply (subst r, subst s, subst kpn)
+            apply (subst power_mult_distrib, simp)
+            apply (subst power_add [symmetric], simp)
+            done
+	  hence ?ths unfolding dvd_def by blast}
 	moreover
 	{assume ds0: "degree s \<noteq> 0"
-	  from ds0 s0 dpn degree_unique[OF s, unfolded linear_pow_mul_degree] oa
-	    have dsn: "degree s < n" by auto 
+	  from ds0 sne dpn s oa
+	    have dsn: "degree s < n" apply auto
+              apply (erule ssubst)
+              apply (simp add: degree_mult_eq degree_linear_power)
+              done
 	    {fix x assume h: "poly s x = 0"
 	      {assume xa: "x = a"
-		from h[unfolded xa poly_linear_divides] sne obtain u where
-		  u: "s = pmult [- a, 1] u" by blast
-		have "poly p = poly (pmult (pexp [- a, 1] (Suc ?op)) u)"
-		  unfolding s u
-		  apply (rule ext)
-		  by (simp add: ring_simps power_mult_distrib[symmetric] poly_mult poly_cmult poly_add poly_exp)
-		with ap(2)[unfolded divides_def] have False by blast}
+		from h[unfolded xa poly_eq_0_iff_dvd] obtain u where
+		  u: "s = [:- a, 1:] * u" by (rule dvdE)
+		have "p = [:- a, 1:] ^ (Suc ?op) * u"
+                  by (subst s, subst u, simp only: power_Suc mult_ac)
+		with ap(2)[unfolded dvd_def] have False by blast}
 	      note xa = this
-	      from h s have "poly p x = 0" by (simp add: poly_mult)
+	      from h have "poly p x = 0" by (subst s, simp)
 	      with pq0 have "poly q x = 0" by blast
 	      with r xa have "poly r x = 0"
-		by (auto simp add: poly_mult poly_add poly_cmult eq_diff_eq[symmetric])}
+                by (auto simp add: uminus_add_conv_diff)}
 	    note impth = this
 	    from IH[rule_format, OF dsn, of s r] impth ds0
-	    have "s divides (pexp r (degree s))" by blast
-	    then obtain u where u: "poly (pexp r (degree s)) = poly (pmult s u)"
-	      unfolding divides_def by blast
+	    have "s dvd (r ^ (degree s))" by blast
+	    then obtain u where u: "r ^ (degree s) = s * u" ..
 	    hence u': "\<And>x. poly s x * poly u x = poly r x ^ degree s"
-	      by (simp add: poly_mult[symmetric] poly_exp[symmetric])
-	    let ?w = "pmult (pmult u (pexp [-a,1] (n - ?op))) (pexp r (n - degree s))"
-	    from u' s r oop[of a] dsn have "poly (pexp q n) = poly (pmult p ?w)"
-	      apply - apply (rule ext)
-	      apply (simp only:  power_mult_distrib power_add[symmetric] poly_add poly_mult poly_exp poly_cmult ring_simps)
-	      
-	      apply (simp add:  power_mult_distrib power_add[symmetric] poly_add poly_mult poly_exp poly_cmult mult_assoc[symmetric])
-	      done
-	    hence ?ths unfolding divides_def by blast}
+              by (simp only: poly_mult[symmetric] poly_power[symmetric])
+	    let ?w = "(u * ([:-a,1:] ^ (n - ?op))) * (r ^ (n - degree s))"
+	    from oop[of a] dsn have "q ^ n = p * ?w"
+              apply -
+              apply (subst s, subst r)
+              apply (simp only: power_mult_distrib)
+              apply (subst mult_assoc [where b=s])
+              apply (subst mult_assoc [where a=u])
+              apply (subst mult_assoc [where b=u, symmetric])
+              apply (subst u [symmetric])
+              apply (simp add: mult_ac power_add [symmetric])
+              done
+	    hence ?ths unfolding dvd_def by blast}
       ultimately have ?ths by blast }
       ultimately have ?ths by blast}
-    ultimately have ?ths using a order_root by blast}
+    then have ?ths using a order_root pne by blast}
   moreover
   {assume exa: "\<not> (\<exists>a. poly p a = 0)"
-    from fundamental_theorem_of_algebra_alt[of p] exa obtain c cs where
-      ccs: "c\<noteq>0" "list_all (\<lambda>c. c = 0) cs" "p = c#cs" by blast
+    from fundamental_theorem_of_algebra_alt[of p] exa obtain c where
+      ccs: "c\<noteq>0" "p = pCons c 0" by blast
     
-    from poly_0[OF ccs(2)] ccs(3) 
-    have pp: "\<And>x. poly p x =  c" by simp
-    let ?w = "pmult [1/c] (pexp q n)"
-    from pp ccs(1) 
-    have "poly (pexp q n) = poly (pmult p ?w) "
-      apply - apply (rule ext)
-      unfolding poly_mult_assoc[symmetric] by (simp add: poly_mult)
-    hence ?ths unfolding divides_def by blast}
+    then have pp: "\<And>x. poly p x =  c" by simp
+    let ?w = "[:1/c:] * (q ^ n)"
+    from ccs
+    have "(q ^ n) = (p * ?w) "
+      by (simp add: smult_smult)
+    hence ?ths unfolding dvd_def by blast}
   ultimately show ?ths by blast
 qed
 
 lemma nullstellensatz_univariate:
   "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow> 
-    p divides (q %^ (degree p)) \<or> (poly p = poly [] \<and> poly q = poly [])"
+    p dvd (q ^ (degree p)) \<or> (p = 0 \<and> q = 0)"
 proof-
-  {assume pe: "poly p = poly []"
-    hence eq: "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow> poly q = poly []"
+  {assume pe: "p = 0"
+    hence eq: "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow> q = 0"
       apply auto
+      apply (rule poly_zero [THEN iffD1])
       by (rule ext, simp)
-    {assume "p divides (pexp q (degree p))"
-      then obtain r where r: "poly (pexp q (degree p)) = poly (pmult p r)" 
-	unfolding divides_def by blast
-      from cong[OF r refl] pe degree_unique[OF pe]
-      have False by (simp add: poly_mult degree_def)}
+    {assume "p dvd (q ^ (degree p))"
+      then obtain r where r: "q ^ (degree p) = p * r" ..
+      from r pe have False by simp}
     with eq pe have ?thesis by blast}
   moreover
-  {assume pe: "poly p \<noteq> poly []"
-    have p0: "poly [0] = poly []" by (rule ext, simp)
+  {assume pe: "p \<noteq> 0"
     {assume dp: "degree p = 0"
-      then obtain k where "pnormalize p = [k]" using pe poly_normalize[of p]
-	unfolding degree_def by (cases "pnormalize p", auto)
-      hence k: "pnormalize p = [k]" "poly p = poly [k]" "k\<noteq>0"
-	using pe poly_normalize[of p] by (auto simp add: p0)
+      then obtain k where k: "p = [:k:]" "k\<noteq>0" using pe
+        by (cases p, simp split: if_splits)
       hence th1: "\<forall>x. poly p x \<noteq> 0" by simp
-      from k(2,3) dp have "poly (pexp q (degree p)) = poly (pmult p [1/k]) "
-	by - (rule ext, simp add: poly_mult poly_exp)
-      hence th2: "p divides (pexp q (degree p))" unfolding divides_def by blast
+      from k dp have "q ^ (degree p) = p * [:1/k:]"
+        by (simp add: one_poly_def)
+      hence th2: "p dvd (q ^ (degree p))" ..
       from th1 th2 pe have ?thesis by blast}
     moreover
     {assume dp: "degree p \<noteq> 0"
       then obtain n where n: "degree p = Suc n " by (cases "degree p", auto)
-      {assume "p divides (pexp q (Suc n))"
-	then obtain u where u: "poly (pexp q (Suc n)) = poly (pmult p u)"
-	  unfolding divides_def by blast
-	hence u' :"\<And>x. poly (pexp q (Suc n)) x = poly (pmult p u) x" by simp_all
+      {assume "p dvd (q ^ (Suc n))"
+	then obtain u where u: "q ^ (Suc n) = p * u" ..
 	{fix x assume h: "poly p x = 0" "poly q x \<noteq> 0"
-	  hence "poly (pexp q (Suc n)) x \<noteq> 0" by (simp only: poly_exp) simp	  
-	  hence False using u' h(1) by (simp only: poly_mult poly_exp) simp}}
+	  hence "poly (q ^ (Suc n)) x \<noteq> 0" by simp
+	  hence False using u h(1) by (simp only: poly_mult) simp}}
 	with n nullstellensatz_lemma[of p q "degree p"] dp 
 	have ?thesis by auto}
     ultimately have ?thesis by blast}
@@ -1091,230 +1213,176 @@
 
 text{* Useful lemma *}
 
-lemma (in idom_char_0) constant_degree: "constant (poly p) \<longleftrightarrow> degree p = 0" (is "?lhs = ?rhs")
+lemma constant_degree:
+  fixes p :: "'a::{idom,ring_char_0} poly"
+  shows "constant (poly p) \<longleftrightarrow> degree p = 0" (is "?lhs = ?rhs")
 proof
   assume l: ?lhs
-  from l[unfolded constant_def, rule_format, of _ "zero"]
-  have th: "poly p = poly [poly p 0]" apply - by (rule ext, simp)
-  from degree_unique[OF th] show ?rhs by (simp add: degree_def)
+  from l[unfolded constant_def, rule_format, of _ "0"]
+  have th: "poly p = poly [:poly p 0:]" apply - by (rule ext, simp)
+  then have "p = [:poly p 0:]" by (simp add: poly_eq_iff)
+  then have "degree p = degree [:poly p 0:]" by simp
+  then show ?rhs by simp
 next
   assume r: ?rhs
-  from r have "pnormalize p = [] \<or> (\<exists>k. pnormalize p = [k])"
-    unfolding degree_def by (cases "pnormalize p", auto)
-  then show ?lhs unfolding constant_def poly_normalize[of p, symmetric]
-    by (auto simp del: poly_normalize)
+  then obtain k where "p = [:k:]"
+    by (cases p, simp split: if_splits)
+  then show ?lhs unfolding constant_def by auto
 qed
 
-(* It would be nicer to prove this without using algebraic closure...        *)
-
-lemma divides_degree_lemma: assumes dpn: "degree (p::complex list) = n"
-  shows "n \<le> degree (p *** q) \<or> poly (p *** q) = poly []"
-  using dpn
-proof(induct n arbitrary: p q)
-  case 0 thus ?case by simp
-next
-  case (Suc n p q)
-  from Suc.prems fundamental_theorem_of_algebra[of p] constant_degree[of p]
-  obtain a where a: "poly p a = 0" by auto
-  then obtain r where r: "p = pmult [-a, 1] r" unfolding poly_linear_divides
-    using Suc.prems by (auto simp add: degree_def)
-  {assume h: "poly (pmult r q) = poly []"
-    hence "poly (pmult p q) = poly []" using r
-      apply - apply (rule ext)  by (auto simp add: poly_entire poly_mult poly_add poly_cmult) hence ?case by blast}
-  moreover
-  {assume h: "poly (pmult r q) \<noteq> poly []" 
-    hence r0: "poly r \<noteq> poly []" and q0: "poly q \<noteq> poly []"
-      by (auto simp add: poly_entire)
-    have eq: "poly (pmult p q) = poly (pmult [-a, 1] (pmult r q))"
-      apply - apply (rule ext)
-      by (simp add: r poly_mult poly_add poly_cmult ring_simps)
-    from linear_mul_degree[OF h, of "- a"]
-    have dqe: "degree (pmult p q) = degree (pmult r q) + 1"
-      unfolding degree_unique[OF eq] .
-    from linear_mul_degree[OF r0, of "- a", unfolded r[symmetric]] r Suc.prems 
-    have dr: "degree r = n" by auto
-    from  Suc.hyps[OF dr, of q] have "Suc n \<le> degree (pmult p q)"
-      unfolding dqe using h by (auto simp del: poly.simps) 
-    hence ?case by blast}
-  ultimately show ?case by blast
-qed
-
-lemma divides_degree: assumes pq: "p divides (q:: complex list)"
-  shows "degree p \<le> degree q \<or> poly q = poly []"
-using pq  divides_degree_lemma[OF refl, of p]
-apply (auto simp add: divides_def poly_entire)
-apply atomize
-apply (erule_tac x="qa" in allE, auto)
-apply (subgoal_tac "degree q = degree (p *** qa)", simp)
-apply (rule degree_unique, simp)
+lemma divides_degree: assumes pq: "p dvd (q:: complex poly)"
+  shows "degree p \<le> degree q \<or> q = 0"
+apply (cases "q = 0", simp_all)
+apply (erule dvd_imp_degree_le [OF pq])
 done
 
 (* Arithmetic operations on multivariate polynomials.                        *)
 
 lemma mpoly_base_conv: 
-  "(0::complex) \<equiv> poly [] x" "c \<equiv> poly [c] x" "x \<equiv> poly [0,1] x" by simp_all
+  "(0::complex) \<equiv> poly 0 x" "c \<equiv> poly [:c:] x" "x \<equiv> poly [:0,1:] x" by simp_all
 
 lemma mpoly_norm_conv: 
-  "poly [0] (x::complex) \<equiv> poly [] x" "poly [poly [] y] x \<equiv> poly [] x" by simp_all
+  "poly [:0:] (x::complex) \<equiv> poly 0 x" "poly [:poly 0 y:] x \<equiv> poly 0 x" by simp_all
 
 lemma mpoly_sub_conv: 
   "poly p (x::complex) - poly q x \<equiv> poly p x + -1 * poly q x"
   by (simp add: diff_def)
 
-lemma poly_pad_rule: "poly p x = 0 ==> poly (0#p) x = (0::complex)" by simp
+lemma poly_pad_rule: "poly p x = 0 ==> poly (pCons 0 p) x = (0::complex)" by simp
 
 lemma poly_cancel_eq_conv: "p = (0::complex) \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> (q = 0) \<equiv> (a * q - b * p = 0)" apply (atomize (full)) by auto
 
-lemma resolve_eq_raw:  "poly [] x \<equiv> 0" "poly [c] x \<equiv> (c::complex)" by auto
+lemma resolve_eq_raw:  "poly 0 x \<equiv> 0" "poly [:c:] x \<equiv> (c::complex)" by auto
 lemma  resolve_eq_then: "(P \<Longrightarrow> (Q \<equiv> Q1)) \<Longrightarrow> (\<not>P \<Longrightarrow> (Q \<equiv> Q2))
   \<Longrightarrow> Q \<equiv> P \<and> Q1 \<or> \<not>P\<and> Q2" apply (atomize (full)) by blast 
-lemma expand_ex_beta_conv: "list_ex P [c] \<equiv> P c" by simp
 
 lemma poly_divides_pad_rule: 
-  fixes p q :: "complex list"
-  assumes pq: "p divides q"
-  shows "p divides ((0::complex)#q)"
+  fixes p q :: "complex poly"
+  assumes pq: "p dvd q"
+  shows "p dvd (pCons (0::complex) q)"
 proof-
-  from pq obtain r where r: "poly q = poly (p *** r)" unfolding divides_def by blast
-  hence "poly (0#q) = poly (p *** ([0,1] *** r))" 
-    by - (rule ext, simp add: poly_mult poly_cmult poly_add)
-  thus ?thesis unfolding divides_def by blast
+  have "pCons 0 q = q * [:0,1:]" by simp
+  then have "q dvd (pCons 0 q)" ..
+  with pq show ?thesis by (rule dvd_trans)
 qed
 
 lemma poly_divides_pad_const_rule: 
-  fixes p q :: "complex list"
-  assumes pq: "p divides q"
-  shows "p divides (a %* q)"
+  fixes p q :: "complex poly"
+  assumes pq: "p dvd q"
+  shows "p dvd (smult a q)"
 proof-
-  from pq obtain r where r: "poly q = poly (p *** r)" unfolding divides_def by blast
-  hence "poly (a %* q) = poly (p *** (a %* r))" 
-    by - (rule ext, simp add: poly_mult poly_cmult poly_add)
-  thus ?thesis unfolding divides_def by blast
+  have "smult a q = q * [:a:]" by simp
+  then have "q dvd smult a q" ..
+  with pq show ?thesis by (rule dvd_trans)
 qed
 
 
 lemma poly_divides_conv0:  
-  fixes p :: "complex list"
-  assumes lgpq: "length q < length p" and lq:"last p \<noteq> 0"
-  shows "p divides q \<equiv> (\<not> (list_ex (\<lambda>c. c \<noteq> 0) q))" (is "?lhs \<equiv> ?rhs")
+  fixes p :: "complex poly"
+  assumes lgpq: "degree q < degree p" and lq:"p \<noteq> 0"
+  shows "p dvd q \<equiv> q = 0" (is "?lhs \<equiv> ?rhs")
 proof-
   {assume r: ?rhs 
-    hence eq: "poly q = poly []" unfolding poly_zero 
-      by (simp add: list_all_iff list_ex_iff)
-    hence "poly q = poly (p *** [])" by - (rule ext, simp add: poly_mult)
-    hence ?lhs unfolding divides_def  by blast}
+    hence "q = p * 0" by simp
+    hence ?lhs ..}
   moreover
   {assume l: ?lhs
-    have ath: "\<And>lq lp dq::nat. lq < lp ==> lq \<noteq> 0 \<Longrightarrow> dq <= lq - 1 ==> dq < lp - 1"
-      by arith
-    {assume q0: "length q = 0"
-      hence "q = []" by simp
+    {assume q0: "q = 0"
       hence ?rhs by simp}
     moreover
-    {assume lgq0: "length q \<noteq> 0"
-      from pnormalize_length[of q] have dql: "degree q \<le> length q - 1" 
-	unfolding degree_def by simp
-      from ath[OF lgpq lgq0 dql, unfolded pnormal_degree[OF lq, symmetric]] divides_degree[OF l] have "poly q = poly []" by auto
-      hence ?rhs unfolding poly_zero by (simp add: list_all_iff list_ex_iff)}
+    {assume q0: "q \<noteq> 0"
+      from l q0 have "degree p \<le> degree q"
+        by (rule dvd_imp_degree_le)
+      with lgpq have ?rhs by simp }
     ultimately have ?rhs by blast }
   ultimately show "?lhs \<equiv> ?rhs" by - (atomize (full), blast) 
 qed
 
 lemma poly_divides_conv1: 
-  assumes a0: "a\<noteq> (0::complex)" and pp': "(p::complex list) divides p'"
-  and qrp': "\<And>x. a * poly q x - poly p' x \<equiv> poly r x"
-  shows "p divides q \<equiv> p divides (r::complex list)" (is "?lhs \<equiv> ?rhs")
+  assumes a0: "a\<noteq> (0::complex)" and pp': "(p::complex poly) dvd p'"
+  and qrp': "smult a q - p' \<equiv> r"
+  shows "p dvd q \<equiv> p dvd (r::complex poly)" (is "?lhs \<equiv> ?rhs")
 proof-
   {
-  from pp' obtain t where t: "poly p' = poly (p *** t)" 
-    unfolding divides_def by blast
+  from pp' obtain t where t: "p' = p * t" ..
   {assume l: ?lhs
-    then obtain u where u: "poly q = poly (p *** u)" unfolding divides_def by blast
-     have "poly r = poly (p *** ((a %* u) +++ (-- t)))"
-       using u qrp' t
-       by - (rule ext, 
-	 simp add: poly_add poly_mult poly_cmult poly_minus ring_simps)
-     then have ?rhs unfolding divides_def by blast}
+    then obtain u where u: "q = p * u" ..
+     have "r = p * (smult a u - t)"
+       using u qrp' [symmetric] t by (simp add: ring_simps mult_smult_right)
+     then have ?rhs ..}
   moreover
   {assume r: ?rhs
-    then obtain u where u: "poly r = poly (p *** u)" unfolding divides_def by blast
-    from u t qrp' a0 have "poly q = poly (p *** ((1/a) %* (u +++ t)))"
-      by - (rule ext, atomize (full), simp add: poly_mult poly_add poly_cmult field_simps)
-    hence ?lhs  unfolding divides_def by blast}
+    then obtain u where u: "r = p * u" ..
+    from u [symmetric] t qrp' [symmetric] a0
+    have "q = p * smult (1/a) (u + t)"
+      by (simp add: ring_simps mult_smult_right smult_smult)
+    hence ?lhs ..}
   ultimately have "?lhs = ?rhs" by blast }
 thus "?lhs \<equiv> ?rhs"  by - (atomize(full), blast) 
 qed
 
 lemma basic_cqe_conv1:
-  "(\<exists>x. poly p x = 0 \<and> poly [] x \<noteq> 0) \<equiv> False"
-  "(\<exists>x. poly [] x \<noteq> 0) \<equiv> False"
-  "(\<exists>x. poly [c] x \<noteq> 0) \<equiv> c\<noteq>0"
-  "(\<exists>x. poly [] x = 0) \<equiv> True"
-  "(\<exists>x. poly [c] x = 0) \<equiv> c = 0" by simp_all
+  "(\<exists>x. poly p x = 0 \<and> poly 0 x \<noteq> 0) \<equiv> False"
+  "(\<exists>x. poly 0 x \<noteq> 0) \<equiv> False"
+  "(\<exists>x. poly [:c:] x \<noteq> 0) \<equiv> c\<noteq>0"
+  "(\<exists>x. poly 0 x = 0) \<equiv> True"
+  "(\<exists>x. poly [:c:] x = 0) \<equiv> c = 0" by simp_all
 
 lemma basic_cqe_conv2: 
-  assumes l:"last (a#b#p) \<noteq> 0" 
-  shows "(\<exists>x. poly (a#b#p) x = (0::complex)) \<equiv> True"
+  assumes l:"p \<noteq> 0" 
+  shows "(\<exists>x. poly (pCons a (pCons b p)) x = (0::complex)) \<equiv> True"
 proof-
   {fix h t
-    assume h: "h\<noteq>0" "list_all (\<lambda>c. c=(0::complex)) t"  "a#b#p = h#t"
-    hence "list_all (\<lambda>c. c= 0) (b#p)" by simp
-    moreover have "last (b#p) \<in> set (b#p)" by simp
-    ultimately have "last (b#p) = 0" by (simp add: list_all_iff)
+    assume h: "h\<noteq>0" "t=0"  "pCons a (pCons b p) = pCons h t"
     with l have False by simp}
-  hence th: "\<not> (\<exists> h t. h\<noteq>0 \<and> list_all (\<lambda>c. c=0) t \<and> a#b#p = h#t)"
+  hence th: "\<not> (\<exists> h t. h\<noteq>0 \<and> t=0 \<and> pCons a (pCons b p) = pCons h t)"
     by blast
   from fundamental_theorem_of_algebra_alt[OF th] 
-  show "(\<exists>x. poly (a#b#p) x = (0::complex)) \<equiv> True" by auto
+  show "(\<exists>x. poly (pCons a (pCons b p)) x = (0::complex)) \<equiv> True" by auto
 qed
 
-lemma  basic_cqe_conv_2b: "(\<exists>x. poly p x \<noteq> (0::complex)) \<equiv> (list_ex (\<lambda>c. c \<noteq> 0) p)"
+lemma  basic_cqe_conv_2b: "(\<exists>x. poly p x \<noteq> (0::complex)) \<equiv> (p \<noteq> 0)"
 proof-
-  have "\<not> (list_ex (\<lambda>c. c \<noteq> 0) p) \<longleftrightarrow> poly p = poly []" 
-    by (simp add: poly_zero list_all_iff list_ex_iff)
+  have "p = 0 \<longleftrightarrow> poly p = poly 0"
+    by (simp add: poly_zero)
   also have "\<dots> \<longleftrightarrow> (\<not> (\<exists>x. poly p x \<noteq> 0))" by (auto intro: ext)
-  finally show "(\<exists>x. poly p x \<noteq> (0::complex)) \<equiv> (list_ex (\<lambda>c. c \<noteq> 0) p)"
+  finally show "(\<exists>x. poly p x \<noteq> (0::complex)) \<equiv> p \<noteq> 0"
     by - (atomize (full), blast)
 qed
 
 lemma basic_cqe_conv3:
-  fixes p q :: "complex list"
-  assumes l: "last (a#p) \<noteq> 0" 
-  shows "(\<exists>x. poly (a#p) x =0 \<and> poly q x \<noteq> 0) \<equiv> \<not> ((a#p) divides (q %^ (length p)))"
+  fixes p q :: "complex poly"
+  assumes l: "p \<noteq> 0" 
+  shows "(\<exists>x. poly (pCons a p) x = 0 \<and> poly q x \<noteq> 0) \<equiv> \<not> ((pCons a p) dvd (q ^ (psize p)))"
 proof-
-  note np = pnormalize_eq[OF l]
-  {assume "poly (a#p) = poly []" hence False using l
-      unfolding poly_zero apply (auto simp add: list_all_iff del: last.simps)
-      apply (cases p, simp_all) done}
-  then have p0: "poly (a#p) \<noteq> poly []"  by blast
-  from np have dp:"degree (a#p) = length p" by (simp add: degree_def)
-  from nullstellensatz_univariate[of "a#p" q] p0 dp
-  show "(\<exists>x. poly (a#p) x =0 \<and> poly q x \<noteq> 0) \<equiv> \<not> ((a#p) divides (q %^ (length p)))"
+  from l have dp:"degree (pCons a p) = psize p" by (simp add: psize_def)
+  from nullstellensatz_univariate[of "pCons a p" q] l
+  show "(\<exists>x. poly (pCons a p) x = 0 \<and> poly q x \<noteq> 0) \<equiv> \<not> ((pCons a p) dvd (q ^ (psize p)))"
+    unfolding dp
     by - (atomize (full), auto)
 qed
 
 lemma basic_cqe_conv4:
-  fixes p q :: "complex list"
-  assumes h: "\<And>x. poly (q %^ n) x \<equiv> poly r x"
-  shows "p divides (q %^ n) \<equiv> p divides r"
+  fixes p q :: "complex poly"
+  assumes h: "\<And>x. poly (q ^ n) x \<equiv> poly r x"
+  shows "p dvd (q ^ n) \<equiv> p dvd r"
 proof-
-  from h have "poly (q %^ n) = poly r" by (auto intro: ext)  
-  thus "p divides (q %^ n) \<equiv> p divides r" unfolding divides_def by simp
+  from h have "poly (q ^ n) = poly r" by (auto intro: ext)
+  then have "(q ^ n) = r" by (simp add: poly_eq_iff)
+  thus "p dvd (q ^ n) \<equiv> p dvd r" by simp
 qed
 
-lemma pmult_Cons_Cons: "((a::complex)#b#p) *** q = (a %*q) +++ (0#((b#p) *** q))"
+lemma pmult_Cons_Cons: "(pCons (a::complex) (pCons b p) * q = (smult a q) + (pCons 0 (pCons b p * q)))"
   by simp
 
 lemma elim_neg_conv: "- z \<equiv> (-1) * (z::complex)" by simp
 lemma eqT_intr: "PROP P \<Longrightarrow> (True \<Longrightarrow> PROP P )" "PROP P \<Longrightarrow> True" by blast+
 lemma negate_negate_rule: "Trueprop P \<equiv> \<not> P \<equiv> False" by (atomize (full), auto)
-lemma last_simps: "last [x] = x" "last (x#y#ys) = last (y#ys)" by simp_all
-lemma length_simps: "length [] = 0" "length (x#y#xs) = length xs + 2" "length [x] = 1" by simp_all
 
 lemma complex_entire: "(z::complex) \<noteq> 0 \<and> w \<noteq> 0 \<equiv> z*w \<noteq> 0" by simp
 lemma resolve_eq_ne: "(P \<equiv> True) \<equiv> (\<not>P \<equiv> False)" "(P \<equiv> False) \<equiv> (\<not>P \<equiv> True)" 
   by (atomize (full)) simp_all
-lemma cqe_conv1: "poly [] x = 0 \<longleftrightarrow> True"  by simp
+lemma cqe_conv1: "poly 0 x = 0 \<longleftrightarrow> True"  by simp
 lemma cqe_conv2: "(p \<Longrightarrow> (q \<equiv> r)) \<equiv> ((p \<and> q) \<equiv> (p \<and> r))"  (is "?l \<equiv> ?r")
 proof
   assume "p \<Longrightarrow> q \<equiv> r" thus "p \<and> q \<equiv> p \<and> r" apply - apply (atomize (full)) by blast
@@ -1322,6 +1390,6 @@
   assume "p \<and> q \<equiv> p \<and> r" "p"
   thus "q \<equiv> r" apply - apply (atomize (full)) apply blast done
 qed
-lemma poly_const_conv: "poly [c] (x::complex) = y \<longleftrightarrow> c = y" by simp
+lemma poly_const_conv: "poly [:c:] (x::complex) = y \<longleftrightarrow> c = y" by simp
 
-end
\ No newline at end of file
+end
--- a/src/HOL/HOL.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/HOL.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -35,7 +35,7 @@
   "~~/src/Tools/code/code_ml.ML"
   "~~/src/Tools/code/code_haskell.ML"
   "~~/src/Tools/nbe.ML"
-  ("~~/src/HOL/Tools/recfun_codegen.ML")
+  ("Tools/recfun_codegen.ML")
 begin
 
 subsection {* Primitive logic *}
@@ -1690,7 +1690,7 @@
 
 text {* Module setup *}
 
-use "~~/src/HOL/Tools/recfun_codegen.ML"
+use "Tools/recfun_codegen.ML"
 
 setup {*
   Code_ML.setup
--- a/src/HOL/HahnBanach/FunctionNorm.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/HahnBanach/FunctionNorm.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -204,7 +204,7 @@
   shows "\<bar>f x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>V * \<parallel>x\<parallel>"
 proof -
   interpret continuous V norm f by fact
-  interpret linearform V f .
+  interpret linearform V f by fact
   show ?thesis
   proof cases
     assume "x = 0"
--- a/src/HOL/HahnBanach/HahnBanach.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/HahnBanach/HahnBanach.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -492,7 +492,7 @@
 	then have "\<bar>f x\<bar> = \<bar>g x\<bar>" by (simp only:)
 	also from g_cont
 	have "\<dots> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>"
-	proof (rule fn_norm_le_cong [of g, folded B_def fn_norm_def])
+	proof (rule fn_norm_le_cong [OF _ linearformE, folded B_def fn_norm_def])
 	  from FE x show "x \<in> E" ..
 	qed
 	finally show "\<bar>f x\<bar> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>" .
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Imperative_HOL/Array.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -0,0 +1,209 @@
+(*  Title:      HOL/Library/Array.thy
+    ID:         $Id$
+    Author:     John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen
+*)
+
+header {* Monadic arrays *}
+
+theory Array
+imports Heap_Monad Code_Index
+begin
+
+subsection {* Primitives *}
+
+definition
+  new :: "nat \<Rightarrow> 'a\<Colon>heap \<Rightarrow> 'a array Heap" where
+  [code del]: "new n x = Heap_Monad.heap (Heap.array n x)"
+
+definition
+  of_list :: "'a\<Colon>heap list \<Rightarrow> 'a array Heap" where
+  [code del]: "of_list xs = Heap_Monad.heap (Heap.array_of_list xs)"
+
+definition
+  length :: "'a\<Colon>heap array \<Rightarrow> nat Heap" where
+  [code del]: "length arr = Heap_Monad.heap (\<lambda>h. (Heap.length arr h, h))"
+
+definition
+  nth :: "'a\<Colon>heap array \<Rightarrow> nat \<Rightarrow> 'a Heap"
+where
+  [code del]: "nth a i = (do len \<leftarrow> length a;
+                 (if i < len
+                     then Heap_Monad.heap (\<lambda>h. (get_array a h ! i, h))
+                     else raise (''array lookup: index out of range''))
+              done)"
+
+-- {* FIXME adjustion for List theory *}
+no_syntax
+  nth  :: "'a list \<Rightarrow> nat \<Rightarrow> 'a" (infixl "!" 100)
+
+abbreviation
+  nth_list :: "'a list \<Rightarrow> nat \<Rightarrow> 'a" (infixl "!" 100)
+where
+  "nth_list \<equiv> List.nth"
+
+definition
+  upd :: "nat \<Rightarrow> 'a \<Rightarrow> 'a\<Colon>heap array \<Rightarrow> 'a\<Colon>heap array Heap"
+where
+  [code del]: "upd i x a = (do len \<leftarrow> length a;
+                      (if i < len
+                           then Heap_Monad.heap (\<lambda>h. (a, Heap.upd a i x h))
+                           else raise (''array update: index out of range''))
+                   done)" 
+
+lemma upd_return:
+  "upd i x a \<guillemotright> return a = upd i x a"
+proof (rule Heap_eqI)
+  fix h
+  obtain len h' where "Heap_Monad.execute (Array.length a) h = (len, h')"
+    by (cases "Heap_Monad.execute (Array.length a) h")
+  then show "Heap_Monad.execute (upd i x a \<guillemotright> return a) h = Heap_Monad.execute (upd i x a) h"
+    by (auto simp add: upd_def bindM_def split: sum.split)
+qed
+
+
+subsection {* Derivates *}
+
+definition
+  map_entry :: "nat \<Rightarrow> ('a\<Colon>heap \<Rightarrow> 'a) \<Rightarrow> 'a array \<Rightarrow> 'a array Heap"
+where
+  "map_entry i f a = (do
+     x \<leftarrow> nth a i;
+     upd i (f x) a
+   done)"
+
+definition
+  swap :: "nat \<Rightarrow> 'a \<Rightarrow> 'a\<Colon>heap array \<Rightarrow> 'a Heap"
+where
+  "swap i x a = (do
+     y \<leftarrow> nth a i;
+     upd i x a;
+     return y
+   done)"
+
+definition
+  make :: "nat \<Rightarrow> (nat \<Rightarrow> 'a\<Colon>heap) \<Rightarrow> 'a array Heap"
+where
+  "make n f = of_list (map f [0 ..< n])"
+
+definition
+  freeze :: "'a\<Colon>heap array \<Rightarrow> 'a list Heap"
+where
+  "freeze a = (do
+     n \<leftarrow> length a;
+     mapM (nth a) [0..<n]
+   done)"
+
+definition
+   map :: "('a\<Colon>heap \<Rightarrow> 'a) \<Rightarrow> 'a array \<Rightarrow> 'a array Heap"
+where
+  "map f a = (do
+     n \<leftarrow> length a;
+     mapM (\<lambda>n. map_entry n f a) [0..<n];
+     return a
+   done)"
+
+hide (open) const new map -- {* avoid clashed with some popular names *}
+
+
+subsection {* Properties *}
+
+lemma array_make [code]:
+  "Array.new n x = make n (\<lambda>_. x)"
+  by (induct n) (simp_all add: make_def new_def Heap_Monad.heap_def
+    monad_simp array_of_list_replicate [symmetric]
+    map_replicate_trivial replicate_append_same
+    of_list_def)
+
+lemma array_of_list_make [code]:
+  "of_list xs = make (List.length xs) (\<lambda>n. xs ! n)"
+  unfolding make_def map_nth ..
+
+
+subsection {* Code generator setup *}
+
+subsubsection {* Logical intermediate layer *}
+
+definition new' where
+  [code del]: "new' = Array.new o nat_of_index"
+hide (open) const new'
+lemma [code]:
+  "Array.new = Array.new' o index_of_nat"
+  by (simp add: new'_def o_def)
+
+definition of_list' where
+  [code del]: "of_list' i xs = Array.of_list (take (nat_of_index i) xs)"
+hide (open) const of_list'
+lemma [code]:
+  "Array.of_list xs = Array.of_list' (index_of_nat (List.length xs)) xs"
+  by (simp add: of_list'_def)
+
+definition make' where
+  [code del]: "make' i f = Array.make (nat_of_index i) (f o index_of_nat)"
+hide (open) const make'
+lemma [code]:
+  "Array.make n f = Array.make' (index_of_nat n) (f o nat_of_index)"
+  by (simp add: make'_def o_def)
+
+definition length' where
+  [code del]: "length' = Array.length \<guillemotright>== liftM index_of_nat"
+hide (open) const length'
+lemma [code]:
+  "Array.length = Array.length' \<guillemotright>== liftM nat_of_index"
+  by (simp add: length'_def monad_simp',
+    simp add: liftM_def comp_def monad_simp,
+    simp add: monad_simp')
+
+definition nth' where
+  [code del]: "nth' a = Array.nth a o nat_of_index"
+hide (open) const nth'
+lemma [code]:
+  "Array.nth a n = Array.nth' a (index_of_nat n)"
+  by (simp add: nth'_def)
+
+definition upd' where
+  [code del]: "upd' a i x = Array.upd (nat_of_index i) x a \<guillemotright> return ()"
+hide (open) const upd'
+lemma [code]:
+  "Array.upd i x a = Array.upd' a (index_of_nat i) x \<guillemotright> return a"
+  by (simp add: upd'_def monad_simp upd_return)
+
+
+subsubsection {* SML *}
+
+code_type array (SML "_/ array")
+code_const Array (SML "raise/ (Fail/ \"bare Array\")")
+code_const Array.new' (SML "(fn/ ()/ =>/ Array.array/ ((_),/ (_)))")
+code_const Array.of_list (SML "(fn/ ()/ =>/ Array.fromList/ _)")
+code_const Array.make' (SML "(fn/ ()/ =>/ Array.tabulate/ ((_),/ (_)))")
+code_const Array.length' (SML "(fn/ ()/ =>/ Array.length/ _)")
+code_const Array.nth' (SML "(fn/ ()/ =>/ Array.sub/ ((_),/ (_)))")
+code_const Array.upd' (SML "(fn/ ()/ =>/ Array.update/ ((_),/ (_),/ (_)))")
+
+code_reserved SML Array
+
+
+subsubsection {* OCaml *}
+
+code_type array (OCaml "_/ array")
+code_const Array (OCaml "failwith/ \"bare Array\"")
+code_const Array.new' (OCaml "(fun/ ()/ ->/ Array.make/ _/ _)")
+code_const Array.of_list (OCaml "(fun/ ()/ ->/ Array.of'_list/ _)")
+code_const Array.make' (OCaml "(fun/ ()/ ->/ Array.init/ _/ _)")
+code_const Array.length' (OCaml "(fun/ ()/ ->/ Array.length/ _)")
+code_const Array.nth' (OCaml "(fun/ ()/ ->/ Array.get/ _/ _)")
+code_const Array.upd' (OCaml "(fun/ ()/ ->/ Array.set/ _/ _/ _)")
+
+code_reserved OCaml Array
+
+
+subsubsection {* Haskell *}
+
+code_type array (Haskell "STArray/ RealWorld/ _")
+code_const Array (Haskell "error/ \"bare Array\"")
+code_const Array.new' (Haskell "newArray/ (0,/ _)")
+code_const Array.of_list' (Haskell "newListArray/ (0,/ _)")
+code_const Array.length' (Haskell "lengthArray")
+code_const Array.nth' (Haskell "readArray")
+code_const Array.upd' (Haskell "writeArray")
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Imperative_HOL/Heap.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -0,0 +1,434 @@
+(*  Title:      HOL/Library/Heap.thy
+    ID:         $Id$
+    Author:     John Matthews, Galois Connections; Alexander Krauss, TU Muenchen
+*)
+
+header {* A polymorphic heap based on cantor encodings *}
+
+theory Heap
+imports Plain "~~/src/HOL/List" Countable Typerep
+begin
+
+subsection {* Representable types *}
+
+text {* The type class of representable types *}
+
+class heap = typerep + countable
+
+text {* Instances for common HOL types *}
+
+instance nat :: heap ..
+
+instance "*" :: (heap, heap) heap ..
+
+instance "+" :: (heap, heap) heap ..
+
+instance list :: (heap) heap ..
+
+instance option :: (heap) heap ..
+
+instance int :: heap ..
+
+instance message_string :: countable
+  by (rule countable_classI [of "message_string_case to_nat"])
+   (auto split: message_string.splits)
+
+instance message_string :: heap ..
+
+text {* Reflected types themselves are heap-representable *}
+
+instantiation typerep :: countable
+begin
+
+fun to_nat_typerep :: "typerep \<Rightarrow> nat" where
+  "to_nat_typerep (Typerep.Typerep c ts) = to_nat (to_nat c, to_nat (map to_nat_typerep ts))"
+
+instance
+proof (rule countable_classI)
+  fix t t' :: typerep and ts
+  have "(\<forall>t'. to_nat_typerep t = to_nat_typerep t' \<longrightarrow> t = t')
+    \<and> (\<forall>ts'. map to_nat_typerep ts = map to_nat_typerep ts' \<longrightarrow> ts = ts')"
+  proof (induct rule: typerep.induct)
+    case (Typerep c ts) show ?case
+    proof (rule allI, rule impI)
+      fix t'
+      assume hyp: "to_nat_typerep (Typerep.Typerep c ts) = to_nat_typerep t'"
+      then obtain c' ts' where t': "t' = (Typerep.Typerep c' ts')"
+        by (cases t') auto
+      with Typerep hyp have "c = c'" and "ts = ts'" by simp_all
+      with t' show "Typerep.Typerep c ts = t'" by simp
+    qed
+  next
+    case Nil_typerep then show ?case by simp
+  next
+    case (Cons_typerep t ts) then show ?case by auto
+  qed
+  then have "to_nat_typerep t = to_nat_typerep t' \<Longrightarrow> t = t'" by auto
+  moreover assume "to_nat_typerep t = to_nat_typerep t'"
+  ultimately show "t = t'" by simp
+qed
+
+end
+
+instance typerep :: heap ..
+
+
+subsection {* A polymorphic heap with dynamic arrays and references *}
+
+types addr = nat -- "untyped heap references"
+
+datatype 'a array = Array addr
+datatype 'a ref = Ref addr -- "note the phantom type 'a "
+
+primrec addr_of_array :: "'a array \<Rightarrow> addr" where
+  "addr_of_array (Array x) = x"
+
+primrec addr_of_ref :: "'a ref \<Rightarrow> addr" where
+  "addr_of_ref (Ref x) = x"
+
+lemma addr_of_array_inj [simp]:
+  "addr_of_array a = addr_of_array a' \<longleftrightarrow> a = a'"
+  by (cases a, cases a') simp_all
+
+lemma addr_of_ref_inj [simp]:
+  "addr_of_ref r = addr_of_ref r' \<longleftrightarrow> r = r'"
+  by (cases r, cases r') simp_all
+
+instance array :: (type) countable
+  by (rule countable_classI [of addr_of_array]) simp
+
+instance ref :: (type) countable
+  by (rule countable_classI [of addr_of_ref]) simp
+
+setup {*
+  Sign.add_const_constraint (@{const_name Array}, SOME @{typ "nat \<Rightarrow> 'a\<Colon>heap array"})
+  #> Sign.add_const_constraint (@{const_name Ref}, SOME @{typ "nat \<Rightarrow> 'a\<Colon>heap ref"})
+  #> Sign.add_const_constraint (@{const_name addr_of_array}, SOME @{typ "'a\<Colon>heap array \<Rightarrow> nat"})
+  #> Sign.add_const_constraint (@{const_name addr_of_ref}, SOME @{typ "'a\<Colon>heap ref \<Rightarrow> nat"})
+*}
+
+types heap_rep = nat -- "representable values"
+
+record heap =
+  arrays :: "typerep \<Rightarrow> addr \<Rightarrow> heap_rep list"
+  refs :: "typerep \<Rightarrow> addr \<Rightarrow> heap_rep"
+  lim  :: addr
+
+definition empty :: heap where
+  "empty = \<lparr>arrays = (\<lambda>_. undefined), refs = (\<lambda>_. undefined), lim = 0\<rparr>" -- "why undefined?"
+
+
+subsection {* Imperative references and arrays *}
+
+text {*
+  References and arrays are developed in parallel,
+  but keeping them separate makes some later proofs simpler.
+*}
+
+subsubsection {* Primitive operations *}
+
+definition
+  new_ref :: "heap \<Rightarrow> ('a\<Colon>heap) ref \<times> heap" where
+  "new_ref h = (let l = lim h in (Ref l, h\<lparr>lim := l + 1\<rparr>))"
+
+definition
+  new_array :: "heap \<Rightarrow> ('a\<Colon>heap) array \<times> heap" where
+  "new_array h = (let l = lim h in (Array l, h\<lparr>lim := l + 1\<rparr>))"
+
+definition
+  ref_present :: "'a\<Colon>heap ref \<Rightarrow> heap \<Rightarrow> bool" where
+  "ref_present r h \<longleftrightarrow> addr_of_ref r < lim h"
+
+definition 
+  array_present :: "'a\<Colon>heap array \<Rightarrow> heap \<Rightarrow> bool" where
+  "array_present a h \<longleftrightarrow> addr_of_array a < lim h"
+
+definition
+  get_ref :: "'a\<Colon>heap ref \<Rightarrow> heap \<Rightarrow> 'a" where
+  "get_ref r h = from_nat (refs h (TYPEREP('a)) (addr_of_ref r))"
+
+definition
+  get_array :: "'a\<Colon>heap array \<Rightarrow> heap \<Rightarrow> 'a list" where
+  "get_array a h = map from_nat (arrays h (TYPEREP('a)) (addr_of_array a))"
+
+definition
+  set_ref :: "'a\<Colon>heap ref \<Rightarrow> 'a \<Rightarrow> heap \<Rightarrow> heap" where
+  "set_ref r x = 
+  refs_update (\<lambda>h. h(TYPEREP('a) := ((h (TYPEREP('a))) (addr_of_ref r:=to_nat x))))"
+
+definition
+  set_array :: "'a\<Colon>heap array \<Rightarrow> 'a list \<Rightarrow> heap \<Rightarrow> heap" where
+  "set_array a x = 
+  arrays_update (\<lambda>h. h(TYPEREP('a) := ((h(TYPEREP('a))) (addr_of_array a:=map to_nat x))))"
+
+subsubsection {* Interface operations *}
+
+definition
+  ref :: "'a \<Rightarrow> heap \<Rightarrow> 'a\<Colon>heap ref \<times> heap" where
+  "ref x h = (let (r, h') = new_ref h;
+                   h''    = set_ref r x h'
+         in (r, h''))"
+
+definition
+  array :: "nat \<Rightarrow> 'a \<Rightarrow> heap \<Rightarrow> 'a\<Colon>heap array \<times> heap" where
+  "array n x h = (let (r, h') = new_array h;
+                       h'' = set_array r (replicate n x) h'
+        in (r, h''))"
+
+definition
+  array_of_list :: "'a list \<Rightarrow> heap \<Rightarrow> 'a\<Colon>heap array \<times> heap" where
+  "array_of_list xs h = (let (r, h') = new_array h;
+           h'' = set_array r xs h'
+        in (r, h''))"  
+
+definition
+  upd :: "'a\<Colon>heap array \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> heap \<Rightarrow> heap" where
+  "upd a i x h = set_array a ((get_array a h)[i:=x]) h"
+
+definition
+  length :: "'a\<Colon>heap array \<Rightarrow> heap \<Rightarrow> nat" where
+  "length a h = size (get_array a h)"
+
+definition
+  array_ran :: "('a\<Colon>heap) option array \<Rightarrow> heap \<Rightarrow> 'a set" where
+  "array_ran a h = {e. Some e \<in> set (get_array a h)}"
+    -- {*FIXME*}
+
+
+subsubsection {* Reference equality *}
+
+text {* 
+  The following relations are useful for comparing arrays and references.
+*}
+
+definition
+  noteq_refs :: "('a\<Colon>heap) ref \<Rightarrow> ('b\<Colon>heap) ref \<Rightarrow> bool" (infix "=!=" 70)
+where
+  "r =!= s \<longleftrightarrow> TYPEREP('a) \<noteq> TYPEREP('b) \<or> addr_of_ref r \<noteq> addr_of_ref s"
+
+definition
+  noteq_arrs :: "('a\<Colon>heap) array \<Rightarrow> ('b\<Colon>heap) array \<Rightarrow> bool" (infix "=!!=" 70)
+where
+  "r =!!= s \<longleftrightarrow> TYPEREP('a) \<noteq> TYPEREP('b) \<or> addr_of_array r \<noteq> addr_of_array s"
+
+lemma noteq_refs_sym: "r =!= s \<Longrightarrow> s =!= r"
+  and noteq_arrs_sym: "a =!!= b \<Longrightarrow> b =!!= a"
+  and unequal_refs [simp]: "r \<noteq> r' \<longleftrightarrow> r =!= r'" -- "same types!"
+  and unequal_arrs [simp]: "a \<noteq> a' \<longleftrightarrow> a =!!= a'"
+unfolding noteq_refs_def noteq_arrs_def by auto
+
+lemma present_new_ref: "ref_present r h \<Longrightarrow> r =!= fst (ref v h)"
+  by (simp add: ref_present_def new_ref_def ref_def Let_def noteq_refs_def)
+
+lemma present_new_arr: "array_present a h \<Longrightarrow> a =!!= fst (array v x h)"
+  by (simp add: array_present_def noteq_arrs_def new_array_def array_def Let_def)
+
+
+subsubsection {* Properties of heap containers *}
+
+text {* Properties of imperative arrays *}
+
+text {* FIXME: Does there exist a "canonical" array axiomatisation in
+the literature?  *}
+
+lemma array_get_set_eq [simp]: "get_array r (set_array r x h) = x"
+  by (simp add: get_array_def set_array_def)
+
+lemma array_get_set_neq [simp]: "r =!!= s \<Longrightarrow> get_array r (set_array s x h) = get_array r h"
+  by (simp add: noteq_arrs_def get_array_def set_array_def)
+
+lemma set_array_same [simp]:
+  "set_array r x (set_array r y h) = set_array r x h"
+  by (simp add: set_array_def)
+
+lemma array_set_set_swap:
+  "r =!!= r' \<Longrightarrow> set_array r x (set_array r' x' h) = set_array r' x' (set_array r x h)"
+  by (simp add: Let_def expand_fun_eq noteq_arrs_def set_array_def)
+
+lemma array_ref_set_set_swap:
+  "set_array r x (set_ref r' x' h) = set_ref r' x' (set_array r x h)"
+  by (simp add: Let_def expand_fun_eq set_array_def set_ref_def)
+
+lemma get_array_upd_eq [simp]:
+  "get_array a (upd a i v h) = (get_array a h) [i := v]"
+  by (simp add: upd_def)
+
+lemma nth_upd_array_neq_array [simp]:
+  "a =!!= b \<Longrightarrow> get_array a (upd b j v h) ! i = get_array a h ! i"
+  by (simp add: upd_def noteq_arrs_def)
+
+lemma get_arry_array_upd_elem_neqIndex [simp]:
+  "i \<noteq> j \<Longrightarrow> get_array a (upd a j v h) ! i = get_array a h ! i"
+  by simp
+
+lemma length_upd_eq [simp]: 
+  "length a (upd a i v h) = length a h" 
+  by (simp add: length_def upd_def)
+
+lemma length_upd_neq [simp]: 
+  "length a (upd b i v h) = length a h"
+  by (simp add: upd_def length_def set_array_def get_array_def)
+
+lemma upd_swap_neqArray:
+  "a =!!= a' \<Longrightarrow> 
+  upd a i v (upd a' i' v' h) 
+  = upd a' i' v' (upd a i v h)"
+apply (unfold upd_def)
+apply simp
+apply (subst array_set_set_swap, assumption)
+apply (subst array_get_set_neq)
+apply (erule noteq_arrs_sym)
+apply (simp)
+done
+
+lemma upd_swap_neqIndex:
+  "\<lbrakk> i \<noteq> i' \<rbrakk> \<Longrightarrow> upd a i v (upd a i' v' h) = upd a i' v' (upd a i v h)"
+by (auto simp add: upd_def array_set_set_swap list_update_swap)
+
+lemma get_array_init_array_list:
+  "get_array (fst (array_of_list ls h)) (snd (array_of_list ls' h)) = ls'"
+  by (simp add: Let_def split_def array_of_list_def)
+
+lemma set_array:
+  "set_array (fst (array_of_list ls h))
+     new_ls (snd (array_of_list ls h))
+       = snd (array_of_list new_ls h)"
+  by (simp add: Let_def split_def array_of_list_def)
+
+lemma array_present_upd [simp]: 
+  "array_present a (upd b i v h) = array_present a h"
+  by (simp add: upd_def array_present_def set_array_def get_array_def)
+
+lemma array_of_list_replicate:
+  "array_of_list (replicate n x) = array n x"
+  by (simp add: expand_fun_eq array_of_list_def array_def)
+
+
+text {* Properties of imperative references *}
+
+lemma next_ref_fresh [simp]:
+  assumes "(r, h') = new_ref h"
+  shows "\<not> ref_present r h"
+  using assms by (cases h) (auto simp add: new_ref_def ref_present_def Let_def)
+
+lemma next_ref_present [simp]:
+  assumes "(r, h') = new_ref h"
+  shows "ref_present r h'"
+  using assms by (cases h) (auto simp add: new_ref_def ref_present_def Let_def)
+
+lemma ref_get_set_eq [simp]: "get_ref r (set_ref r x h) = x"
+  by (simp add: get_ref_def set_ref_def)
+
+lemma ref_get_set_neq [simp]: "r =!= s \<Longrightarrow> get_ref r (set_ref s x h) = get_ref r h"
+  by (simp add: noteq_refs_def get_ref_def set_ref_def)
+
+(* FIXME: We need some infrastructure to infer that locally generated
+  new refs (by new_ref(_no_init), new_array(')) are distinct
+  from all existing refs.
+*)
+
+lemma ref_set_get: "set_ref r (get_ref r h) h = h"
+apply (simp add: set_ref_def get_ref_def)
+oops
+
+lemma set_ref_same[simp]:
+  "set_ref r x (set_ref r y h) = set_ref r x h"
+  by (simp add: set_ref_def)
+
+lemma ref_set_set_swap:
+  "r =!= r' \<Longrightarrow> set_ref r x (set_ref r' x' h) = set_ref r' x' (set_ref r x h)"
+  by (simp add: Let_def expand_fun_eq noteq_refs_def set_ref_def)
+
+lemma ref_new_set: "fst (ref v (set_ref r v' h)) = fst (ref v h)"
+  by (simp add: ref_def new_ref_def set_ref_def Let_def)
+
+lemma ref_get_new [simp]:
+  "get_ref (fst (ref v h)) (snd (ref v' h)) = v'"
+  by (simp add: ref_def Let_def split_def)
+
+lemma ref_set_new [simp]:
+  "set_ref (fst (ref v h)) new_v (snd (ref v h)) = snd (ref new_v h)"
+  by (simp add: ref_def Let_def split_def)
+
+lemma ref_get_new_neq: "r =!= (fst (ref v h)) \<Longrightarrow> 
+  get_ref r (snd (ref v h)) = get_ref r h"
+  by (simp add: get_ref_def set_ref_def ref_def Let_def new_ref_def noteq_refs_def)
+
+lemma lim_set_ref [simp]:
+  "lim (set_ref r v h) = lim h"
+  by (simp add: set_ref_def)
+
+lemma ref_present_new_ref [simp]: 
+  "ref_present r h \<Longrightarrow> ref_present r (snd (ref v h))"
+  by (simp add: new_ref_def ref_present_def ref_def Let_def)
+
+lemma ref_present_set_ref [simp]:
+  "ref_present r (set_ref r' v h) = ref_present r h"
+  by (simp add: set_ref_def ref_present_def)
+
+lemma array_ranI: "\<lbrakk> Some b = get_array a h ! i; i < Heap.length a h \<rbrakk> \<Longrightarrow> b \<in> array_ran a h"
+unfolding array_ran_def Heap.length_def by simp
+
+lemma array_ran_upd_array_Some:
+  assumes "cl \<in> array_ran a (Heap.upd a i (Some b) h)"
+  shows "cl \<in> array_ran a h \<or> cl = b"
+proof -
+  have "set (get_array a h[i := Some b]) \<subseteq> insert (Some b) (set (get_array a h))" by (rule set_update_subset_insert)
+  with assms show ?thesis 
+    unfolding array_ran_def Heap.upd_def by fastsimp
+qed
+
+lemma array_ran_upd_array_None:
+  assumes "cl \<in> array_ran a (Heap.upd a i None h)"
+  shows "cl \<in> array_ran a h"
+proof -
+  have "set (get_array a h[i := None]) \<subseteq> insert None (set (get_array a h))" by (rule set_update_subset_insert)
+  with assms show ?thesis
+    unfolding array_ran_def Heap.upd_def by auto
+qed
+
+
+text {* Non-interaction between imperative array and imperative references *}
+
+lemma get_array_set_ref [simp]: "get_array a (set_ref r v h) = get_array a h"
+  by (simp add: get_array_def set_ref_def)
+
+lemma nth_set_ref [simp]: "get_array a (set_ref r v h) ! i = get_array a h ! i"
+  by simp
+
+lemma get_ref_upd [simp]: "get_ref r (upd a i v h) = get_ref r h"
+  by (simp add: get_ref_def set_array_def upd_def)
+
+lemma new_ref_upd: "fst (ref v (upd a i v' h)) = fst (ref v h)"
+  by (simp add: set_array_def get_array_def Let_def ref_new_set upd_def ref_def  new_ref_def)
+
+text {*not actually true ???*}
+lemma upd_set_ref_swap: "upd a i v (set_ref r v' h) = set_ref r v' (upd a i v h)"
+apply (case_tac a)
+apply (simp add: Let_def upd_def)
+apply auto
+oops
+
+lemma length_new_ref[simp]: 
+  "length a (snd (ref v h)) = length a h"
+  by (simp add: get_array_def set_ref_def length_def new_ref_def ref_def Let_def)
+
+lemma get_array_new_ref [simp]: 
+  "get_array a (snd (ref v h)) = get_array a h"
+  by (simp add: new_ref_def ref_def set_ref_def get_array_def Let_def)
+
+lemma ref_present_upd [simp]: 
+  "ref_present r (upd a i v h) = ref_present r h"
+  by (simp add: upd_def ref_present_def set_array_def get_array_def)
+
+lemma array_present_set_ref [simp]:
+  "array_present a (set_ref r v h) = array_present a h"
+  by (simp add: array_present_def set_ref_def)
+
+lemma array_present_new_ref [simp]:
+  "array_present a h \<Longrightarrow> array_present a (snd (ref v h))"
+  by (simp add: array_present_def new_ref_def ref_def Let_def)
+
+hide (open) const empty array array_of_list upd length ref
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Imperative_HOL/Heap_Monad.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -0,0 +1,425 @@
+(*  Title:      HOL/Library/Heap_Monad.thy
+    ID:         $Id$
+    Author:     John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen
+*)
+
+header {* A monad with a polymorphic heap *}
+
+theory Heap_Monad
+imports Heap
+begin
+
+subsection {* The monad *}
+
+subsubsection {* Monad combinators *}
+
+datatype exception = Exn
+
+text {* Monadic heap actions either produce values
+  and transform the heap, or fail *}
+datatype 'a Heap = Heap "heap \<Rightarrow> ('a + exception) \<times> heap"
+
+primrec
+  execute :: "'a Heap \<Rightarrow> heap \<Rightarrow> ('a + exception) \<times> heap" where
+  "execute (Heap f) = f"
+lemmas [code del] = execute.simps
+
+lemma Heap_execute [simp]:
+  "Heap (execute f) = f" by (cases f) simp_all
+
+lemma Heap_eqI:
+  "(\<And>h. execute f h = execute g h) \<Longrightarrow> f = g"
+    by (cases f, cases g) (auto simp: expand_fun_eq)
+
+lemma Heap_eqI':
+  "(\<And>h. (\<lambda>x. execute (f x) h) = (\<lambda>y. execute (g y) h)) \<Longrightarrow> f = g"
+    by (auto simp: expand_fun_eq intro: Heap_eqI)
+
+lemma Heap_strip: "(\<And>f. PROP P f) \<equiv> (\<And>g. PROP P (Heap g))"
+proof
+  fix g :: "heap \<Rightarrow> ('a + exception) \<times> heap" 
+  assume "\<And>f. PROP P f"
+  then show "PROP P (Heap g)" .
+next
+  fix f :: "'a Heap" 
+  assume assm: "\<And>g. PROP P (Heap g)"
+  then have "PROP P (Heap (execute f))" .
+  then show "PROP P f" by simp
+qed
+
+definition
+  heap :: "(heap \<Rightarrow> 'a \<times> heap) \<Rightarrow> 'a Heap" where
+  [code del]: "heap f = Heap (\<lambda>h. apfst Inl (f h))"
+
+lemma execute_heap [simp]:
+  "execute (heap f) h = apfst Inl (f h)"
+  by (simp add: heap_def)
+
+definition
+  bindM :: "'a Heap \<Rightarrow> ('a \<Rightarrow> 'b Heap) \<Rightarrow> 'b Heap" (infixl ">>=" 54) where
+  [code del]: "f >>= g = Heap (\<lambda>h. case execute f h of
+                  (Inl x, h') \<Rightarrow> execute (g x) h'
+                | r \<Rightarrow> r)"
+
+notation
+  bindM (infixl "\<guillemotright>=" 54)
+
+abbreviation
+  chainM :: "'a Heap \<Rightarrow> 'b Heap \<Rightarrow> 'b Heap"  (infixl ">>" 54) where
+  "f >> g \<equiv> f >>= (\<lambda>_. g)"
+
+notation
+  chainM (infixl "\<guillemotright>" 54)
+
+definition
+  return :: "'a \<Rightarrow> 'a Heap" where
+  [code del]: "return x = heap (Pair x)"
+
+lemma execute_return [simp]:
+  "execute (return x) h = apfst Inl (x, h)"
+  by (simp add: return_def)
+
+definition
+  raise :: "string \<Rightarrow> 'a Heap" where -- {* the string is just decoration *}
+  [code del]: "raise s = Heap (Pair (Inr Exn))"
+
+notation (latex output)
+  "raise" ("\<^raw:{\textsf{raise}}>")
+
+lemma execute_raise [simp]:
+  "execute (raise s) h = (Inr Exn, h)"
+  by (simp add: raise_def)
+
+
+subsubsection {* do-syntax *}
+
+text {*
+  We provide a convenient do-notation for monadic expressions
+  well-known from Haskell.  @{const Let} is printed
+  specially in do-expressions.
+*}
+
+nonterminals do_expr
+
+syntax
+  "_do" :: "do_expr \<Rightarrow> 'a"
+    ("(do (_)//done)" [12] 100)
+  "_bindM" :: "pttrn \<Rightarrow> 'a \<Rightarrow> do_expr \<Rightarrow> do_expr"
+    ("_ <- _;//_" [1000, 13, 12] 12)
+  "_chainM" :: "'a \<Rightarrow> do_expr \<Rightarrow> do_expr"
+    ("_;//_" [13, 12] 12)
+  "_let" :: "pttrn \<Rightarrow> 'a \<Rightarrow> do_expr \<Rightarrow> do_expr"
+    ("let _ = _;//_" [1000, 13, 12] 12)
+  "_nil" :: "'a \<Rightarrow> do_expr"
+    ("_" [12] 12)
+
+syntax (xsymbols)
+  "_bindM" :: "pttrn \<Rightarrow> 'a \<Rightarrow> do_expr \<Rightarrow> do_expr"
+    ("_ \<leftarrow> _;//_" [1000, 13, 12] 12)
+syntax (latex output)
+  "_do" :: "do_expr \<Rightarrow> 'a"
+    ("(\<^raw:{\textsf{do}}> (_))" [12] 100)
+  "_let" :: "pttrn \<Rightarrow> 'a \<Rightarrow> do_expr \<Rightarrow> do_expr"
+    ("\<^raw:\textsf{let}> _ = _;//_" [1000, 13, 12] 12)
+notation (latex output)
+  "return" ("\<^raw:{\textsf{return}}>")
+
+translations
+  "_do f" => "f"
+  "_bindM x f g" => "f \<guillemotright>= (\<lambda>x. g)"
+  "_chainM f g" => "f \<guillemotright> g"
+  "_let x t f" => "CONST Let t (\<lambda>x. f)"
+  "_nil f" => "f"
+
+print_translation {*
+let
+  fun dest_abs_eta (Abs (abs as (_, ty, _))) =
+        let
+          val (v, t) = Syntax.variant_abs abs;
+        in (Free (v, ty), t) end
+    | dest_abs_eta t =
+        let
+          val (v, t) = Syntax.variant_abs ("", dummyT, t $ Bound 0);
+        in (Free (v, dummyT), t) end;
+  fun unfold_monad (Const (@{const_syntax bindM}, _) $ f $ g) =
+        let
+          val (v, g') = dest_abs_eta g;
+          val vs = fold_aterms (fn Free (v, _) => insert (op =) v | _ => I) v [];
+          val v_used = fold_aterms
+            (fn Free (w, _) => (fn s => s orelse member (op =) vs w) | _ => I) g' false;
+        in if v_used then
+          Const ("_bindM", dummyT) $ v $ f $ unfold_monad g'
+        else
+          Const ("_chainM", dummyT) $ f $ unfold_monad g'
+        end
+    | unfold_monad (Const (@{const_syntax chainM}, _) $ f $ g) =
+        Const ("_chainM", dummyT) $ f $ unfold_monad g
+    | unfold_monad (Const (@{const_syntax Let}, _) $ f $ g) =
+        let
+          val (v, g') = dest_abs_eta g;
+        in Const ("_let", dummyT) $ v $ f $ unfold_monad g' end
+    | unfold_monad (Const (@{const_syntax Pair}, _) $ f) =
+        Const (@{const_syntax return}, dummyT) $ f
+    | unfold_monad f = f;
+  fun contains_bindM (Const (@{const_syntax bindM}, _) $ _ $ _) = true
+    | contains_bindM (Const (@{const_syntax Let}, _) $ _ $ Abs (_, _, t)) =
+        contains_bindM t;
+  fun bindM_monad_tr' (f::g::ts) = list_comb
+    (Const ("_do", dummyT) $ unfold_monad (Const (@{const_syntax bindM}, dummyT) $ f $ g), ts);
+  fun Let_monad_tr' (f :: (g as Abs (_, _, g')) :: ts) = if contains_bindM g' then list_comb
+      (Const ("_do", dummyT) $ unfold_monad (Const (@{const_syntax Let}, dummyT) $ f $ g), ts)
+    else raise Match;
+in [
+  (@{const_syntax bindM}, bindM_monad_tr'),
+  (@{const_syntax Let}, Let_monad_tr')
+] end;
+*}
+
+
+subsection {* Monad properties *}
+
+subsubsection {* Monad laws *}
+
+lemma return_bind: "return x \<guillemotright>= f = f x"
+  by (simp add: bindM_def return_def)
+
+lemma bind_return: "f \<guillemotright>= return = f"
+proof (rule Heap_eqI)
+  fix h
+  show "execute (f \<guillemotright>= return) h = execute f h"
+    by (auto simp add: bindM_def return_def split: sum.splits prod.splits)
+qed
+
+lemma bind_bind: "(f \<guillemotright>= g) \<guillemotright>= h = f \<guillemotright>= (\<lambda>x. g x \<guillemotright>= h)"
+  by (rule Heap_eqI) (auto simp add: bindM_def split: split: sum.splits prod.splits)
+
+lemma bind_bind': "f \<guillemotright>= (\<lambda>x. g x \<guillemotright>= h x) = f \<guillemotright>= (\<lambda>x. g x \<guillemotright>= (\<lambda>y. return (x, y))) \<guillemotright>= (\<lambda>(x, y). h x y)"
+  by (rule Heap_eqI) (auto simp add: bindM_def split: split: sum.splits prod.splits)
+
+lemma raise_bind: "raise e \<guillemotright>= f = raise e"
+  by (simp add: raise_def bindM_def)
+
+
+lemmas monad_simp = return_bind bind_return bind_bind raise_bind
+
+
+subsection {* Generic combinators *}
+
+definition
+  liftM :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b Heap"
+where
+  "liftM f = return o f"
+
+definition
+  compM :: "('a \<Rightarrow> 'b Heap) \<Rightarrow> ('b \<Rightarrow> 'c Heap) \<Rightarrow> 'a \<Rightarrow> 'c Heap" (infixl ">>==" 54)
+where
+  "(f >>== g) = (\<lambda>x. f x \<guillemotright>= g)"
+
+notation
+  compM (infixl "\<guillemotright>==" 54)
+
+lemma liftM_collapse: "liftM f x = return (f x)"
+  by (simp add: liftM_def)
+
+lemma liftM_compM: "liftM f \<guillemotright>== g = g o f"
+  by (auto intro: Heap_eqI' simp add: expand_fun_eq liftM_def compM_def bindM_def)
+
+lemma compM_return: "f \<guillemotright>== return = f"
+  by (simp add: compM_def monad_simp)
+
+lemma compM_compM: "(f \<guillemotright>== g) \<guillemotright>== h = f \<guillemotright>== (g \<guillemotright>== h)"
+  by (simp add: compM_def monad_simp)
+
+lemma liftM_bind:
+  "(\<lambda>x. liftM f x \<guillemotright>= liftM g) = liftM (\<lambda>x. g (f x))"
+  by (rule Heap_eqI') (simp add: monad_simp liftM_def bindM_def)
+
+lemma liftM_comp:
+  "liftM f o g = liftM (f o g)"
+  by (rule Heap_eqI') (simp add: liftM_def)
+
+lemmas monad_simp' = monad_simp liftM_compM compM_return
+  compM_compM liftM_bind liftM_comp
+
+primrec 
+  mapM :: "('a \<Rightarrow> 'b Heap) \<Rightarrow> 'a list \<Rightarrow> 'b list Heap"
+where
+  "mapM f [] = return []"
+  | "mapM f (x#xs) = do y \<leftarrow> f x;
+                        ys \<leftarrow> mapM f xs;
+                        return (y # ys)
+                     done"
+
+primrec
+  foldM :: "('a \<Rightarrow> 'b \<Rightarrow> 'b Heap) \<Rightarrow> 'a list \<Rightarrow> 'b \<Rightarrow> 'b Heap"
+where
+  "foldM f [] s = return s"
+  | "foldM f (x#xs) s = f x s \<guillemotright>= foldM f xs"
+
+definition
+  assert :: "('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a Heap"
+where
+  "assert P x = (if P x then return x else raise (''assert''))"
+
+lemma assert_cong [fundef_cong]:
+  assumes "P = P'"
+  assumes "\<And>x. P' x \<Longrightarrow> f x = f' x"
+  shows "(assert P x >>= f) = (assert P' x >>= f')"
+  using assms by (auto simp add: assert_def return_bind raise_bind)
+
+hide (open) const heap execute
+
+
+subsection {* Code generator setup *}
+
+subsubsection {* Logical intermediate layer *}
+
+definition
+  Fail :: "message_string \<Rightarrow> exception"
+where
+  [code del]: "Fail s = Exn"
+
+definition
+  raise_exc :: "exception \<Rightarrow> 'a Heap"
+where
+  [code del]: "raise_exc e = raise []"
+
+lemma raise_raise_exc [code, code inline]:
+  "raise s = raise_exc (Fail (STR s))"
+  unfolding Fail_def raise_exc_def raise_def ..
+
+hide (open) const Fail raise_exc
+
+
+subsubsection {* SML and OCaml *}
+
+code_type Heap (SML "unit/ ->/ _")
+code_const Heap (SML "raise/ (Fail/ \"bare Heap\")")
+code_const "op \<guillemotright>=" (SML "!(fn/ f'_/ =>/ fn/ ()/ =>/ f'_/ (_/ ())/ ())")
+code_const return (SML "!(fn/ ()/ =>/ _)")
+code_const "Heap_Monad.Fail" (SML "Fail")
+code_const "Heap_Monad.raise_exc" (SML "!(fn/ ()/ =>/ raise/ _)")
+
+code_type Heap (OCaml "_")
+code_const Heap (OCaml "failwith/ \"bare Heap\"")
+code_const "op \<guillemotright>=" (OCaml "!(fun/ f'_/ ()/ ->/ f'_/ (_/ ())/ ())")
+code_const return (OCaml "!(fun/ ()/ ->/ _)")
+code_const "Heap_Monad.Fail" (OCaml "Failure")
+code_const "Heap_Monad.raise_exc" (OCaml "!(fun/ ()/ ->/ raise/ _)")
+
+setup {* let
+  open Code_Thingol;
+
+  fun lookup naming = the o Code_Thingol.lookup_const naming;
+
+  fun imp_monad_bind'' bind' return' unit' ts =
+    let
+      val dummy_name = "";
+      val dummy_type = ITyVar dummy_name;
+      val dummy_case_term = IVar dummy_name;
+      (*assumption: dummy values are not relevant for serialization*)
+      val unitt = IConst (unit', ([], []));
+      fun dest_abs ((v, ty) `|-> t, _) = ((v, ty), t)
+        | dest_abs (t, ty) =
+            let
+              val vs = Code_Thingol.fold_varnames cons t [];
+              val v = Name.variant vs "x";
+              val ty' = (hd o fst o Code_Thingol.unfold_fun) ty;
+            in ((v, ty'), t `$ IVar v) end;
+      fun force (t as IConst (c, _) `$ t') = if c = return'
+            then t' else t `$ unitt
+        | force t = t `$ unitt;
+      fun tr_bind' [(t1, _), (t2, ty2)] =
+        let
+          val ((v, ty), t) = dest_abs (t2, ty2);
+        in ICase (((force t1, ty), [(IVar v, tr_bind'' t)]), dummy_case_term) end
+      and tr_bind'' t = case Code_Thingol.unfold_app t
+           of (IConst (c, (_, ty1 :: ty2 :: _)), [x1, x2]) => if c = bind'
+                then tr_bind' [(x1, ty1), (x2, ty2)]
+                else force t
+            | _ => force t;
+    in (dummy_name, dummy_type) `|-> ICase (((IVar dummy_name, dummy_type),
+      [(unitt, tr_bind' ts)]), dummy_case_term) end
+  and imp_monad_bind' bind' return' unit' (const as (c, (_, tys))) ts = if c = bind' then case (ts, tys)
+     of ([t1, t2], ty1 :: ty2 :: _) => imp_monad_bind'' bind' return' unit' [(t1, ty1), (t2, ty2)]
+      | ([t1, t2, t3], ty1 :: ty2 :: _) => imp_monad_bind'' bind' return' unit' [(t1, ty1), (t2, ty2)] `$ t3
+      | (ts, _) => imp_monad_bind bind' return' unit' (eta_expand 2 (const, ts))
+    else IConst const `$$ map (imp_monad_bind bind' return' unit') ts
+  and imp_monad_bind bind' return' unit' (IConst const) = imp_monad_bind' bind' return' unit' const []
+    | imp_monad_bind bind' return' unit' (t as IVar _) = t
+    | imp_monad_bind bind' return' unit' (t as _ `$ _) = (case unfold_app t
+       of (IConst const, ts) => imp_monad_bind' bind' return' unit' const ts
+        | (t, ts) => imp_monad_bind bind' return' unit' t `$$ map (imp_monad_bind bind' return' unit') ts)
+    | imp_monad_bind bind' return' unit' (v_ty `|-> t) = v_ty `|-> imp_monad_bind bind' return' unit' t
+    | imp_monad_bind bind' return' unit' (ICase (((t, ty), pats), t0)) = ICase
+        (((imp_monad_bind bind' return' unit' t, ty), (map o pairself) (imp_monad_bind bind' return' unit') pats), imp_monad_bind bind' return' unit' t0);
+
+   fun imp_program naming = (Graph.map_nodes o map_terms_stmt)
+     (imp_monad_bind (lookup naming @{const_name bindM})
+       (lookup naming @{const_name return})
+       (lookup naming @{const_name Unity}));
+
+in
+
+  Code_Target.extend_target ("SML_imp", ("SML", imp_program))
+  #> Code_Target.extend_target ("OCaml_imp", ("OCaml", imp_program))
+
+end
+*}
+
+
+code_reserved OCaml Failure raise
+
+
+subsubsection {* Haskell *}
+
+text {* Adaption layer *}
+
+code_include Haskell "STMonad"
+{*import qualified Control.Monad;
+import qualified Control.Monad.ST;
+import qualified Data.STRef;
+import qualified Data.Array.ST;
+
+type RealWorld = Control.Monad.ST.RealWorld;
+type ST s a = Control.Monad.ST.ST s a;
+type STRef s a = Data.STRef.STRef s a;
+type STArray s a = Data.Array.ST.STArray s Int a;
+
+runST :: (forall s. ST s a) -> a;
+runST s = Control.Monad.ST.runST s;
+
+newSTRef = Data.STRef.newSTRef;
+readSTRef = Data.STRef.readSTRef;
+writeSTRef = Data.STRef.writeSTRef;
+
+newArray :: (Int, Int) -> a -> ST s (STArray s a);
+newArray = Data.Array.ST.newArray;
+
+newListArray :: (Int, Int) -> [a] -> ST s (STArray s a);
+newListArray = Data.Array.ST.newListArray;
+
+lengthArray :: STArray s a -> ST s Int;
+lengthArray a = Control.Monad.liftM snd (Data.Array.ST.getBounds a);
+
+readArray :: STArray s a -> Int -> ST s a;
+readArray = Data.Array.ST.readArray;
+
+writeArray :: STArray s a -> Int -> a -> ST s ();
+writeArray = Data.Array.ST.writeArray;*}
+
+code_reserved Haskell RealWorld ST STRef Array
+  runST
+  newSTRef reasSTRef writeSTRef
+  newArray newListArray lengthArray readArray writeArray
+
+text {* Monad *}
+
+code_type Heap (Haskell "ST/ RealWorld/ _")
+code_const Heap (Haskell "error/ \"bare Heap\"")
+code_monad "op \<guillemotright>=" Haskell
+code_const return (Haskell "return")
+code_const "Heap_Monad.Fail" (Haskell "_")
+code_const "Heap_Monad.raise_exc" (Haskell "error")
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Imperative_HOL/Imperative_HOL.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -0,0 +1,12 @@
+(*  Title:      HOL/Library/Imperative_HOL.thy
+    ID:         $Id$
+    Author:     John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen
+*)
+
+header {* Entry point into monadic imperative HOL *}
+
+theory Imperative_HOL
+imports Array Ref Relational
+begin
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Imperative_HOL/ROOT.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -0,0 +1,2 @@
+
+use_thy "Imperative_HOL";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Imperative_HOL/Ref.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -0,0 +1,91 @@
+(*  Title:      HOL/Library/Ref.thy
+    ID:         $Id$
+    Author:     John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen
+*)
+
+header {* Monadic references *}
+
+theory Ref
+imports Heap_Monad
+begin
+
+text {*
+  Imperative reference operations; modeled after their ML counterparts.
+  See http://caml.inria.fr/pub/docs/manual-caml-light/node14.15.html
+  and http://www.smlnj.org/doc/Conversion/top-level-comparison.html
+*}
+
+subsection {* Primitives *}
+
+definition
+  new :: "'a\<Colon>heap \<Rightarrow> 'a ref Heap" where
+  [code del]: "new v = Heap_Monad.heap (Heap.ref v)"
+
+definition
+  lookup :: "'a\<Colon>heap ref \<Rightarrow> 'a Heap" ("!_" 61) where
+  [code del]: "lookup r = Heap_Monad.heap (\<lambda>h. (get_ref r h, h))"
+
+definition
+  update :: "'a ref \<Rightarrow> ('a\<Colon>heap) \<Rightarrow> unit Heap" ("_ := _" 62) where
+  [code del]: "update r e = Heap_Monad.heap (\<lambda>h. ((), set_ref r e h))"
+
+
+subsection {* Derivates *}
+
+definition
+  change :: "('a\<Colon>heap \<Rightarrow> 'a) \<Rightarrow> 'a ref \<Rightarrow> 'a Heap"
+where
+  "change f r = (do x \<leftarrow> ! r;
+                    let y = f x;
+                    r := y;
+                    return y
+                 done)"
+
+hide (open) const new lookup update change
+
+
+subsection {* Properties *}
+
+lemma lookup_chain:
+  "(!r \<guillemotright> f) = f"
+  by (cases f)
+    (auto simp add: Let_def bindM_def lookup_def expand_fun_eq)
+
+lemma update_change [code]:
+  "r := e = Ref.change (\<lambda>_. e) r \<guillemotright> return ()"
+  by (auto simp add: monad_simp change_def lookup_chain)
+
+
+subsection {* Code generator setup *}
+
+subsubsection {* SML *}
+
+code_type ref (SML "_/ ref")
+code_const Ref (SML "raise/ (Fail/ \"bare Ref\")")
+code_const Ref.new (SML "(fn/ ()/ =>/ ref/ _)")
+code_const Ref.lookup (SML "(fn/ ()/ =>/ !/ _)")
+code_const Ref.update (SML "(fn/ ()/ =>/ _/ :=/ _)")
+
+code_reserved SML ref
+
+
+subsubsection {* OCaml *}
+
+code_type ref (OCaml "_/ ref")
+code_const Ref (OCaml "failwith/ \"bare Ref\")")
+code_const Ref.new (OCaml "(fn/ ()/ =>/ ref/ _)")
+code_const Ref.lookup (OCaml "(fn/ ()/ =>/ !/ _)")
+code_const Ref.update (OCaml "(fn/ ()/ =>/ _/ :=/ _)")
+
+code_reserved OCaml ref
+
+
+subsubsection {* Haskell *}
+
+code_type ref (Haskell "STRef/ RealWorld/ _")
+code_const Ref (Haskell "error/ \"bare Ref\"")
+code_const Ref.new (Haskell "newSTRef")
+code_const Ref.lookup (Haskell "readSTRef")
+code_const Ref.update (Haskell "writeSTRef")
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Imperative_HOL/Relational.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -0,0 +1,700 @@
+theory Relational 
+imports Array Ref
+begin
+
+section{* Definition of the Relational framework *}
+
+text {* The crel predicate states that when a computation c runs with the heap h
+  will result in return value r and a heap h' (if no exception occurs). *}  
+
+definition crel :: "'a Heap \<Rightarrow> heap \<Rightarrow> heap \<Rightarrow> 'a \<Rightarrow> bool"
+where
+  crel_def': "crel c h h' r \<longleftrightarrow> Heap_Monad.execute c h = (Inl r, h')"
+
+lemma crel_def: -- FIXME
+  "crel c h h' r \<longleftrightarrow> (Inl r, h') = Heap_Monad.execute c h"
+  unfolding crel_def' by auto
+
+lemma crel_deterministic: "\<lbrakk> crel f h h' a; crel f h h'' b \<rbrakk> \<Longrightarrow> (a = b) \<and> (h' = h'')"
+unfolding crel_def' by auto
+
+section {* Elimination rules *}
+
+text {* For all commands, we define simple elimination rules. *}
+(* FIXME: consumes 1 necessary ?? *)
+
+subsection {* Elimination rules for basic monadic commands *}
+
+lemma crelE[consumes 1]:
+  assumes "crel (f >>= g) h h'' r'"
+  obtains h' r where "crel f h h' r" "crel (g r) h' h'' r'"
+  using assms
+  by (auto simp add: crel_def bindM_def Let_def prod_case_beta split_def Pair_fst_snd_eq split add: sum.split_asm)
+
+lemma crelE'[consumes 1]:
+  assumes "crel (f >> g) h h'' r'"
+  obtains h' r where "crel f h h' r" "crel g h' h'' r'"
+  using assms
+  by (elim crelE) auto
+
+lemma crel_return[consumes 1]:
+  assumes "crel (return x) h h' r"
+  obtains "r = x" "h = h'"
+  using assms
+  unfolding crel_def return_def by simp
+
+lemma crel_raise[consumes 1]:
+  assumes "crel (raise x) h h' r"
+  obtains "False"
+  using assms
+  unfolding crel_def raise_def by simp
+
+lemma crel_if:
+  assumes "crel (if c then t else e) h h' r"
+  obtains "c" "crel t h h' r"
+        | "\<not>c" "crel e h h' r"
+  using assms
+  unfolding crel_def by auto
+
+lemma crel_option_case:
+  assumes "crel (case x of None \<Rightarrow> n | Some y \<Rightarrow> s y) h h' r"
+  obtains "x = None" "crel n h h' r"
+        | y where "x = Some y" "crel (s y) h h' r" 
+  using assms
+  unfolding crel_def by auto
+
+lemma crel_mapM:
+  assumes "crel (mapM f xs) h h' r"
+  assumes "\<And>h h'. P f [] h h' []"
+  assumes "\<And>h h1 h' x xs y ys. \<lbrakk> crel (f x) h h1 y; crel (mapM f xs) h1 h' ys; P f xs h1 h' ys \<rbrakk> \<Longrightarrow> P f (x#xs) h h' (y#ys)"
+  shows "P f xs h h' r"
+using assms(1)
+proof (induct xs arbitrary: h h' r)
+  case Nil  with assms(2) show ?case
+    by (auto elim: crel_return)
+next
+  case (Cons x xs)
+  from Cons(2) obtain h1 y ys where crel_f: "crel (f x) h h1 y"
+    and crel_mapM: "crel (mapM f xs) h1 h' ys"
+    and r_def: "r = y#ys"
+    unfolding mapM.simps
+    by (auto elim!: crelE crel_return)
+  from Cons(1)[OF crel_mapM] crel_mapM crel_f assms(3) r_def
+  show ?case by auto
+qed
+
+lemma crel_heap:
+  assumes "crel (Heap_Monad.heap f) h h' r"
+  obtains "h' = snd (f h)" "r = fst (f h)"
+  using assms
+  unfolding heap_def crel_def apfst_def split_def prod_fun_def by simp_all
+
+subsection {* Elimination rules for array commands *}
+
+lemma crel_length:
+  assumes "crel (length a) h h' r"
+  obtains "h = h'" "r = Heap.length a h'"
+  using assms
+  unfolding length_def
+  by (elim crel_heap) simp
+
+(* Strong version of the lemma for this operation is missing *) 
+lemma crel_new_weak:
+  assumes "crel (Array.new n v) h h' r"
+  obtains "get_array r h' = List.replicate n v"
+  using assms unfolding  Array.new_def
+  by (elim crel_heap) (auto simp:Heap.array_def Let_def split_def)
+
+lemma crel_nth[consumes 1]:
+  assumes "crel (nth a i) h h' r"
+  obtains "r = (get_array a h) ! i" "h = h'" "i < Heap.length a h"
+  using assms
+  unfolding nth_def
+  by (auto elim!: crelE crel_if crel_raise crel_length crel_heap)
+
+lemma crel_upd[consumes 1]:
+  assumes "crel (upd i v a) h h' r"
+  obtains "r = a" "h' = Heap.upd a i v h"
+  using assms
+  unfolding upd_def
+  by (elim crelE crel_if crel_return crel_raise
+    crel_length crel_heap) auto
+
+(* Strong version of the lemma for this operation is missing *)
+lemma crel_of_list_weak:
+  assumes "crel (Array.of_list xs) h h' r"
+  obtains "get_array r h' = xs"
+  using assms
+  unfolding of_list_def 
+  by (elim crel_heap) (simp add:get_array_init_array_list)
+
+lemma crel_map_entry:
+  assumes "crel (Array.map_entry i f a) h h' r"
+  obtains "r = a" "h' = Heap.upd a i (f (get_array a h ! i)) h"
+  using assms
+  unfolding Array.map_entry_def
+  by (elim crelE crel_upd crel_nth) auto
+
+lemma crel_swap:
+  assumes "crel (Array.swap i x a) h h' r"
+  obtains "r = get_array a h ! i" "h' = Heap.upd a i x h"
+  using assms
+  unfolding Array.swap_def
+  by (elim crelE crel_upd crel_nth crel_return) auto
+
+(* Strong version of the lemma for this operation is missing *)
+lemma crel_make_weak:
+  assumes "crel (Array.make n f) h h' r"
+  obtains "i < n \<Longrightarrow> get_array r h' ! i = f i"
+  using assms
+  unfolding Array.make_def
+  by (elim crel_of_list_weak) auto
+
+lemma upt_conv_Cons':
+  assumes "Suc a \<le> b"
+  shows "[b - Suc a..<b] = (b - Suc a)#[b - a..<b]"
+proof -
+  from assms have l: "b - Suc a < b" by arith
+  from assms have "Suc (b - Suc a) = b - a" by arith
+  with l show ?thesis by (simp add: upt_conv_Cons)
+qed
+
+lemma crel_mapM_nth:
+  assumes
+  "crel (mapM (Array.nth a) [Heap.length a h - n..<Heap.length a h]) h h' xs"
+  assumes "n \<le> Heap.length a h"
+  shows "h = h' \<and> xs = drop (Heap.length a h - n) (get_array a h)"
+using assms
+proof (induct n arbitrary: xs h h')
+  case 0 thus ?case
+    by (auto elim!: crel_return simp add: Heap.length_def)
+next
+  case (Suc n)
+  from Suc(3) have "[Heap.length a h - Suc n..<Heap.length a h] = (Heap.length a h - Suc n)#[Heap.length a h - n..<Heap.length a h]"
+    by (simp add: upt_conv_Cons')
+  with Suc(2) obtain r where
+    crel_mapM: "crel (mapM (Array.nth a) [Heap.length a h - n..<Heap.length a h]) h h' r"
+    and xs_def: "xs = get_array a h ! (Heap.length a h - Suc n) # r"
+    by (auto elim!: crelE crel_nth crel_return)
+  from Suc(3) have "Heap.length a h - n = Suc (Heap.length a h - Suc n)" 
+    by arith
+  with Suc.hyps[OF crel_mapM] xs_def show ?case
+    unfolding Heap.length_def
+    by (auto simp add: nth_drop')
+qed
+
+lemma crel_freeze:
+  assumes "crel (Array.freeze a) h h' xs"
+  obtains "h = h'" "xs = get_array a h"
+proof 
+  from assms have "crel (mapM (Array.nth a) [0..<Heap.length a h]) h h' xs"
+    unfolding freeze_def
+    by (auto elim: crelE crel_length)
+  hence "crel (mapM (Array.nth a) [(Heap.length a h - Heap.length a h)..<Heap.length a h]) h h' xs"
+    by simp
+  from crel_mapM_nth[OF this] show "h = h'" and "xs = get_array a h" by auto
+qed
+
+lemma crel_mapM_map_entry_remains:
+  assumes "crel (mapM (\<lambda>n. map_entry n f a) [Heap.length a h - n..<Heap.length a h]) h h' r"
+  assumes "i < Heap.length a h - n"
+  shows "get_array a h ! i = get_array a h' ! i"
+using assms
+proof (induct n arbitrary: h h' r)
+  case 0
+  thus ?case
+    by (auto elim: crel_return)
+next
+  case (Suc n)
+  let ?h1 = "Heap.upd a (Heap.length a h - Suc n) (f (get_array a h ! (Heap.length a h - Suc n))) h"
+  from Suc(3) have "[Heap.length a h - Suc n..<Heap.length a h] = (Heap.length a h - Suc n)#[Heap.length a h - n..<Heap.length a h]"
+    by (simp add: upt_conv_Cons')
+  from Suc(2) this obtain r where
+    crel_mapM: "crel (mapM (\<lambda>n. map_entry n f a) [Heap.length a h - n..<Heap.length a h]) ?h1 h' r"
+    by (auto simp add: elim!: crelE crel_map_entry crel_return)
+  have length_remains: "Heap.length a ?h1 = Heap.length a h" by simp
+  from crel_mapM have crel_mapM': "crel (mapM (\<lambda>n. map_entry n f a) [Heap.length a ?h1 - n..<Heap.length a ?h1]) ?h1 h' r"
+    by simp
+  from Suc(1)[OF this] length_remains Suc(3) show ?case by simp
+qed
+
+lemma crel_mapM_map_entry_changes:
+  assumes "crel (mapM (\<lambda>n. map_entry n f a) [Heap.length a h - n..<Heap.length a h]) h h' r"
+  assumes "n \<le> Heap.length a h"  
+  assumes "i \<ge> Heap.length a h - n"
+  assumes "i < Heap.length a h"
+  shows "get_array a h' ! i = f (get_array a h ! i)"
+using assms
+proof (induct n arbitrary: h h' r)
+  case 0
+  thus ?case
+    by (auto elim!: crel_return)
+next
+  case (Suc n)
+  let ?h1 = "Heap.upd a (Heap.length a h - Suc n) (f (get_array a h ! (Heap.length a h - Suc n))) h"
+  from Suc(3) have "[Heap.length a h - Suc n..<Heap.length a h] = (Heap.length a h - Suc n)#[Heap.length a h - n..<Heap.length a h]"
+    by (simp add: upt_conv_Cons')
+  from Suc(2) this obtain r where
+    crel_mapM: "crel (mapM (\<lambda>n. map_entry n f a) [Heap.length a h - n..<Heap.length a h]) ?h1 h' r"
+    by (auto simp add: elim!: crelE crel_map_entry crel_return)
+  have length_remains: "Heap.length a ?h1 = Heap.length a h" by simp
+  from Suc(3) have less: "Heap.length a h - Suc n < Heap.length a h - n" by arith
+  from Suc(3) have less2: "Heap.length a h - Suc n < Heap.length a h" by arith
+  from Suc(4) length_remains have cases: "i = Heap.length a ?h1 - Suc n \<or> i \<ge> Heap.length a ?h1 - n" by arith
+  from crel_mapM have crel_mapM': "crel (mapM (\<lambda>n. map_entry n f a) [Heap.length a ?h1 - n..<Heap.length a ?h1]) ?h1 h' r"
+    by simp
+  from Suc(1)[OF this] cases Suc(3) Suc(5) length_remains
+    crel_mapM_map_entry_remains[OF this, of "Heap.length a h - Suc n", symmetric] less less2
+  show ?case
+    by (auto simp add: nth_list_update_eq Heap.length_def)
+qed
+
+lemma crel_mapM_map_entry_length:
+  assumes "crel (mapM (\<lambda>n. map_entry n f a) [Heap.length a h - n..<Heap.length a h]) h h' r"
+  assumes "n \<le> Heap.length a h"
+  shows "Heap.length a h' = Heap.length a h"
+using assms
+proof (induct n arbitrary: h h' r)
+  case 0
+  thus ?case by (auto elim!: crel_return)
+next
+  case (Suc n)
+  let ?h1 = "Heap.upd a (Heap.length a h - Suc n) (f (get_array a h ! (Heap.length a h - Suc n))) h"
+  from Suc(3) have "[Heap.length a h - Suc n..<Heap.length a h] = (Heap.length a h - Suc n)#[Heap.length a h - n..<Heap.length a h]"
+    by (simp add: upt_conv_Cons')
+  from Suc(2) this obtain r where 
+    crel_mapM: "crel (mapM (\<lambda>n. map_entry n f a) [Heap.length a h - n..<Heap.length a h]) ?h1 h' r"
+    by (auto elim!: crelE crel_map_entry crel_return)
+  have length_remains: "Heap.length a ?h1 = Heap.length a h" by simp
+  from crel_mapM have crel_mapM': "crel (mapM (\<lambda>n. map_entry n f a) [Heap.length a ?h1 - n..<Heap.length a ?h1]) ?h1 h' r"
+    by simp
+  from Suc(1)[OF this] length_remains Suc(3) show ?case by simp
+qed
+
+lemma crel_mapM_map_entry:
+assumes "crel (mapM (\<lambda>n. map_entry n f a) [0..<Heap.length a h]) h h' r"
+  shows "get_array a h' = List.map f (get_array a h)"
+proof -
+  from assms have "crel (mapM (\<lambda>n. map_entry n f a) [Heap.length a h - Heap.length a h..<Heap.length a h]) h h' r" by simp
+  from crel_mapM_map_entry_length[OF this]
+  crel_mapM_map_entry_changes[OF this] show ?thesis
+    unfolding Heap.length_def
+    by (auto intro: nth_equalityI)
+qed
+
+lemma crel_map_weak:
+  assumes crel_map: "crel (Array.map f a) h h' r"
+  obtains "r = a" "get_array a h' = List.map f (get_array a h)"
+proof
+  from assms crel_mapM_map_entry show  "get_array a h' = List.map f (get_array a h)"
+    unfolding Array.map_def
+    by (fastsimp elim!: crelE crel_length crel_return)
+  from assms show "r = a"
+    unfolding Array.map_def
+    by (elim crelE crel_return)
+qed
+
+subsection {* Elimination rules for reference commands *}
+
+(* TODO:
+maybe introduce a new predicate "extends h' h x"
+which means h' extends h with a new reference x.
+Then crel_new: would be
+assumes "crel (Ref.new v) h h' x"
+obtains "get_ref x h' = v"
+and "extends h' h x"
+
+and we would need further rules for extends:
+extends h' h x \<Longrightarrow> \<not> ref_present x h
+extends h' h x \<Longrightarrow>  ref_present x h'
+extends h' h x \<Longrightarrow> ref_present y h \<Longrightarrow> ref_present y h'
+extends h' h x \<Longrightarrow> ref_present y h \<Longrightarrow> get_ref y h = get_ref y h'
+extends h' h x \<Longrightarrow> lim h' = Suc (lim h)
+*)
+
+lemma crel_Ref_new:
+  assumes "crel (Ref.new v) h h' x"
+  obtains "get_ref x h' = v"
+  and "\<not> ref_present x h"
+  and "ref_present x h'"
+  and "\<forall>y. ref_present y h \<longrightarrow> get_ref y h = get_ref y h'"
+ (* and "lim h' = Suc (lim h)" *)
+  and "\<forall>y. ref_present y h \<longrightarrow> ref_present y h'"
+  using assms
+  unfolding Ref.new_def
+  apply (elim crel_heap)
+  unfolding Heap.ref_def
+  apply (simp add: Let_def)
+  unfolding Heap.new_ref_def
+  apply (simp add: Let_def)
+  unfolding ref_present_def
+  apply auto
+  unfolding get_ref_def set_ref_def
+  apply auto
+  done
+
+lemma crel_lookup:
+  assumes "crel (!ref) h h' r"
+  obtains "h = h'" "r = get_ref ref h"
+using assms
+unfolding Ref.lookup_def
+by (auto elim: crel_heap)
+
+lemma crel_update:
+  assumes "crel (ref := v) h h' r"
+  obtains "h' = set_ref ref v h" "r = ()"
+using assms
+unfolding Ref.update_def
+by (auto elim: crel_heap)
+
+lemma crel_change:
+  assumes "crel (Ref.change f ref) h h' r"
+  obtains "h' = set_ref ref (f (get_ref ref h)) h" "r = f (get_ref ref h)"
+using assms
+unfolding Ref.change_def Let_def
+by (auto elim!: crelE crel_lookup crel_update crel_return)
+
+subsection {* Elimination rules for the assert command *}
+
+lemma crel_assert[consumes 1]:
+  assumes "crel (assert P x) h h' r"
+  obtains "P x" "r = x" "h = h'"
+  using assms
+  unfolding assert_def
+  by (elim crel_if crel_return crel_raise) auto
+
+lemma crel_assert_eq: "(\<And>h h' r. crel f h h' r \<Longrightarrow> P r) \<Longrightarrow> f \<guillemotright>= assert P = f"
+unfolding crel_def bindM_def Let_def assert_def
+  raise_def return_def prod_case_beta
+apply (cases f)
+apply simp
+apply (simp add: expand_fun_eq split_def)
+apply auto
+apply (case_tac "fst (fun x)")
+apply (simp_all add: Pair_fst_snd_eq)
+apply (erule_tac x="x" in meta_allE)
+apply fastsimp
+done
+
+section {* Introduction rules *}
+
+subsection {* Introduction rules for basic monadic commands *}
+
+lemma crelI:
+  assumes "crel f h h' r" "crel (g r) h' h'' r'"
+  shows "crel (f >>= g) h h'' r'"
+  using assms by (simp add: crel_def' bindM_def)
+
+lemma crelI':
+  assumes "crel f h h' r" "crel g h' h'' r'"
+  shows "crel (f >> g) h h'' r'"
+  using assms by (intro crelI) auto
+
+lemma crel_returnI:
+  shows "crel (return x) h h x"
+  unfolding crel_def return_def by simp
+
+lemma crel_raiseI:
+  shows "\<not> (crel (raise x) h h' r)"
+  unfolding crel_def raise_def by simp
+
+lemma crel_ifI:
+  assumes "c \<longrightarrow> crel t h h' r"
+  "\<not>c \<longrightarrow> crel e h h' r"
+  shows "crel (if c then t else e) h h' r"
+  using assms
+  unfolding crel_def by auto
+
+lemma crel_option_caseI:
+  assumes "\<And>y. x = Some y \<Longrightarrow> crel (s y) h h' r"
+  assumes "x = None \<Longrightarrow> crel n h h' r"
+  shows "crel (case x of None \<Rightarrow> n | Some y \<Rightarrow> s y) h h' r"
+using assms
+by (auto split: option.split)
+
+lemma crel_heapI:
+  shows "crel (Heap_Monad.heap f) h (snd (f h)) (fst (f h))"
+  by (simp add: crel_def apfst_def split_def prod_fun_def)
+
+lemma crel_heapI':
+  assumes "h' = snd (f h)" "r = fst (f h)"
+  shows "crel (Heap_Monad.heap f) h h' r"
+  using assms
+  by (simp add: crel_def split_def apfst_def prod_fun_def)
+
+lemma crelI2:
+  assumes "\<exists>h' rs'. crel f h h' rs' \<and> (\<exists>h'' rs. crel (g rs') h' h'' rs)"
+  shows "\<exists>h'' rs. crel (f\<guillemotright>= g) h h'' rs"
+  oops
+
+lemma crel_ifI2:
+  assumes "c \<Longrightarrow> \<exists>h' r. crel t h h' r"
+  "\<not> c \<Longrightarrow> \<exists>h' r. crel e h h' r"
+  shows "\<exists> h' r. crel (if c then t else e) h h' r"
+  oops
+
+subsection {* Introduction rules for array commands *}
+
+lemma crel_lengthI:
+  shows "crel (length a) h h (Heap.length a h)"
+  unfolding length_def
+  by (rule crel_heapI') auto
+
+(* thm crel_newI for Array.new is missing *)
+
+lemma crel_nthI:
+  assumes "i < Heap.length a h"
+  shows "crel (nth a i) h h ((get_array a h) ! i)"
+  using assms
+  unfolding nth_def
+  by (auto intro!: crelI crel_ifI crel_raiseI crel_lengthI crel_heapI')
+
+lemma crel_updI:
+  assumes "i < Heap.length a h"
+  shows "crel (upd i v a) h (Heap.upd a i v h) a"
+  using assms
+  unfolding upd_def
+  by (auto intro!: crelI crel_ifI crel_returnI crel_raiseI
+    crel_lengthI crel_heapI')
+
+(* thm crel_of_listI is missing *)
+
+(* thm crel_map_entryI is missing *)
+
+(* thm crel_swapI is missing *)
+
+(* thm crel_makeI is missing *)
+
+(* thm crel_freezeI is missing *)
+
+(* thm crel_mapI is missing *)
+
+subsection {* Introduction rules for reference commands *}
+
+lemma crel_lookupI:
+  shows "crel (!ref) h h (get_ref ref h)"
+  unfolding lookup_def by (auto intro!: crel_heapI')
+
+lemma crel_updateI:
+  shows "crel (ref := v) h (set_ref ref v h) ()"
+  unfolding update_def by (auto intro!: crel_heapI')
+
+lemma crel_changeI: 
+  shows "crel (Ref.change f ref) h (set_ref ref (f (get_ref ref h)) h) (f (get_ref ref h))"
+unfolding change_def Let_def by (auto intro!: crelI crel_returnI crel_lookupI crel_updateI)
+
+subsection {* Introduction rules for the assert command *}
+
+lemma crel_assertI:
+  assumes "P x"
+  shows "crel (assert P x) h h x"
+  using assms
+  unfolding assert_def
+  by (auto intro!: crel_ifI crel_returnI crel_raiseI)
+ 
+section {* Defintion of the noError predicate *}
+
+text {* We add a simple definitional setting for crel intro rules
+  where we only would like to show that the computation does not result in a exception for heap h,
+  but we do not care about statements about the resulting heap and return value.*}
+
+definition noError :: "'a Heap \<Rightarrow> heap \<Rightarrow> bool"
+where
+  "noError c h \<longleftrightarrow> (\<exists>r h'. (Inl r, h') = Heap_Monad.execute c h)"
+
+lemma noError_def': -- FIXME
+  "noError c h \<longleftrightarrow> (\<exists>r h'. Heap_Monad.execute c h = (Inl r, h'))"
+  unfolding noError_def apply auto proof -
+  fix r h'
+  assume "(Inl r, h') = Heap_Monad.execute c h"
+  then have "Heap_Monad.execute c h = (Inl r, h')" ..
+  then show "\<exists>r h'. Heap_Monad.execute c h = (Inl r, h')" by blast
+qed
+
+subsection {* Introduction rules for basic monadic commands *}
+
+lemma noErrorI:
+  assumes "noError f h"
+  assumes "\<And>h' r. crel f h h' r \<Longrightarrow> noError (g r) h'"
+  shows "noError (f \<guillemotright>= g) h"
+  using assms
+  by (auto simp add: noError_def' crel_def' bindM_def)
+
+lemma noErrorI':
+  assumes "noError f h"
+  assumes "\<And>h' r. crel f h h' r \<Longrightarrow> noError g h'"
+  shows "noError (f \<guillemotright> g) h"
+  using assms
+  by (auto simp add: noError_def' crel_def' bindM_def)
+
+lemma noErrorI2:
+"\<lbrakk>crel f h h' r ; noError f h; noError (g r) h'\<rbrakk>
+\<Longrightarrow> noError (f \<guillemotright>= g) h"
+by (auto simp add: noError_def' crel_def' bindM_def)
+
+lemma noError_return: 
+  shows "noError (return x) h"
+  unfolding noError_def return_def
+  by auto
+
+lemma noError_if:
+  assumes "c \<Longrightarrow> noError t h" "\<not> c \<Longrightarrow> noError e h"
+  shows "noError (if c then t else e) h"
+  using assms
+  unfolding noError_def
+  by auto
+
+lemma noError_option_case:
+  assumes "\<And>y. x = Some y \<Longrightarrow> noError (s y) h"
+  assumes "noError n h"
+  shows "noError (case x of None \<Rightarrow> n | Some y \<Rightarrow> s y) h"
+using assms
+by (auto split: option.split)
+
+lemma noError_mapM: 
+assumes "\<forall>x \<in> set xs. noError (f x) h \<and> crel (f x) h h (r x)" 
+shows "noError (mapM f xs) h"
+using assms
+proof (induct xs)
+  case Nil
+  thus ?case
+    unfolding mapM.simps by (intro noError_return)
+next
+  case (Cons x xs)
+  thus ?case
+    unfolding mapM.simps
+    by (auto intro: noErrorI2[of "f x"] noErrorI noError_return)
+qed
+
+lemma noError_heap:
+  shows "noError (Heap_Monad.heap f) h"
+  by (simp add: noError_def' apfst_def prod_fun_def split_def)
+
+subsection {* Introduction rules for array commands *}
+
+lemma noError_length:
+  shows "noError (Array.length a) h"
+  unfolding length_def
+  by (intro noError_heap)
+
+lemma noError_new:
+  shows "noError (Array.new n v) h"
+unfolding Array.new_def by (intro noError_heap)
+
+lemma noError_upd:
+  assumes "i < Heap.length a h"
+  shows "noError (Array.upd i v a) h"
+  using assms
+  unfolding upd_def
+  by (auto intro!: noErrorI noError_if noError_return noError_length noError_heap) (auto elim: crel_length)
+
+lemma noError_nth:
+assumes "i < Heap.length a h"
+  shows "noError (Array.nth a i) h"
+  using assms
+  unfolding nth_def
+  by (auto intro!: noErrorI noError_if noError_return noError_length noError_heap) (auto elim: crel_length)
+
+lemma noError_of_list:
+  shows "noError (of_list ls) h"
+  unfolding of_list_def by (rule noError_heap)
+
+lemma noError_map_entry:
+  assumes "i < Heap.length a h"
+  shows "noError (map_entry i f a) h"
+  using assms
+  unfolding map_entry_def
+  by (auto elim: crel_nth intro!: noErrorI noError_nth noError_upd)
+
+lemma noError_swap:
+  assumes "i < Heap.length a h"
+  shows "noError (swap i x a) h"
+  using assms
+  unfolding swap_def
+  by (auto elim: crel_nth intro!: noErrorI noError_return noError_nth noError_upd)
+
+lemma noError_make:
+  shows "noError (make n f) h"
+  unfolding make_def
+  by (auto intro: noError_of_list)
+
+(*TODO: move to HeapMonad *)
+lemma mapM_append:
+  "mapM f (xs @ ys) = mapM f xs \<guillemotright>= (\<lambda>xs. mapM f ys \<guillemotright>= (\<lambda>ys. return (xs @ ys)))"
+  by (induct xs) (simp_all add: monad_simp)
+
+lemma noError_freeze:
+  shows "noError (freeze a) h"
+unfolding freeze_def
+by (auto intro!: noErrorI noError_length noError_mapM[of _ _ _ "\<lambda>x. get_array a h ! x"]
+  noError_nth crel_nthI elim: crel_length)
+
+lemma noError_mapM_map_entry:
+  assumes "n \<le> Heap.length a h"
+  shows "noError (mapM (\<lambda>n. map_entry n f a) [Heap.length a h - n..<Heap.length a h]) h"
+using assms
+proof (induct n arbitrary: h)
+  case 0
+  thus ?case by (auto intro: noError_return)
+next
+  case (Suc n)
+  from Suc.prems have "[Heap.length a h - Suc n..<Heap.length a h] = (Heap.length a h - Suc n)#[Heap.length a h - n..<Heap.length a h]"
+    by (simp add: upt_conv_Cons')
+  with Suc.hyps[of "(Heap.upd a (Heap.length a h - Suc n) (f (get_array a h ! (Heap.length a h - Suc n))) h)"] Suc.prems show ?case
+    by (auto simp add: intro!: noErrorI noError_return noError_map_entry elim: crel_map_entry)
+qed
+
+lemma noError_map:
+  shows "noError (Array.map f a) h"
+using noError_mapM_map_entry[of "Heap.length a h" a h]
+unfolding Array.map_def
+by (auto intro: noErrorI noError_length noError_return elim!: crel_length)
+
+subsection {* Introduction rules for the reference commands *}
+
+lemma noError_Ref_new:
+  shows "noError (Ref.new v) h"
+unfolding Ref.new_def by (intro noError_heap)
+
+lemma noError_lookup:
+  shows "noError (!ref) h"
+  unfolding lookup_def by (intro noError_heap)
+
+lemma noError_update:
+  shows "noError (ref := v) h"
+  unfolding update_def by (intro noError_heap)
+
+lemma noError_change:
+  shows "noError (Ref.change f ref) h"
+  unfolding Ref.change_def Let_def by (intro noErrorI noError_lookup noError_update noError_return)
+
+subsection {* Introduction rules for the assert command *}
+
+lemma noError_assert:
+  assumes "P x"
+  shows "noError (assert P x) h"
+  using assms
+  unfolding assert_def
+  by (auto intro: noError_if noError_return)
+
+section {* Cumulative lemmas *}
+
+lemmas crel_elim_all =
+  crelE crelE' crel_return crel_raise crel_if crel_option_case
+  crel_length crel_new_weak crel_nth crel_upd crel_of_list_weak crel_map_entry crel_swap crel_make_weak crel_freeze crel_map_weak
+  crel_Ref_new crel_lookup crel_update crel_change
+  crel_assert
+
+lemmas crel_intro_all =
+  crelI crelI' crel_returnI crel_raiseI crel_ifI crel_option_caseI
+  crel_lengthI (* crel_newI *) crel_nthI crel_updI (* crel_of_listI crel_map_entryI crel_swapI crel_makeI crel_freezeI crel_mapI *)
+  (* crel_Ref_newI *) crel_lookupI crel_updateI crel_changeI
+  crel_assert
+
+lemmas noError_intro_all = 
+  noErrorI noErrorI' noError_return noError_if noError_option_case
+  noError_length noError_new noError_nth noError_upd noError_of_list noError_map_entry noError_swap noError_make noError_freeze noError_map
+  noError_Ref_new noError_lookup noError_update noError_change
+  noError_assert
+
+end
\ No newline at end of file
--- a/src/HOL/Import/proof_kernel.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Import/proof_kernel.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,6 +1,5 @@
 (*  Title:      HOL/Import/proof_kernel.ML
-    ID:         $Id$
-    Author:     Sebastian Skalberg (TU Muenchen), Steven Obua
+    Author:     Sebastian Skalberg and Steven Obua, TU Muenchen
 *)
 
 signature ProofKernel =
@@ -281,14 +280,12 @@
           | itr (a::rst) = i=a orelse itr rst
     in itr L end;
 
-fun insert i L = if i mem L then L else i::L
-
 fun mk_set [] = []
-  | mk_set (a::rst) = insert a (mk_set rst)
+  | mk_set (a::rst) = insert (op =) a (mk_set rst)
 
 fun [] union S = S
   | S union [] = S
-  | (a::rst) union S2 = rst union (insert a S2)
+  | (a::rst) union S2 = rst union (insert (op =) a S2)
 
 fun implode_subst [] = []
   | implode_subst (x::r::rest) = ((x,r)::(implode_subst rest))
@@ -1149,9 +1146,9 @@
       val ct = cprop_of thm
       val t = term_of ct
       val thy = theory_of_cterm ct
-      val frees = term_frees t
-      val freenames = add_term_free_names (t, [])
-      fun is_old_name n = n mem_string freenames
+      val frees = OldTerm.term_frees t
+      val freenames = Term.add_free_names t []
+      val is_old_name = member (op =) freenames
       fun name_of (Free (n, _)) = n
         | name_of _ = ERR "name_of"
       fun new_name' bump map n =
@@ -1214,24 +1211,23 @@
               | NONE => NONE
         end
 
-fun non_trivial_term_consts tm =
-    List.filter (fn c => not (c = "Trueprop" orelse
-                         c = "All" orelse
-                         c = "op -->" orelse
-                         c = "op &" orelse
-                         c = "op =")) (Term.term_consts tm)
+fun non_trivial_term_consts t = fold_aterms
+  (fn Const (c, _) =>
+      if c = "Trueprop" orelse c = "All" orelse c = "op -->" orelse c = "op &" orelse c = "op ="
+      then I else insert (op =) c
+    | _ => I) t [];
 
 fun match_consts t (* th *) =
     let
         fun add_consts (Const (c, _), cs) =
             (case c of
-                 "op =" => Library.insert (op =) "==" cs
-               | "op -->" => Library.insert (op =) "==>" cs
+                 "op =" => insert (op =) "==" cs
+               | "op -->" => insert (op =) "==>" cs
                | "All" => cs
                | "all" => cs
                | "op &" => cs
                | "Trueprop" => cs
-               | _ => Library.insert (op =) c cs)
+               | _ => insert (op =) c cs)
           | add_consts (t $ u, cs) = add_consts (t, add_consts (u, cs))
           | add_consts (Abs (_, _, t), cs) = add_consts (t, cs)
           | add_consts (_, cs) = cs
@@ -1298,7 +1294,7 @@
             let
                 val _ = (message "Looking for conclusion:";
                          if_debug prin isaconc)
-                val cs = non_trivial_term_consts isaconc
+                val cs = non_trivial_term_consts isaconc;
                 val _ = (message "Looking for consts:";
                          message (commas cs))
                 val pot_thms = Shuffler.find_potential thy isaconc
@@ -1424,9 +1420,9 @@
     let
         val _ = message "INST_TYPE:"
         val _ = if_debug pth hth
-        val tys_before = add_term_tfrees (prop_of th,[])
+        val tys_before = OldTerm.add_term_tfrees (prop_of th,[])
         val th1 = Thm.varifyT th
-        val tys_after = add_term_tvars (prop_of th1,[])
+        val tys_after = OldTerm.add_term_tvars (prop_of th1,[])
         val tyinst = map (fn (bef, iS) =>
                              (case try (Lib.assoc (TFree bef)) lambda of
                                   SOME ty => (ctyp_of thy (TVar iS), ctyp_of thy ty)
@@ -2093,7 +2089,7 @@
             val c = case concl_of th2 of
                         _ $ (Const("Ex",_) $ Abs(_,_,Const("op :",_) $ _ $ c)) => c
                       | _ => raise ERR "new_type_definition" "Bad type definition theorem"
-            val tfrees = term_tfrees c
+            val tfrees = OldTerm.term_tfrees c
             val tnames = map fst tfrees
             val tsyn = mk_syn thy tycname
             val typ = (tycname,tnames,tsyn)
@@ -2179,7 +2175,7 @@
             val c = case concl_of th2 of
                         _ $ (Const("Ex",_) $ Abs(_,_,Const("op :",_) $ _ $ c)) => c
                       | _ => raise ERR "type_introduction" "Bad type definition theorem"
-            val tfrees = term_tfrees c
+            val tfrees = OldTerm.term_tfrees c
             val tnames = sort string_ord (map fst tfrees)
             val tsyn = mk_syn thy tycname
             val typ = (tycname,tnames,tsyn)
--- a/src/HOL/Import/shuffler.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Import/shuffler.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Import/shuffler.ML
-    ID:         $Id$
     Author:     Sebastian Skalberg, TU Muenchen
 
 Package for proving two terms equal by normalizing (hence the
@@ -248,8 +247,8 @@
 
 fun freeze_thaw_term t =
     let
-        val tvars = term_tvars t
-        val tfree_names = add_term_tfree_names(t,[])
+        val tvars = OldTerm.term_tvars t
+        val tfree_names = OldTerm.add_term_tfree_names(t,[])
         val (type_inst,_) =
             Library.foldl (fn ((inst,used),(w as (v,_),S)) =>
                       let
@@ -268,7 +267,7 @@
   | inst_tfrees thy ((name,U)::rest) thm =
     let
         val cU = ctyp_of thy U
-        val tfrees = add_term_tfrees (prop_of thm,[])
+        val tfrees = OldTerm.add_term_tfrees (prop_of thm,[])
         val (rens, thm') = Thm.varifyT'
     (remove (op = o apsnd fst) name tfrees) thm
         val mid =
@@ -322,7 +321,7 @@
               then
                   let
                       val cert = cterm_of thy
-                      val v = Free(Name.variant (add_term_free_names(t,[])) "v",xT)
+                      val v = Free (Name.variant (Term.add_free_names t []) "v", xT)
                       val cv = cert v
                       val ct = cert t
                       val th = (assume ct)
@@ -354,7 +353,7 @@
 
 fun equals_fun thy assume t =
     case t of
-        Const("op ==",_) $ u $ v => if Term.term_ord (u,v) = LESS then SOME (meta_sym_rew) else NONE
+        Const("op ==",_) $ u $ v => if TermOrd.term_ord (u,v) = LESS then SOME (meta_sym_rew) else NONE
       | _ => NONE
 
 fun eta_expand thy assumes origt =
@@ -385,7 +384,7 @@
                       Type("fun",[aT,bT]) =>
                       let
                           val cert = cterm_of thy
-                          val vname = Name.variant (add_term_free_names(t,[])) "v"
+                          val vname = Name.variant (Term.add_free_names t []) "v"
                           val v = Free(vname,aT)
                           val cv = cert v
                           val ct = cert t
@@ -520,7 +519,7 @@
     let
         val thy = Thm.theory_of_thm th
         val c = prop_of th
-        val vars = add_term_frees (c,add_term_vars(c,[]))
+        val vars = OldTerm.add_term_frees (c, OldTerm.add_term_vars(c,[]))
     in
         Drule.forall_intr_list (map (cterm_of thy) vars) th
     end
@@ -573,8 +572,8 @@
     fold_rev (fn thm => fn cs =>
               let
                   val (lhs,rhs) = Logic.dest_equals (prop_of thm)
-                  val ignore_lhs = term_consts lhs \\ term_consts rhs
-                  val ignore_rhs = term_consts rhs \\ term_consts lhs
+                  val ignore_lhs = Term.add_const_names lhs [] \\ Term.add_const_names rhs []
+                  val ignore_rhs = Term.add_const_names rhs [] \\ Term.add_const_names lhs []
               in
                   fold_rev (insert (op =)) cs (ignore_lhs @ ignore_rhs)
               end)
@@ -585,7 +584,7 @@
 
 fun set_prop thy t =
     let
-        val vars = add_term_frees (t,add_term_vars (t,[]))
+        val vars = OldTerm.add_term_frees (t, OldTerm.add_term_vars (t,[]))
         val closed_t = fold_rev Logic.all vars t
         val rew_th = norm_term thy closed_t
         val rhs = Thm.rhs_of rew_th
--- a/src/HOL/Inductive.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Inductive.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Inductive.thy
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 *)
 
@@ -363,7 +362,7 @@
 let
   fun fun_tr ctxt [cs] =
     let
-      val x = Free (Name.variant (add_term_free_names (cs, [])) "x", dummyT);
+      val x = Free (Name.variant (Term.add_free_names cs []) "x", dummyT);
       val ft = DatatypeCase.case_tr true DatatypePackage.datatype_of_constr
                  ctxt [x, cs]
     in lambda x ft end
--- a/src/HOL/IntDiv.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/IntDiv.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -747,12 +747,12 @@
 lemma zdiv_zmult_self1 [simp]: "b \<noteq> (0::int) ==> (a*b) div b = a"
 by (simp add: zdiv_zmult1_eq)
 
-lemma mod_div_trivial [simp]: "(a mod b) div b = (0::int)"
+lemma zmod_zdiv_trivial: "(a mod b) div b = (0::int)"
 apply (case_tac "b = 0", simp)
 apply (auto simp add: linorder_neq_iff div_pos_pos_trivial div_neg_neg_trivial)
 done
 
-lemma mod_mod_trivial [simp]: "(a mod b) mod b = a mod (b::int)"
+lemma zmod_zmod_trivial: "(a mod b) mod b = a mod (b::int)"
 apply (case_tac "b = 0", simp)
 apply (force simp add: linorder_neq_iff mod_pos_pos_trivial mod_neg_neg_trivial)
 done
@@ -776,72 +776,50 @@
 apply (blast intro: zadd1_lemma [OF quorem_div_mod quorem_div_mod] quorem_mod)
 done
 
-lemma zdiv_zadd_self1[simp]: "a \<noteq> (0::int) ==> (a+b) div a = b div a + 1"
-by (simp add: zdiv_zadd1_eq)
-
-lemma zdiv_zadd_self2[simp]: "a \<noteq> (0::int) ==> (b+a) div a = b div a + 1"
-by (simp add: zdiv_zadd1_eq)
-
-instance int :: semiring_div
+instance int :: ring_div
 proof
   fix a b c :: int
   assume not0: "b \<noteq> 0"
   show "(a + c * b) div b = c + a div b"
     unfolding zdiv_zadd1_eq [of a "c * b"] using not0 
-      by (simp add: zmod_zmult1_eq)
+      by (simp add: zmod_zmult1_eq zmod_zdiv_trivial)
 qed auto
 
-lemma zdiv_zmult_self2 [simp]: "b \<noteq> (0::int) ==> (b*a) div b = a"
-by (subst mult_commute, erule zdiv_zmult_self1)
+lemma zdiv_zadd_self1: "a \<noteq> (0::int) ==> (a+b) div a = b div a + 1"
+by (rule div_add_self1) (* already declared [simp] *)
+
+lemma zdiv_zadd_self2: "a \<noteq> (0::int) ==> (b+a) div a = b div a + 1"
+by (rule div_add_self2) (* already declared [simp] *)
 
-lemma zmod_zmult_self1 [simp]: "(a*b) mod b = (0::int)"
-by (simp add: zmod_zmult1_eq)
+lemma zdiv_zmult_self2: "b \<noteq> (0::int) ==> (b*a) div b = a"
+by (rule div_mult_self1_is_id) (* already declared [simp] *)
 
-lemma zmod_zmult_self2 [simp]: "(b*a) mod b = (0::int)"
-by (simp add: mult_commute zmod_zmult1_eq)
+lemma zmod_zmult_self1: "(a*b) mod b = (0::int)"
+by (rule mod_mult_self2_is_0) (* already declared [simp] *)
+
+lemma zmod_zmult_self2: "(b*a) mod b = (0::int)"
+by (rule mod_mult_self1_is_0) (* already declared [simp] *)
 
 lemma zmod_eq_0_iff: "(m mod d = 0) = (EX q::int. m = d*q)"
-proof
-  assume "m mod d = 0"
-  with zmod_zdiv_equality[of m d] show "EX q::int. m = d*q" by auto
-next
-  assume "EX q::int. m = d*q"
-  thus "m mod d = 0" by auto
-qed
+by (simp add: dvd_eq_mod_eq_0 [symmetric] dvd_def)
 
+(* REVISIT: should this be generalized to all semiring_div types? *)
 lemmas zmod_eq_0D [dest!] = zmod_eq_0_iff [THEN iffD1]
 
 lemma zmod_zadd_left_eq: "(a+b) mod (c::int) = ((a mod c) + b) mod c"
-apply (rule trans [symmetric])
-apply (rule zmod_zadd1_eq, simp)
-apply (rule zmod_zadd1_eq [symmetric])
-done
+by (rule mod_add_left_eq)
 
 lemma zmod_zadd_right_eq: "(a+b) mod (c::int) = (a + (b mod c)) mod c"
-apply (rule trans [symmetric])
-apply (rule zmod_zadd1_eq, simp)
-apply (rule zmod_zadd1_eq [symmetric])
-done
+by (rule mod_add_right_eq)
 
-lemma zmod_zadd_self1[simp]: "(a+b) mod a = b mod (a::int)"
-apply (case_tac "a = 0", simp)
-apply (simp add: zmod_zadd1_eq)
-done
+lemma zmod_zadd_self1: "(a+b) mod a = b mod (a::int)"
+by (rule mod_add_self1) (* already declared [simp] *)
 
-lemma zmod_zadd_self2[simp]: "(b+a) mod a = b mod (a::int)"
-apply (case_tac "a = 0", simp)
-apply (simp add: zmod_zadd1_eq)
-done
-
+lemma zmod_zadd_self2: "(b+a) mod a = b mod (a::int)"
+by (rule mod_add_self2) (* already declared [simp] *)
 
-lemma zmod_zdiff1_eq: fixes a::int
-  shows "(a - b) mod c = (a mod c - b mod c) mod c" (is "?l = ?r")
-proof -
-  have "?l = (c + (a mod c - b mod c)) mod c"
-    using zmod_zadd1_eq[of a "-b" c] by(simp add:ring_simps zmod_zminus1_eq_if)
-  also have "\<dots> = ?r" by simp
-  finally show ?thesis .
-qed
+lemma zmod_zdiff1_eq: "(a - b) mod c = (a mod c - b mod c) mod (c::int)"
+by (rule mod_diff_eq)
 
 subsection{*Proving  @{term "a div (b*c) = (a div b) div c"} *}
 
@@ -954,18 +932,8 @@
 apply (auto simp add: mult_commute)
 done
 
-lemma zmod_zmod_cancel:
-assumes "n dvd m" shows "(k::int) mod m mod n = k mod n"
-proof -
-  from `n dvd m` obtain r where "m = n*r" by(auto simp:dvd_def)
-  have "k mod n = (m * (k div m) + k mod m) mod n"
-    using zmod_zdiv_equality[of k m] by simp
-  also have "\<dots> = (m * (k div m) mod n + k mod m mod n) mod n"
-    by(subst zmod_zadd1_eq, rule refl)
-  also have "m * (k div m) mod n = 0" using `m = n*r`
-    by(simp add:mult_ac)
-  finally show ?thesis by simp
-qed
+lemma zmod_zmod_cancel: "n dvd m \<Longrightarrow> (k::int) mod m mod n = k mod n"
+by (rule mod_mod_cancel)
 
 
 subsection {*Splitting Rules for div and mod*}
@@ -1169,50 +1137,31 @@
 subsection {* The Divides Relation *}
 
 lemma zdvd_iff_zmod_eq_0: "(m dvd n) = (n mod m = (0::int))"
-  by (simp add: dvd_def zmod_eq_0_iff)
+  by (rule dvd_eq_mod_eq_0)
 
 lemmas zdvd_iff_zmod_eq_0_number_of [simp] =
   zdvd_iff_zmod_eq_0 [of "number_of x" "number_of y", standard]
 
-lemma zdvd_0_right [iff]: "(m::int) dvd 0"
-  by (simp add: dvd_def)
+lemma zdvd_0_right: "(m::int) dvd 0"
+  by (rule dvd_0_right) (* already declared [iff] *)
 
-lemma zdvd_0_left [iff,noatp]: "(0 dvd (m::int)) = (m = 0)"
-  by (simp add: dvd_def)
+lemma zdvd_0_left: "(0 dvd (m::int)) = (m = 0)"
+  by (rule dvd_0_left_iff) (* already declared [noatp,simp] *)
 
-lemma zdvd_1_left [iff]: "1 dvd (m::int)"
-  by (simp add: dvd_def)
+lemma zdvd_1_left: "1 dvd (m::int)"
+  by (rule one_dvd) (* already declared [simp] *)
 
 lemma zdvd_refl [simp]: "m dvd (m::int)"
-  by (auto simp add: dvd_def intro: zmult_1_right [symmetric])
+  by (rule dvd_refl) (* TODO: declare generic dvd_refl [simp] *)
 
 lemma zdvd_trans: "m dvd n ==> n dvd k ==> m dvd (k::int)"
-  by (auto simp add: dvd_def intro: mult_assoc)
+  by (rule dvd_trans)
 
 lemma zdvd_zminus_iff: "m dvd -n \<longleftrightarrow> m dvd (n::int)"
-proof
-  assume "m dvd - n"
-  then obtain k where "- n = m * k" ..
-  then have "n = m * - k" by simp
-  then show "m dvd n" ..
-next
-  assume "m dvd n"
-  then have "m dvd n * -1" by (rule dvd_mult2)
-  then show "m dvd - n" by simp
-qed
+  by (rule dvd_minus_iff)
 
 lemma zdvd_zminus2_iff: "-m dvd n \<longleftrightarrow> m dvd (n::int)"
-proof
-  assume "- m dvd n"
-  then obtain k where "n = - m * k" ..
-  then have "n = m * - k" by simp
-  then show "m dvd n" ..
-next
-  assume "m dvd n"
-  then obtain k where "n = m * k" ..
-  then have "n = - m * - k" by simp
-  then show "- m dvd n" ..
-qed
+  by (rule minus_dvd_iff)
 
 lemma zdvd_abs1: "( \<bar>i::int\<bar> dvd j) = (i dvd j)" 
   by (cases "i > 0") (simp_all add: zdvd_zminus2_iff)
@@ -1227,9 +1176,7 @@
   done
 
 lemma zdvd_zadd: "k dvd m ==> k dvd n ==> k dvd (m + n :: int)"
-  apply (simp add: dvd_def)
-  apply (blast intro: right_distrib [symmetric])
-  done
+  by (rule dvd_add)
 
 lemma zdvd_dvd_eq: assumes anz:"a \<noteq> 0" and ab: "(a::int) dvd b" and ba:"b dvd a" 
   shows "\<bar>a\<bar> = \<bar>b\<bar>"
@@ -1244,9 +1191,7 @@
 qed
 
 lemma zdvd_zdiff: "k dvd m ==> k dvd n ==> k dvd (m - n :: int)"
-  apply (simp add: dvd_def)
-  apply (blast intro: right_diff_distrib [symmetric])
-  done
+  by (rule Ring_and_Field.dvd_diff)
 
 lemma zdvd_zdiffD: "k dvd m - n ==> k dvd n ==> k dvd (m::int)"
   apply (subgoal_tac "m = n + (m - n)")
@@ -1255,34 +1200,22 @@
   done
 
 lemma zdvd_zmult: "k dvd (n::int) ==> k dvd m * n"
-  apply (simp add: dvd_def)
-  apply (blast intro: mult_left_commute)
-  done
+  by (rule dvd_mult)
 
 lemma zdvd_zmult2: "k dvd (m::int) ==> k dvd m * n"
-  apply (subst mult_commute)
-  apply (erule zdvd_zmult)
-  done
+  by (rule dvd_mult2)
 
-lemma zdvd_triv_right [iff]: "(k::int) dvd m * k"
-  apply (rule zdvd_zmult)
-  apply (rule zdvd_refl)
-  done
+lemma zdvd_triv_right: "(k::int) dvd m * k"
+  by (rule dvd_triv_right) (* already declared [simp] *)
 
-lemma zdvd_triv_left [iff]: "(k::int) dvd k * m"
-  apply (rule zdvd_zmult2)
-  apply (rule zdvd_refl)
-  done
+lemma zdvd_triv_left: "(k::int) dvd k * m"
+  by (rule dvd_triv_left) (* already declared [simp] *)
 
 lemma zdvd_zmultD2: "j * k dvd n ==> j dvd (n::int)"
-  apply (simp add: dvd_def)
-  apply (simp add: mult_assoc, blast)
-  done
+  by (rule dvd_mult_left)
 
 lemma zdvd_zmultD: "j * k dvd n ==> k dvd (n::int)"
-  apply (rule zdvd_zmultD2)
-  apply (subst mult_commute, assumption)
-  done
+  by (rule dvd_mult_right)
 
 lemma zdvd_zmult_mono: "i dvd m ==> j dvd (n::int) ==> i * j dvd m * n"
   by (rule mult_dvd_mono)
@@ -1333,7 +1266,7 @@
   {assume "n \<noteq> m*h" hence "k* n \<noteq> k* (m*h)" using kz by simp
     with h have False by (simp add: mult_assoc)}
   hence "n = m * h" by blast
-  thus ?thesis by blast
+  thus ?thesis by simp
 qed
 
 lemma zdvd_zmult_cancel_disj[simp]:
@@ -1371,8 +1304,8 @@
       then show ?thesis by (simp only: negative_eq_positive) auto
     qed
   qed
-  then show ?thesis by (auto elim!: dvdE simp only: zmult_int [symmetric])
-qed 
+  then show ?thesis by (auto elim!: dvdE simp only: zdvd_triv_left int_mult)
+qed
 
 lemma zdvd1_eq[simp]: "(x::int) dvd 1 = ( \<bar>x\<bar> = 1)"
 proof
@@ -1404,10 +1337,10 @@
   by (auto simp add: dvd_int_iff)
 
 lemma zminus_dvd_iff [iff]: "(-z dvd w) = (z dvd (w::int))"
-  by (simp add: zdvd_zminus2_iff)
+  by (rule minus_dvd_iff)
 
 lemma dvd_zminus_iff [iff]: "(z dvd -w) = (z dvd (w::int))"
-  by (simp add: zdvd_zminus_iff)
+  by (rule dvd_minus_iff)
 
 lemma zdvd_imp_le: "[| z dvd n; 0 < n |] ==> z \<le> (n::int)"
   apply (rule_tac z=n in int_cases)
@@ -1449,20 +1382,13 @@
 
 text {* by Brian Huffman *}
 lemma zminus_zmod: "- ((x::int) mod m) mod m = - x mod m"
-by (simp only: zmod_zminus1_eq_if mod_mod_trivial)
+by (rule mod_minus_eq [symmetric])
 
 lemma zdiff_zmod_left: "(x mod m - y) mod m = (x - y) mod (m::int)"
-by (simp only: diff_def zmod_zadd_left_eq [symmetric])
+by (rule mod_diff_left_eq [symmetric])
 
 lemma zdiff_zmod_right: "(x - y mod m) mod m = (x - y) mod (m::int)"
-proof -
-  have "(x + - (y mod m) mod m) mod m = (x + - y mod m) mod m"
-    by (simp only: zminus_zmod)
-  hence "(x + - (y mod m)) mod m = (x + - y) mod m"
-    by (simp only: zmod_zadd_right_eq [symmetric])
-  thus "(x - y mod m) mod m = (x - y) mod m"
-    by (simp only: diff_def)
-qed
+by (rule mod_diff_right_eq [symmetric])
 
 lemmas zmod_simps =
   IntDiv.zmod_zadd_left_eq  [symmetric]
--- a/src/HOL/Integration.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Integration.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -6,7 +6,7 @@
 header{*Theory of Integration*}
 
 theory Integration
-imports MacLaurin
+imports Deriv ATP_Linkup
 begin
 
 text{*We follow John Harrison in formalizing the Gauge integral.*}
@@ -55,14 +55,37 @@
                                          \<bar>rsum(D,p) f - k\<bar> < e)))"
 
 
+lemma psize_unique:
+  assumes 1: "\<forall>n < N. D(n) < D(Suc n)"
+  assumes 2: "\<forall>n \<ge> N. D(n) = D(N)"
+  shows "psize D = N"
+unfolding psize_def
+proof (rule some_equality)
+  show "(\<forall>n<N. D(n) < D(Suc n)) \<and> (\<forall>n\<ge>N. D(n) = D(N))" using prems ..
+next
+  fix M assume "(\<forall>n<M. D(n) < D(Suc n)) \<and> (\<forall>n\<ge>M. D(n) = D(M))"
+  hence 3: "\<forall>n<M. D(n) < D(Suc n)" and 4: "\<forall>n\<ge>M. D(n) = D(M)" by fast+
+  show "M = N"
+  proof (rule linorder_cases)
+    assume "M < N"
+    hence "D(M) < D(Suc M)" by (rule 1 [rule_format])
+    also have "D(Suc M) = D(M)" by (rule 4 [rule_format], simp)
+    finally show "M = N" by simp
+  next
+    assume "N < M"
+    hence "D(N) < D(Suc N)" by (rule 3 [rule_format])
+    also have "D(Suc N) = D(N)" by (rule 2 [rule_format], simp)
+    finally show "M = N" by simp
+  next
+    assume "M = N" thus "M = N" .
+  qed
+qed
+
 lemma partition_zero [simp]: "a = b ==> psize (%n. if n = 0 then a else b) = 0"
-by (auto simp add: psize_def)
+by (rule psize_unique, simp_all)
 
 lemma partition_one [simp]: "a < b ==> psize (%n. if n = 0 then a else b) = 1"
-apply (simp add: psize_def)
-apply (rule some_equality, auto)
-apply (drule_tac x = 1 in spec, auto)
-done
+by (rule psize_unique, simp_all)
 
 lemma partition_single [simp]:
      "a \<le> b ==> partition(a,b)(%n. if n = 0 then a else b)"
@@ -76,20 +99,11 @@
         ((D 0 = a) &
          (\<forall>n < psize D. D n < D(Suc n)) &
          (\<forall>n \<ge> psize D. D n = b))"
-apply (simp add: partition_def, auto)
-apply (subgoal_tac [!] "psize D = N", auto)
-apply (simp_all (no_asm) add: psize_def)
-apply (rule_tac [!] some_equality, blast)
-  prefer 2 apply blast
-apply (rule_tac [!] ccontr)
-apply (simp_all add: linorder_neq_iff, safe)
-apply (drule_tac x = Na in spec)
-apply (rotate_tac 3)
-apply (drule_tac x = "Suc Na" in spec, simp)
-apply (rotate_tac 2)
-apply (drule_tac x = N in spec, simp)
-apply (drule_tac x = Na in spec)
-apply (drule_tac x = "Suc Na" and P = "%n. Na\<le>n \<longrightarrow> D n = D Na" in spec, auto)
+apply (simp add: partition_def)
+apply (rule iffI, clarify)
+apply (subgoal_tac "psize D = N", simp)
+apply (rule psize_unique, assumption, simp)
+apply (simp, rule_tac x="psize D" in exI, simp)
 done
 
 lemma partition_rhs: "partition(a,b) D ==> (D(psize D) = b)"
@@ -211,26 +225,10 @@
 lemma lemma_partition_append2:
      "[| partition (a, b) D1; partition (b, c) D2 |]
       ==> psize (%n. if n < psize D1 then D1 n else D2 (n - psize D1)) =
-          psize D1 + psize D2" 
-apply (unfold psize_def 
-         [of "%n. if n < psize D1 then D1 n else D2 (n - psize D1)"])
-apply (rule some1_equality)
- prefer 2 apply (blast intro!: lemma_partition_append1)
-apply (rule ex1I, rule lemma_partition_append1) 
-apply (simp_all split: split_if_asm)
- txt{*The case @{term "N < psize D1"}*} 
- apply (drule_tac x = "psize D1 + psize D2" and P="%n. ?P n & ?Q n" in spec) 
- apply (force dest: lemma_psize1)
-apply (rule order_antisym);
- txt{*The case @{term "psize D1 \<le> N"}: 
-       proving @{term "N \<le> psize D1 + psize D2"}*}
- apply (drule_tac x = "psize D1 + psize D2" in spec)
- apply (simp add: partition_rhs2)
-apply (case_tac "N - psize D1 < psize D2") 
- prefer 2 apply arith
- txt{*Proving @{term "psize D1 + psize D2 \<le> N"}*}
-apply (drule_tac x = "psize D1 + psize D2" and P="%n. N\<le>n --> ?P n" in spec, simp)
-apply (drule_tac a = b and b = c in partition_gt, assumption, simp)
+          psize D1 + psize D2"
+apply (rule psize_unique)
+apply (erule (1) lemma_partition_append1 [THEN conjunct1])
+apply (erule (1) lemma_partition_append1 [THEN conjunct2])
 done
 
 lemma tpart_eq_lhs_rhs: "[|psize D = 0; tpart(a,b) (D,p)|] ==> a = b"
@@ -566,20 +564,12 @@
 lemma lemma_additivity4_psize_eq:
      "[| a \<le> D n; D n < b; partition (a, b) D |]
       ==> psize (%x. if D x < D n then D(x) else D n) = n"
-apply (unfold psize_def)
-apply (frule lemma_additivity1)
-apply (assumption, assumption)
-apply (rule some_equality)
-apply (auto intro: partition_lt_Suc)
-apply (drule_tac n = n in partition_lt_gen, assumption)
-apply (arith, arith)
-apply (cut_tac x = na and y = "psize D" in less_linear)
-apply (auto dest: partition_lt_cancel)
-apply (rule_tac x=N and y=n in linorder_cases)
-apply (drule_tac x = n and P="%m. N \<le> m --> ?f m = ?g m" in spec, simp)
-apply (drule_tac n = n in partition_lt_gen, auto)
-apply (drule_tac x = n in spec)
-apply (simp split: split_if_asm)
+apply (frule (2) lemma_additivity1)
+apply (rule psize_unique, auto)
+apply (erule partition_lt_Suc, erule (1) less_trans)
+apply (erule notE)
+apply (erule (1) partition_lt_gen, erule less_imp_le)
+apply (drule (1) partition_lt_cancel, simp)
 done
 
 lemma lemma_psize_left_less_psize:
--- a/src/HOL/IsaMakefile	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/IsaMakefile	Mon Jan 19 21:20:18 2009 +0100
@@ -6,7 +6,7 @@
 
 default: HOL
 generate: HOL-Generate-HOL HOL-Generate-HOLLight
-images: HOL-Plain HOL-Main HOL HOL-Algebra HOL-Nominal HOL-NSA HOL-Word TLA HOL4
+images: HOL HOL-Base HOL-Plain HOL-Main HOL-Algebra HOL-Nominal HOL-NSA HOL-Word TLA HOL4
 
 #Note: keep targets sorted (except for HOL-Library and HOL-ex)
 test: \
@@ -23,6 +23,7 @@
   HOL-IMP \
   HOL-IMPP \
   HOL-IOA \
+  HOL-Imperative_HOL \
   HOL-Induct \
   HOL-Isar_examples \
   HOL-Lambda \
@@ -65,6 +66,8 @@
 
 HOL: Pure $(OUT)/HOL
 
+HOL-Base: Pure $(OUT)/HOL-Base
+
 HOL-Plain: Pure $(OUT)/HOL-Plain
 
 HOL-Main: Pure $(OUT)/HOL-Main
@@ -74,15 +77,50 @@
 
 $(OUT)/Pure: Pure
 
-PLAIN_DEPENDENCIES = $(OUT)/Pure \
+BASE_DEPENDENCIES = $(OUT)/Pure \
   Code_Setup.thy \
+  HOL.thy \
+  Tools/hologic.ML \
+  Tools/recfun_codegen.ML \
+  Tools/simpdata.ML \
+  $(SRC)/Tools/atomize_elim.ML \
+  $(SRC)/Tools/code/code_funcgr.ML \
+  $(SRC)/Tools/code/code_funcgr.ML \
+  $(SRC)/Tools/code/code_name.ML \
+  $(SRC)/Tools/code/code_printer.ML \
+  $(SRC)/Tools/code/code_target.ML \
+  $(SRC)/Tools/code/code_ml.ML \
+  $(SRC)/Tools/code/code_haskell.ML \
+  $(SRC)/Tools/code/code_thingol.ML \
+  $(SRC)/Tools/induct.ML \
+  $(SRC)/Tools/induct_tacs.ML \
+  $(SRC)/Tools/IsaPlanner/isand.ML \
+  $(SRC)/Tools/IsaPlanner/rw_inst.ML \
+  $(SRC)/Tools/IsaPlanner/rw_tools.ML \
+  $(SRC)/Tools/IsaPlanner/zipper.ML \
+  $(SRC)/Tools/nbe.ML \
+  $(SRC)/Tools/random_word.ML \
+  $(SRC)/Tools/value.ML \
+  $(SRC)/Provers/blast.ML \
+  $(SRC)/Provers/clasimp.ML \
+  $(SRC)/Provers/classical.ML \
+  $(SRC)/Provers/coherent.ML \
+  $(SRC)/Provers/eqsubst.ML \
+  $(SRC)/Provers/hypsubst.ML \
+  $(SRC)/Provers/project_rule.ML \
+  $(SRC)/Provers/quantifier1.ML \
+  $(SRC)/Provers/splitter.ML \
+
+$(OUT)/HOL-Base: base.ML $(BASE_DEPENDENCIES)
+	@$(ISABELLE_TOOL) usedir -b -f base.ML -d false -g false $(OUT)/Pure HOL-Base
+
+PLAIN_DEPENDENCIES = $(BASE_DEPENDENCIES)\
   Datatype.thy \
   Divides.thy \
   Extraction.thy \
   Finite_Set.thy \
   Fun.thy \
   FunDef.thy \
-  HOL.thy \
   Inductive.thy \
   Lattices.thy \
   Nat.thy \
@@ -130,7 +168,6 @@
   Tools/function_package/size.ML \
   Tools/function_package/sum_tree.ML \
   Tools/function_package/termination.ML \
-  Tools/hologic.ML \
   Tools/inductive_codegen.ML \
   Tools/inductive_package.ML \
   Tools/inductive_realizer.ML \
@@ -139,14 +176,12 @@
   Tools/old_primrec_package.ML \
   Tools/primrec_package.ML \
   Tools/prop_logic.ML \
-  Tools/recfun_codegen.ML \
   Tools/record_package.ML \
   Tools/refute.ML \
   Tools/refute_isar.ML \
   Tools/rewrite_hol_proof.ML \
   Tools/sat_funcs.ML \
   Tools/sat_solver.ML \
-  Tools/simpdata.ML \
   Tools/split_rule.ML \
   Tools/typecopy_package.ML \
   Tools/typedef_codegen.ML \
@@ -158,35 +193,8 @@
   $(SRC)/Provers/Arith/cancel_div_mod.ML \
   $(SRC)/Provers/Arith/cancel_sums.ML \
   $(SRC)/Provers/Arith/fast_lin_arith.ML \
-  $(SRC)/Provers/blast.ML \
-  $(SRC)/Provers/clasimp.ML \
-  $(SRC)/Provers/classical.ML \
-  $(SRC)/Provers/coherent.ML \
-  $(SRC)/Provers/eqsubst.ML \
-  $(SRC)/Provers/hypsubst.ML \
   $(SRC)/Provers/order.ML \
-  $(SRC)/Provers/project_rule.ML \
-  $(SRC)/Provers/quantifier1.ML \
-  $(SRC)/Provers/splitter.ML \
   $(SRC)/Provers/trancl.ML \
-  $(SRC)/Tools/IsaPlanner/isand.ML \
-  $(SRC)/Tools/IsaPlanner/rw_inst.ML \
-  $(SRC)/Tools/IsaPlanner/rw_tools.ML \
-  $(SRC)/Tools/IsaPlanner/zipper.ML \
-  $(SRC)/Tools/atomize_elim.ML \
-  $(SRC)/Tools/code/code_funcgr.ML \
-  $(SRC)/Tools/code/code_funcgr.ML \
-  $(SRC)/Tools/code/code_name.ML \
-  $(SRC)/Tools/code/code_printer.ML \
-  $(SRC)/Tools/code/code_target.ML \
-  $(SRC)/Tools/code/code_ml.ML \
-  $(SRC)/Tools/code/code_haskell.ML \
-  $(SRC)/Tools/code/code_thingol.ML \
-  $(SRC)/Tools/induct.ML \
-  $(SRC)/Tools/induct_tacs.ML \
-  $(SRC)/Tools/value.ML \
-  $(SRC)/Tools/nbe.ML \
-  $(SRC)/Tools/random_word.ML \
   $(SRC)/Tools/rat.ML
 
 $(OUT)/HOL-Plain: plain.ML $(PLAIN_DEPENDENCIES)
@@ -279,8 +287,8 @@
   GCD.thy \
   Order_Relation.thy \
   Parity.thy \
-  Univ_Poly.thy \
   Lubs.thy \
+  Polynomial.thy \
   PReal.thy \
   Rational.thy \
   RComplete.thy \
@@ -325,9 +333,7 @@
   Library/Code_Char_chr.thy Library/Code_Integer.thy			\
   Library/Numeral_Type.thy			\
   Library/Boolean_Algebra.thy Library/Countable.thy	\
-  Library/Heap.thy Library/Heap_Monad.thy Library/Array.thy		\
-  Library/Relational.thy Library/Sublist.thy Library/Subarray.thy	\
-  Library/Ref.thy Library/Imperative_HOL.thy Library/RBT.thy		\
+  Library/RBT.thy	Library/Univ_Poly.thy	\
   Library/Enum.thy Library/Float.thy $(SRC)/Tools/float.ML $(SRC)/HOL/Tools/float_arith.ML
 	@cd Library; $(ISABELLE_TOOL) usedir $(OUT)/HOL Library
 
@@ -625,6 +631,17 @@
 	@$(ISABELLE_TOOL) usedir $(OUT)/HOL Modelcheck
 
 
+## HOL-Imperative_HOL
+
+HOL-Imperative_HOL: HOL $(LOG)/HOL-Imperative_HOL.gz
+
+$(LOG)/HOL-Imperative_HOL.gz: $(OUT)/HOL Imperative_HOL/Heap.thy \
+  Imperative_HOL/Heap_Monad.thy Imperative_HOL/Array.thy \
+  Imperative_HOL/Relational.thy \
+  Imperative_HOL/Ref.thy Imperative_HOL/Imperative_HOL.thy
+	@$(ISABELLE_TOOL) usedir $(OUT)/HOL Imperative_HOL
+
+
 ## HOL-SizeChange
 
 HOL-SizeChange: HOL $(LOG)/HOL-SizeChange.gz
@@ -796,8 +813,9 @@
   ex/ReflectionEx.thy ex/ROOT.ML ex/Recdefs.thy ex/Records.thy		\
   ex/Reflected_Presburger.thy ex/coopertac.ML				\
   ex/Refute_Examples.thy ex/SAT_Examples.thy ex/SVC_Oracle.thy		\
-  ex/Sudoku.thy ex/Tarski.thy ex/Termination.thy ex/Term_Of_Syntax.thy			\
-  ex/Unification.thy ex/document/root.bib			\
+  ex/Subarray.thy ex/Sublist.thy                                        \
+  ex/Sudoku.thy ex/Tarski.thy ex/Termination.thy ex/Term_Of_Syntax.thy	\
+  ex/Unification.thy ex/document/root.bib			        \
   ex/document/root.tex ex/Meson_Test.thy ex/reflection.ML ex/set.thy	\
   ex/svc_funcs.ML ex/svc_test.thy	\
   ex/ImperativeQuicksort.thy	\
--- a/src/HOL/Lattices.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Lattices.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -300,8 +300,7 @@
   by auto
 qed (auto simp add: min_def max_def not_le less_imp_le)
 
-class_interpretation min_max:
-  distrib_lattice ["op \<le> \<Colon> 'a\<Colon>linorder \<Rightarrow> 'a \<Rightarrow> bool" "op <" min max]
+interpretation min_max!: distrib_lattice "op \<le> :: 'a::linorder \<Rightarrow> 'a \<Rightarrow> bool" "op <" min max
   by (rule distrib_lattice_min_max)
 
 lemma inf_min: "inf = (min \<Colon> 'a\<Colon>{lower_semilattice, linorder} \<Rightarrow> 'a \<Rightarrow> 'a)"
--- a/src/HOL/Library/Array.thy	Mon Jan 19 20:37:08 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,209 +0,0 @@
-(*  Title:      HOL/Library/Array.thy
-    ID:         $Id$
-    Author:     John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen
-*)
-
-header {* Monadic arrays *}
-
-theory Array
-imports Heap_Monad Code_Index
-begin
-
-subsection {* Primitives *}
-
-definition
-  new :: "nat \<Rightarrow> 'a\<Colon>heap \<Rightarrow> 'a array Heap" where
-  [code del]: "new n x = Heap_Monad.heap (Heap.array n x)"
-
-definition
-  of_list :: "'a\<Colon>heap list \<Rightarrow> 'a array Heap" where
-  [code del]: "of_list xs = Heap_Monad.heap (Heap.array_of_list xs)"
-
-definition
-  length :: "'a\<Colon>heap array \<Rightarrow> nat Heap" where
-  [code del]: "length arr = Heap_Monad.heap (\<lambda>h. (Heap.length arr h, h))"
-
-definition
-  nth :: "'a\<Colon>heap array \<Rightarrow> nat \<Rightarrow> 'a Heap"
-where
-  [code del]: "nth a i = (do len \<leftarrow> length a;
-                 (if i < len
-                     then Heap_Monad.heap (\<lambda>h. (get_array a h ! i, h))
-                     else raise (''array lookup: index out of range''))
-              done)"
-
--- {* FIXME adjustion for List theory *}
-no_syntax
-  nth  :: "'a list \<Rightarrow> nat \<Rightarrow> 'a" (infixl "!" 100)
-
-abbreviation
-  nth_list :: "'a list \<Rightarrow> nat \<Rightarrow> 'a" (infixl "!" 100)
-where
-  "nth_list \<equiv> List.nth"
-
-definition
-  upd :: "nat \<Rightarrow> 'a \<Rightarrow> 'a\<Colon>heap array \<Rightarrow> 'a\<Colon>heap array Heap"
-where
-  [code del]: "upd i x a = (do len \<leftarrow> length a;
-                      (if i < len
-                           then Heap_Monad.heap (\<lambda>h. (a, Heap.upd a i x h))
-                           else raise (''array update: index out of range''))
-                   done)" 
-
-lemma upd_return:
-  "upd i x a \<guillemotright> return a = upd i x a"
-proof (rule Heap_eqI)
-  fix h
-  obtain len h' where "Heap_Monad.execute (Array.length a) h = (len, h')"
-    by (cases "Heap_Monad.execute (Array.length a) h")
-  then show "Heap_Monad.execute (upd i x a \<guillemotright> return a) h = Heap_Monad.execute (upd i x a) h"
-    by (auto simp add: upd_def bindM_def split: sum.split)
-qed
-
-
-subsection {* Derivates *}
-
-definition
-  map_entry :: "nat \<Rightarrow> ('a\<Colon>heap \<Rightarrow> 'a) \<Rightarrow> 'a array \<Rightarrow> 'a array Heap"
-where
-  "map_entry i f a = (do
-     x \<leftarrow> nth a i;
-     upd i (f x) a
-   done)"
-
-definition
-  swap :: "nat \<Rightarrow> 'a \<Rightarrow> 'a\<Colon>heap array \<Rightarrow> 'a Heap"
-where
-  "swap i x a = (do
-     y \<leftarrow> nth a i;
-     upd i x a;
-     return y
-   done)"
-
-definition
-  make :: "nat \<Rightarrow> (nat \<Rightarrow> 'a\<Colon>heap) \<Rightarrow> 'a array Heap"
-where
-  "make n f = of_list (map f [0 ..< n])"
-
-definition
-  freeze :: "'a\<Colon>heap array \<Rightarrow> 'a list Heap"
-where
-  "freeze a = (do
-     n \<leftarrow> length a;
-     mapM (nth a) [0..<n]
-   done)"
-
-definition
-   map :: "('a\<Colon>heap \<Rightarrow> 'a) \<Rightarrow> 'a array \<Rightarrow> 'a array Heap"
-where
-  "map f a = (do
-     n \<leftarrow> length a;
-     mapM (\<lambda>n. map_entry n f a) [0..<n];
-     return a
-   done)"
-
-hide (open) const new map -- {* avoid clashed with some popular names *}
-
-
-subsection {* Properties *}
-
-lemma array_make [code]:
-  "Array.new n x = make n (\<lambda>_. x)"
-  by (induct n) (simp_all add: make_def new_def Heap_Monad.heap_def
-    monad_simp array_of_list_replicate [symmetric]
-    map_replicate_trivial replicate_append_same
-    of_list_def)
-
-lemma array_of_list_make [code]:
-  "of_list xs = make (List.length xs) (\<lambda>n. xs ! n)"
-  unfolding make_def map_nth ..
-
-
-subsection {* Code generator setup *}
-
-subsubsection {* Logical intermediate layer *}
-
-definition new' where
-  [code del]: "new' = Array.new o nat_of_index"
-hide (open) const new'
-lemma [code]:
-  "Array.new = Array.new' o index_of_nat"
-  by (simp add: new'_def o_def)
-
-definition of_list' where
-  [code del]: "of_list' i xs = Array.of_list (take (nat_of_index i) xs)"
-hide (open) const of_list'
-lemma [code]:
-  "Array.of_list xs = Array.of_list' (index_of_nat (List.length xs)) xs"
-  by (simp add: of_list'_def)
-
-definition make' where
-  [code del]: "make' i f = Array.make (nat_of_index i) (f o index_of_nat)"
-hide (open) const make'
-lemma [code]:
-  "Array.make n f = Array.make' (index_of_nat n) (f o nat_of_index)"
-  by (simp add: make'_def o_def)
-
-definition length' where
-  [code del]: "length' = Array.length \<guillemotright>== liftM index_of_nat"
-hide (open) const length'
-lemma [code]:
-  "Array.length = Array.length' \<guillemotright>== liftM nat_of_index"
-  by (simp add: length'_def monad_simp',
-    simp add: liftM_def comp_def monad_simp,
-    simp add: monad_simp')
-
-definition nth' where
-  [code del]: "nth' a = Array.nth a o nat_of_index"
-hide (open) const nth'
-lemma [code]:
-  "Array.nth a n = Array.nth' a (index_of_nat n)"
-  by (simp add: nth'_def)
-
-definition upd' where
-  [code del]: "upd' a i x = Array.upd (nat_of_index i) x a \<guillemotright> return ()"
-hide (open) const upd'
-lemma [code]:
-  "Array.upd i x a = Array.upd' a (index_of_nat i) x \<guillemotright> return a"
-  by (simp add: upd'_def monad_simp upd_return)
-
-
-subsubsection {* SML *}
-
-code_type array (SML "_/ array")
-code_const Array (SML "raise/ (Fail/ \"bare Array\")")
-code_const Array.new' (SML "(fn/ ()/ =>/ Array.array/ ((_),/ (_)))")
-code_const Array.of_list (SML "(fn/ ()/ =>/ Array.fromList/ _)")
-code_const Array.make' (SML "(fn/ ()/ =>/ Array.tabulate/ ((_),/ (_)))")
-code_const Array.length' (SML "(fn/ ()/ =>/ Array.length/ _)")
-code_const Array.nth' (SML "(fn/ ()/ =>/ Array.sub/ ((_),/ (_)))")
-code_const Array.upd' (SML "(fn/ ()/ =>/ Array.update/ ((_),/ (_),/ (_)))")
-
-code_reserved SML Array
-
-
-subsubsection {* OCaml *}
-
-code_type array (OCaml "_/ array")
-code_const Array (OCaml "failwith/ \"bare Array\"")
-code_const Array.new' (OCaml "(fun/ ()/ ->/ Array.make/ _/ _)")
-code_const Array.of_list (OCaml "(fun/ ()/ ->/ Array.of'_list/ _)")
-code_const Array.make' (OCaml "(fun/ ()/ ->/ Array.init/ _/ _)")
-code_const Array.length' (OCaml "(fun/ ()/ ->/ Array.length/ _)")
-code_const Array.nth' (OCaml "(fun/ ()/ ->/ Array.get/ _/ _)")
-code_const Array.upd' (OCaml "(fun/ ()/ ->/ Array.set/ _/ _/ _)")
-
-code_reserved OCaml Array
-
-
-subsubsection {* Haskell *}
-
-code_type array (Haskell "STArray/ RealWorld/ _")
-code_const Array (Haskell "error/ \"bare Array\"")
-code_const Array.new' (Haskell "newArray/ (0,/ _)")
-code_const Array.of_list' (Haskell "newListArray/ (0,/ _)")
-code_const Array.length' (Haskell "lengthArray")
-code_const Array.nth' (Haskell "readArray")
-code_const Array.upd' (Haskell "writeArray")
-
-end
--- a/src/HOL/Library/Countable.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Library/Countable.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Library/Countable.thy
-    ID:         $Id$
     Author:     Alexander Krauss, TU Muenchen
 *)
 
--- a/src/HOL/Library/Efficient_Nat.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Library/Efficient_Nat.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Library/Efficient_Nat.thy
-    ID:         $Id$
     Author:     Stefan Berghofer, Florian Haftmann, TU Muenchen
 *)
 
@@ -127,7 +126,7 @@
 fun remove_suc thy thms =
   let
     val vname = Name.variant (map fst
-      (fold (Term.add_varnames o Thm.full_prop_of) thms [])) "x";
+      (fold (Term.add_var_names o Thm.full_prop_of) thms [])) "x";
     val cv = cterm_of thy (Var ((vname, 0), HOLogic.natT));
     fun lhs_of th = snd (Thm.dest_comb
       (fst (Thm.dest_comb (snd (Thm.dest_comb (cprop_of th))))));
@@ -170,7 +169,7 @@
 fun eqn_suc_preproc thy ths =
   let
     val dest = fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of;
-    fun contains_suc t = member (op =) (term_consts t) @{const_name Suc};
+    val contains_suc = exists_Const (fn (c, _) => c = @{const_name Suc});
   in
     if forall (can dest) ths andalso
       exists (contains_suc o dest) ths
@@ -180,7 +179,7 @@
 fun remove_suc_clause thy thms =
   let
     val vname = Name.variant (map fst
-      (fold (Term.add_varnames o Thm.full_prop_of) thms [])) "x";
+      (fold (Term.add_var_names o Thm.full_prop_of) thms [])) "x";
     fun find_var (t as Const (@{const_name Suc}, _) $ (v as Var _)) = SOME (t, v)
       | find_var (t $ u) = (case find_var t of NONE => find_var u | x => x)
       | find_var _ = NONE;
@@ -212,8 +211,8 @@
     val dest = fst o HOLogic.dest_mem o HOLogic.dest_Trueprop
   in
     if forall (can (dest o concl_of)) ths andalso
-      exists (fn th => member (op =) (foldr add_term_consts
-        [] (map_filter (try dest) (concl_of th :: prems_of th))) "Suc") ths
+      exists (fn th => exists (exists_Const (fn (c, _) => c = @{const_name Suc}))
+        (map_filter (try dest) (concl_of th :: prems_of th))) ths
     then remove_suc_clause thy ths else ths
   end;
 
--- a/src/HOL/Library/Heap.thy	Mon Jan 19 20:37:08 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,434 +0,0 @@
-(*  Title:      HOL/Library/Heap.thy
-    ID:         $Id$
-    Author:     John Matthews, Galois Connections; Alexander Krauss, TU Muenchen
-*)
-
-header {* A polymorphic heap based on cantor encodings *}
-
-theory Heap
-imports Plain "~~/src/HOL/List" Countable Typerep
-begin
-
-subsection {* Representable types *}
-
-text {* The type class of representable types *}
-
-class heap = typerep + countable
-
-text {* Instances for common HOL types *}
-
-instance nat :: heap ..
-
-instance "*" :: (heap, heap) heap ..
-
-instance "+" :: (heap, heap) heap ..
-
-instance list :: (heap) heap ..
-
-instance option :: (heap) heap ..
-
-instance int :: heap ..
-
-instance message_string :: countable
-  by (rule countable_classI [of "message_string_case to_nat"])
-   (auto split: message_string.splits)
-
-instance message_string :: heap ..
-
-text {* Reflected types themselves are heap-representable *}
-
-instantiation typerep :: countable
-begin
-
-fun to_nat_typerep :: "typerep \<Rightarrow> nat" where
-  "to_nat_typerep (Typerep.Typerep c ts) = to_nat (to_nat c, to_nat (map to_nat_typerep ts))"
-
-instance
-proof (rule countable_classI)
-  fix t t' :: typerep and ts
-  have "(\<forall>t'. to_nat_typerep t = to_nat_typerep t' \<longrightarrow> t = t')
-    \<and> (\<forall>ts'. map to_nat_typerep ts = map to_nat_typerep ts' \<longrightarrow> ts = ts')"
-  proof (induct rule: typerep.induct)
-    case (Typerep c ts) show ?case
-    proof (rule allI, rule impI)
-      fix t'
-      assume hyp: "to_nat_typerep (Typerep.Typerep c ts) = to_nat_typerep t'"
-      then obtain c' ts' where t': "t' = (Typerep.Typerep c' ts')"
-        by (cases t') auto
-      with Typerep hyp have "c = c'" and "ts = ts'" by simp_all
-      with t' show "Typerep.Typerep c ts = t'" by simp
-    qed
-  next
-    case Nil_typerep then show ?case by simp
-  next
-    case (Cons_typerep t ts) then show ?case by auto
-  qed
-  then have "to_nat_typerep t = to_nat_typerep t' \<Longrightarrow> t = t'" by auto
-  moreover assume "to_nat_typerep t = to_nat_typerep t'"
-  ultimately show "t = t'" by simp
-qed
-
-end
-
-instance typerep :: heap ..
-
-
-subsection {* A polymorphic heap with dynamic arrays and references *}
-
-types addr = nat -- "untyped heap references"
-
-datatype 'a array = Array addr
-datatype 'a ref = Ref addr -- "note the phantom type 'a "
-
-primrec addr_of_array :: "'a array \<Rightarrow> addr" where
-  "addr_of_array (Array x) = x"
-
-primrec addr_of_ref :: "'a ref \<Rightarrow> addr" where
-  "addr_of_ref (Ref x) = x"
-
-lemma addr_of_array_inj [simp]:
-  "addr_of_array a = addr_of_array a' \<longleftrightarrow> a = a'"
-  by (cases a, cases a') simp_all
-
-lemma addr_of_ref_inj [simp]:
-  "addr_of_ref r = addr_of_ref r' \<longleftrightarrow> r = r'"
-  by (cases r, cases r') simp_all
-
-instance array :: (type) countable
-  by (rule countable_classI [of addr_of_array]) simp
-
-instance ref :: (type) countable
-  by (rule countable_classI [of addr_of_ref]) simp
-
-setup {*
-  Sign.add_const_constraint (@{const_name Array}, SOME @{typ "nat \<Rightarrow> 'a\<Colon>heap array"})
-  #> Sign.add_const_constraint (@{const_name Ref}, SOME @{typ "nat \<Rightarrow> 'a\<Colon>heap ref"})
-  #> Sign.add_const_constraint (@{const_name addr_of_array}, SOME @{typ "'a\<Colon>heap array \<Rightarrow> nat"})
-  #> Sign.add_const_constraint (@{const_name addr_of_ref}, SOME @{typ "'a\<Colon>heap ref \<Rightarrow> nat"})
-*}
-
-types heap_rep = nat -- "representable values"
-
-record heap =
-  arrays :: "typerep \<Rightarrow> addr \<Rightarrow> heap_rep list"
-  refs :: "typerep \<Rightarrow> addr \<Rightarrow> heap_rep"
-  lim  :: addr
-
-definition empty :: heap where
-  "empty = \<lparr>arrays = (\<lambda>_. undefined), refs = (\<lambda>_. undefined), lim = 0\<rparr>" -- "why undefined?"
-
-
-subsection {* Imperative references and arrays *}
-
-text {*
-  References and arrays are developed in parallel,
-  but keeping them separate makes some later proofs simpler.
-*}
-
-subsubsection {* Primitive operations *}
-
-definition
-  new_ref :: "heap \<Rightarrow> ('a\<Colon>heap) ref \<times> heap" where
-  "new_ref h = (let l = lim h in (Ref l, h\<lparr>lim := l + 1\<rparr>))"
-
-definition
-  new_array :: "heap \<Rightarrow> ('a\<Colon>heap) array \<times> heap" where
-  "new_array h = (let l = lim h in (Array l, h\<lparr>lim := l + 1\<rparr>))"
-
-definition
-  ref_present :: "'a\<Colon>heap ref \<Rightarrow> heap \<Rightarrow> bool" where
-  "ref_present r h \<longleftrightarrow> addr_of_ref r < lim h"
-
-definition 
-  array_present :: "'a\<Colon>heap array \<Rightarrow> heap \<Rightarrow> bool" where
-  "array_present a h \<longleftrightarrow> addr_of_array a < lim h"
-
-definition
-  get_ref :: "'a\<Colon>heap ref \<Rightarrow> heap \<Rightarrow> 'a" where
-  "get_ref r h = from_nat (refs h (TYPEREP('a)) (addr_of_ref r))"
-
-definition
-  get_array :: "'a\<Colon>heap array \<Rightarrow> heap \<Rightarrow> 'a list" where
-  "get_array a h = map from_nat (arrays h (TYPEREP('a)) (addr_of_array a))"
-
-definition
-  set_ref :: "'a\<Colon>heap ref \<Rightarrow> 'a \<Rightarrow> heap \<Rightarrow> heap" where
-  "set_ref r x = 
-  refs_update (\<lambda>h. h(TYPEREP('a) := ((h (TYPEREP('a))) (addr_of_ref r:=to_nat x))))"
-
-definition
-  set_array :: "'a\<Colon>heap array \<Rightarrow> 'a list \<Rightarrow> heap \<Rightarrow> heap" where
-  "set_array a x = 
-  arrays_update (\<lambda>h. h(TYPEREP('a) := ((h(TYPEREP('a))) (addr_of_array a:=map to_nat x))))"
-
-subsubsection {* Interface operations *}
-
-definition
-  ref :: "'a \<Rightarrow> heap \<Rightarrow> 'a\<Colon>heap ref \<times> heap" where
-  "ref x h = (let (r, h') = new_ref h;
-                   h''    = set_ref r x h'
-         in (r, h''))"
-
-definition
-  array :: "nat \<Rightarrow> 'a \<Rightarrow> heap \<Rightarrow> 'a\<Colon>heap array \<times> heap" where
-  "array n x h = (let (r, h') = new_array h;
-                       h'' = set_array r (replicate n x) h'
-        in (r, h''))"
-
-definition
-  array_of_list :: "'a list \<Rightarrow> heap \<Rightarrow> 'a\<Colon>heap array \<times> heap" where
-  "array_of_list xs h = (let (r, h') = new_array h;
-           h'' = set_array r xs h'
-        in (r, h''))"  
-
-definition
-  upd :: "'a\<Colon>heap array \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> heap \<Rightarrow> heap" where
-  "upd a i x h = set_array a ((get_array a h)[i:=x]) h"
-
-definition
-  length :: "'a\<Colon>heap array \<Rightarrow> heap \<Rightarrow> nat" where
-  "length a h = size (get_array a h)"
-
-definition
-  array_ran :: "('a\<Colon>heap) option array \<Rightarrow> heap \<Rightarrow> 'a set" where
-  "array_ran a h = {e. Some e \<in> set (get_array a h)}"
-    -- {*FIXME*}
-
-
-subsubsection {* Reference equality *}
-
-text {* 
-  The following relations are useful for comparing arrays and references.
-*}
-
-definition
-  noteq_refs :: "('a\<Colon>heap) ref \<Rightarrow> ('b\<Colon>heap) ref \<Rightarrow> bool" (infix "=!=" 70)
-where
-  "r =!= s \<longleftrightarrow> TYPEREP('a) \<noteq> TYPEREP('b) \<or> addr_of_ref r \<noteq> addr_of_ref s"
-
-definition
-  noteq_arrs :: "('a\<Colon>heap) array \<Rightarrow> ('b\<Colon>heap) array \<Rightarrow> bool" (infix "=!!=" 70)
-where
-  "r =!!= s \<longleftrightarrow> TYPEREP('a) \<noteq> TYPEREP('b) \<or> addr_of_array r \<noteq> addr_of_array s"
-
-lemma noteq_refs_sym: "r =!= s \<Longrightarrow> s =!= r"
-  and noteq_arrs_sym: "a =!!= b \<Longrightarrow> b =!!= a"
-  and unequal_refs [simp]: "r \<noteq> r' \<longleftrightarrow> r =!= r'" -- "same types!"
-  and unequal_arrs [simp]: "a \<noteq> a' \<longleftrightarrow> a =!!= a'"
-unfolding noteq_refs_def noteq_arrs_def by auto
-
-lemma present_new_ref: "ref_present r h \<Longrightarrow> r =!= fst (ref v h)"
-  by (simp add: ref_present_def new_ref_def ref_def Let_def noteq_refs_def)
-
-lemma present_new_arr: "array_present a h \<Longrightarrow> a =!!= fst (array v x h)"
-  by (simp add: array_present_def noteq_arrs_def new_array_def array_def Let_def)
-
-
-subsubsection {* Properties of heap containers *}
-
-text {* Properties of imperative arrays *}
-
-text {* FIXME: Does there exist a "canonical" array axiomatisation in
-the literature?  *}
-
-lemma array_get_set_eq [simp]: "get_array r (set_array r x h) = x"
-  by (simp add: get_array_def set_array_def)
-
-lemma array_get_set_neq [simp]: "r =!!= s \<Longrightarrow> get_array r (set_array s x h) = get_array r h"
-  by (simp add: noteq_arrs_def get_array_def set_array_def)
-
-lemma set_array_same [simp]:
-  "set_array r x (set_array r y h) = set_array r x h"
-  by (simp add: set_array_def)
-
-lemma array_set_set_swap:
-  "r =!!= r' \<Longrightarrow> set_array r x (set_array r' x' h) = set_array r' x' (set_array r x h)"
-  by (simp add: Let_def expand_fun_eq noteq_arrs_def set_array_def)
-
-lemma array_ref_set_set_swap:
-  "set_array r x (set_ref r' x' h) = set_ref r' x' (set_array r x h)"
-  by (simp add: Let_def expand_fun_eq set_array_def set_ref_def)
-
-lemma get_array_upd_eq [simp]:
-  "get_array a (upd a i v h) = (get_array a h) [i := v]"
-  by (simp add: upd_def)
-
-lemma nth_upd_array_neq_array [simp]:
-  "a =!!= b \<Longrightarrow> get_array a (upd b j v h) ! i = get_array a h ! i"
-  by (simp add: upd_def noteq_arrs_def)
-
-lemma get_arry_array_upd_elem_neqIndex [simp]:
-  "i \<noteq> j \<Longrightarrow> get_array a (upd a j v h) ! i = get_array a h ! i"
-  by simp
-
-lemma length_upd_eq [simp]: 
-  "length a (upd a i v h) = length a h" 
-  by (simp add: length_def upd_def)
-
-lemma length_upd_neq [simp]: 
-  "length a (upd b i v h) = length a h"
-  by (simp add: upd_def length_def set_array_def get_array_def)
-
-lemma upd_swap_neqArray:
-  "a =!!= a' \<Longrightarrow> 
-  upd a i v (upd a' i' v' h) 
-  = upd a' i' v' (upd a i v h)"
-apply (unfold upd_def)
-apply simp
-apply (subst array_set_set_swap, assumption)
-apply (subst array_get_set_neq)
-apply (erule noteq_arrs_sym)
-apply (simp)
-done
-
-lemma upd_swap_neqIndex:
-  "\<lbrakk> i \<noteq> i' \<rbrakk> \<Longrightarrow> upd a i v (upd a i' v' h) = upd a i' v' (upd a i v h)"
-by (auto simp add: upd_def array_set_set_swap list_update_swap)
-
-lemma get_array_init_array_list:
-  "get_array (fst (array_of_list ls h)) (snd (array_of_list ls' h)) = ls'"
-  by (simp add: Let_def split_def array_of_list_def)
-
-lemma set_array:
-  "set_array (fst (array_of_list ls h))
-     new_ls (snd (array_of_list ls h))
-       = snd (array_of_list new_ls h)"
-  by (simp add: Let_def split_def array_of_list_def)
-
-lemma array_present_upd [simp]: 
-  "array_present a (upd b i v h) = array_present a h"
-  by (simp add: upd_def array_present_def set_array_def get_array_def)
-
-lemma array_of_list_replicate:
-  "array_of_list (replicate n x) = array n x"
-  by (simp add: expand_fun_eq array_of_list_def array_def)
-
-
-text {* Properties of imperative references *}
-
-lemma next_ref_fresh [simp]:
-  assumes "(r, h') = new_ref h"
-  shows "\<not> ref_present r h"
-  using assms by (cases h) (auto simp add: new_ref_def ref_present_def Let_def)
-
-lemma next_ref_present [simp]:
-  assumes "(r, h') = new_ref h"
-  shows "ref_present r h'"
-  using assms by (cases h) (auto simp add: new_ref_def ref_present_def Let_def)
-
-lemma ref_get_set_eq [simp]: "get_ref r (set_ref r x h) = x"
-  by (simp add: get_ref_def set_ref_def)
-
-lemma ref_get_set_neq [simp]: "r =!= s \<Longrightarrow> get_ref r (set_ref s x h) = get_ref r h"
-  by (simp add: noteq_refs_def get_ref_def set_ref_def)
-
-(* FIXME: We need some infrastructure to infer that locally generated
-  new refs (by new_ref(_no_init), new_array(')) are distinct
-  from all existing refs.
-*)
-
-lemma ref_set_get: "set_ref r (get_ref r h) h = h"
-apply (simp add: set_ref_def get_ref_def)
-oops
-
-lemma set_ref_same[simp]:
-  "set_ref r x (set_ref r y h) = set_ref r x h"
-  by (simp add: set_ref_def)
-
-lemma ref_set_set_swap:
-  "r =!= r' \<Longrightarrow> set_ref r x (set_ref r' x' h) = set_ref r' x' (set_ref r x h)"
-  by (simp add: Let_def expand_fun_eq noteq_refs_def set_ref_def)
-
-lemma ref_new_set: "fst (ref v (set_ref r v' h)) = fst (ref v h)"
-  by (simp add: ref_def new_ref_def set_ref_def Let_def)
-
-lemma ref_get_new [simp]:
-  "get_ref (fst (ref v h)) (snd (ref v' h)) = v'"
-  by (simp add: ref_def Let_def split_def)
-
-lemma ref_set_new [simp]:
-  "set_ref (fst (ref v h)) new_v (snd (ref v h)) = snd (ref new_v h)"
-  by (simp add: ref_def Let_def split_def)
-
-lemma ref_get_new_neq: "r =!= (fst (ref v h)) \<Longrightarrow> 
-  get_ref r (snd (ref v h)) = get_ref r h"
-  by (simp add: get_ref_def set_ref_def ref_def Let_def new_ref_def noteq_refs_def)
-
-lemma lim_set_ref [simp]:
-  "lim (set_ref r v h) = lim h"
-  by (simp add: set_ref_def)
-
-lemma ref_present_new_ref [simp]: 
-  "ref_present r h \<Longrightarrow> ref_present r (snd (ref v h))"
-  by (simp add: new_ref_def ref_present_def ref_def Let_def)
-
-lemma ref_present_set_ref [simp]:
-  "ref_present r (set_ref r' v h) = ref_present r h"
-  by (simp add: set_ref_def ref_present_def)
-
-lemma array_ranI: "\<lbrakk> Some b = get_array a h ! i; i < Heap.length a h \<rbrakk> \<Longrightarrow> b \<in> array_ran a h"
-unfolding array_ran_def Heap.length_def by simp
-
-lemma array_ran_upd_array_Some:
-  assumes "cl \<in> array_ran a (Heap.upd a i (Some b) h)"
-  shows "cl \<in> array_ran a h \<or> cl = b"
-proof -
-  have "set (get_array a h[i := Some b]) \<subseteq> insert (Some b) (set (get_array a h))" by (rule set_update_subset_insert)
-  with assms show ?thesis 
-    unfolding array_ran_def Heap.upd_def by fastsimp
-qed
-
-lemma array_ran_upd_array_None:
-  assumes "cl \<in> array_ran a (Heap.upd a i None h)"
-  shows "cl \<in> array_ran a h"
-proof -
-  have "set (get_array a h[i := None]) \<subseteq> insert None (set (get_array a h))" by (rule set_update_subset_insert)
-  with assms show ?thesis
-    unfolding array_ran_def Heap.upd_def by auto
-qed
-
-
-text {* Non-interaction between imperative array and imperative references *}
-
-lemma get_array_set_ref [simp]: "get_array a (set_ref r v h) = get_array a h"
-  by (simp add: get_array_def set_ref_def)
-
-lemma nth_set_ref [simp]: "get_array a (set_ref r v h) ! i = get_array a h ! i"
-  by simp
-
-lemma get_ref_upd [simp]: "get_ref r (upd a i v h) = get_ref r h"
-  by (simp add: get_ref_def set_array_def upd_def)
-
-lemma new_ref_upd: "fst (ref v (upd a i v' h)) = fst (ref v h)"
-  by (simp add: set_array_def get_array_def Let_def ref_new_set upd_def ref_def  new_ref_def)
-
-text {*not actually true ???*}
-lemma upd_set_ref_swap: "upd a i v (set_ref r v' h) = set_ref r v' (upd a i v h)"
-apply (case_tac a)
-apply (simp add: Let_def upd_def)
-apply auto
-oops
-
-lemma length_new_ref[simp]: 
-  "length a (snd (ref v h)) = length a h"
-  by (simp add: get_array_def set_ref_def length_def new_ref_def ref_def Let_def)
-
-lemma get_array_new_ref [simp]: 
-  "get_array a (snd (ref v h)) = get_array a h"
-  by (simp add: new_ref_def ref_def set_ref_def get_array_def Let_def)
-
-lemma ref_present_upd [simp]: 
-  "ref_present r (upd a i v h) = ref_present r h"
-  by (simp add: upd_def ref_present_def set_array_def get_array_def)
-
-lemma array_present_set_ref [simp]:
-  "array_present a (set_ref r v h) = array_present a h"
-  by (simp add: array_present_def set_ref_def)
-
-lemma array_present_new_ref [simp]:
-  "array_present a h \<Longrightarrow> array_present a (snd (ref v h))"
-  by (simp add: array_present_def new_ref_def ref_def Let_def)
-
-hide (open) const empty array array_of_list upd length ref
-
-end
--- a/src/HOL/Library/Heap_Monad.thy	Mon Jan 19 20:37:08 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,425 +0,0 @@
-(*  Title:      HOL/Library/Heap_Monad.thy
-    ID:         $Id$
-    Author:     John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen
-*)
-
-header {* A monad with a polymorphic heap *}
-
-theory Heap_Monad
-imports Heap
-begin
-
-subsection {* The monad *}
-
-subsubsection {* Monad combinators *}
-
-datatype exception = Exn
-
-text {* Monadic heap actions either produce values
-  and transform the heap, or fail *}
-datatype 'a Heap = Heap "heap \<Rightarrow> ('a + exception) \<times> heap"
-
-primrec
-  execute :: "'a Heap \<Rightarrow> heap \<Rightarrow> ('a + exception) \<times> heap" where
-  "execute (Heap f) = f"
-lemmas [code del] = execute.simps
-
-lemma Heap_execute [simp]:
-  "Heap (execute f) = f" by (cases f) simp_all
-
-lemma Heap_eqI:
-  "(\<And>h. execute f h = execute g h) \<Longrightarrow> f = g"
-    by (cases f, cases g) (auto simp: expand_fun_eq)
-
-lemma Heap_eqI':
-  "(\<And>h. (\<lambda>x. execute (f x) h) = (\<lambda>y. execute (g y) h)) \<Longrightarrow> f = g"
-    by (auto simp: expand_fun_eq intro: Heap_eqI)
-
-lemma Heap_strip: "(\<And>f. PROP P f) \<equiv> (\<And>g. PROP P (Heap g))"
-proof
-  fix g :: "heap \<Rightarrow> ('a + exception) \<times> heap" 
-  assume "\<And>f. PROP P f"
-  then show "PROP P (Heap g)" .
-next
-  fix f :: "'a Heap" 
-  assume assm: "\<And>g. PROP P (Heap g)"
-  then have "PROP P (Heap (execute f))" .
-  then show "PROP P f" by simp
-qed
-
-definition
-  heap :: "(heap \<Rightarrow> 'a \<times> heap) \<Rightarrow> 'a Heap" where
-  [code del]: "heap f = Heap (\<lambda>h. apfst Inl (f h))"
-
-lemma execute_heap [simp]:
-  "execute (heap f) h = apfst Inl (f h)"
-  by (simp add: heap_def)
-
-definition
-  bindM :: "'a Heap \<Rightarrow> ('a \<Rightarrow> 'b Heap) \<Rightarrow> 'b Heap" (infixl ">>=" 54) where
-  [code del]: "f >>= g = Heap (\<lambda>h. case execute f h of
-                  (Inl x, h') \<Rightarrow> execute (g x) h'
-                | r \<Rightarrow> r)"
-
-notation
-  bindM (infixl "\<guillemotright>=" 54)
-
-abbreviation
-  chainM :: "'a Heap \<Rightarrow> 'b Heap \<Rightarrow> 'b Heap"  (infixl ">>" 54) where
-  "f >> g \<equiv> f >>= (\<lambda>_. g)"
-
-notation
-  chainM (infixl "\<guillemotright>" 54)
-
-definition
-  return :: "'a \<Rightarrow> 'a Heap" where
-  [code del]: "return x = heap (Pair x)"
-
-lemma execute_return [simp]:
-  "execute (return x) h = apfst Inl (x, h)"
-  by (simp add: return_def)
-
-definition
-  raise :: "string \<Rightarrow> 'a Heap" where -- {* the string is just decoration *}
-  [code del]: "raise s = Heap (Pair (Inr Exn))"
-
-notation (latex output)
-  "raise" ("\<^raw:{\textsf{raise}}>")
-
-lemma execute_raise [simp]:
-  "execute (raise s) h = (Inr Exn, h)"
-  by (simp add: raise_def)
-
-
-subsubsection {* do-syntax *}
-
-text {*
-  We provide a convenient do-notation for monadic expressions
-  well-known from Haskell.  @{const Let} is printed
-  specially in do-expressions.
-*}
-
-nonterminals do_expr
-
-syntax
-  "_do" :: "do_expr \<Rightarrow> 'a"
-    ("(do (_)//done)" [12] 100)
-  "_bindM" :: "pttrn \<Rightarrow> 'a \<Rightarrow> do_expr \<Rightarrow> do_expr"
-    ("_ <- _;//_" [1000, 13, 12] 12)
-  "_chainM" :: "'a \<Rightarrow> do_expr \<Rightarrow> do_expr"
-    ("_;//_" [13, 12] 12)
-  "_let" :: "pttrn \<Rightarrow> 'a \<Rightarrow> do_expr \<Rightarrow> do_expr"
-    ("let _ = _;//_" [1000, 13, 12] 12)
-  "_nil" :: "'a \<Rightarrow> do_expr"
-    ("_" [12] 12)
-
-syntax (xsymbols)
-  "_bindM" :: "pttrn \<Rightarrow> 'a \<Rightarrow> do_expr \<Rightarrow> do_expr"
-    ("_ \<leftarrow> _;//_" [1000, 13, 12] 12)
-syntax (latex output)
-  "_do" :: "do_expr \<Rightarrow> 'a"
-    ("(\<^raw:{\textsf{do}}> (_))" [12] 100)
-  "_let" :: "pttrn \<Rightarrow> 'a \<Rightarrow> do_expr \<Rightarrow> do_expr"
-    ("\<^raw:\textsf{let}> _ = _;//_" [1000, 13, 12] 12)
-notation (latex output)
-  "return" ("\<^raw:{\textsf{return}}>")
-
-translations
-  "_do f" => "f"
-  "_bindM x f g" => "f \<guillemotright>= (\<lambda>x. g)"
-  "_chainM f g" => "f \<guillemotright> g"
-  "_let x t f" => "CONST Let t (\<lambda>x. f)"
-  "_nil f" => "f"
-
-print_translation {*
-let
-  fun dest_abs_eta (Abs (abs as (_, ty, _))) =
-        let
-          val (v, t) = Syntax.variant_abs abs;
-        in (Free (v, ty), t) end
-    | dest_abs_eta t =
-        let
-          val (v, t) = Syntax.variant_abs ("", dummyT, t $ Bound 0);
-        in (Free (v, dummyT), t) end;
-  fun unfold_monad (Const (@{const_syntax bindM}, _) $ f $ g) =
-        let
-          val (v, g') = dest_abs_eta g;
-          val vs = fold_aterms (fn Free (v, _) => insert (op =) v | _ => I) v [];
-          val v_used = fold_aterms
-            (fn Free (w, _) => (fn s => s orelse member (op =) vs w) | _ => I) g' false;
-        in if v_used then
-          Const ("_bindM", dummyT) $ v $ f $ unfold_monad g'
-        else
-          Const ("_chainM", dummyT) $ f $ unfold_monad g'
-        end
-    | unfold_monad (Const (@{const_syntax chainM}, _) $ f $ g) =
-        Const ("_chainM", dummyT) $ f $ unfold_monad g
-    | unfold_monad (Const (@{const_syntax Let}, _) $ f $ g) =
-        let
-          val (v, g') = dest_abs_eta g;
-        in Const ("_let", dummyT) $ v $ f $ unfold_monad g' end
-    | unfold_monad (Const (@{const_syntax Pair}, _) $ f) =
-        Const (@{const_syntax return}, dummyT) $ f
-    | unfold_monad f = f;
-  fun contains_bindM (Const (@{const_syntax bindM}, _) $ _ $ _) = true
-    | contains_bindM (Const (@{const_syntax Let}, _) $ _ $ Abs (_, _, t)) =
-        contains_bindM t;
-  fun bindM_monad_tr' (f::g::ts) = list_comb
-    (Const ("_do", dummyT) $ unfold_monad (Const (@{const_syntax bindM}, dummyT) $ f $ g), ts);
-  fun Let_monad_tr' (f :: (g as Abs (_, _, g')) :: ts) = if contains_bindM g' then list_comb
-      (Const ("_do", dummyT) $ unfold_monad (Const (@{const_syntax Let}, dummyT) $ f $ g), ts)
-    else raise Match;
-in [
-  (@{const_syntax bindM}, bindM_monad_tr'),
-  (@{const_syntax Let}, Let_monad_tr')
-] end;
-*}
-
-
-subsection {* Monad properties *}
-
-subsubsection {* Monad laws *}
-
-lemma return_bind: "return x \<guillemotright>= f = f x"
-  by (simp add: bindM_def return_def)
-
-lemma bind_return: "f \<guillemotright>= return = f"
-proof (rule Heap_eqI)
-  fix h
-  show "execute (f \<guillemotright>= return) h = execute f h"
-    by (auto simp add: bindM_def return_def split: sum.splits prod.splits)
-qed
-
-lemma bind_bind: "(f \<guillemotright>= g) \<guillemotright>= h = f \<guillemotright>= (\<lambda>x. g x \<guillemotright>= h)"
-  by (rule Heap_eqI) (auto simp add: bindM_def split: split: sum.splits prod.splits)
-
-lemma bind_bind': "f \<guillemotright>= (\<lambda>x. g x \<guillemotright>= h x) = f \<guillemotright>= (\<lambda>x. g x \<guillemotright>= (\<lambda>y. return (x, y))) \<guillemotright>= (\<lambda>(x, y). h x y)"
-  by (rule Heap_eqI) (auto simp add: bindM_def split: split: sum.splits prod.splits)
-
-lemma raise_bind: "raise e \<guillemotright>= f = raise e"
-  by (simp add: raise_def bindM_def)
-
-
-lemmas monad_simp = return_bind bind_return bind_bind raise_bind
-
-
-subsection {* Generic combinators *}
-
-definition
-  liftM :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b Heap"
-where
-  "liftM f = return o f"
-
-definition
-  compM :: "('a \<Rightarrow> 'b Heap) \<Rightarrow> ('b \<Rightarrow> 'c Heap) \<Rightarrow> 'a \<Rightarrow> 'c Heap" (infixl ">>==" 54)
-where
-  "(f >>== g) = (\<lambda>x. f x \<guillemotright>= g)"
-
-notation
-  compM (infixl "\<guillemotright>==" 54)
-
-lemma liftM_collapse: "liftM f x = return (f x)"
-  by (simp add: liftM_def)
-
-lemma liftM_compM: "liftM f \<guillemotright>== g = g o f"
-  by (auto intro: Heap_eqI' simp add: expand_fun_eq liftM_def compM_def bindM_def)
-
-lemma compM_return: "f \<guillemotright>== return = f"
-  by (simp add: compM_def monad_simp)
-
-lemma compM_compM: "(f \<guillemotright>== g) \<guillemotright>== h = f \<guillemotright>== (g \<guillemotright>== h)"
-  by (simp add: compM_def monad_simp)
-
-lemma liftM_bind:
-  "(\<lambda>x. liftM f x \<guillemotright>= liftM g) = liftM (\<lambda>x. g (f x))"
-  by (rule Heap_eqI') (simp add: monad_simp liftM_def bindM_def)
-
-lemma liftM_comp:
-  "liftM f o g = liftM (f o g)"
-  by (rule Heap_eqI') (simp add: liftM_def)
-
-lemmas monad_simp' = monad_simp liftM_compM compM_return
-  compM_compM liftM_bind liftM_comp
-
-primrec 
-  mapM :: "('a \<Rightarrow> 'b Heap) \<Rightarrow> 'a list \<Rightarrow> 'b list Heap"
-where
-  "mapM f [] = return []"
-  | "mapM f (x#xs) = do y \<leftarrow> f x;
-                        ys \<leftarrow> mapM f xs;
-                        return (y # ys)
-                     done"
-
-primrec
-  foldM :: "('a \<Rightarrow> 'b \<Rightarrow> 'b Heap) \<Rightarrow> 'a list \<Rightarrow> 'b \<Rightarrow> 'b Heap"
-where
-  "foldM f [] s = return s"
-  | "foldM f (x#xs) s = f x s \<guillemotright>= foldM f xs"
-
-definition
-  assert :: "('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a Heap"
-where
-  "assert P x = (if P x then return x else raise (''assert''))"
-
-lemma assert_cong [fundef_cong]:
-  assumes "P = P'"
-  assumes "\<And>x. P' x \<Longrightarrow> f x = f' x"
-  shows "(assert P x >>= f) = (assert P' x >>= f')"
-  using assms by (auto simp add: assert_def return_bind raise_bind)
-
-hide (open) const heap execute
-
-
-subsection {* Code generator setup *}
-
-subsubsection {* Logical intermediate layer *}
-
-definition
-  Fail :: "message_string \<Rightarrow> exception"
-where
-  [code del]: "Fail s = Exn"
-
-definition
-  raise_exc :: "exception \<Rightarrow> 'a Heap"
-where
-  [code del]: "raise_exc e = raise []"
-
-lemma raise_raise_exc [code, code inline]:
-  "raise s = raise_exc (Fail (STR s))"
-  unfolding Fail_def raise_exc_def raise_def ..
-
-hide (open) const Fail raise_exc
-
-
-subsubsection {* SML and OCaml *}
-
-code_type Heap (SML "unit/ ->/ _")
-code_const Heap (SML "raise/ (Fail/ \"bare Heap\")")
-code_const "op \<guillemotright>=" (SML "!(fn/ f'_/ =>/ fn/ ()/ =>/ f'_/ (_/ ())/ ())")
-code_const return (SML "!(fn/ ()/ =>/ _)")
-code_const "Heap_Monad.Fail" (SML "Fail")
-code_const "Heap_Monad.raise_exc" (SML "!(fn/ ()/ =>/ raise/ _)")
-
-code_type Heap (OCaml "_")
-code_const Heap (OCaml "failwith/ \"bare Heap\"")
-code_const "op \<guillemotright>=" (OCaml "!(fun/ f'_/ ()/ ->/ f'_/ (_/ ())/ ())")
-code_const return (OCaml "!(fun/ ()/ ->/ _)")
-code_const "Heap_Monad.Fail" (OCaml "Failure")
-code_const "Heap_Monad.raise_exc" (OCaml "!(fun/ ()/ ->/ raise/ _)")
-
-setup {* let
-  open Code_Thingol;
-
-  fun lookup naming = the o Code_Thingol.lookup_const naming;
-
-  fun imp_monad_bind'' bind' return' unit' ts =
-    let
-      val dummy_name = "";
-      val dummy_type = ITyVar dummy_name;
-      val dummy_case_term = IVar dummy_name;
-      (*assumption: dummy values are not relevant for serialization*)
-      val unitt = IConst (unit', ([], []));
-      fun dest_abs ((v, ty) `|-> t, _) = ((v, ty), t)
-        | dest_abs (t, ty) =
-            let
-              val vs = Code_Thingol.fold_varnames cons t [];
-              val v = Name.variant vs "x";
-              val ty' = (hd o fst o Code_Thingol.unfold_fun) ty;
-            in ((v, ty'), t `$ IVar v) end;
-      fun force (t as IConst (c, _) `$ t') = if c = return'
-            then t' else t `$ unitt
-        | force t = t `$ unitt;
-      fun tr_bind' [(t1, _), (t2, ty2)] =
-        let
-          val ((v, ty), t) = dest_abs (t2, ty2);
-        in ICase (((force t1, ty), [(IVar v, tr_bind'' t)]), dummy_case_term) end
-      and tr_bind'' t = case Code_Thingol.unfold_app t
-           of (IConst (c, (_, ty1 :: ty2 :: _)), [x1, x2]) => if c = bind'
-                then tr_bind' [(x1, ty1), (x2, ty2)]
-                else force t
-            | _ => force t;
-    in (dummy_name, dummy_type) `|-> ICase (((IVar dummy_name, dummy_type),
-      [(unitt, tr_bind' ts)]), dummy_case_term) end
-  and imp_monad_bind' bind' return' unit' (const as (c, (_, tys))) ts = if c = bind' then case (ts, tys)
-     of ([t1, t2], ty1 :: ty2 :: _) => imp_monad_bind'' bind' return' unit' [(t1, ty1), (t2, ty2)]
-      | ([t1, t2, t3], ty1 :: ty2 :: _) => imp_monad_bind'' bind' return' unit' [(t1, ty1), (t2, ty2)] `$ t3
-      | (ts, _) => imp_monad_bind bind' return' unit' (eta_expand 2 (const, ts))
-    else IConst const `$$ map (imp_monad_bind bind' return' unit') ts
-  and imp_monad_bind bind' return' unit' (IConst const) = imp_monad_bind' bind' return' unit' const []
-    | imp_monad_bind bind' return' unit' (t as IVar _) = t
-    | imp_monad_bind bind' return' unit' (t as _ `$ _) = (case unfold_app t
-       of (IConst const, ts) => imp_monad_bind' bind' return' unit' const ts
-        | (t, ts) => imp_monad_bind bind' return' unit' t `$$ map (imp_monad_bind bind' return' unit') ts)
-    | imp_monad_bind bind' return' unit' (v_ty `|-> t) = v_ty `|-> imp_monad_bind bind' return' unit' t
-    | imp_monad_bind bind' return' unit' (ICase (((t, ty), pats), t0)) = ICase
-        (((imp_monad_bind bind' return' unit' t, ty), (map o pairself) (imp_monad_bind bind' return' unit') pats), imp_monad_bind bind' return' unit' t0);
-
-   fun imp_program naming = (Graph.map_nodes o map_terms_stmt)
-     (imp_monad_bind (lookup naming @{const_name bindM})
-       (lookup naming @{const_name return})
-       (lookup naming @{const_name Unity}));
-
-in
-
-  Code_Target.extend_target ("SML_imp", ("SML", imp_program))
-  #> Code_Target.extend_target ("OCaml_imp", ("OCaml", imp_program))
-
-end
-*}
-
-
-code_reserved OCaml Failure raise
-
-
-subsubsection {* Haskell *}
-
-text {* Adaption layer *}
-
-code_include Haskell "STMonad"
-{*import qualified Control.Monad;
-import qualified Control.Monad.ST;
-import qualified Data.STRef;
-import qualified Data.Array.ST;
-
-type RealWorld = Control.Monad.ST.RealWorld;
-type ST s a = Control.Monad.ST.ST s a;
-type STRef s a = Data.STRef.STRef s a;
-type STArray s a = Data.Array.ST.STArray s Int a;
-
-runST :: (forall s. ST s a) -> a;
-runST s = Control.Monad.ST.runST s;
-
-newSTRef = Data.STRef.newSTRef;
-readSTRef = Data.STRef.readSTRef;
-writeSTRef = Data.STRef.writeSTRef;
-
-newArray :: (Int, Int) -> a -> ST s (STArray s a);
-newArray = Data.Array.ST.newArray;
-
-newListArray :: (Int, Int) -> [a] -> ST s (STArray s a);
-newListArray = Data.Array.ST.newListArray;
-
-lengthArray :: STArray s a -> ST s Int;
-lengthArray a = Control.Monad.liftM snd (Data.Array.ST.getBounds a);
-
-readArray :: STArray s a -> Int -> ST s a;
-readArray = Data.Array.ST.readArray;
-
-writeArray :: STArray s a -> Int -> a -> ST s ();
-writeArray = Data.Array.ST.writeArray;*}
-
-code_reserved Haskell RealWorld ST STRef Array
-  runST
-  newSTRef reasSTRef writeSTRef
-  newArray newListArray lengthArray readArray writeArray
-
-text {* Monad *}
-
-code_type Heap (Haskell "ST/ RealWorld/ _")
-code_const Heap (Haskell "error/ \"bare Heap\"")
-code_monad "op \<guillemotright>=" Haskell
-code_const return (Haskell "return")
-code_const "Heap_Monad.Fail" (Haskell "_")
-code_const "Heap_Monad.raise_exc" (Haskell "error")
-
-end
--- a/src/HOL/Library/Imperative_HOL.thy	Mon Jan 19 20:37:08 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,12 +0,0 @@
-(*  Title:      HOL/Library/Imperative_HOL.thy
-    ID:         $Id$
-    Author:     John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen
-*)
-
-header {* Entry point into monadic imperative HOL *}
-
-theory Imperative_HOL
-imports Array Ref Relational
-begin
-
-end
--- a/src/HOL/Library/LaTeXsugar.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Library/LaTeXsugar.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Library/LaTeXsugar.thy
-    ID:         $Id$
     Author:     Gerwin Klain, Tobias Nipkow, Norbert Schirmer
     Copyright   2005 NICTA and TUM
 *)
--- a/src/HOL/Library/Library.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Library/Library.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -1,4 +1,3 @@
-(* $Id$ *)
 (*<*)
 theory Library
 imports
@@ -22,7 +21,6 @@
   Executable_Set
   Float
   FuncSet
-  Imperative_HOL
   Infinite_Set
   ListVector
   Multiset
@@ -39,6 +37,7 @@
   Ramsey
   RBT
   State_Monad
+  Univ_Poly
   While_Combinator
   Word
   Zorn
--- a/src/HOL/Library/Multiset.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Library/Multiset.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Library/Multiset.thy
-    ID:         $Id$
     Author:     Tobias Nipkow, Markus Wenzel, Lawrence C Paulson, Norbert Voelker
 *)
 
@@ -1080,16 +1079,16 @@
 apply simp
 done
 
-class_interpretation mset_order: order ["op \<le>#" "op <#"]
+interpretation mset_order!: order "op \<le>#" "op <#"
 proof qed (auto intro: order.intro mset_le_refl mset_le_antisym
   mset_le_trans simp: mset_less_def)
 
-class_interpretation mset_order_cancel_semigroup:
-  pordered_cancel_ab_semigroup_add ["op +" "op \<le>#" "op <#"]
+interpretation mset_order_cancel_semigroup!:
+  pordered_cancel_ab_semigroup_add "op +" "op \<le>#" "op <#"
 proof qed (erule mset_le_mono_add [OF mset_le_refl])
 
-class_interpretation mset_order_semigroup_cancel:
-  pordered_ab_semigroup_add_imp_le ["op +" "op \<le>#" "op <#"]
+interpretation mset_order_semigroup_cancel!:
+  pordered_ab_semigroup_add_imp_le "op +" "op \<le>#" "op <#"
 proof qed simp
 
 
@@ -1156,7 +1155,7 @@
   then show ?case using T by simp
 qed
 
-lemmas mset_less_trans = mset_order.less_eq_less.less_trans
+lemmas mset_less_trans = mset_order.less_trans
 
 lemma mset_less_diff_self: "c \<in># B \<Longrightarrow> B - {#c#} \<subset># B"
 by (auto simp: mset_le_def mset_less_def multi_drop_mem_not_eq)
--- a/src/HOL/Library/Nat_Infinity.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Library/Nat_Infinity.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Library/Nat_Infinity.thy
-    ID:         $Id$
     Author:     David von Oheimb, TU Muenchen;  Florian Haftmann, TU Muenchen
 *)
 
@@ -9,6 +8,17 @@
 imports Plain "~~/src/HOL/Presburger"
 begin
 
+text {* FIXME: move to Nat.thy *}
+
+instantiation nat :: bot
+begin
+
+definition bot_nat :: nat where
+  "bot_nat = 0"
+
+instance proof
+qed (simp add: bot_nat_def)
+
 subsection {* Type definition *}
 
 text {*
@@ -16,6 +26,8 @@
   infinity.
 *}
 
+end
+
 datatype inat = Fin nat | Infty
 
 notation (xsymbols)
@@ -353,6 +365,20 @@
 apply (erule (1) le_less_trans)
 done
 
+instantiation inat :: "{bot, top}"
+begin
+
+definition bot_inat :: inat where
+  "bot_inat = 0"
+
+definition top_inat :: inat where
+  "top_inat = \<infinity>"
+
+instance proof
+qed (simp_all add: bot_inat_def top_inat_def)
+
+end
+
 
 subsection {* Well-ordering *}
 
--- a/src/HOL/Library/OptionalSugar.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Library/OptionalSugar.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Library/OptionalSugar.thy
-    ID:         $Id$
     Author:     Gerwin Klain, Tobias Nipkow
     Copyright   2005 NICTA and TUM
 *)
@@ -37,6 +36,36 @@
   "_bind (p#DUMMY) e" <= "_bind p (hd e)"
   "_bind (DUMMY#p) e" <= "_bind p (tl e)"
 
+(* type constraints with spacing *)
+setup {*
+let
+  val typ = SimpleSyntax.read_typ;
+  val typeT = Syntax.typeT;
+  val spropT = Syntax.spropT;
+in
+  Sign.del_modesyntax_i (Symbol.xsymbolsN, false) [
+    ("_constrain", typ "logic => type => logic", Mixfix ("_\<Colon>_", [4, 0], 3)),
+    ("_constrain", [spropT, typeT] ---> spropT, Mixfix ("_\<Colon>_", [4, 0], 3))]
+  #> Sign.add_modesyntax_i (Symbol.xsymbolsN, false) [
+      ("_constrain", typ "logic => type => logic", Mixfix ("_ \<Colon>  _", [4, 0], 3)),
+      ("_constrain", [spropT, typeT] ---> spropT, Mixfix ("_ \<Colon> _", [4, 0], 3))]
+end
+*}
+
+(* sorts as intersections *)
+setup {*
+let
+  val sortT = Type ("sort", []); (*FIXME*)
+  val classesT = Type ("classes", []); (*FIXME*)
+in
+  Sign.add_modesyntax_i (Symbol.xsymbolsN, false) [
+    ("_topsort", sortT, Mixfix ("\<top>", [], Syntax.max_pri)),
+    ("_sort", classesT --> sortT, Mixfix ("'(_')", [], Syntax.max_pri)),
+    ("_classes", Lexicon.idT --> classesT --> classesT, Mixfix ("_ \<inter> _", [], Syntax.max_pri)),
+    ("_classes", Lexicon.longidT --> classesT --> classesT, Mixfix ("_ \<inter> _", [], Syntax.max_pri))
+  ]
+end
+*}
 
 end
 (*>*)
--- a/src/HOL/Library/Ref.thy	Mon Jan 19 20:37:08 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,91 +0,0 @@
-(*  Title:      HOL/Library/Ref.thy
-    ID:         $Id$
-    Author:     John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen
-*)
-
-header {* Monadic references *}
-
-theory Ref
-imports Heap_Monad
-begin
-
-text {*
-  Imperative reference operations; modeled after their ML counterparts.
-  See http://caml.inria.fr/pub/docs/manual-caml-light/node14.15.html
-  and http://www.smlnj.org/doc/Conversion/top-level-comparison.html
-*}
-
-subsection {* Primitives *}
-
-definition
-  new :: "'a\<Colon>heap \<Rightarrow> 'a ref Heap" where
-  [code del]: "new v = Heap_Monad.heap (Heap.ref v)"
-
-definition
-  lookup :: "'a\<Colon>heap ref \<Rightarrow> 'a Heap" ("!_" 61) where
-  [code del]: "lookup r = Heap_Monad.heap (\<lambda>h. (get_ref r h, h))"
-
-definition
-  update :: "'a ref \<Rightarrow> ('a\<Colon>heap) \<Rightarrow> unit Heap" ("_ := _" 62) where
-  [code del]: "update r e = Heap_Monad.heap (\<lambda>h. ((), set_ref r e h))"
-
-
-subsection {* Derivates *}
-
-definition
-  change :: "('a\<Colon>heap \<Rightarrow> 'a) \<Rightarrow> 'a ref \<Rightarrow> 'a Heap"
-where
-  "change f r = (do x \<leftarrow> ! r;
-                    let y = f x;
-                    r := y;
-                    return y
-                 done)"
-
-hide (open) const new lookup update change
-
-
-subsection {* Properties *}
-
-lemma lookup_chain:
-  "(!r \<guillemotright> f) = f"
-  by (cases f)
-    (auto simp add: Let_def bindM_def lookup_def expand_fun_eq)
-
-lemma update_change [code]:
-  "r := e = Ref.change (\<lambda>_. e) r \<guillemotright> return ()"
-  by (auto simp add: monad_simp change_def lookup_chain)
-
-
-subsection {* Code generator setup *}
-
-subsubsection {* SML *}
-
-code_type ref (SML "_/ ref")
-code_const Ref (SML "raise/ (Fail/ \"bare Ref\")")
-code_const Ref.new (SML "(fn/ ()/ =>/ ref/ _)")
-code_const Ref.lookup (SML "(fn/ ()/ =>/ !/ _)")
-code_const Ref.update (SML "(fn/ ()/ =>/ _/ :=/ _)")
-
-code_reserved SML ref
-
-
-subsubsection {* OCaml *}
-
-code_type ref (OCaml "_/ ref")
-code_const Ref (OCaml "failwith/ \"bare Ref\")")
-code_const Ref.new (OCaml "(fn/ ()/ =>/ ref/ _)")
-code_const Ref.lookup (OCaml "(fn/ ()/ =>/ !/ _)")
-code_const Ref.update (OCaml "(fn/ ()/ =>/ _/ :=/ _)")
-
-code_reserved OCaml ref
-
-
-subsubsection {* Haskell *}
-
-code_type ref (Haskell "STRef/ RealWorld/ _")
-code_const Ref (Haskell "error/ \"bare Ref\"")
-code_const Ref.new (Haskell "newSTRef")
-code_const Ref.lookup (Haskell "readSTRef")
-code_const Ref.update (Haskell "writeSTRef")
-
-end
\ No newline at end of file
--- a/src/HOL/Library/Relational.thy	Mon Jan 19 20:37:08 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,700 +0,0 @@
-theory Relational 
-imports Array Ref
-begin
-
-section{* Definition of the Relational framework *}
-
-text {* The crel predicate states that when a computation c runs with the heap h
-  will result in return value r and a heap h' (if no exception occurs). *}  
-
-definition crel :: "'a Heap \<Rightarrow> heap \<Rightarrow> heap \<Rightarrow> 'a \<Rightarrow> bool"
-where
-  crel_def': "crel c h h' r \<longleftrightarrow> Heap_Monad.execute c h = (Inl r, h')"
-
-lemma crel_def: -- FIXME
-  "crel c h h' r \<longleftrightarrow> (Inl r, h') = Heap_Monad.execute c h"
-  unfolding crel_def' by auto
-
-lemma crel_deterministic: "\<lbrakk> crel f h h' a; crel f h h'' b \<rbrakk> \<Longrightarrow> (a = b) \<and> (h' = h'')"
-unfolding crel_def' by auto
-
-section {* Elimination rules *}
-
-text {* For all commands, we define simple elimination rules. *}
-(* FIXME: consumes 1 necessary ?? *)
-
-subsection {* Elimination rules for basic monadic commands *}
-
-lemma crelE[consumes 1]:
-  assumes "crel (f >>= g) h h'' r'"
-  obtains h' r where "crel f h h' r" "crel (g r) h' h'' r'"
-  using assms
-  by (auto simp add: crel_def bindM_def Let_def prod_case_beta split_def Pair_fst_snd_eq split add: sum.split_asm)
-
-lemma crelE'[consumes 1]:
-  assumes "crel (f >> g) h h'' r'"
-  obtains h' r where "crel f h h' r" "crel g h' h'' r'"
-  using assms
-  by (elim crelE) auto
-
-lemma crel_return[consumes 1]:
-  assumes "crel (return x) h h' r"
-  obtains "r = x" "h = h'"
-  using assms
-  unfolding crel_def return_def by simp
-
-lemma crel_raise[consumes 1]:
-  assumes "crel (raise x) h h' r"
-  obtains "False"
-  using assms
-  unfolding crel_def raise_def by simp
-
-lemma crel_if:
-  assumes "crel (if c then t else e) h h' r"
-  obtains "c" "crel t h h' r"
-        | "\<not>c" "crel e h h' r"
-  using assms
-  unfolding crel_def by auto
-
-lemma crel_option_case:
-  assumes "crel (case x of None \<Rightarrow> n | Some y \<Rightarrow> s y) h h' r"
-  obtains "x = None" "crel n h h' r"
-        | y where "x = Some y" "crel (s y) h h' r" 
-  using assms
-  unfolding crel_def by auto
-
-lemma crel_mapM:
-  assumes "crel (mapM f xs) h h' r"
-  assumes "\<And>h h'. P f [] h h' []"
-  assumes "\<And>h h1 h' x xs y ys. \<lbrakk> crel (f x) h h1 y; crel (mapM f xs) h1 h' ys; P f xs h1 h' ys \<rbrakk> \<Longrightarrow> P f (x#xs) h h' (y#ys)"
-  shows "P f xs h h' r"
-using assms(1)
-proof (induct xs arbitrary: h h' r)
-  case Nil  with assms(2) show ?case
-    by (auto elim: crel_return)
-next
-  case (Cons x xs)
-  from Cons(2) obtain h1 y ys where crel_f: "crel (f x) h h1 y"
-    and crel_mapM: "crel (mapM f xs) h1 h' ys"
-    and r_def: "r = y#ys"
-    unfolding mapM.simps
-    by (auto elim!: crelE crel_return)
-  from Cons(1)[OF crel_mapM] crel_mapM crel_f assms(3) r_def
-  show ?case by auto
-qed
-
-lemma crel_heap:
-  assumes "crel (Heap_Monad.heap f) h h' r"
-  obtains "h' = snd (f h)" "r = fst (f h)"
-  using assms
-  unfolding heap_def crel_def apfst_def split_def prod_fun_def by simp_all
-
-subsection {* Elimination rules for array commands *}
-
-lemma crel_length:
-  assumes "crel (length a) h h' r"
-  obtains "h = h'" "r = Heap.length a h'"
-  using assms
-  unfolding length_def
-  by (elim crel_heap) simp
-
-(* Strong version of the lemma for this operation is missing *) 
-lemma crel_new_weak:
-  assumes "crel (Array.new n v) h h' r"
-  obtains "get_array r h' = List.replicate n v"
-  using assms unfolding  Array.new_def
-  by (elim crel_heap) (auto simp:Heap.array_def Let_def split_def)
-
-lemma crel_nth[consumes 1]:
-  assumes "crel (nth a i) h h' r"
-  obtains "r = (get_array a h) ! i" "h = h'" "i < Heap.length a h"
-  using assms
-  unfolding nth_def
-  by (auto elim!: crelE crel_if crel_raise crel_length crel_heap)
-
-lemma crel_upd[consumes 1]:
-  assumes "crel (upd i v a) h h' r"
-  obtains "r = a" "h' = Heap.upd a i v h"
-  using assms
-  unfolding upd_def
-  by (elim crelE crel_if crel_return crel_raise
-    crel_length crel_heap) auto
-
-(* Strong version of the lemma for this operation is missing *)
-lemma crel_of_list_weak:
-  assumes "crel (Array.of_list xs) h h' r"
-  obtains "get_array r h' = xs"
-  using assms
-  unfolding of_list_def 
-  by (elim crel_heap) (simp add:get_array_init_array_list)
-
-lemma crel_map_entry:
-  assumes "crel (Array.map_entry i f a) h h' r"
-  obtains "r = a" "h' = Heap.upd a i (f (get_array a h ! i)) h"
-  using assms
-  unfolding Array.map_entry_def
-  by (elim crelE crel_upd crel_nth) auto
-
-lemma crel_swap:
-  assumes "crel (Array.swap i x a) h h' r"
-  obtains "r = get_array a h ! i" "h' = Heap.upd a i x h"
-  using assms
-  unfolding Array.swap_def
-  by (elim crelE crel_upd crel_nth crel_return) auto
-
-(* Strong version of the lemma for this operation is missing *)
-lemma crel_make_weak:
-  assumes "crel (Array.make n f) h h' r"
-  obtains "i < n \<Longrightarrow> get_array r h' ! i = f i"
-  using assms
-  unfolding Array.make_def
-  by (elim crel_of_list_weak) auto
-
-lemma upt_conv_Cons':
-  assumes "Suc a \<le> b"
-  shows "[b - Suc a..<b] = (b - Suc a)#[b - a..<b]"
-proof -
-  from assms have l: "b - Suc a < b" by arith
-  from assms have "Suc (b - Suc a) = b - a" by arith
-  with l show ?thesis by (simp add: upt_conv_Cons)
-qed
-
-lemma crel_mapM_nth:
-  assumes
-  "crel (mapM (Array.nth a) [Heap.length a h - n..<Heap.length a h]) h h' xs"
-  assumes "n \<le> Heap.length a h"
-  shows "h = h' \<and> xs = drop (Heap.length a h - n) (get_array a h)"
-using assms
-proof (induct n arbitrary: xs h h')
-  case 0 thus ?case
-    by (auto elim!: crel_return simp add: Heap.length_def)
-next
-  case (Suc n)
-  from Suc(3) have "[Heap.length a h - Suc n..<Heap.length a h] = (Heap.length a h - Suc n)#[Heap.length a h - n..<Heap.length a h]"
-    by (simp add: upt_conv_Cons')
-  with Suc(2) obtain r where
-    crel_mapM: "crel (mapM (Array.nth a) [Heap.length a h - n..<Heap.length a h]) h h' r"
-    and xs_def: "xs = get_array a h ! (Heap.length a h - Suc n) # r"
-    by (auto elim!: crelE crel_nth crel_return)
-  from Suc(3) have "Heap.length a h - n = Suc (Heap.length a h - Suc n)" 
-    by arith
-  with Suc.hyps[OF crel_mapM] xs_def show ?case
-    unfolding Heap.length_def
-    by (auto simp add: nth_drop')
-qed
-
-lemma crel_freeze:
-  assumes "crel (Array.freeze a) h h' xs"
-  obtains "h = h'" "xs = get_array a h"
-proof 
-  from assms have "crel (mapM (Array.nth a) [0..<Heap.length a h]) h h' xs"
-    unfolding freeze_def
-    by (auto elim: crelE crel_length)
-  hence "crel (mapM (Array.nth a) [(Heap.length a h - Heap.length a h)..<Heap.length a h]) h h' xs"
-    by simp
-  from crel_mapM_nth[OF this] show "h = h'" and "xs = get_array a h" by auto
-qed
-
-lemma crel_mapM_map_entry_remains:
-  assumes "crel (mapM (\<lambda>n. map_entry n f a) [Heap.length a h - n..<Heap.length a h]) h h' r"
-  assumes "i < Heap.length a h - n"
-  shows "get_array a h ! i = get_array a h' ! i"
-using assms
-proof (induct n arbitrary: h h' r)
-  case 0
-  thus ?case
-    by (auto elim: crel_return)
-next
-  case (Suc n)
-  let ?h1 = "Heap.upd a (Heap.length a h - Suc n) (f (get_array a h ! (Heap.length a h - Suc n))) h"
-  from Suc(3) have "[Heap.length a h - Suc n..<Heap.length a h] = (Heap.length a h - Suc n)#[Heap.length a h - n..<Heap.length a h]"
-    by (simp add: upt_conv_Cons')
-  from Suc(2) this obtain r where
-    crel_mapM: "crel (mapM (\<lambda>n. map_entry n f a) [Heap.length a h - n..<Heap.length a h]) ?h1 h' r"
-    by (auto simp add: elim!: crelE crel_map_entry crel_return)
-  have length_remains: "Heap.length a ?h1 = Heap.length a h" by simp
-  from crel_mapM have crel_mapM': "crel (mapM (\<lambda>n. map_entry n f a) [Heap.length a ?h1 - n..<Heap.length a ?h1]) ?h1 h' r"
-    by simp
-  from Suc(1)[OF this] length_remains Suc(3) show ?case by simp
-qed
-
-lemma crel_mapM_map_entry_changes:
-  assumes "crel (mapM (\<lambda>n. map_entry n f a) [Heap.length a h - n..<Heap.length a h]) h h' r"
-  assumes "n \<le> Heap.length a h"  
-  assumes "i \<ge> Heap.length a h - n"
-  assumes "i < Heap.length a h"
-  shows "get_array a h' ! i = f (get_array a h ! i)"
-using assms
-proof (induct n arbitrary: h h' r)
-  case 0
-  thus ?case
-    by (auto elim!: crel_return)
-next
-  case (Suc n)
-  let ?h1 = "Heap.upd a (Heap.length a h - Suc n) (f (get_array a h ! (Heap.length a h - Suc n))) h"
-  from Suc(3) have "[Heap.length a h - Suc n..<Heap.length a h] = (Heap.length a h - Suc n)#[Heap.length a h - n..<Heap.length a h]"
-    by (simp add: upt_conv_Cons')
-  from Suc(2) this obtain r where
-    crel_mapM: "crel (mapM (\<lambda>n. map_entry n f a) [Heap.length a h - n..<Heap.length a h]) ?h1 h' r"
-    by (auto simp add: elim!: crelE crel_map_entry crel_return)
-  have length_remains: "Heap.length a ?h1 = Heap.length a h" by simp
-  from Suc(3) have less: "Heap.length a h - Suc n < Heap.length a h - n" by arith
-  from Suc(3) have less2: "Heap.length a h - Suc n < Heap.length a h" by arith
-  from Suc(4) length_remains have cases: "i = Heap.length a ?h1 - Suc n \<or> i \<ge> Heap.length a ?h1 - n" by arith
-  from crel_mapM have crel_mapM': "crel (mapM (\<lambda>n. map_entry n f a) [Heap.length a ?h1 - n..<Heap.length a ?h1]) ?h1 h' r"
-    by simp
-  from Suc(1)[OF this] cases Suc(3) Suc(5) length_remains
-    crel_mapM_map_entry_remains[OF this, of "Heap.length a h - Suc n", symmetric] less less2
-  show ?case
-    by (auto simp add: nth_list_update_eq Heap.length_def)
-qed
-
-lemma crel_mapM_map_entry_length:
-  assumes "crel (mapM (\<lambda>n. map_entry n f a) [Heap.length a h - n..<Heap.length a h]) h h' r"
-  assumes "n \<le> Heap.length a h"
-  shows "Heap.length a h' = Heap.length a h"
-using assms
-proof (induct n arbitrary: h h' r)
-  case 0
-  thus ?case by (auto elim!: crel_return)
-next
-  case (Suc n)
-  let ?h1 = "Heap.upd a (Heap.length a h - Suc n) (f (get_array a h ! (Heap.length a h - Suc n))) h"
-  from Suc(3) have "[Heap.length a h - Suc n..<Heap.length a h] = (Heap.length a h - Suc n)#[Heap.length a h - n..<Heap.length a h]"
-    by (simp add: upt_conv_Cons')
-  from Suc(2) this obtain r where 
-    crel_mapM: "crel (mapM (\<lambda>n. map_entry n f a) [Heap.length a h - n..<Heap.length a h]) ?h1 h' r"
-    by (auto elim!: crelE crel_map_entry crel_return)
-  have length_remains: "Heap.length a ?h1 = Heap.length a h" by simp
-  from crel_mapM have crel_mapM': "crel (mapM (\<lambda>n. map_entry n f a) [Heap.length a ?h1 - n..<Heap.length a ?h1]) ?h1 h' r"
-    by simp
-  from Suc(1)[OF this] length_remains Suc(3) show ?case by simp
-qed
-
-lemma crel_mapM_map_entry:
-assumes "crel (mapM (\<lambda>n. map_entry n f a) [0..<Heap.length a h]) h h' r"
-  shows "get_array a h' = List.map f (get_array a h)"
-proof -
-  from assms have "crel (mapM (\<lambda>n. map_entry n f a) [Heap.length a h - Heap.length a h..<Heap.length a h]) h h' r" by simp
-  from crel_mapM_map_entry_length[OF this]
-  crel_mapM_map_entry_changes[OF this] show ?thesis
-    unfolding Heap.length_def
-    by (auto intro: nth_equalityI)
-qed
-
-lemma crel_map_weak:
-  assumes crel_map: "crel (Array.map f a) h h' r"
-  obtains "r = a" "get_array a h' = List.map f (get_array a h)"
-proof
-  from assms crel_mapM_map_entry show  "get_array a h' = List.map f (get_array a h)"
-    unfolding Array.map_def
-    by (fastsimp elim!: crelE crel_length crel_return)
-  from assms show "r = a"
-    unfolding Array.map_def
-    by (elim crelE crel_return)
-qed
-
-subsection {* Elimination rules for reference commands *}
-
-(* TODO:
-maybe introduce a new predicate "extends h' h x"
-which means h' extends h with a new reference x.
-Then crel_new: would be
-assumes "crel (Ref.new v) h h' x"
-obtains "get_ref x h' = v"
-and "extends h' h x"
-
-and we would need further rules for extends:
-extends h' h x \<Longrightarrow> \<not> ref_present x h
-extends h' h x \<Longrightarrow>  ref_present x h'
-extends h' h x \<Longrightarrow> ref_present y h \<Longrightarrow> ref_present y h'
-extends h' h x \<Longrightarrow> ref_present y h \<Longrightarrow> get_ref y h = get_ref y h'
-extends h' h x \<Longrightarrow> lim h' = Suc (lim h)
-*)
-
-lemma crel_Ref_new:
-  assumes "crel (Ref.new v) h h' x"
-  obtains "get_ref x h' = v"
-  and "\<not> ref_present x h"
-  and "ref_present x h'"
-  and "\<forall>y. ref_present y h \<longrightarrow> get_ref y h = get_ref y h'"
- (* and "lim h' = Suc (lim h)" *)
-  and "\<forall>y. ref_present y h \<longrightarrow> ref_present y h'"
-  using assms
-  unfolding Ref.new_def
-  apply (elim crel_heap)
-  unfolding Heap.ref_def
-  apply (simp add: Let_def)
-  unfolding Heap.new_ref_def
-  apply (simp add: Let_def)
-  unfolding ref_present_def
-  apply auto
-  unfolding get_ref_def set_ref_def
-  apply auto
-  done
-
-lemma crel_lookup:
-  assumes "crel (!ref) h h' r"
-  obtains "h = h'" "r = get_ref ref h"
-using assms
-unfolding Ref.lookup_def
-by (auto elim: crel_heap)
-
-lemma crel_update:
-  assumes "crel (ref := v) h h' r"
-  obtains "h' = set_ref ref v h" "r = ()"
-using assms
-unfolding Ref.update_def
-by (auto elim: crel_heap)
-
-lemma crel_change:
-  assumes "crel (Ref.change f ref) h h' r"
-  obtains "h' = set_ref ref (f (get_ref ref h)) h" "r = f (get_ref ref h)"
-using assms
-unfolding Ref.change_def Let_def
-by (auto elim!: crelE crel_lookup crel_update crel_return)
-
-subsection {* Elimination rules for the assert command *}
-
-lemma crel_assert[consumes 1]:
-  assumes "crel (assert P x) h h' r"
-  obtains "P x" "r = x" "h = h'"
-  using assms
-  unfolding assert_def
-  by (elim crel_if crel_return crel_raise) auto
-
-lemma crel_assert_eq: "(\<And>h h' r. crel f h h' r \<Longrightarrow> P r) \<Longrightarrow> f \<guillemotright>= assert P = f"
-unfolding crel_def bindM_def Let_def assert_def
-  raise_def return_def prod_case_beta
-apply (cases f)
-apply simp
-apply (simp add: expand_fun_eq split_def)
-apply auto
-apply (case_tac "fst (fun x)")
-apply (simp_all add: Pair_fst_snd_eq)
-apply (erule_tac x="x" in meta_allE)
-apply fastsimp
-done
-
-section {* Introduction rules *}
-
-subsection {* Introduction rules for basic monadic commands *}
-
-lemma crelI:
-  assumes "crel f h h' r" "crel (g r) h' h'' r'"
-  shows "crel (f >>= g) h h'' r'"
-  using assms by (simp add: crel_def' bindM_def)
-
-lemma crelI':
-  assumes "crel f h h' r" "crel g h' h'' r'"
-  shows "crel (f >> g) h h'' r'"
-  using assms by (intro crelI) auto
-
-lemma crel_returnI:
-  shows "crel (return x) h h x"
-  unfolding crel_def return_def by simp
-
-lemma crel_raiseI:
-  shows "\<not> (crel (raise x) h h' r)"
-  unfolding crel_def raise_def by simp
-
-lemma crel_ifI:
-  assumes "c \<longrightarrow> crel t h h' r"
-  "\<not>c \<longrightarrow> crel e h h' r"
-  shows "crel (if c then t else e) h h' r"
-  using assms
-  unfolding crel_def by auto
-
-lemma crel_option_caseI:
-  assumes "\<And>y. x = Some y \<Longrightarrow> crel (s y) h h' r"
-  assumes "x = None \<Longrightarrow> crel n h h' r"
-  shows "crel (case x of None \<Rightarrow> n | Some y \<Rightarrow> s y) h h' r"
-using assms
-by (auto split: option.split)
-
-lemma crel_heapI:
-  shows "crel (Heap_Monad.heap f) h (snd (f h)) (fst (f h))"
-  by (simp add: crel_def apfst_def split_def prod_fun_def)
-
-lemma crel_heapI':
-  assumes "h' = snd (f h)" "r = fst (f h)"
-  shows "crel (Heap_Monad.heap f) h h' r"
-  using assms
-  by (simp add: crel_def split_def apfst_def prod_fun_def)
-
-lemma crelI2:
-  assumes "\<exists>h' rs'. crel f h h' rs' \<and> (\<exists>h'' rs. crel (g rs') h' h'' rs)"
-  shows "\<exists>h'' rs. crel (f\<guillemotright>= g) h h'' rs"
-  oops
-
-lemma crel_ifI2:
-  assumes "c \<Longrightarrow> \<exists>h' r. crel t h h' r"
-  "\<not> c \<Longrightarrow> \<exists>h' r. crel e h h' r"
-  shows "\<exists> h' r. crel (if c then t else e) h h' r"
-  oops
-
-subsection {* Introduction rules for array commands *}
-
-lemma crel_lengthI:
-  shows "crel (length a) h h (Heap.length a h)"
-  unfolding length_def
-  by (rule crel_heapI') auto
-
-(* thm crel_newI for Array.new is missing *)
-
-lemma crel_nthI:
-  assumes "i < Heap.length a h"
-  shows "crel (nth a i) h h ((get_array a h) ! i)"
-  using assms
-  unfolding nth_def
-  by (auto intro!: crelI crel_ifI crel_raiseI crel_lengthI crel_heapI')
-
-lemma crel_updI:
-  assumes "i < Heap.length a h"
-  shows "crel (upd i v a) h (Heap.upd a i v h) a"
-  using assms
-  unfolding upd_def
-  by (auto intro!: crelI crel_ifI crel_returnI crel_raiseI
-    crel_lengthI crel_heapI')
-
-(* thm crel_of_listI is missing *)
-
-(* thm crel_map_entryI is missing *)
-
-(* thm crel_swapI is missing *)
-
-(* thm crel_makeI is missing *)
-
-(* thm crel_freezeI is missing *)
-
-(* thm crel_mapI is missing *)
-
-subsection {* Introduction rules for reference commands *}
-
-lemma crel_lookupI:
-  shows "crel (!ref) h h (get_ref ref h)"
-  unfolding lookup_def by (auto intro!: crel_heapI')
-
-lemma crel_updateI:
-  shows "crel (ref := v) h (set_ref ref v h) ()"
-  unfolding update_def by (auto intro!: crel_heapI')
-
-lemma crel_changeI: 
-  shows "crel (Ref.change f ref) h (set_ref ref (f (get_ref ref h)) h) (f (get_ref ref h))"
-unfolding change_def Let_def by (auto intro!: crelI crel_returnI crel_lookupI crel_updateI)
-
-subsection {* Introduction rules for the assert command *}
-
-lemma crel_assertI:
-  assumes "P x"
-  shows "crel (assert P x) h h x"
-  using assms
-  unfolding assert_def
-  by (auto intro!: crel_ifI crel_returnI crel_raiseI)
- 
-section {* Defintion of the noError predicate *}
-
-text {* We add a simple definitional setting for crel intro rules
-  where we only would like to show that the computation does not result in a exception for heap h,
-  but we do not care about statements about the resulting heap and return value.*}
-
-definition noError :: "'a Heap \<Rightarrow> heap \<Rightarrow> bool"
-where
-  "noError c h \<longleftrightarrow> (\<exists>r h'. (Inl r, h') = Heap_Monad.execute c h)"
-
-lemma noError_def': -- FIXME
-  "noError c h \<longleftrightarrow> (\<exists>r h'. Heap_Monad.execute c h = (Inl r, h'))"
-  unfolding noError_def apply auto proof -
-  fix r h'
-  assume "(Inl r, h') = Heap_Monad.execute c h"
-  then have "Heap_Monad.execute c h = (Inl r, h')" ..
-  then show "\<exists>r h'. Heap_Monad.execute c h = (Inl r, h')" by blast
-qed
-
-subsection {* Introduction rules for basic monadic commands *}
-
-lemma noErrorI:
-  assumes "noError f h"
-  assumes "\<And>h' r. crel f h h' r \<Longrightarrow> noError (g r) h'"
-  shows "noError (f \<guillemotright>= g) h"
-  using assms
-  by (auto simp add: noError_def' crel_def' bindM_def)
-
-lemma noErrorI':
-  assumes "noError f h"
-  assumes "\<And>h' r. crel f h h' r \<Longrightarrow> noError g h'"
-  shows "noError (f \<guillemotright> g) h"
-  using assms
-  by (auto simp add: noError_def' crel_def' bindM_def)
-
-lemma noErrorI2:
-"\<lbrakk>crel f h h' r ; noError f h; noError (g r) h'\<rbrakk>
-\<Longrightarrow> noError (f \<guillemotright>= g) h"
-by (auto simp add: noError_def' crel_def' bindM_def)
-
-lemma noError_return: 
-  shows "noError (return x) h"
-  unfolding noError_def return_def
-  by auto
-
-lemma noError_if:
-  assumes "c \<Longrightarrow> noError t h" "\<not> c \<Longrightarrow> noError e h"
-  shows "noError (if c then t else e) h"
-  using assms
-  unfolding noError_def
-  by auto
-
-lemma noError_option_case:
-  assumes "\<And>y. x = Some y \<Longrightarrow> noError (s y) h"
-  assumes "noError n h"
-  shows "noError (case x of None \<Rightarrow> n | Some y \<Rightarrow> s y) h"
-using assms
-by (auto split: option.split)
-
-lemma noError_mapM: 
-assumes "\<forall>x \<in> set xs. noError (f x) h \<and> crel (f x) h h (r x)" 
-shows "noError (mapM f xs) h"
-using assms
-proof (induct xs)
-  case Nil
-  thus ?case
-    unfolding mapM.simps by (intro noError_return)
-next
-  case (Cons x xs)
-  thus ?case
-    unfolding mapM.simps
-    by (auto intro: noErrorI2[of "f x"] noErrorI noError_return)
-qed
-
-lemma noError_heap:
-  shows "noError (Heap_Monad.heap f) h"
-  by (simp add: noError_def' apfst_def prod_fun_def split_def)
-
-subsection {* Introduction rules for array commands *}
-
-lemma noError_length:
-  shows "noError (Array.length a) h"
-  unfolding length_def
-  by (intro noError_heap)
-
-lemma noError_new:
-  shows "noError (Array.new n v) h"
-unfolding Array.new_def by (intro noError_heap)
-
-lemma noError_upd:
-  assumes "i < Heap.length a h"
-  shows "noError (Array.upd i v a) h"
-  using assms
-  unfolding upd_def
-  by (auto intro!: noErrorI noError_if noError_return noError_length noError_heap) (auto elim: crel_length)
-
-lemma noError_nth:
-assumes "i < Heap.length a h"
-  shows "noError (Array.nth a i) h"
-  using assms
-  unfolding nth_def
-  by (auto intro!: noErrorI noError_if noError_return noError_length noError_heap) (auto elim: crel_length)
-
-lemma noError_of_list:
-  shows "noError (of_list ls) h"
-  unfolding of_list_def by (rule noError_heap)
-
-lemma noError_map_entry:
-  assumes "i < Heap.length a h"
-  shows "noError (map_entry i f a) h"
-  using assms
-  unfolding map_entry_def
-  by (auto elim: crel_nth intro!: noErrorI noError_nth noError_upd)
-
-lemma noError_swap:
-  assumes "i < Heap.length a h"
-  shows "noError (swap i x a) h"
-  using assms
-  unfolding swap_def
-  by (auto elim: crel_nth intro!: noErrorI noError_return noError_nth noError_upd)
-
-lemma noError_make:
-  shows "noError (make n f) h"
-  unfolding make_def
-  by (auto intro: noError_of_list)
-
-(*TODO: move to HeapMonad *)
-lemma mapM_append:
-  "mapM f (xs @ ys) = mapM f xs \<guillemotright>= (\<lambda>xs. mapM f ys \<guillemotright>= (\<lambda>ys. return (xs @ ys)))"
-  by (induct xs) (simp_all add: monad_simp)
-
-lemma noError_freeze:
-  shows "noError (freeze a) h"
-unfolding freeze_def
-by (auto intro!: noErrorI noError_length noError_mapM[of _ _ _ "\<lambda>x. get_array a h ! x"]
-  noError_nth crel_nthI elim: crel_length)
-
-lemma noError_mapM_map_entry:
-  assumes "n \<le> Heap.length a h"
-  shows "noError (mapM (\<lambda>n. map_entry n f a) [Heap.length a h - n..<Heap.length a h]) h"
-using assms
-proof (induct n arbitrary: h)
-  case 0
-  thus ?case by (auto intro: noError_return)
-next
-  case (Suc n)
-  from Suc.prems have "[Heap.length a h - Suc n..<Heap.length a h] = (Heap.length a h - Suc n)#[Heap.length a h - n..<Heap.length a h]"
-    by (simp add: upt_conv_Cons')
-  with Suc.hyps[of "(Heap.upd a (Heap.length a h - Suc n) (f (get_array a h ! (Heap.length a h - Suc n))) h)"] Suc.prems show ?case
-    by (auto simp add: intro!: noErrorI noError_return noError_map_entry elim: crel_map_entry)
-qed
-
-lemma noError_map:
-  shows "noError (Array.map f a) h"
-using noError_mapM_map_entry[of "Heap.length a h" a h]
-unfolding Array.map_def
-by (auto intro: noErrorI noError_length noError_return elim!: crel_length)
-
-subsection {* Introduction rules for the reference commands *}
-
-lemma noError_Ref_new:
-  shows "noError (Ref.new v) h"
-unfolding Ref.new_def by (intro noError_heap)
-
-lemma noError_lookup:
-  shows "noError (!ref) h"
-  unfolding lookup_def by (intro noError_heap)
-
-lemma noError_update:
-  shows "noError (ref := v) h"
-  unfolding update_def by (intro noError_heap)
-
-lemma noError_change:
-  shows "noError (Ref.change f ref) h"
-  unfolding Ref.change_def Let_def by (intro noErrorI noError_lookup noError_update noError_return)
-
-subsection {* Introduction rules for the assert command *}
-
-lemma noError_assert:
-  assumes "P x"
-  shows "noError (assert P x) h"
-  using assms
-  unfolding assert_def
-  by (auto intro: noError_if noError_return)
-
-section {* Cumulative lemmas *}
-
-lemmas crel_elim_all =
-  crelE crelE' crel_return crel_raise crel_if crel_option_case
-  crel_length crel_new_weak crel_nth crel_upd crel_of_list_weak crel_map_entry crel_swap crel_make_weak crel_freeze crel_map_weak
-  crel_Ref_new crel_lookup crel_update crel_change
-  crel_assert
-
-lemmas crel_intro_all =
-  crelI crelI' crel_returnI crel_raiseI crel_ifI crel_option_caseI
-  crel_lengthI (* crel_newI *) crel_nthI crel_updI (* crel_of_listI crel_map_entryI crel_swapI crel_makeI crel_freezeI crel_mapI *)
-  (* crel_Ref_newI *) crel_lookupI crel_updateI crel_changeI
-  crel_assert
-
-lemmas noError_intro_all = 
-  noErrorI noErrorI' noError_return noError_if noError_option_case
-  noError_length noError_new noError_nth noError_upd noError_of_list noError_map_entry noError_swap noError_make noError_freeze noError_map
-  noError_Ref_new noError_lookup noError_update noError_change
-  noError_assert
-
-end
\ No newline at end of file
--- a/src/HOL/Library/SetsAndFunctions.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Library/SetsAndFunctions.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -107,26 +107,26 @@
   apply simp
   done
 
-class_interpretation set_semigroup_add: semigroup_add ["op \<oplus> :: ('a::semigroup_add) set => 'a set => 'a set"]
+interpretation set_semigroup_add!: semigroup_add "op \<oplus> :: ('a::semigroup_add) set => 'a set => 'a set"
   apply default
   apply (unfold set_plus_def)
   apply (force simp add: add_assoc)
   done
 
-class_interpretation set_semigroup_mult: semigroup_mult ["op \<otimes> :: ('a::semigroup_mult) set => 'a set => 'a set"]
+interpretation set_semigroup_mult!: semigroup_mult "op \<otimes> :: ('a::semigroup_mult) set => 'a set => 'a set"
   apply default
   apply (unfold set_times_def)
   apply (force simp add: mult_assoc)
   done
 
-class_interpretation set_comm_monoid_add: comm_monoid_add ["{0}" "op \<oplus> :: ('a::comm_monoid_add) set => 'a set => 'a set"]
+interpretation set_comm_monoid_add!: comm_monoid_add "{0}" "op \<oplus> :: ('a::comm_monoid_add) set => 'a set => 'a set"
   apply default
    apply (unfold set_plus_def)
    apply (force simp add: add_ac)
   apply force
   done
 
-class_interpretation set_comm_monoid_mult: comm_monoid_mult ["{1}" "op \<otimes> :: ('a::comm_monoid_mult) set => 'a set => 'a set"]
+interpretation set_comm_monoid_mult!: comm_monoid_mult "{1}" "op \<otimes> :: ('a::comm_monoid_mult) set => 'a set => 'a set"
   apply default
    apply (unfold set_times_def)
    apply (force simp add: mult_ac)
--- a/src/HOL/Library/Subarray.thy	Mon Jan 19 20:37:08 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,66 +0,0 @@
-theory Subarray
-imports Array Sublist
-begin
-
-definition subarray :: "nat \<Rightarrow> nat \<Rightarrow> ('a::heap) array \<Rightarrow> heap \<Rightarrow> 'a list"
-where
-  "subarray n m a h \<equiv> sublist' n m (get_array a h)"
-
-lemma subarray_upd: "i \<ge> m \<Longrightarrow> subarray n m a (Heap.upd a i v h) = subarray n m a h"
-apply (simp add: subarray_def Heap.upd_def)
-apply (simp add: sublist'_update1)
-done
-
-lemma subarray_upd2: " i < n  \<Longrightarrow> subarray n m a (Heap.upd a i v h) = subarray n m a h"
-apply (simp add: subarray_def Heap.upd_def)
-apply (subst sublist'_update2)
-apply fastsimp
-apply simp
-done
-
-lemma subarray_upd3: "\<lbrakk> n \<le> i; i < m\<rbrakk> \<Longrightarrow> subarray n m a (Heap.upd a i v h) = subarray n m a h[i - n := v]"
-unfolding subarray_def Heap.upd_def
-by (simp add: sublist'_update3)
-
-lemma subarray_Nil: "n \<ge> m \<Longrightarrow> subarray n m a h = []"
-by (simp add: subarray_def sublist'_Nil')
-
-lemma subarray_single: "\<lbrakk> n < Heap.length a h \<rbrakk> \<Longrightarrow> subarray n (Suc n) a h = [get_array a h ! n]" 
-by (simp add: subarray_def Heap.length_def sublist'_single)
-
-lemma length_subarray: "m \<le> Heap.length a h \<Longrightarrow> List.length (subarray n m a h) = m - n"
-by (simp add: subarray_def Heap.length_def length_sublist')
-
-lemma length_subarray_0: "m \<le> Heap.length a h \<Longrightarrow> List.length (subarray 0 m a h) = m"
-by (simp add: length_subarray)
-
-lemma subarray_nth_array_Cons: "\<lbrakk> i < Heap.length a h; i < j \<rbrakk> \<Longrightarrow> (get_array a h ! i) # subarray (Suc i) j a h = subarray i j a h"
-unfolding Heap.length_def subarray_def
-by (simp add: sublist'_front)
-
-lemma subarray_nth_array_back: "\<lbrakk> i < j; j \<le> Heap.length a h\<rbrakk> \<Longrightarrow> subarray i j a h = subarray i (j - 1) a h @ [get_array a h ! (j - 1)]"
-unfolding Heap.length_def subarray_def
-by (simp add: sublist'_back)
-
-lemma subarray_append: "\<lbrakk> i < j; j < k \<rbrakk> \<Longrightarrow> subarray i j a h @ subarray j k a h = subarray i k a h"
-unfolding subarray_def
-by (simp add: sublist'_append)
-
-lemma subarray_all: "subarray 0 (Heap.length a h) a h = get_array a h"
-unfolding Heap.length_def subarray_def
-by (simp add: sublist'_all)
-
-lemma nth_subarray: "\<lbrakk> k < j - i; j \<le> Heap.length a h \<rbrakk> \<Longrightarrow> subarray i j a h ! k = get_array a h ! (i + k)"
-unfolding Heap.length_def subarray_def
-by (simp add: nth_sublist')
-
-lemma subarray_eq_samelength_iff: "Heap.length a h = Heap.length a h' \<Longrightarrow> (subarray i j a h = subarray i j a h') = (\<forall>i'. i \<le> i' \<and> i' < j \<longrightarrow> get_array a h ! i' = get_array a h' ! i')"
-unfolding Heap.length_def subarray_def by (rule sublist'_eq_samelength_iff)
-
-lemma all_in_set_subarray_conv: "(\<forall>j. j \<in> set (subarray l r a h) \<longrightarrow> P j) = (\<forall>k. l \<le> k \<and> k < r \<and> k < Heap.length a h \<longrightarrow> P (get_array a h ! k))"
-unfolding subarray_def Heap.length_def by (rule all_in_set_sublist'_conv)
-
-lemma ball_in_set_subarray_conv: "(\<forall>j \<in> set (subarray l r a h). P j) = (\<forall>k. l \<le> k \<and> k < r \<and> k < Heap.length a h \<longrightarrow> P (get_array a h ! k))"
-unfolding subarray_def Heap.length_def by (rule ball_in_set_sublist'_conv)
-
-end
\ No newline at end of file
--- a/src/HOL/Library/Sublist.thy	Mon Jan 19 20:37:08 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,507 +0,0 @@
-(* $Id$ *)
-
-header {* Slices of lists *}
-
-theory Sublist
-imports Multiset
-begin
-
-
-lemma sublist_split: "i \<le> j \<and> j \<le> k \<Longrightarrow> sublist xs {i..<j} @ sublist xs {j..<k} = sublist xs {i..<k}" 
-apply (induct xs arbitrary: i j k)
-apply simp
-apply (simp only: sublist_Cons)
-apply simp
-apply safe
-apply simp
-apply (erule_tac x="0" in meta_allE)
-apply (erule_tac x="j - 1" in meta_allE)
-apply (erule_tac x="k - 1" in meta_allE)
-apply (subgoal_tac "0 \<le> j - 1 \<and> j - 1 \<le> k - 1")
-apply simp
-apply (subgoal_tac "{ja. Suc ja < j} = {0..<j - Suc 0}")
-apply (subgoal_tac "{ja. j \<le> Suc ja \<and> Suc ja < k} = {j - Suc 0..<k - Suc 0}")
-apply (subgoal_tac "{j. Suc j < k} = {0..<k - Suc 0}")
-apply simp
-apply fastsimp
-apply fastsimp
-apply fastsimp
-apply fastsimp
-apply (erule_tac x="i - 1" in meta_allE)
-apply (erule_tac x="j - 1" in meta_allE)
-apply (erule_tac x="k - 1" in meta_allE)
-apply (subgoal_tac " {ja. i \<le> Suc ja \<and> Suc ja < j} = {i - 1 ..<j - 1}")
-apply (subgoal_tac " {ja. j \<le> Suc ja \<and> Suc ja < k} = {j - 1..<k - 1}")
-apply (subgoal_tac "{j. i \<le> Suc j \<and> Suc j < k} = {i - 1..<k - 1}")
-apply (subgoal_tac " i - 1 \<le> j - 1 \<and> j - 1 \<le> k - 1")
-apply simp
-apply fastsimp
-apply fastsimp
-apply fastsimp
-apply fastsimp
-done
-
-lemma sublist_update1: "i \<notin> inds \<Longrightarrow> sublist (xs[i := v]) inds = sublist xs inds"
-apply (induct xs arbitrary: i inds)
-apply simp
-apply (case_tac i)
-apply (simp add: sublist_Cons)
-apply (simp add: sublist_Cons)
-done
-
-lemma sublist_update2: "i \<in> inds \<Longrightarrow> sublist (xs[i := v]) inds = (sublist xs inds)[(card {k \<in> inds. k < i}):= v]"
-proof (induct xs arbitrary: i inds)
-  case Nil thus ?case by simp
-next
-  case (Cons x xs)
-  thus ?case
-  proof (cases i)
-    case 0 with Cons show ?thesis by (simp add: sublist_Cons)
-  next
-    case (Suc i')
-    with Cons show ?thesis
-      apply simp
-      apply (simp add: sublist_Cons)
-      apply auto
-      apply (auto simp add: nat.split)
-      apply (simp add: card_less)
-      apply (simp add: card_less)
-      apply (simp add: card_less_Suc[symmetric])
-      apply (simp add: card_less_Suc2)
-      done
-  qed
-qed
-
-lemma sublist_update: "sublist (xs[i := v]) inds = (if i \<in> inds then (sublist xs inds)[(card {k \<in> inds. k < i}) := v] else sublist xs inds)"
-by (simp add: sublist_update1 sublist_update2)
-
-lemma sublist_take: "sublist xs {j. j < m} = take m xs"
-apply (induct xs arbitrary: m)
-apply simp
-apply (case_tac m)
-apply simp
-apply (simp add: sublist_Cons)
-done
-
-lemma sublist_take': "sublist xs {0..<m} = take m xs"
-apply (induct xs arbitrary: m)
-apply simp
-apply (case_tac m)
-apply simp
-apply (simp add: sublist_Cons sublist_take)
-done
-
-lemma sublist_all[simp]: "sublist xs {j. j < length xs} = xs"
-apply (induct xs)
-apply simp
-apply (simp add: sublist_Cons)
-done
-
-lemma sublist_all'[simp]: "sublist xs {0..<length xs} = xs"
-apply (induct xs)
-apply simp
-apply (simp add: sublist_Cons)
-done
-
-lemma sublist_single: "a < length xs \<Longrightarrow> sublist xs {a} = [xs ! a]"
-apply (induct xs arbitrary: a)
-apply simp
-apply(case_tac aa)
-apply simp
-apply (simp add: sublist_Cons)
-apply simp
-apply (simp add: sublist_Cons)
-done
-
-lemma sublist_is_Nil: "\<forall>i \<in> inds. i \<ge> length xs \<Longrightarrow> sublist xs inds = []" 
-apply (induct xs arbitrary: inds)
-apply simp
-apply (simp add: sublist_Cons)
-apply auto
-apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
-apply auto
-done
-
-lemma sublist_Nil': "sublist xs inds = [] \<Longrightarrow> \<forall>i \<in> inds. i \<ge> length xs"
-apply (induct xs arbitrary: inds)
-apply simp
-apply (simp add: sublist_Cons)
-apply (auto split: if_splits)
-apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
-apply (case_tac x, auto)
-done
-
-lemma sublist_Nil[simp]: "(sublist xs inds = []) = (\<forall>i \<in> inds. i \<ge> length xs)"
-apply (induct xs arbitrary: inds)
-apply simp
-apply (simp add: sublist_Cons)
-apply auto
-apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
-apply (case_tac x, auto)
-done
-
-lemma sublist_eq_subseteq: " \<lbrakk> inds' \<subseteq> inds; sublist xs inds = sublist ys inds \<rbrakk> \<Longrightarrow> sublist xs inds' = sublist ys inds'"
-apply (induct xs arbitrary: ys inds inds')
-apply simp
-apply (drule sym, rule sym)
-apply (simp add: sublist_Nil, fastsimp)
-apply (case_tac ys)
-apply (simp add: sublist_Nil, fastsimp)
-apply (auto simp add: sublist_Cons)
-apply (erule_tac x="list" in meta_allE)
-apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
-apply (erule_tac x="{j. Suc j \<in> inds'}" in meta_allE)
-apply fastsimp
-apply (erule_tac x="list" in meta_allE)
-apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
-apply (erule_tac x="{j. Suc j \<in> inds'}" in meta_allE)
-apply fastsimp
-done
-
-lemma sublist_eq: "\<lbrakk> \<forall>i \<in> inds. ((i < length xs) \<and> (i < length ys)) \<or> ((i \<ge> length xs ) \<and> (i \<ge> length ys)); \<forall>i \<in> inds. xs ! i = ys ! i \<rbrakk> \<Longrightarrow> sublist xs inds = sublist ys inds"
-apply (induct xs arbitrary: ys inds)
-apply simp
-apply (rule sym, simp add: sublist_Nil)
-apply (case_tac ys)
-apply (simp add: sublist_Nil)
-apply (auto simp add: sublist_Cons)
-apply (erule_tac x="list" in meta_allE)
-apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
-apply fastsimp
-apply (erule_tac x="list" in meta_allE)
-apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
-apply fastsimp
-done
-
-lemma sublist_eq_samelength: "\<lbrakk> length xs = length ys; \<forall>i \<in> inds. xs ! i = ys ! i \<rbrakk> \<Longrightarrow> sublist xs inds = sublist ys inds"
-by (rule sublist_eq, auto)
-
-lemma sublist_eq_samelength_iff: "length xs = length ys \<Longrightarrow> (sublist xs inds = sublist ys inds) = (\<forall>i \<in> inds. xs ! i = ys ! i)"
-apply (induct xs arbitrary: ys inds)
-apply simp
-apply (rule sym, simp add: sublist_Nil)
-apply (case_tac ys)
-apply (simp add: sublist_Nil)
-apply (auto simp add: sublist_Cons)
-apply (case_tac i)
-apply auto
-apply (case_tac i)
-apply auto
-done
-
-section {* Another sublist function *}
-
-function sublist' :: "nat \<Rightarrow> nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
-where
-  "sublist' n m [] = []"
-| "sublist' n 0 xs = []"
-| "sublist' 0 (Suc m) (x#xs) = (x#sublist' 0 m xs)"
-| "sublist' (Suc n) (Suc m) (x#xs) = sublist' n m xs"
-by pat_completeness auto
-termination by lexicographic_order
-
-subsection {* Proving equivalence to the other sublist command *}
-
-lemma sublist'_sublist: "sublist' n m xs = sublist xs {j. n \<le> j \<and> j < m}"
-apply (induct xs arbitrary: n m)
-apply simp
-apply (case_tac n)
-apply (case_tac m)
-apply simp
-apply (simp add: sublist_Cons)
-apply (case_tac m)
-apply simp
-apply (simp add: sublist_Cons)
-done
-
-
-lemma "sublist' n m xs = sublist xs {n..<m}"
-apply (induct xs arbitrary: n m)
-apply simp
-apply (case_tac n, case_tac m)
-apply simp
-apply simp
-apply (simp add: sublist_take')
-apply (case_tac m)
-apply simp
-apply (simp add: sublist_Cons sublist'_sublist)
-done
-
-
-subsection {* Showing equivalence to use of drop and take for definition *}
-
-lemma "sublist' n m xs = take (m - n) (drop n xs)"
-apply (induct xs arbitrary: n m)
-apply simp
-apply (case_tac m)
-apply simp
-apply (case_tac n)
-apply simp
-apply simp
-done
-
-subsection {* General lemma about sublist *}
-
-lemma sublist'_Nil[simp]: "sublist' i j [] = []"
-by simp
-
-lemma sublist'_Cons[simp]: "sublist' i (Suc j) (x#xs) = (case i of 0 \<Rightarrow> (x # sublist' 0 j xs) | Suc i' \<Rightarrow>  sublist' i' j xs)"
-by (cases i) auto
-
-lemma sublist'_Cons2[simp]: "sublist' i j (x#xs) = (if (j = 0) then [] else ((if (i = 0) then [x] else []) @ sublist' (i - 1) (j - 1) xs))"
-apply (cases j)
-apply auto
-apply (cases i)
-apply auto
-done
-
-lemma sublist_n_0: "sublist' n 0 xs = []"
-by (induct xs, auto)
-
-lemma sublist'_Nil': "n \<ge> m \<Longrightarrow> sublist' n m xs = []"
-apply (induct xs arbitrary: n m)
-apply simp
-apply (case_tac m)
-apply simp
-apply (case_tac n)
-apply simp
-apply simp
-done
-
-lemma sublist'_Nil2: "n \<ge> length xs \<Longrightarrow> sublist' n m xs = []"
-apply (induct xs arbitrary: n m)
-apply simp
-apply (case_tac m)
-apply simp
-apply (case_tac n)
-apply simp
-apply simp
-done
-
-lemma sublist'_Nil3: "(sublist' n m xs = []) = ((n \<ge> m) \<or> (n \<ge> length xs))"
-apply (induct xs arbitrary: n m)
-apply simp
-apply (case_tac m)
-apply simp
-apply (case_tac n)
-apply simp
-apply simp
-done
-
-lemma sublist'_notNil: "\<lbrakk> n < length xs; n < m \<rbrakk> \<Longrightarrow> sublist' n m xs \<noteq> []"
-apply (induct xs arbitrary: n m)
-apply simp
-apply (case_tac m)
-apply simp
-apply (case_tac n)
-apply simp
-apply simp
-done
-
-lemma sublist'_single: "n < length xs \<Longrightarrow> sublist' n (Suc n) xs = [xs ! n]"
-apply (induct xs arbitrary: n)
-apply simp
-apply simp
-apply (case_tac n)
-apply (simp add: sublist_n_0)
-apply simp
-done
-
-lemma sublist'_update1: "i \<ge> m \<Longrightarrow> sublist' n m (xs[i:=v]) = sublist' n m xs"
-apply (induct xs arbitrary: n m i)
-apply simp
-apply simp
-apply (case_tac i)
-apply simp
-apply simp
-done
-
-lemma sublist'_update2: "i < n \<Longrightarrow> sublist' n m (xs[i:=v]) = sublist' n m xs"
-apply (induct xs arbitrary: n m i)
-apply simp
-apply simp
-apply (case_tac i)
-apply simp
-apply simp
-done
-
-lemma sublist'_update3: "\<lbrakk>n \<le> i; i < m\<rbrakk> \<Longrightarrow> sublist' n m (xs[i := v]) = (sublist' n m xs)[i - n := v]"
-proof (induct xs arbitrary: n m i)
-  case Nil thus ?case by auto
-next
-  case (Cons x xs)
-  thus ?case
-    apply -
-    apply auto
-    apply (cases i)
-    apply auto
-    apply (cases i)
-    apply auto
-    done
-qed
-
-lemma "\<lbrakk> sublist' i j xs = sublist' i j ys; n \<ge> i; m \<le> j \<rbrakk> \<Longrightarrow> sublist' n m xs = sublist' n m ys"
-proof (induct xs arbitrary: i j ys n m)
-  case Nil
-  thus ?case
-    apply -
-    apply (rule sym, drule sym)
-    apply (simp add: sublist'_Nil)
-    apply (simp add: sublist'_Nil3)
-    apply arith
-    done
-next
-  case (Cons x xs i j ys n m)
-  note c = this
-  thus ?case
-  proof (cases m)
-    case 0 thus ?thesis by (simp add: sublist_n_0)
-  next
-    case (Suc m')
-    note a = this
-    thus ?thesis
-    proof (cases n)
-      case 0 note b = this
-      show ?thesis
-      proof (cases ys)
-	case Nil  with a b Cons.prems show ?thesis by (simp add: sublist'_Nil3)
-      next
-	case (Cons y ys)
-	show ?thesis
-	proof (cases j)
-	  case 0 with a b Cons.prems show ?thesis by simp
-	next
-	  case (Suc j') with a b Cons.prems Cons show ?thesis 
-	    apply -
-	    apply (simp, rule Cons.hyps [of "0" "j'" "ys" "0" "m'"], auto)
-	    done
-	qed
-      qed
-    next
-      case (Suc n')
-      show ?thesis
-      proof (cases ys)
-	case Nil with Suc a Cons.prems show ?thesis by (auto simp add: sublist'_Nil3)
-      next
-	case (Cons y ys) with Suc a Cons.prems show ?thesis
-	  apply -
-	  apply simp
-	  apply (cases j)
-	  apply simp
-	  apply (cases i)
-	  apply simp
-	  apply (rule_tac j="nat" in Cons.hyps [of "0" _ "ys" "n'" "m'"])
-	  apply simp
-	  apply simp
-	  apply simp
-	  apply simp
-	  apply (rule_tac i="nata" and j="nat" in Cons.hyps [of _ _ "ys" "n'" "m'"])
-	  apply simp
-	  apply simp
-	  apply simp
-	  done
-      qed
-    qed
-  qed
-qed
-
-lemma length_sublist': "j \<le> length xs \<Longrightarrow> length (sublist' i j xs) = j - i"
-by (induct xs arbitrary: i j, auto)
-
-lemma sublist'_front: "\<lbrakk> i < j; i < length xs \<rbrakk> \<Longrightarrow> sublist' i j xs = xs ! i # sublist' (Suc i) j xs"
-apply (induct xs arbitrary: a i j)
-apply simp
-apply (case_tac j)
-apply simp
-apply (case_tac i)
-apply simp
-apply simp
-done
-
-lemma sublist'_back: "\<lbrakk> i < j; j \<le> length xs \<rbrakk> \<Longrightarrow> sublist' i j xs = sublist' i (j - 1) xs @ [xs ! (j - 1)]"
-apply (induct xs arbitrary: a i j)
-apply simp
-apply simp
-apply (case_tac j)
-apply simp
-apply auto
-apply (case_tac nat)
-apply auto
-done
-
-(* suffices that j \<le> length xs and length ys *) 
-lemma sublist'_eq_samelength_iff: "length xs = length ys \<Longrightarrow> (sublist' i j xs  = sublist' i j ys) = (\<forall>i'. i \<le> i' \<and> i' < j \<longrightarrow> xs ! i' = ys ! i')"
-proof (induct xs arbitrary: ys i j)
-  case Nil thus ?case by simp
-next
-  case (Cons x xs)
-  thus ?case
-    apply -
-    apply (cases ys)
-    apply simp
-    apply simp
-    apply auto
-    apply (case_tac i', auto)
-    apply (erule_tac x="Suc i'" in allE, auto)
-    apply (erule_tac x="i' - 1" in allE, auto)
-    apply (case_tac i', auto)
-    apply (erule_tac x="Suc i'" in allE, auto)
-    done
-qed
-
-lemma sublist'_all[simp]: "sublist' 0 (length xs) xs = xs"
-by (induct xs, auto)
-
-lemma sublist'_sublist': "sublist' n m (sublist' i j xs) = sublist' (i + n) (min (i + m) j) xs" 
-by (induct xs arbitrary: i j n m) (auto simp add: min_diff)
-
-lemma sublist'_append: "\<lbrakk> i \<le> j; j \<le> k \<rbrakk> \<Longrightarrow>(sublist' i j xs) @ (sublist' j k xs) = sublist' i k xs"
-by (induct xs arbitrary: i j k) auto
-
-lemma nth_sublist': "\<lbrakk> k < j - i; j \<le> length xs \<rbrakk> \<Longrightarrow> (sublist' i j xs) ! k = xs ! (i + k)"
-apply (induct xs arbitrary: i j k)
-apply auto
-apply (case_tac k)
-apply auto
-apply (case_tac i)
-apply auto
-done
-
-lemma set_sublist': "set (sublist' i j xs) = {x. \<exists>k. i \<le> k \<and> k < j \<and> k < List.length xs \<and> x = xs ! k}"
-apply (simp add: sublist'_sublist)
-apply (simp add: set_sublist)
-apply auto
-done
-
-lemma all_in_set_sublist'_conv: "(\<forall>j. j \<in> set (sublist' l r xs) \<longrightarrow> P j) = (\<forall>k. l \<le> k \<and> k < r \<and> k < List.length xs \<longrightarrow> P (xs ! k))"
-unfolding set_sublist' by blast
-
-lemma ball_in_set_sublist'_conv: "(\<forall>j \<in> set (sublist' l r xs). P j) = (\<forall>k. l \<le> k \<and> k < r \<and> k < List.length xs \<longrightarrow> P (xs ! k))"
-unfolding set_sublist' by blast
-
-
-lemma multiset_of_sublist:
-assumes l_r: "l \<le> r \<and> r \<le> List.length xs"
-assumes left: "\<forall> i. i < l \<longrightarrow> (xs::'a list) ! i = ys ! i"
-assumes right: "\<forall> i. i \<ge> r \<longrightarrow> (xs::'a list) ! i = ys ! i"
-assumes multiset: "multiset_of xs = multiset_of ys"
-  shows "multiset_of (sublist' l r xs) = multiset_of (sublist' l r ys)"
-proof -
-  from l_r have xs_def: "xs = (sublist' 0 l xs) @ (sublist' l r xs) @ (sublist' r (List.length xs) xs)" (is "_ = ?xs_long") 
-    by (simp add: sublist'_append)
-  from multiset have length_eq: "List.length xs = List.length ys" by (rule multiset_of_eq_length)
-  with l_r have ys_def: "ys = (sublist' 0 l ys) @ (sublist' l r ys) @ (sublist' r (List.length ys) ys)" (is "_ = ?ys_long") 
-    by (simp add: sublist'_append)
-  from xs_def ys_def multiset have "multiset_of ?xs_long = multiset_of ?ys_long" by simp
-  moreover
-  from left l_r length_eq have "sublist' 0 l xs = sublist' 0 l ys"
-    by (auto simp add: length_sublist' nth_sublist' intro!: nth_equalityI)
-  moreover
-  from right l_r length_eq have "sublist' r (List.length xs) xs = sublist' r (List.length ys) ys"
-    by (auto simp add: length_sublist' nth_sublist' intro!: nth_equalityI)
-  moreover
-  ultimately show ?thesis by (simp add: multiset_of_append)
-qed
-
-
-end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Univ_Poly.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -0,0 +1,1050 @@
+(*  Title:       Univ_Poly.thy
+    Author:      Amine Chaieb
+*)
+
+header {* Univariate Polynomials *}
+
+theory Univ_Poly
+imports Plain List
+begin
+
+text{* Application of polynomial as a function. *}
+
+primrec (in semiring_0) poly :: "'a list => 'a  => 'a" where
+  poly_Nil:  "poly [] x = 0"
+| poly_Cons: "poly (h#t) x = h + x * poly t x"
+
+
+subsection{*Arithmetic Operations on Polynomials*}
+
+text{*addition*}
+
+primrec (in semiring_0) padd :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"  (infixl "+++" 65) 
+where
+  padd_Nil:  "[] +++ l2 = l2"
+| padd_Cons: "(h#t) +++ l2 = (if l2 = [] then h#t
+                            else (h + hd l2)#(t +++ tl l2))"
+
+text{*Multiplication by a constant*}
+primrec (in semiring_0) cmult :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list"  (infixl "%*" 70) where
+   cmult_Nil:  "c %* [] = []"
+|  cmult_Cons: "c %* (h#t) = (c * h)#(c %* t)"
+
+text{*Multiplication by a polynomial*}
+primrec (in semiring_0) pmult :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"  (infixl "***" 70)
+where
+   pmult_Nil:  "[] *** l2 = []"
+|  pmult_Cons: "(h#t) *** l2 = (if t = [] then h %* l2
+                              else (h %* l2) +++ ((0) # (t *** l2)))"
+
+text{*Repeated multiplication by a polynomial*}
+primrec (in semiring_0) mulexp :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a  list \<Rightarrow> 'a list" where
+   mulexp_zero:  "mulexp 0 p q = q"
+|  mulexp_Suc:   "mulexp (Suc n) p q = p *** mulexp n p q"
+
+text{*Exponential*}
+primrec (in semiring_1) pexp :: "'a list \<Rightarrow> nat \<Rightarrow> 'a list"  (infixl "%^" 80) where
+   pexp_0:   "p %^ 0 = [1]"
+|  pexp_Suc: "p %^ (Suc n) = p *** (p %^ n)"
+
+text{*Quotient related value of dividing a polynomial by x + a*}
+(* Useful for divisor properties in inductive proofs *)
+primrec (in field) "pquot" :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a list" where
+   pquot_Nil:  "pquot [] a= []"
+|  pquot_Cons: "pquot (h#t) a = (if t = [] then [h]
+                   else (inverse(a) * (h - hd( pquot t a)))#(pquot t a))"
+
+text{*normalization of polynomials (remove extra 0 coeff)*}
+primrec (in semiring_0) pnormalize :: "'a list \<Rightarrow> 'a list" where
+  pnormalize_Nil:  "pnormalize [] = []"
+| pnormalize_Cons: "pnormalize (h#p) = (if ( (pnormalize p) = [])
+                                     then (if (h = 0) then [] else [h])
+                                     else (h#(pnormalize p)))"
+
+definition (in semiring_0) "pnormal p = ((pnormalize p = p) \<and> p \<noteq> [])"
+definition (in semiring_0) "nonconstant p = (pnormal p \<and> (\<forall>x. p \<noteq> [x]))"
+text{*Other definitions*}
+
+definition (in ring_1)
+  poly_minus :: "'a list => 'a list" ("-- _" [80] 80) where
+  "-- p = (- 1) %* p"
+
+definition (in semiring_0)
+  divides :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"  (infixl "divides" 70) where
+  [code del]: "p1 divides p2 = (\<exists>q. poly p2 = poly(p1 *** q))"
+
+    --{*order of a polynomial*}
+definition (in ring_1) order :: "'a => 'a list => nat" where
+  "order a p = (SOME n. ([-a, 1] %^ n) divides p &
+                      ~ (([-a, 1] %^ (Suc n)) divides p))"
+
+     --{*degree of a polynomial*}
+definition (in semiring_0) degree :: "'a list => nat" where 
+  "degree p = length (pnormalize p) - 1"
+
+     --{*squarefree polynomials --- NB with respect to real roots only.*}
+definition (in ring_1)
+  rsquarefree :: "'a list => bool" where
+  "rsquarefree p = (poly p \<noteq> poly [] &
+                     (\<forall>a. (order a p = 0) | (order a p = 1)))"
+
+context semiring_0
+begin
+
+lemma padd_Nil2[simp]: "p +++ [] = p"
+by (induct p) auto
+
+lemma padd_Cons_Cons: "(h1 # p1) +++ (h2 # p2) = (h1 + h2) # (p1 +++ p2)"
+by auto
+
+lemma pminus_Nil[simp]: "-- [] = []"
+by (simp add: poly_minus_def)
+
+lemma pmult_singleton: "[h1] *** p1 = h1 %* p1" by simp
+end
+
+lemma (in semiring_1) poly_ident_mult[simp]: "1 %* t = t" by (induct "t", auto)
+
+lemma (in semiring_0) poly_simple_add_Cons[simp]: "[a] +++ ((0)#t) = (a#t)"
+by simp
+
+text{*Handy general properties*}
+
+lemma (in comm_semiring_0) padd_commut: "b +++ a = a +++ b"
+proof(induct b arbitrary: a)
+  case Nil thus ?case by auto
+next
+  case (Cons b bs a) thus ?case by (cases a, simp_all add: add_commute)
+qed
+
+lemma (in comm_semiring_0) padd_assoc: "\<forall>b c. (a +++ b) +++ c = a +++ (b +++ c)"
+apply (induct a arbitrary: b c)
+apply (simp, clarify)
+apply (case_tac b, simp_all add: add_ac)
+done
+
+lemma (in semiring_0) poly_cmult_distr: "a %* ( p +++ q) = (a %* p +++ a %* q)"
+apply (induct p arbitrary: q,simp)
+apply (case_tac q, simp_all add: right_distrib)
+done
+
+lemma (in ring_1) pmult_by_x[simp]: "[0, 1] *** t = ((0)#t)"
+apply (induct "t", simp)
+apply (auto simp add: mult_zero_left poly_ident_mult padd_commut)
+apply (case_tac t, auto)
+done
+
+text{*properties of evaluation of polynomials.*}
+
+lemma (in semiring_0) poly_add: "poly (p1 +++ p2) x = poly p1 x + poly p2 x"
+proof(induct p1 arbitrary: p2)
+  case Nil thus ?case by simp
+next
+  case (Cons a as p2) thus ?case 
+    by (cases p2, simp_all  add: add_ac right_distrib)
+qed
+
+lemma (in comm_semiring_0) poly_cmult: "poly (c %* p) x = c * poly p x"
+apply (induct "p") 
+apply (case_tac [2] "x=zero")
+apply (auto simp add: right_distrib mult_ac)
+done
+
+lemma (in comm_semiring_0) poly_cmult_map: "poly (map (op * c) p) x = c*poly p x"
+  by (induct p, auto simp add: right_distrib mult_ac)
+
+lemma (in comm_ring_1) poly_minus: "poly (-- p) x = - (poly p x)"
+apply (simp add: poly_minus_def)
+apply (auto simp add: poly_cmult minus_mult_left[symmetric])
+done
+
+lemma (in comm_semiring_0) poly_mult: "poly (p1 *** p2) x = poly p1 x * poly p2 x"
+proof(induct p1 arbitrary: p2)
+  case Nil thus ?case by simp
+next
+  case (Cons a as p2)
+  thus ?case by (cases as, 
+    simp_all add: poly_cmult poly_add left_distrib right_distrib mult_ac)
+qed
+
+class recpower_semiring = semiring + recpower
+class recpower_semiring_1 = semiring_1 + recpower
+class recpower_semiring_0 = semiring_0 + recpower
+class recpower_ring = ring + recpower
+class recpower_ring_1 = ring_1 + recpower
+subclass (in recpower_ring_1) recpower_ring ..
+class recpower_comm_semiring_1 = recpower + comm_semiring_1
+class recpower_comm_ring_1 = recpower + comm_ring_1
+subclass (in recpower_comm_ring_1) recpower_comm_semiring_1 ..
+class recpower_idom = recpower + idom
+subclass (in recpower_idom) recpower_comm_ring_1 ..
+class idom_char_0 = idom + ring_char_0
+class recpower_idom_char_0 = recpower + idom_char_0
+subclass (in recpower_idom_char_0) recpower_idom ..
+
+lemma (in recpower_comm_ring_1) poly_exp: "poly (p %^ n) x = (poly p x) ^ n"
+apply (induct "n")
+apply (auto simp add: poly_cmult poly_mult power_Suc)
+done
+
+text{*More Polynomial Evaluation Lemmas*}
+
+lemma  (in semiring_0) poly_add_rzero[simp]: "poly (a +++ []) x = poly a x"
+by simp
+
+lemma (in comm_semiring_0) poly_mult_assoc: "poly ((a *** b) *** c) x = poly (a *** (b *** c)) x"
+  by (simp add: poly_mult mult_assoc)
+
+lemma (in semiring_0) poly_mult_Nil2[simp]: "poly (p *** []) x = 0"
+by (induct "p", auto)
+
+lemma (in comm_semiring_1) poly_exp_add: "poly (p %^ (n + d)) x = poly( p %^ n *** p %^ d) x"
+apply (induct "n")
+apply (auto simp add: poly_mult mult_assoc)
+done
+
+subsection{*Key Property: if @{term "f(a) = 0"} then @{term "(x - a)"} divides
+ @{term "p(x)"} *}
+
+lemma (in comm_ring_1) lemma_poly_linear_rem: "\<forall>h. \<exists>q r. h#t = [r] +++ [-a, 1] *** q"
+proof(induct t)
+  case Nil
+  {fix h have "[h] = [h] +++ [- a, 1] *** []" by simp}
+  thus ?case by blast
+next
+  case (Cons  x xs)
+  {fix h 
+    from Cons.hyps[rule_format, of x] 
+    obtain q r where qr: "x#xs = [r] +++ [- a, 1] *** q" by blast
+    have "h#x#xs = [a*r + h] +++ [-a, 1] *** (r#q)" 
+      using qr by(cases q, simp_all add: ring_simps diff_def[symmetric] 
+	minus_mult_left[symmetric] right_minus)
+    hence "\<exists>q r. h#x#xs = [r] +++ [-a, 1] *** q" by blast}
+  thus ?case by blast
+qed
+
+lemma (in comm_ring_1) poly_linear_rem: "\<exists>q r. h#t = [r] +++ [-a, 1] *** q"
+by (cut_tac t = t and a = a in lemma_poly_linear_rem, auto)
+
+
+lemma (in comm_ring_1) poly_linear_divides: "(poly p a = 0) = ((p = []) | (\<exists>q. p = [-a, 1] *** q))"
+proof-
+  {assume p: "p = []" hence ?thesis by simp}
+  moreover
+  {fix x xs assume p: "p = x#xs"
+    {fix q assume "p = [-a, 1] *** q" hence "poly p a = 0" 
+	by (simp add: poly_add poly_cmult minus_mult_left[symmetric])}
+    moreover
+    {assume p0: "poly p a = 0"
+      from poly_linear_rem[of x xs a] obtain q r 
+      where qr: "x#xs = [r] +++ [- a, 1] *** q" by blast
+      have "r = 0" using p0 by (simp only: p qr poly_mult poly_add) simp
+      hence "\<exists>q. p = [- a, 1] *** q" using p qr  apply - apply (rule exI[where x=q])apply auto apply (cases q) apply auto done}
+    ultimately have ?thesis using p by blast}
+  ultimately show ?thesis by (cases p, auto)
+qed
+
+lemma (in semiring_0) lemma_poly_length_mult[simp]: "\<forall>h k a. length (k %* p +++  (h # (a %* p))) = Suc (length p)"
+by (induct "p", auto)
+
+lemma (in semiring_0) lemma_poly_length_mult2[simp]: "\<forall>h k. length (k %* p +++  (h # p)) = Suc (length p)"
+by (induct "p", auto)
+
+lemma (in ring_1) poly_length_mult[simp]: "length([-a,1] *** q) = Suc (length q)"
+by auto
+
+subsection{*Polynomial length*}
+
+lemma (in semiring_0) poly_cmult_length[simp]: "length (a %* p) = length p"
+by (induct "p", auto)
+
+lemma (in semiring_0) poly_add_length: "length (p1 +++ p2) = max (length p1) (length p2)"
+apply (induct p1 arbitrary: p2, simp_all)
+apply arith
+done
+
+lemma (in semiring_0) poly_root_mult_length[simp]: "length([a,b] *** p) = Suc (length p)"
+by (simp add: poly_add_length)
+
+lemma (in idom) poly_mult_not_eq_poly_Nil[simp]: 
+ "poly (p *** q) x \<noteq> poly [] x \<longleftrightarrow> poly p x \<noteq> poly [] x \<and> poly q x \<noteq> poly [] x"
+by (auto simp add: poly_mult)
+
+lemma (in idom) poly_mult_eq_zero_disj: "poly (p *** q) x = 0 \<longleftrightarrow> poly p x = 0 \<or> poly q x = 0"
+by (auto simp add: poly_mult)
+
+text{*Normalisation Properties*}
+
+lemma (in semiring_0) poly_normalized_nil: "(pnormalize p = []) --> (poly p x = 0)"
+by (induct "p", auto)
+
+text{*A nontrivial polynomial of degree n has no more than n roots*}
+lemma (in idom) poly_roots_index_lemma:
+   assumes p: "poly p x \<noteq> poly [] x" and n: "length p = n" 
+  shows "\<exists>i. \<forall>x. poly p x = 0 \<longrightarrow> (\<exists>m\<le>n. x = i m)"
+  using p n
+proof(induct n arbitrary: p x)
+  case 0 thus ?case by simp 
+next
+  case (Suc n p x)
+  {assume C: "\<And>i. \<exists>x. poly p x = 0 \<and> (\<forall>m\<le>Suc n. x \<noteq> i m)"
+    from Suc.prems have p0: "poly p x \<noteq> 0" "p\<noteq> []" by auto
+    from p0(1)[unfolded poly_linear_divides[of p x]] 
+    have "\<forall>q. p \<noteq> [- x, 1] *** q" by blast
+    from C obtain a where a: "poly p a = 0" by blast
+    from a[unfolded poly_linear_divides[of p a]] p0(2) 
+    obtain q where q: "p = [-a, 1] *** q" by blast
+    have lg: "length q = n" using q Suc.prems(2) by simp
+    from q p0 have qx: "poly q x \<noteq> poly [] x" 
+      by (auto simp add: poly_mult poly_add poly_cmult)
+    from Suc.hyps[OF qx lg] obtain i where 
+      i: "\<forall>x. poly q x = 0 \<longrightarrow> (\<exists>m\<le>n. x = i m)" by blast
+    let ?i = "\<lambda>m. if m = Suc n then a else i m"
+    from C[of ?i] obtain y where y: "poly p y = 0" "\<forall>m\<le> Suc n. y \<noteq> ?i m" 
+      by blast
+    from y have "y = a \<or> poly q y = 0" 
+      by (simp only: q poly_mult_eq_zero_disj poly_add) (simp add: ring_simps)
+    with i[rule_format, of y] y(1) y(2) have False apply auto 
+      apply (erule_tac x="m" in allE)
+      apply auto
+      done}
+  thus ?case by blast
+qed
+
+
+lemma (in idom) poly_roots_index_length: "poly p x \<noteq> poly [] x ==>
+      \<exists>i. \<forall>x. (poly p x = 0) --> (\<exists>n. n \<le> length p & x = i n)"
+by (blast intro: poly_roots_index_lemma)
+
+lemma (in idom) poly_roots_finite_lemma1: "poly p x \<noteq> poly [] x ==>
+      \<exists>N i. \<forall>x. (poly p x = 0) --> (\<exists>n. (n::nat) < N & x = i n)"
+apply (drule poly_roots_index_length, safe)
+apply (rule_tac x = "Suc (length p)" in exI)
+apply (rule_tac x = i in exI) 
+apply (simp add: less_Suc_eq_le)
+done
+
+
+lemma (in idom) idom_finite_lemma:
+  assumes P: "\<forall>x. P x --> (\<exists>n. n < length j & x = j!n)"
+  shows "finite {x. P x}"
+proof-
+  let ?M = "{x. P x}"
+  let ?N = "set j"
+  have "?M \<subseteq> ?N" using P by auto
+  thus ?thesis using finite_subset by auto
+qed
+
+
+lemma (in idom) poly_roots_finite_lemma2: "poly p x \<noteq> poly [] x ==>
+      \<exists>i. \<forall>x. (poly p x = 0) --> x \<in> set i"
+apply (drule poly_roots_index_length, safe)
+apply (rule_tac x="map (\<lambda>n. i n) [0 ..< Suc (length p)]" in exI)
+apply (auto simp add: image_iff)
+apply (erule_tac x="x" in allE, clarsimp)
+by (case_tac "n=length p", auto simp add: order_le_less)
+
+lemma UNIV_nat_infinite: "\<not> finite (UNIV :: nat set)"
+  unfolding finite_conv_nat_seg_image
+proof(auto simp add: expand_set_eq image_iff)
+  fix n::nat and f:: "nat \<Rightarrow> nat"
+  let ?N = "{i. i < n}"
+  let ?fN = "f ` ?N"
+  let ?y = "Max ?fN + 1"
+  from nat_seg_image_imp_finite[of "?fN" "f" n] 
+  have thfN: "finite ?fN" by simp
+  {assume "n =0" hence "\<exists>x. \<forall>xa<n. x \<noteq> f xa" by auto}
+  moreover
+  {assume nz: "n \<noteq> 0"
+    hence thne: "?fN \<noteq> {}" by (auto simp add: neq0_conv)
+    have "\<forall>x\<in> ?fN. Max ?fN \<ge> x" using nz Max_ge_iff[OF thfN thne] by auto
+    hence "\<forall>x\<in> ?fN. ?y > x" by auto
+    hence "?y \<notin> ?fN" by auto
+    hence "\<exists>x. \<forall>xa<n. x \<noteq> f xa" by auto }
+  ultimately show "\<exists>x. \<forall>xa<n. x \<noteq> f xa" by blast
+qed
+
+lemma (in ring_char_0) UNIV_ring_char_0_infinte: 
+  "\<not> (finite (UNIV:: 'a set))" 
+proof
+  assume F: "finite (UNIV :: 'a set)"
+  have "finite (UNIV :: nat set)"
+  proof (rule finite_imageD)
+    have "of_nat ` UNIV \<subseteq> UNIV" by simp
+    then show "finite (of_nat ` UNIV :: 'a set)" using F by (rule finite_subset)
+    show "inj (of_nat :: nat \<Rightarrow> 'a)" by (simp add: inj_on_def)
+  qed
+  with UNIV_nat_infinite show False ..
+qed
+
+lemma (in idom_char_0) poly_roots_finite: "(poly p \<noteq> poly []) = 
+  finite {x. poly p x = 0}"
+proof
+  assume H: "poly p \<noteq> poly []"
+  show "finite {x. poly p x = (0::'a)}"
+    using H
+    apply -
+    apply (erule contrapos_np, rule ext)
+    apply (rule ccontr)
+    apply (clarify dest!: poly_roots_finite_lemma2)
+    using finite_subset
+  proof-
+    fix x i
+    assume F: "\<not> finite {x. poly p x = (0\<Colon>'a)}" 
+      and P: "\<forall>x. poly p x = (0\<Colon>'a) \<longrightarrow> x \<in> set i"
+    let ?M= "{x. poly p x = (0\<Colon>'a)}"
+    from P have "?M \<subseteq> set i" by auto
+    with finite_subset F show False by auto
+  qed
+next
+  assume F: "finite {x. poly p x = (0\<Colon>'a)}"
+  show "poly p \<noteq> poly []" using F UNIV_ring_char_0_infinte by auto  
+qed
+
+text{*Entirety and Cancellation for polynomials*}
+
+lemma (in idom_char_0) poly_entire_lemma2: 
+  assumes p0: "poly p \<noteq> poly []" and q0: "poly q \<noteq> poly []"
+  shows "poly (p***q) \<noteq> poly []"
+proof-
+  let ?S = "\<lambda>p. {x. poly p x = 0}"
+  have "?S (p *** q) = ?S p \<union> ?S q" by (auto simp add: poly_mult)
+  with p0 q0 show ?thesis  unfolding poly_roots_finite by auto
+qed
+
+lemma (in idom_char_0) poly_entire: 
+  "poly (p *** q) = poly [] \<longleftrightarrow> poly p = poly [] \<or> poly q = poly []"
+using poly_entire_lemma2[of p q] 
+by auto (rule ext, simp add: poly_mult)+
+
+lemma (in idom_char_0) poly_entire_neg: "(poly (p *** q) \<noteq> poly []) = ((poly p \<noteq> poly []) & (poly q \<noteq> poly []))"
+by (simp add: poly_entire)
+
+lemma fun_eq: " (f = g) = (\<forall>x. f x = g x)"
+by (auto intro!: ext)
+
+lemma (in comm_ring_1) poly_add_minus_zero_iff: "(poly (p +++ -- q) = poly []) = (poly p = poly q)"
+by (auto simp add: ring_simps poly_add poly_minus_def fun_eq poly_cmult minus_mult_left[symmetric])
+
+lemma (in comm_ring_1) poly_add_minus_mult_eq: "poly (p *** q +++ --(p *** r)) = poly (p *** (q +++ -- r))"
+by (auto simp add: poly_add poly_minus_def fun_eq poly_mult poly_cmult right_distrib minus_mult_left[symmetric] minus_mult_right[symmetric])
+
+subclass (in idom_char_0) comm_ring_1 ..
+lemma (in idom_char_0) poly_mult_left_cancel: "(poly (p *** q) = poly (p *** r)) = (poly p = poly [] | poly q = poly r)"
+proof-
+  have "poly (p *** q) = poly (p *** r) \<longleftrightarrow> poly (p *** q +++ -- (p *** r)) = poly []" by (simp only: poly_add_minus_zero_iff)
+  also have "\<dots> \<longleftrightarrow> poly p = poly [] | poly q = poly r"
+    by (auto intro: ext simp add: poly_add_minus_mult_eq poly_entire poly_add_minus_zero_iff)
+  finally show ?thesis .
+qed
+
+lemma (in recpower_idom) poly_exp_eq_zero[simp]:
+     "(poly (p %^ n) = poly []) = (poly p = poly [] & n \<noteq> 0)"
+apply (simp only: fun_eq add: all_simps [symmetric]) 
+apply (rule arg_cong [where f = All]) 
+apply (rule ext)
+apply (induct n)
+apply (auto simp add: poly_exp poly_mult)
+done
+
+lemma (in semiring_1) one_neq_zero[simp]: "1 \<noteq> 0" using zero_neq_one by blast
+lemma (in comm_ring_1) poly_prime_eq_zero[simp]: "poly [a,1] \<noteq> poly []"
+apply (simp add: fun_eq)
+apply (rule_tac x = "minus one a" in exI)
+apply (unfold diff_minus)
+apply (subst add_commute)
+apply (subst add_assoc)
+apply simp
+done 
+
+lemma (in recpower_idom) poly_exp_prime_eq_zero: "(poly ([a, 1] %^ n) \<noteq> poly [])"
+by auto
+
+text{*A more constructive notion of polynomials being trivial*}
+
+lemma (in idom_char_0) poly_zero_lemma': "poly (h # t) = poly [] ==> h = 0 & poly t = poly []"
+apply(simp add: fun_eq)
+apply (case_tac "h = zero")
+apply (drule_tac [2] x = zero in spec, auto) 
+apply (cases "poly t = poly []", simp) 
+proof-
+  fix x
+  assume H: "\<forall>x. x = (0\<Colon>'a) \<or> poly t x = (0\<Colon>'a)"  and pnz: "poly t \<noteq> poly []"
+  let ?S = "{x. poly t x = 0}"
+  from H have "\<forall>x. x \<noteq>0 \<longrightarrow> poly t x = 0" by blast
+  hence th: "?S \<supseteq> UNIV - {0}" by auto
+  from poly_roots_finite pnz have th': "finite ?S" by blast
+  from finite_subset[OF th th'] UNIV_ring_char_0_infinte
+  show "poly t x = (0\<Colon>'a)" by simp
+  qed
+
+lemma (in idom_char_0) poly_zero: "(poly p = poly []) = list_all (%c. c = 0) p"
+apply (induct "p", simp)
+apply (rule iffI)
+apply (drule poly_zero_lemma', auto)
+done
+
+lemma (in idom_char_0) poly_0: "list_all (\<lambda>c. c = 0) p \<Longrightarrow> poly p x = 0"
+  unfolding poly_zero[symmetric] by simp
+
+
+
+text{*Basics of divisibility.*}
+
+lemma (in idom) poly_primes: "([a, 1] divides (p *** q)) = ([a, 1] divides p | [a, 1] divides q)"
+apply (auto simp add: divides_def fun_eq poly_mult poly_add poly_cmult left_distrib [symmetric])
+apply (drule_tac x = "uminus a" in spec)
+apply (simp add: poly_linear_divides poly_add poly_cmult left_distrib [symmetric])
+apply (cases "p = []")
+apply (rule exI[where x="[]"])
+apply simp
+apply (cases "q = []")
+apply (erule allE[where x="[]"], simp)
+
+apply clarsimp
+apply (cases "\<exists>q\<Colon>'a list. p = a %* q +++ ((0\<Colon>'a) # q)")
+apply (clarsimp simp add: poly_add poly_cmult)
+apply (rule_tac x="qa" in exI)
+apply (simp add: left_distrib [symmetric])
+apply clarsimp
+
+apply (auto simp add: right_minus poly_linear_divides poly_add poly_cmult left_distrib [symmetric])
+apply (rule_tac x = "pmult qa q" in exI)
+apply (rule_tac [2] x = "pmult p qa" in exI)
+apply (auto simp add: poly_add poly_mult poly_cmult mult_ac)
+done
+
+lemma (in comm_semiring_1) poly_divides_refl[simp]: "p divides p"
+apply (simp add: divides_def)
+apply (rule_tac x = "[one]" in exI)
+apply (auto simp add: poly_mult fun_eq)
+done
+
+lemma (in comm_semiring_1) poly_divides_trans: "[| p divides q; q divides r |] ==> p divides r"
+apply (simp add: divides_def, safe)
+apply (rule_tac x = "pmult qa qaa" in exI)
+apply (auto simp add: poly_mult fun_eq mult_assoc)
+done
+
+
+lemma (in recpower_comm_semiring_1) poly_divides_exp: "m \<le> n ==> (p %^ m) divides (p %^ n)"
+apply (auto simp add: le_iff_add)
+apply (induct_tac k)
+apply (rule_tac [2] poly_divides_trans)
+apply (auto simp add: divides_def)
+apply (rule_tac x = p in exI)
+apply (auto simp add: poly_mult fun_eq mult_ac)
+done
+
+lemma (in recpower_comm_semiring_1) poly_exp_divides: "[| (p %^ n) divides q;  m\<le>n |] ==> (p %^ m) divides q"
+by (blast intro: poly_divides_exp poly_divides_trans)
+
+lemma (in comm_semiring_0) poly_divides_add:
+   "[| p divides q; p divides r |] ==> p divides (q +++ r)"
+apply (simp add: divides_def, auto)
+apply (rule_tac x = "padd qa qaa" in exI)
+apply (auto simp add: poly_add fun_eq poly_mult right_distrib)
+done
+
+lemma (in comm_ring_1) poly_divides_diff:
+   "[| p divides q; p divides (q +++ r) |] ==> p divides r"
+apply (simp add: divides_def, auto)
+apply (rule_tac x = "padd qaa (poly_minus qa)" in exI)
+apply (auto simp add: poly_add fun_eq poly_mult poly_minus right_diff_distrib compare_rls add_ac)
+done
+
+lemma (in comm_ring_1) poly_divides_diff2: "[| p divides r; p divides (q +++ r) |] ==> p divides q"
+apply (erule poly_divides_diff)
+apply (auto simp add: poly_add fun_eq poly_mult divides_def add_ac)
+done
+
+lemma (in semiring_0) poly_divides_zero: "poly p = poly [] ==> q divides p"
+apply (simp add: divides_def)
+apply (rule exI[where x="[]"])
+apply (auto simp add: fun_eq poly_mult)
+done
+
+lemma (in semiring_0) poly_divides_zero2[simp]: "q divides []"
+apply (simp add: divides_def)
+apply (rule_tac x = "[]" in exI)
+apply (auto simp add: fun_eq)
+done
+
+text{*At last, we can consider the order of a root.*}
+
+lemma (in idom_char_0)  poly_order_exists_lemma:
+  assumes lp: "length p = d" and p: "poly p \<noteq> poly []"
+  shows "\<exists>n q. p = mulexp n [-a, 1] q \<and> poly q a \<noteq> 0"
+using lp p
+proof(induct d arbitrary: p)
+  case 0 thus ?case by simp
+next
+  case (Suc n p)
+  {assume p0: "poly p a = 0"
+    from Suc.prems have h: "length p = Suc n" "poly p \<noteq> poly []" by auto
+    hence pN: "p \<noteq> []" by auto
+    from p0[unfolded poly_linear_divides] pN  obtain q where 
+      q: "p = [-a, 1] *** q" by blast
+    from q h p0 have qh: "length q = n" "poly q \<noteq> poly []" 
+      apply -
+      apply simp
+      apply (simp only: fun_eq)
+      apply (rule ccontr)
+      apply (simp add: fun_eq poly_add poly_cmult minus_mult_left[symmetric])
+      done
+    from Suc.hyps[OF qh] obtain m r where 
+      mr: "q = mulexp m [-a,1] r" "poly r a \<noteq> 0" by blast    
+    from mr q have "p = mulexp (Suc m) [-a,1] r \<and> poly r a \<noteq> 0" by simp
+    hence ?case by blast}
+  moreover
+  {assume p0: "poly p a \<noteq> 0"
+    hence ?case using Suc.prems apply simp by (rule exI[where x="0::nat"], simp)}
+  ultimately show ?case by blast
+qed
+
+
+lemma (in recpower_comm_semiring_1) poly_mulexp: "poly (mulexp n p q) x = (poly p x) ^ n * poly q x"
+by(induct n, auto simp add: poly_mult power_Suc mult_ac)
+
+lemma (in comm_semiring_1) divides_left_mult:
+  assumes d:"(p***q) divides r" shows "p divides r \<and> q divides r"
+proof-
+  from d obtain t where r:"poly r = poly (p***q *** t)"
+    unfolding divides_def by blast
+  hence "poly r = poly (p *** (q *** t))"
+    "poly r = poly (q *** (p***t))" by(auto simp add: fun_eq poly_mult mult_ac)
+  thus ?thesis unfolding divides_def by blast
+qed
+
+
+
+(* FIXME: Tidy up *)
+
+lemma (in recpower_semiring_1) 
+  zero_power_iff: "0 ^ n = (if n = 0 then 1 else 0)"
+  by (induct n, simp_all add: power_Suc)
+
+lemma (in recpower_idom_char_0) poly_order_exists:
+  assumes lp: "length p = d" and p0: "poly p \<noteq> poly []"
+  shows "\<exists>n. ([-a, 1] %^ n) divides p & ~(([-a, 1] %^ (Suc n)) divides p)"
+proof-
+let ?poly = poly
+let ?mulexp = mulexp
+let ?pexp = pexp
+from lp p0
+show ?thesis
+apply -
+apply (drule poly_order_exists_lemma [where a=a], assumption, clarify)  
+apply (rule_tac x = n in exI, safe)
+apply (unfold divides_def)
+apply (rule_tac x = q in exI)
+apply (induct_tac "n", simp)
+apply (simp (no_asm_simp) add: poly_add poly_cmult poly_mult right_distrib mult_ac)
+apply safe
+apply (subgoal_tac "?poly (?mulexp n [uminus a, one] q) \<noteq> ?poly (pmult (?pexp [uminus a, one] (Suc n)) qa)") 
+apply simp 
+apply (induct_tac "n")
+apply (simp del: pmult_Cons pexp_Suc)
+apply (erule_tac Q = "?poly q a = zero" in contrapos_np)
+apply (simp add: poly_add poly_cmult minus_mult_left[symmetric])
+apply (rule pexp_Suc [THEN ssubst])
+apply (rule ccontr)
+apply (simp add: poly_mult_left_cancel poly_mult_assoc del: pmult_Cons pexp_Suc)
+done
+qed
+
+
+lemma (in semiring_1) poly_one_divides[simp]: "[1] divides p"
+by (simp add: divides_def, auto)
+
+lemma (in recpower_idom_char_0) poly_order: "poly p \<noteq> poly []
+      ==> EX! n. ([-a, 1] %^ n) divides p &
+                 ~(([-a, 1] %^ (Suc n)) divides p)"
+apply (auto intro: poly_order_exists simp add: less_linear simp del: pmult_Cons pexp_Suc)
+apply (cut_tac x = y and y = n in less_linear)
+apply (drule_tac m = n in poly_exp_divides)
+apply (auto dest: Suc_le_eq [THEN iffD2, THEN [2] poly_exp_divides]
+            simp del: pmult_Cons pexp_Suc)
+done
+
+text{*Order*}
+
+lemma some1_equalityD: "[| n = (@n. P n); EX! n. P n |] ==> P n"
+by (blast intro: someI2)
+
+lemma (in recpower_idom_char_0) order:
+      "(([-a, 1] %^ n) divides p &
+        ~(([-a, 1] %^ (Suc n)) divides p)) =
+        ((n = order a p) & ~(poly p = poly []))"
+apply (unfold order_def)
+apply (rule iffI)
+apply (blast dest: poly_divides_zero intro!: some1_equality [symmetric] poly_order)
+apply (blast intro!: poly_order [THEN [2] some1_equalityD])
+done
+
+lemma (in recpower_idom_char_0) order2: "[| poly p \<noteq> poly [] |]
+      ==> ([-a, 1] %^ (order a p)) divides p &
+              ~(([-a, 1] %^ (Suc(order a p))) divides p)"
+by (simp add: order del: pexp_Suc)
+
+lemma (in recpower_idom_char_0) order_unique: "[| poly p \<noteq> poly []; ([-a, 1] %^ n) divides p;
+         ~(([-a, 1] %^ (Suc n)) divides p)
+      |] ==> (n = order a p)"
+by (insert order [of a n p], auto) 
+
+lemma (in recpower_idom_char_0) order_unique_lemma: "(poly p \<noteq> poly [] & ([-a, 1] %^ n) divides p &
+         ~(([-a, 1] %^ (Suc n)) divides p))
+      ==> (n = order a p)"
+by (blast intro: order_unique)
+
+lemma (in ring_1) order_poly: "poly p = poly q ==> order a p = order a q"
+by (auto simp add: fun_eq divides_def poly_mult order_def)
+
+lemma (in semiring_1) pexp_one[simp]: "p %^ (Suc 0) = p"
+apply (induct "p")
+apply (auto simp add: numeral_1_eq_1)
+done
+
+lemma (in comm_ring_1) lemma_order_root:
+     " 0 < n & [- a, 1] %^ n divides p & ~ [- a, 1] %^ (Suc n) divides p
+             \<Longrightarrow> poly p a = 0"
+apply (induct n arbitrary: a p, blast)
+apply (auto simp add: divides_def poly_mult simp del: pmult_Cons)
+done
+
+lemma (in recpower_idom_char_0) order_root: "(poly p a = 0) = ((poly p = poly []) | order a p \<noteq> 0)"
+proof-
+  let ?poly = poly
+  show ?thesis 
+apply (case_tac "?poly p = ?poly []", auto)
+apply (simp add: poly_linear_divides del: pmult_Cons, safe)
+apply (drule_tac [!] a = a in order2)
+apply (rule ccontr)
+apply (simp add: divides_def poly_mult fun_eq del: pmult_Cons, blast)
+using neq0_conv
+apply (blast intro: lemma_order_root)
+done
+qed
+
+lemma (in recpower_idom_char_0) order_divides: "(([-a, 1] %^ n) divides p) = ((poly p = poly []) | n \<le> order a p)"
+proof-
+  let ?poly = poly
+  show ?thesis 
+apply (case_tac "?poly p = ?poly []", auto)
+apply (simp add: divides_def fun_eq poly_mult)
+apply (rule_tac x = "[]" in exI)
+apply (auto dest!: order2 [where a=a]
+	    intro: poly_exp_divides simp del: pexp_Suc)
+done
+qed
+
+lemma (in recpower_idom_char_0) order_decomp:
+     "poly p \<noteq> poly []
+      ==> \<exists>q. (poly p = poly (([-a, 1] %^ (order a p)) *** q)) &
+                ~([-a, 1] divides q)"
+apply (unfold divides_def)
+apply (drule order2 [where a = a])
+apply (simp add: divides_def del: pexp_Suc pmult_Cons, safe)
+apply (rule_tac x = q in exI, safe)
+apply (drule_tac x = qa in spec)
+apply (auto simp add: poly_mult fun_eq poly_exp mult_ac simp del: pmult_Cons)
+done
+
+text{*Important composition properties of orders.*}
+lemma order_mult: "poly (p *** q) \<noteq> poly []
+      ==> order a (p *** q) = order a p + order (a::'a::{recpower_idom_char_0}) q"
+apply (cut_tac a = a and p = "p *** q" and n = "order a p + order a q" in order)
+apply (auto simp add: poly_entire simp del: pmult_Cons)
+apply (drule_tac a = a in order2)+
+apply safe
+apply (simp add: divides_def fun_eq poly_exp_add poly_mult del: pmult_Cons, safe)
+apply (rule_tac x = "qa *** qaa" in exI)
+apply (simp add: poly_mult mult_ac del: pmult_Cons)
+apply (drule_tac a = a in order_decomp)+
+apply safe
+apply (subgoal_tac "[-a,1] divides (qa *** qaa) ")
+apply (simp add: poly_primes del: pmult_Cons)
+apply (auto simp add: divides_def simp del: pmult_Cons)
+apply (rule_tac x = qb in exI)
+apply (subgoal_tac "poly ([-a, 1] %^ (order a p) *** (qa *** qaa)) = poly ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))")
+apply (drule poly_mult_left_cancel [THEN iffD1], force)
+apply (subgoal_tac "poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** (qa *** qaa))) = poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))) ")
+apply (drule poly_mult_left_cancel [THEN iffD1], force)
+apply (simp add: fun_eq poly_exp_add poly_mult mult_ac del: pmult_Cons)
+done
+
+lemma (in recpower_idom_char_0) order_mult: 
+  assumes pq0: "poly (p *** q) \<noteq> poly []"
+  shows "order a (p *** q) = order a p + order a q"
+proof-
+  let ?order = order
+  let ?divides = "op divides"
+  let ?poly = poly
+from pq0 
+show ?thesis
+apply (cut_tac a = a and p = "pmult p q" and n = "?order a p + ?order a q" in order)
+apply (auto simp add: poly_entire simp del: pmult_Cons)
+apply (drule_tac a = a in order2)+
+apply safe
+apply (simp add: divides_def fun_eq poly_exp_add poly_mult del: pmult_Cons, safe)
+apply (rule_tac x = "pmult qa qaa" in exI)
+apply (simp add: poly_mult mult_ac del: pmult_Cons)
+apply (drule_tac a = a in order_decomp)+
+apply safe
+apply (subgoal_tac "?divides [uminus a,one ] (pmult qa qaa) ")
+apply (simp add: poly_primes del: pmult_Cons)
+apply (auto simp add: divides_def simp del: pmult_Cons)
+apply (rule_tac x = qb in exI)
+apply (subgoal_tac "?poly (pmult (pexp [uminus a, one] (?order a p)) (pmult qa qaa)) = ?poly (pmult (pexp [uminus a, one] (?order a p)) (pmult [uminus a, one] qb))")
+apply (drule poly_mult_left_cancel [THEN iffD1], force)
+apply (subgoal_tac "?poly (pmult (pexp [uminus a, one ] (order a q)) (pmult (pexp [uminus a, one] (order a p)) (pmult qa qaa))) = ?poly (pmult (pexp [uminus a, one] (order a q)) (pmult (pexp [uminus a, one] (order a p)) (pmult [uminus a, one] qb))) ")
+apply (drule poly_mult_left_cancel [THEN iffD1], force)
+apply (simp add: fun_eq poly_exp_add poly_mult mult_ac del: pmult_Cons)
+done
+qed
+
+lemma (in recpower_idom_char_0) order_root2: "poly p \<noteq> poly [] ==> (poly p a = 0) = (order a p \<noteq> 0)"
+by (rule order_root [THEN ssubst], auto)
+
+lemma (in semiring_1) pmult_one[simp]: "[1] *** p = p" by auto
+
+lemma (in semiring_0) poly_Nil_zero: "poly [] = poly [0]"
+by (simp add: fun_eq)
+
+lemma (in recpower_idom_char_0) rsquarefree_decomp:
+     "[| rsquarefree p; poly p a = 0 |]
+      ==> \<exists>q. (poly p = poly ([-a, 1] *** q)) & poly q a \<noteq> 0"
+apply (simp add: rsquarefree_def, safe)
+apply (frule_tac a = a in order_decomp)
+apply (drule_tac x = a in spec)
+apply (drule_tac a = a in order_root2 [symmetric])
+apply (auto simp del: pmult_Cons)
+apply (rule_tac x = q in exI, safe)
+apply (simp add: poly_mult fun_eq)
+apply (drule_tac p1 = q in poly_linear_divides [THEN iffD1])
+apply (simp add: divides_def del: pmult_Cons, safe)
+apply (drule_tac x = "[]" in spec)
+apply (auto simp add: fun_eq)
+done
+
+
+text{*Normalization of a polynomial.*}
+
+lemma (in semiring_0) poly_normalize[simp]: "poly (pnormalize p) = poly p"
+apply (induct "p")
+apply (auto simp add: fun_eq)
+done
+
+text{*The degree of a polynomial.*}
+
+lemma (in semiring_0) lemma_degree_zero:
+     "list_all (%c. c = 0) p \<longleftrightarrow>  pnormalize p = []"
+by (induct "p", auto)
+
+lemma (in idom_char_0) degree_zero: 
+  assumes pN: "poly p = poly []" shows"degree p = 0"
+proof-
+  let ?pn = pnormalize
+  from pN
+  show ?thesis 
+    apply (simp add: degree_def)
+    apply (case_tac "?pn p = []")
+    apply (auto simp add: poly_zero lemma_degree_zero )
+    done
+qed
+
+lemma (in semiring_0) pnormalize_sing: "(pnormalize [x] = [x]) \<longleftrightarrow> x \<noteq> 0" by simp
+lemma (in semiring_0) pnormalize_pair: "y \<noteq> 0 \<longleftrightarrow> (pnormalize [x, y] = [x, y])" by simp
+lemma (in semiring_0) pnormal_cons: "pnormal p \<Longrightarrow> pnormal (c#p)" 
+  unfolding pnormal_def by simp
+lemma (in semiring_0) pnormal_tail: "p\<noteq>[] \<Longrightarrow> pnormal (c#p) \<Longrightarrow> pnormal p"
+  unfolding pnormal_def 
+  apply (cases "pnormalize p = []", auto)
+  by (cases "c = 0", auto)
+
+
+lemma (in semiring_0) pnormal_last_nonzero: "pnormal p ==> last p \<noteq> 0"
+proof(induct p)
+  case Nil thus ?case by (simp add: pnormal_def)
+next 
+  case (Cons a as) thus ?case
+    apply (simp add: pnormal_def)
+    apply (cases "pnormalize as = []", simp_all)
+    apply (cases "as = []", simp_all)
+    apply (cases "a=0", simp_all)
+    apply (cases "a=0", simp_all)
+    done
+qed
+
+lemma (in semiring_0) pnormal_length: "pnormal p \<Longrightarrow> 0 < length p"
+  unfolding pnormal_def length_greater_0_conv by blast
+
+lemma (in semiring_0) pnormal_last_length: "\<lbrakk>0 < length p ; last p \<noteq> 0\<rbrakk> \<Longrightarrow> pnormal p"
+  apply (induct p, auto)
+  apply (case_tac "p = []", auto)
+  apply (simp add: pnormal_def)
+  by (rule pnormal_cons, auto)
+
+lemma (in semiring_0) pnormal_id: "pnormal p \<longleftrightarrow> (0 < length p \<and> last p \<noteq> 0)"
+  using pnormal_last_length pnormal_length pnormal_last_nonzero by blast
+
+lemma (in idom_char_0) poly_Cons_eq: "poly (c#cs) = poly (d#ds) \<longleftrightarrow> c=d \<and> poly cs = poly ds" (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  assume eq: ?lhs
+  hence "\<And>x. poly ((c#cs) +++ -- (d#ds)) x = 0"
+    by (simp only: poly_minus poly_add ring_simps) simp
+  hence "poly ((c#cs) +++ -- (d#ds)) = poly []" by - (rule ext, simp) 
+  hence "c = d \<and> list_all (\<lambda>x. x=0) ((cs +++ -- ds))"
+    unfolding poly_zero by (simp add: poly_minus_def ring_simps minus_mult_left[symmetric])
+  hence "c = d \<and> (\<forall>x. poly (cs +++ -- ds) x = 0)"
+    unfolding poly_zero[symmetric] by simp 
+  thus ?rhs  apply (simp add: poly_minus poly_add ring_simps) apply (rule ext, simp) done
+next
+  assume ?rhs then show ?lhs  by -  (rule ext,simp)
+qed
+  
+lemma (in idom_char_0) pnormalize_unique: "poly p = poly q \<Longrightarrow> pnormalize p = pnormalize q"
+proof(induct q arbitrary: p)
+  case Nil thus ?case by (simp only: poly_zero lemma_degree_zero) simp
+next
+  case (Cons c cs p)
+  thus ?case
+  proof(induct p)
+    case Nil
+    hence "poly [] = poly (c#cs)" by blast
+    then have "poly (c#cs) = poly [] " by simp 
+    thus ?case by (simp only: poly_zero lemma_degree_zero) simp
+  next
+    case (Cons d ds)
+    hence eq: "poly (d # ds) = poly (c # cs)" by blast
+    hence eq': "\<And>x. poly (d # ds) x = poly (c # cs) x" by simp
+    hence "poly (d # ds) 0 = poly (c # cs) 0" by blast
+    hence dc: "d = c" by auto
+    with eq have "poly ds = poly cs"
+      unfolding  poly_Cons_eq by simp
+    with Cons.prems have "pnormalize ds = pnormalize cs" by blast
+    with dc show ?case by simp
+  qed
+qed
+
+lemma (in idom_char_0) degree_unique: assumes pq: "poly p = poly q"
+  shows "degree p = degree q"
+using pnormalize_unique[OF pq] unfolding degree_def by simp
+
+lemma (in semiring_0) pnormalize_length: "length (pnormalize p) \<le> length p" by (induct p, auto)
+
+lemma (in semiring_0) last_linear_mul_lemma: 
+  "last ((a %* p) +++ (x#(b %* p))) = (if p=[] then x else b*last p)"
+
+apply (induct p arbitrary: a x b, auto)
+apply (subgoal_tac "padd (cmult aa p) (times b a # cmult b p) \<noteq> []", simp)
+apply (induct_tac p, auto)
+done
+
+lemma (in semiring_1) last_linear_mul: assumes p:"p\<noteq>[]" shows "last ([a,1] *** p) = last p"
+proof-
+  from p obtain c cs where cs: "p = c#cs" by (cases p, auto)
+  from cs have eq:"[a,1] *** p = (a %* (c#cs)) +++ (0#(1 %* (c#cs)))"
+    by (simp add: poly_cmult_distr)
+  show ?thesis using cs
+    unfolding eq last_linear_mul_lemma by simp
+qed
+
+lemma (in semiring_0) pnormalize_eq: "last p \<noteq> 0 \<Longrightarrow> pnormalize p = p"
+  apply (induct p, auto)
+  apply (case_tac p, auto)+
+  done
+
+lemma (in semiring_0) last_pnormalize: "pnormalize p \<noteq> [] \<Longrightarrow> last (pnormalize p) \<noteq> 0"
+  by (induct p, auto)
+
+lemma (in semiring_0) pnormal_degree: "last p \<noteq> 0 \<Longrightarrow> degree p = length p - 1"
+  using pnormalize_eq[of p] unfolding degree_def by simp
+
+lemma (in semiring_0) poly_Nil_ext: "poly [] = (\<lambda>x. 0)" by (rule ext) simp
+
+lemma (in idom_char_0) linear_mul_degree: assumes p: "poly p \<noteq> poly []"
+  shows "degree ([a,1] *** p) = degree p + 1"
+proof-
+  from p have pnz: "pnormalize p \<noteq> []"
+    unfolding poly_zero lemma_degree_zero .
+  
+  from last_linear_mul[OF pnz, of a] last_pnormalize[OF pnz]
+  have l0: "last ([a, 1] *** pnormalize p) \<noteq> 0" by simp
+  from last_pnormalize[OF pnz] last_linear_mul[OF pnz, of a]
+    pnormal_degree[OF l0] pnormal_degree[OF last_pnormalize[OF pnz]] pnz
+ 
+
+  have th: "degree ([a,1] *** pnormalize p) = degree (pnormalize p) + 1" 
+    by (auto simp add: poly_length_mult)
+
+  have eqs: "poly ([a,1] *** pnormalize p) = poly ([a,1] *** p)"
+    by (rule ext) (simp add: poly_mult poly_add poly_cmult)
+  from degree_unique[OF eqs] th
+  show ?thesis by (simp add: degree_unique[OF poly_normalize])
+qed
+
+lemma (in idom_char_0) linear_pow_mul_degree:
+  "degree([a,1] %^n *** p) = (if poly p = poly [] then 0 else degree p + n)"
+proof(induct n arbitrary: a p)
+  case (0 a p)
+  {assume p: "poly p = poly []"
+    hence ?case using degree_unique[OF p] by (simp add: degree_def)}
+  moreover
+  {assume p: "poly p \<noteq> poly []" hence ?case by (auto simp add: poly_Nil_ext) }
+  ultimately show ?case by blast
+next
+  case (Suc n a p)
+  have eq: "poly ([a,1] %^(Suc n) *** p) = poly ([a,1] %^ n *** ([a,1] *** p))"
+    apply (rule ext, simp add: poly_mult poly_add poly_cmult)
+    by (simp add: mult_ac add_ac right_distrib)
+  note deq = degree_unique[OF eq]
+  {assume p: "poly p = poly []"
+    with eq have eq': "poly ([a,1] %^(Suc n) *** p) = poly []" 
+      by - (rule ext,simp add: poly_mult poly_cmult poly_add)
+    from degree_unique[OF eq'] p have ?case by (simp add: degree_def)}
+  moreover
+  {assume p: "poly p \<noteq> poly []"
+    from p have ap: "poly ([a,1] *** p) \<noteq> poly []"
+      using poly_mult_not_eq_poly_Nil unfolding poly_entire by auto 
+    have eq: "poly ([a,1] %^(Suc n) *** p) = poly ([a,1]%^n *** ([a,1] *** p))"
+     by (rule ext, simp add: poly_mult poly_add poly_exp poly_cmult mult_ac add_ac right_distrib)
+   from ap have ap': "(poly ([a,1] *** p) = poly []) = False" by blast
+   have  th0: "degree ([a,1]%^n *** ([a,1] *** p)) = degree ([a,1] *** p) + n"
+     apply (simp only: Suc.hyps[of a "pmult [a,one] p"] ap')
+     by simp
+    
+   from degree_unique[OF eq] ap p th0 linear_mul_degree[OF p, of a]
+   have ?case by (auto simp del: poly.simps)}
+  ultimately show ?case by blast
+qed
+
+lemma (in recpower_idom_char_0) order_degree: 
+  assumes p0: "poly p \<noteq> poly []"
+  shows "order a p \<le> degree p"
+proof-
+  from order2[OF p0, unfolded divides_def]
+  obtain q where q: "poly p = poly ([- a, 1]%^ (order a p) *** q)" by blast
+  {assume "poly q = poly []"
+    with q p0 have False by (simp add: poly_mult poly_entire)}
+  with degree_unique[OF q, unfolded linear_pow_mul_degree] 
+  show ?thesis by auto
+qed
+
+text{*Tidier versions of finiteness of roots.*}
+
+lemma (in idom_char_0) poly_roots_finite_set: "poly p \<noteq> poly [] ==> finite {x. poly p x = 0}"
+unfolding poly_roots_finite .
+
+text{*bound for polynomial.*}
+
+lemma poly_mono: "abs(x) \<le> k ==> abs(poly p (x::'a::{ordered_idom})) \<le> poly (map abs p) k"
+apply (induct "p", auto)
+apply (rule_tac y = "abs a + abs (x * poly p x)" in order_trans)
+apply (rule abs_triangle_ineq)
+apply (auto intro!: mult_mono simp add: abs_mult)
+done
+
+lemma (in semiring_0) poly_Sing: "poly [c] x = c" by simp
+
+end
--- a/src/HOL/Library/comm_ring.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Library/comm_ring.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
-(*  ID:         $Id$
-    Author:     Amine Chaieb
+(*  Author:     Amine Chaieb
 
 Tactic for solving equalities over commutative rings.
 *)
@@ -71,7 +70,7 @@
 fun reif_eq thy (eq as Const("op =", Type("fun", [T, _])) $ lhs $ rhs) =
       if Sign.of_sort thy (T, cr_sort) then
         let
-          val fs = term_frees eq;
+          val fs = OldTerm.term_frees eq;
           val cvs = cterm_of thy (HOLogic.mk_list T fs);
           val clhs = cterm_of thy (reif_polex T fs lhs);
           val crhs = cterm_of thy (reif_polex T fs rhs);
--- a/src/HOL/List.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/List.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/List.thy
-    ID:         $Id$
     Author:     Tobias Nipkow
 *)
 
@@ -359,7 +358,7 @@
 
    fun pat_tr ctxt p e opti = (* %x. case x of p => e | _ => [] *)
     let
-      val x = Free (Name.variant (add_term_free_names (p$e, [])) "x", dummyT);
+      val x = Free (Name.variant (fold Term.add_free_names [p, e] []) "x", dummyT);
       val e = if opti then singl e else e;
       val case1 = Syntax.const "_case1" $ p $ e;
       val case2 = Syntax.const "_case1" $ Syntax.const Term.dummy_patternN
@@ -548,9 +547,9 @@
 lemma append_Nil2 [simp]: "xs @ [] = xs"
 by (induct xs) auto
 
-class_interpretation semigroup_append: semigroup_add ["op @"]
+interpretation semigroup_append!: semigroup_add "op @"
   proof qed simp
-class_interpretation monoid_append: monoid_add ["[]" "op @"]
+interpretation monoid_append!: monoid_add "[]" "op @"
   proof qed simp+
 
 lemma append_is_Nil_conv [iff]: "(xs @ ys = []) = (xs = [] \<and> ys = [])"
--- a/src/HOL/Main.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Main.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -1,13 +1,14 @@
-(*  Title:      HOL/Main.thy
-    ID:         $Id$
-*)
-
 header {* Main HOL *}
 
 theory Main
 imports Plain Code_Eval Map Nat_Int_Bij Recdef
 begin
 
+text {*
+  Classical Higher-order Logic -- only ``Main'', excluding real and
+  complex numbers etc.
+*}
+
 text {* See further \cite{Nipkow-et-al:2002:tutorial} *}
 
 end
--- a/src/HOL/Matrix/cplex/matrixlp.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Matrix/cplex/matrixlp.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Matrix/cplex/matrixlp.ML
-    ID:         $Id$
     Author:     Steven Obua
 *)
 
@@ -20,7 +19,7 @@
 fun inst_real thm =
   let val certT = ctyp_of (Thm.theory_of_thm thm) in
     standard (Thm.instantiate
-      ([(certT (TVar (hd (term_tvars (prop_of thm)))), certT HOLogic.realT)], []) thm)
+      ([(certT (TVar (hd (OldTerm.term_tvars (prop_of thm)))), certT HOLogic.realT)], []) thm)
   end
 
 local
@@ -58,7 +57,7 @@
     let
         val certT = Thm.ctyp_of (Thm.theory_of_thm thm);
         val ord = prod_ord (prod_ord string_ord int_ord) (list_ord string_ord)
-        val v = TVar (hd (sort ord (term_tvars (prop_of thm))))
+        val v = TVar (hd (sort ord (OldTerm.term_tvars (prop_of thm))))
     in
         standard (Thm.instantiate ([(certT v, certT ty)], []) thm)
     end
--- a/src/HOL/MetisExamples/BT.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/MetisExamples/BT.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -84,7 +84,7 @@
 lemma depth_reflect: "depth (reflect t) = depth t"
   apply (induct t)
   apply (metis depth.simps(1) reflect.simps(1))
-  apply (metis depth.simps(2) min_max.less_eq_less_sup.sup_commute reflect.simps(2))
+  apply (metis depth.simps(2) min_max.sup_commute reflect.simps(2))
   done
 
 text {*
--- a/src/HOL/MetisExamples/BigO.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/MetisExamples/BigO.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/MetisExamples/BigO.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 
 Testing the metis method
@@ -13,9 +12,7 @@
 
 subsection {* Definitions *}
 
-constdefs 
-
-  bigo :: "('a => 'b::ordered_idom) => ('a => 'b) set"    ("(1O'(_'))")
+definition bigo :: "('a => 'b::ordered_idom) => ('a => 'b) set"    ("(1O'(_'))") where
   "O(f::('a => 'b)) ==   {h. EX c. ALL x. abs (h x) <= c * abs (f x)}"
 
 ML_command{*AtpWrapper.problem_name := "BigO__bigo_pos_const"*}
@@ -362,7 +359,7 @@
   apply (rule add_mono)
 ML_command{*AtpWrapper.problem_name := "BigO__bigo_plus_eq_simpler"*} 
 (*Found by SPASS; SLOW*)
-apply (metis le_maxI2 linorder_linear linorder_not_le min_max.less_eq_less_sup.sup_absorb1 mult_le_cancel_right order_trans)
+apply (metis le_maxI2 linorder_linear linorder_not_le min_max.sup_absorb1 mult_le_cancel_right order_trans)
 apply (metis le_maxI2 linorder_not_le mult_le_cancel_right order_trans)
 done
 
@@ -1164,7 +1161,7 @@
 (*sledgehammer*);  
   apply (case_tac "0 <= k x - g x")
   prefer 2 (*re-order subgoals because I don't know what to put after a structured proof*)
-   apply (metis abs_ge_zero abs_minus_commute linorder_linear min_max.less_eq_less_sup.sup_absorb1 min_max.less_eq_less_sup.sup_commute)
+   apply (metis abs_ge_zero abs_minus_commute linorder_linear min_max.sup_absorb1 min_max.sup_commute)
 proof (neg_clausify)
 fix x
 assume 0: "\<And>A. k A \<le> f A"
@@ -1174,16 +1171,16 @@
 have 3: "\<not> k x - g x < (0\<Colon>'b)"
   by (metis 2 linorder_not_less)
 have 4: "\<And>X1 X2. min X1 (k X2) \<le> f X2"
-  by (metis min_max.less_eq_less_inf.inf_le2 min_max.less_eq_less_inf.le_inf_iff min_max.less_eq_less_inf.le_iff_inf 0)
+  by (metis min_max.inf_le2 min_max.le_inf_iff min_max.le_iff_inf 0)
 have 5: "\<bar>g x - f x\<bar> = f x - g x"
-  by (metis abs_minus_commute combine_common_factor mult_zero_right minus_add_cancel minus_zero abs_if diff_less_eq min_max.less_eq_less_inf.inf_commute 4 linorder_not_le min_max.less_eq_less_inf.le_iff_inf 3 diff_less_0_iff_less linorder_not_less)
+  by (metis abs_minus_commute combine_common_factor mult_zero_right minus_add_cancel minus_zero abs_if diff_less_eq min_max.inf_commute 4 linorder_not_le min_max.le_iff_inf 3 diff_less_0_iff_less linorder_not_less)
 have 6: "max (0\<Colon>'b) (k x - g x) = k x - g x"
-  by (metis min_max.less_eq_less_sup.le_iff_sup 2)
+  by (metis min_max.le_iff_sup 2)
 assume 7: "\<not> max (k x - g x) (0\<Colon>'b) \<le> \<bar>f x - g x\<bar>"
 have 8: "\<not> k x - g x \<le> f x - g x"
-  by (metis 5 abs_minus_commute 7 min_max.less_eq_less_sup.sup_commute 6)
+  by (metis 5 abs_minus_commute 7 min_max.sup_commute 6)
 show "False"
-  by (metis min_max.less_eq_less_sup.sup_commute min_max.less_eq_less_inf.inf_commute min_max.less_eq_less_inf_sup.sup_inf_absorb min_max.less_eq_less_inf.le_iff_inf 0 max_diff_distrib_left 1 linorder_not_le 8)
+  by (metis min_max.sup_commute min_max.inf_commute min_max.sup_inf_absorb min_max.le_iff_inf 0 max_diff_distrib_left 1 linorder_not_le 8)
 qed
 
 ML_command{*AtpWrapper.problem_name := "BigO__bigo_lesso3"*}
@@ -1206,7 +1203,7 @@
 ML_command{*AtpWrapper.problem_name := "BigO__bigo_lesso3_simpler"*}
 apply (metis diff_less_0_iff_less linorder_not_le not_leE uminus_add_conv_diff xt1(12) xt1(6))
 apply (metis add_minus_cancel diff_le_eq le_diff_eq uminus_add_conv_diff)
-apply (metis abs_ge_zero linorder_linear min_max.less_eq_less_sup.sup_absorb1 min_max.less_eq_less_sup.sup_commute)
+apply (metis abs_ge_zero linorder_linear min_max.sup_absorb1 min_max.sup_commute)
 done
 
 lemma bigo_lesso4: "f <o g =o O(k::'a=>'b::ordered_field) ==>
--- a/src/HOL/Modelcheck/mucke_oracle.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Modelcheck/mucke_oracle.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,3 @@
-
-(* $Id$ *)
 
 val trace_mc = ref false; 
 
@@ -249,8 +247,8 @@
 (* replace Vars bei Frees, freeze_thaw shares code of tactic/freeze_thaw
    and thm.instantiate *)
 fun freeze_thaw t =
-  let val used = add_term_names (t, [])
-          and vars = term_vars t;
+  let val used = OldTerm.add_term_names (t, [])
+          and vars = OldTerm.term_vars t;
       fun newName (Var(ix,_), (pairs,used)) = 
           let val v = Name.variant used (string_of_indexname ix)
           in  ((ix,v)::pairs, v::used)  end;
--- a/src/HOL/NatBin.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/NatBin.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -29,14 +29,14 @@
   unfolding nat_number_of_def ..
 
 abbreviation (xsymbols)
-  square :: "'a::power => 'a"  ("(_\<twosuperior>)" [1000] 999) where
+  power2 :: "'a::power => 'a"  ("(_\<twosuperior>)" [1000] 999) where
   "x\<twosuperior> == x^2"
 
 notation (latex output)
-  square  ("(_\<twosuperior>)" [1000] 999)
+  power2  ("(_\<twosuperior>)" [1000] 999)
 
 notation (HTML output)
-  square  ("(_\<twosuperior>)" [1000] 999)
+  power2  ("(_\<twosuperior>)" [1000] 999)
 
 
 subsection {* Predicate for negative binary numbers *}
--- a/src/HOL/Nominal/Examples/SOS.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Nominal/Examples/SOS.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -328,7 +328,7 @@
   have ih2:"\<And>t. e\<^isub>2 \<Down> t \<Longrightarrow> e\<^isub>2' = t" by fact
   have ih3: "\<And>t. e\<^isub>1'[x::=e\<^isub>2'] \<Down> t \<Longrightarrow> e' = t" by fact
   have app: "App e\<^isub>1 e\<^isub>2 \<Down> t\<^isub>2" by fact
-  have vc: "x\<sharp>e\<^isub>1" "x\<sharp>e\<^isub>2" "x\<sharp>t\<^isub>2" by fact
+  have vc: "x\<sharp>e\<^isub>1" "x\<sharp>e\<^isub>2" "x\<sharp>t\<^isub>2" by fact+
   then have "x\<sharp>App e\<^isub>1 e\<^isub>2" by auto
   from app vc obtain f\<^isub>1 f\<^isub>2 where x1: "e\<^isub>1 \<Down> Lam [x]. f\<^isub>1" and x2: "e\<^isub>2 \<Down> f\<^isub>2" and x3: "f\<^isub>1[x::=f\<^isub>2] \<Down> t\<^isub>2"
     by (auto elim!: b_App_elim)
@@ -515,7 +515,7 @@
   have ih:"\<And>\<theta> \<Gamma> T. \<lbrakk>\<theta> Vcloses \<Gamma>; \<Gamma> \<turnstile> e : T\<rbrakk> \<Longrightarrow> \<exists>v. \<theta><e> \<Down> v \<and> v \<in> V T" by fact
   have as\<^isub>1: "\<theta> Vcloses \<Gamma>" by fact
   have as\<^isub>2: "\<Gamma> \<turnstile> Lam [x].e : T" by fact
-  have fs: "x\<sharp>\<Gamma>" "x\<sharp>\<theta>" by fact
+  have fs: "x\<sharp>\<Gamma>" "x\<sharp>\<theta>" by fact+
   from as\<^isub>2 fs obtain T\<^isub>1 T\<^isub>2 
     where "(i)": "(x,T\<^isub>1)#\<Gamma> \<turnstile> e:T\<^isub>2" and "(ii)": "T = T\<^isub>1 \<rightarrow> T\<^isub>2" using fs
     by (auto elim: t_Lam_elim)
--- a/src/HOL/Nominal/nominal_fresh_fun.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Nominal/nominal_fresh_fun.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,10 +1,8 @@
 (*  Title:      HOL/Nominal/nominal_fresh_fun.ML
-    ID:         $Id$
-    Authors:    Stefan Berghofer, Julien Narboux, TU Muenchen
+    Authors:    Stefan Berghofer and Julien Narboux, TU Muenchen
 
 Provides a tactic to generate fresh names and
 a tactic to analyse instances of the fresh_fun.
-
 *)
 
 (* First some functions that should be in the library *)
@@ -83,7 +81,7 @@
    val bvs = map_index (Bound o fst) ps;
 (* select variables of the right class *)
    val vs = filter (fn t => is_of_fs_name (fastype_of1 (Ts, t)))
-     (term_frees goal @ bvs);
+     (OldTerm.term_frees goal @ bvs);
 (* build the tuple *)
    val s = (Library.foldr1 (fn (v, s) =>
      HOLogic.pair_const (fastype_of1 (Ts, v)) (fastype_of1 (Ts, s)) $ v $ s) vs) handle _ => HOLogic.unit ;  (* FIXME avoid handle _ *)
@@ -91,7 +89,7 @@
    val at_name_inst_thm = get_dyn_thm thy ("at_"^atom_basename^"_inst") atom_basename;
    val exists_fresh' = at_name_inst_thm RS at_exists_fresh';
 (* find the variable we want to instantiate *)
-   val x = hd (term_vars (prop_of exists_fresh'));
+   val x = hd (OldTerm.term_vars (prop_of exists_fresh'));
  in 
    (cut_inst_tac_term' [(x,s)] exists_fresh' 1 THEN
    rtac fs_name_thm 1 THEN
@@ -150,7 +148,7 @@
     val fresh_perm_app = PureThy.get_thms thy "fresh_perm_app";
     val ss' = ss addsimps fresh_prod::abs_fresh;
     val ss'' = ss' addsimps fresh_perm_app;
-    val x = hd (tl (term_vars (prop_of exI)));
+    val x = hd (tl (OldTerm.term_vars (prop_of exI)));
     val goal = nth (prems_of thm) (i-1);
     val atom_name_opt = get_inner_fresh_fun goal;
     val n = List.length (Logic.strip_params goal);
@@ -165,7 +163,7 @@
     val pt_name_inst = get_dyn_thm thy ("pt_"^atom_basename^"_inst") atom_basename;
     val at_name_inst = get_dyn_thm thy ("at_"^atom_basename^"_inst") atom_basename;
     fun inst_fresh vars params i st =
-   let val vars' = term_vars (prop_of st);
+   let val vars' = OldTerm.term_vars (prop_of st);
        val thy = theory_of_thm st;
    in case vars' \\ vars of 
      [x] => Seq.single (Thm.instantiate ([],[(cterm_of thy x,cterm_of thy (list_abs (params,Bound 0)))]) st)
@@ -174,7 +172,7 @@
   in
   ((fn st =>
   let 
-    val vars = term_vars (prop_of st);
+    val vars = OldTerm.term_vars (prop_of st);
     val params = Logic.strip_params (nth (prems_of st) (i-1))
     (* The tactics which solve the subgoals generated 
        by the conditionnal rewrite rule. *)
--- a/src/HOL/Nominal/nominal_inductive.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Nominal/nominal_inductive.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Nominal/nominal_inductive.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Infrastructure for proving equivariance and strong induction theorems
@@ -153,7 +152,7 @@
     val elims = map (atomize_induct ctxt) elims;
     val monos = InductivePackage.get_monos ctxt;
     val eqvt_thms = NominalThmDecls.get_eqvt_thms ctxt;
-    val _ = (case names \\ foldl (apfst prop_of #> add_term_consts) [] eqvt_thms of
+    val _ = (case names \\ fold (Term.add_const_names o Thm.prop_of) eqvt_thms [] of
         [] => ()
       | xs => error ("Missing equivariance theorem for predicate(s): " ^
           commas_quote xs));
@@ -200,8 +199,8 @@
     val ind_sort = if null atomTs then HOLogic.typeS
       else Sign.certify_sort thy (map (fn T => Sign.intern_class thy
         ("fs_" ^ Sign.base_name (fst (dest_Type T)))) atomTs);
-    val fs_ctxt_tyname = Name.variant (map fst (term_tfrees raw_induct')) "'n";
-    val fs_ctxt_name = Name.variant (add_term_names (raw_induct', [])) "z";
+    val fs_ctxt_tyname = Name.variant (map fst (OldTerm.term_tfrees raw_induct')) "'n";
+    val fs_ctxt_name = Name.variant (OldTerm.add_term_names (raw_induct', [])) "z";
     val fsT = TFree (fs_ctxt_tyname, ind_sort);
 
     val inductive_forall_def' = Drule.instantiate'
@@ -238,10 +237,10 @@
         val prem = Logic.list_implies
           (map mk_fresh bvars @ mk_distinct bvars @
            map (fn prem =>
-             if null (term_frees prem inter ps) then prem
+             if null (OldTerm.term_frees prem inter ps) then prem
              else lift_prem prem) prems,
            HOLogic.mk_Trueprop (lift_pred p ts));
-        val vs = map (Var o apfst (rpair 0)) (rename_wrt_term prem params')
+        val vs = map (Var o apfst (rpair 0)) (Term.rename_wrt_term prem params')
       in
         (list_all (params', prem), (rev vs, subst_bounds (vs, prem)))
       end) prems);
@@ -264,7 +263,7 @@
     val vc_compat = map (fn (params, bvars, prems, (p, ts)) =>
       map (fn q => list_all (params, incr_boundvars ~1 (Logic.list_implies
           (List.mapPartial (fn prem =>
-             if null (ps inter term_frees prem) then SOME prem
+             if null (ps inter OldTerm.term_frees prem) then SOME prem
              else map_term (split_conj (K o I) names) prem prem) prems, q))))
         (mk_distinct bvars @
          maps (fn (t, T) => map (fn (u, U) => HOLogic.mk_Trueprop
@@ -353,7 +352,7 @@
                          (rev pis' @ pis) th));
                  val (gprems1, gprems2) = split_list
                    (map (fn (th, t) =>
-                      if null (term_frees t inter ps) then (SOME th, mk_pi th)
+                      if null (OldTerm.term_frees t inter ps) then (SOME th, mk_pi th)
                       else
                         (map_thm ctxt (split_conj (K o I) names)
                            (etac conjunct1 1) monos NONE th,
@@ -412,7 +411,7 @@
       let
         val prem :: prems = Logic.strip_imp_prems rule;
         val concl = Logic.strip_imp_concl rule;
-        val used = add_term_free_names (rule, [])
+        val used = Term.add_free_names rule [];
       in
         (prem,
          List.drop (snd (strip_comb (HOLogic.dest_Trueprop prem)), length ind_params),
@@ -614,7 +613,7 @@
       [mk_perm_bool_simproc names,
        NominalPermeq.perm_simproc_app, NominalPermeq.perm_simproc_fun];
     val t = Logic.unvarify (concl_of raw_induct);
-    val pi = Name.variant (add_term_names (t, [])) "pi";
+    val pi = Name.variant (OldTerm.add_term_names (t, [])) "pi";
     val ps = map (fst o HOLogic.dest_imp)
       (HOLogic.dest_conj (HOLogic.dest_Trueprop t));
     fun eqvt_tac pi (intr, vs) st =
--- a/src/HOL/Nominal/nominal_inductive2.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Nominal/nominal_inductive2.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Nominal/nominal_inductive2.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Infrastructure for proving equivariance and strong induction theorems
@@ -139,7 +138,7 @@
   in Option.map prove (map_term f prop (the_default prop opt)) end;
 
 fun abs_params params t =
-  let val vs =  map (Var o apfst (rpair 0)) (rename_wrt_term t params)
+  let val vs =  map (Var o apfst (rpair 0)) (Term.rename_wrt_term t params)
   in (list_all (params, t), (rev vs, subst_bounds (vs, t))) end;
 
 fun inst_params thy (vs, p) th cts =
@@ -159,7 +158,7 @@
     val elims = map (atomize_induct ctxt) elims;
     val monos = InductivePackage.get_monos ctxt;
     val eqvt_thms = NominalThmDecls.get_eqvt_thms ctxt;
-    val _ = (case names \\ foldl (apfst prop_of #> add_term_consts) [] eqvt_thms of
+    val _ = (case names \\ fold (Term.add_const_names o Thm.prop_of) eqvt_thms [] of
         [] => ()
       | xs => error ("Missing equivariance theorem for predicate(s): " ^
           commas_quote xs));
@@ -222,8 +221,8 @@
     val ind_sort = if null atomTs then HOLogic.typeS
       else Sign.certify_sort thy (map (fn a => Sign.intern_class thy
         ("fs_" ^ Sign.base_name a)) atoms);
-    val fs_ctxt_tyname = Name.variant (map fst (term_tfrees raw_induct')) "'n";
-    val fs_ctxt_name = Name.variant (add_term_names (raw_induct', [])) "z";
+    val fs_ctxt_tyname = Name.variant (map fst (OldTerm.term_tfrees raw_induct')) "'n";
+    val fs_ctxt_name = Name.variant (OldTerm.add_term_names (raw_induct', [])) "z";
     val fsT = TFree (fs_ctxt_tyname, ind_sort);
 
     val inductive_forall_def' = Drule.instantiate'
@@ -254,7 +253,7 @@
         val prem = Logic.list_implies
           (map mk_fresh sets @
            map (fn prem =>
-             if null (term_frees prem inter ps) then prem
+             if null (OldTerm.term_frees prem inter ps) then prem
              else lift_prem prem) prems,
            HOLogic.mk_Trueprop (lift_pred p ts));
       in abs_params params' prem end) prems);
@@ -277,7 +276,7 @@
     val (vc_compat, vc_compat') = map (fn (params, sets, prems, (p, ts)) =>
       map (fn q => abs_params params (incr_boundvars ~1 (Logic.list_implies
           (List.mapPartial (fn prem =>
-             if null (ps inter term_frees prem) then SOME prem
+             if null (ps inter OldTerm.term_frees prem) then SOME prem
              else map_term (split_conj (K o I) names) prem prem) prems, q))))
         (maps (fn (t, T) => map (fn (u, U) => HOLogic.mk_Trueprop
            (NominalPackage.fresh_star_const U T $ u $ t)) sets)
@@ -364,7 +363,7 @@
                    fold_rev (NominalPackage.mk_perm []) pis t) sets';
                  val (P, ts) = strip_comb (HOLogic.dest_Trueprop (term_of concl));
                  val gprems1 = List.mapPartial (fn (th, t) =>
-                   if null (term_frees t inter ps) then SOME th
+                   if null (OldTerm.term_frees t inter ps) then SOME th
                    else
                      map_thm ctxt' (split_conj (K o I) names)
                        (etac conjunct1 1) monos NONE th)
@@ -406,7 +405,7 @@
                        (fold_rev (mk_perm_bool o cterm_of thy)
                          (pis' @ pis) th));
                  val gprems2 = map (fn (th, t) =>
-                   if null (term_frees t inter ps) then mk_pi th
+                   if null (OldTerm.term_frees t inter ps) then mk_pi th
                    else
                      mk_pi (the (map_thm ctxt (inst_conj_all names ps (rev pis''))
                        (inst_conj_all_tac (length pis'')) monos (SOME t) th)))
--- a/src/HOL/Nominal/nominal_package.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Nominal/nominal_package.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Nominal/nominal_package.ML
-    ID:         $Id$
     Author:     Stefan Berghofer and Christian Urban, TU Muenchen
 
 Nominal datatype package for Isabelle/HOL.
@@ -572,7 +571,7 @@
         InductivePackage.add_inductive_global (serial_string ())
           {quiet_mode = false, verbose = false, kind = Thm.internalK,
            alt_name = Binding.name big_rep_name, coind = false, no_elim = true, no_ind = false,
-           skip_mono = true}
+           skip_mono = true, fork_mono = false}
           (map (fn (s, T) => ((Binding.name s, T --> HOLogic.boolT), NoSyn))
              (rep_set_names' ~~ recTs'))
           [] (map (fn x => (Attrib.empty_binding, x)) intr_ts) [] thy3;
@@ -1414,7 +1413,7 @@
 
     val _ = warning "defining recursion combinator ...";
 
-    val used = foldr add_typ_tfree_names [] recTs;
+    val used = foldr OldTerm.add_typ_tfree_names [] recTs;
 
     val (rec_result_Ts', rec_fn_Ts') = DatatypeProp.make_primrec_Ts descr' sorts used;
 
@@ -1510,7 +1509,7 @@
         InductivePackage.add_inductive_global (serial_string ())
           {quiet_mode = ! quiet_mode, verbose = false, kind = Thm.internalK,
            alt_name = Binding.name big_rec_name, coind = false, no_elim = false, no_ind = false,
-           skip_mono = true}
+           skip_mono = true, fork_mono = false}
           (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts))
           (map dest_Free rec_fns)
           (map (fn x => (Attrib.empty_binding, x)) rec_intr_ts) [] ||>
--- a/src/HOL/Nominal/nominal_primrec.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Nominal/nominal_primrec.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,6 @@
 (*  Title:      HOL/Nominal/nominal_primrec.ML
-    Author:     Stefan Berghofer, TU Muenchen and Norbert Voelker, FernUni Hagen
+    Author:     Norbert Voelker, FernUni Hagen
+    Author:     Stefan Berghofer, TU Muenchen
 
 Package for defining functions on nominal datatypes by primitive recursion.
 Taken from HOL/Tools/primrec_package.ML
@@ -71,7 +72,7 @@
     else
      (check_vars "repeated variable names in pattern: " (duplicates (op =) lfrees);
       check_vars "extra variables on rhs: "
-        (map dest_Free (term_frees rhs) |> subtract (op =) lfrees
+        (map dest_Free (OldTerm.term_frees rhs) |> subtract (op =) lfrees
           |> filter_out (is_fixed o fst));
       case AList.lookup (op =) rec_fns fname of
         NONE =>
@@ -154,7 +155,7 @@
               val recs = filter (is_rec_type o snd) (cargs' ~~ cargs);
               val rargs = map fst recs;
               val subs = map (rpair dummyT o fst)
-                (rev (rename_wrt_term rhs rargs));
+                (rev (Term.rename_wrt_term rhs rargs));
               val (rhs', (fnames'', fnss'')) = subst (map2 (fn (x, y) => fn z =>
                 (Free x, (body_index y, Free z))) recs subs) rhs (fnames', fnss')
                   handle RecError s => primrec_eq_err lthy s eq
--- a/src/HOL/NumberTheory/IntPrimes.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/NumberTheory/IntPrimes.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -94,7 +94,7 @@
 
 lemma zgcd_zdvd_zgcd_zmult: "zgcd m n dvd zgcd (k * m) n"
   apply (simp add: zgcd_greatest_iff)
-  apply (blast intro: zdvd_trans)
+  apply (blast intro: zdvd_trans dvd_triv_right)
   done
 
 lemma zgcd_zmult_zdvd_zgcd:
--- a/src/HOL/OrderedGroup.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/OrderedGroup.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -1,7 +1,5 @@
 (*  Title:   HOL/OrderedGroup.thy
-    ID:      $Id$
-    Author:  Gertrud Bauer, Steven Obua, Lawrence C Paulson, and Markus Wenzel,
-             with contributions by Jeremy Avigad
+    Author:  Gertrud Bauer, Steven Obua, Lawrence C Paulson, Markus Wenzel, Jeremy Avigad
 *)
 
 header {* Ordered Groups *}
@@ -1233,7 +1231,7 @@
 qed
 
 subclass pordered_ab_group_add_abs
-proof -
+proof
   have abs_ge_zero [simp]: "\<And>a. 0 \<le> \<bar>a\<bar>"
   proof -
     fix a b
@@ -1242,37 +1240,26 @@
   qed
   have abs_leI: "\<And>a b. a \<le> b \<Longrightarrow> - a \<le> b \<Longrightarrow> \<bar>a\<bar> \<le> b"
     by (simp add: abs_lattice le_supI)
-  show ?thesis
-  proof
-    fix a
-    show "0 \<le> \<bar>a\<bar>" by simp
-  next
-    fix a
-    show "a \<le> \<bar>a\<bar>"
-      by (auto simp add: abs_lattice)
-  next
-    fix a
-    show "\<bar>-a\<bar> = \<bar>a\<bar>"
-      by (simp add: abs_lattice sup_commute)
-  next
-    fix a b
-    show "a \<le> b \<Longrightarrow> - a \<le> b \<Longrightarrow> \<bar>a\<bar> \<le> b" by (erule abs_leI)
-  next
-    fix a b
-    show "\<bar>a + b\<bar> \<le> \<bar>a\<bar> + \<bar>b\<bar>"
-    proof -
-      have g:"abs a + abs b = sup (a+b) (sup (-a-b) (sup (-a+b) (a + (-b))))" (is "_=sup ?m ?n")
-        by (simp add: abs_lattice add_sup_inf_distribs sup_ACI diff_minus)
-      have a:"a+b <= sup ?m ?n" by (simp)
-      have b:"-a-b <= ?n" by (simp) 
-      have c:"?n <= sup ?m ?n" by (simp)
-      from b c have d: "-a-b <= sup ?m ?n" by(rule order_trans)
-      have e:"-a-b = -(a+b)" by (simp add: diff_minus)
-      from a d e have "abs(a+b) <= sup ?m ?n" 
-        by (drule_tac abs_leI, auto)
-      with g[symmetric] show ?thesis by simp
-    qed
-  qed auto
+  fix a b
+  show "0 \<le> \<bar>a\<bar>" by simp
+  show "a \<le> \<bar>a\<bar>"
+    by (auto simp add: abs_lattice)
+  show "\<bar>-a\<bar> = \<bar>a\<bar>"
+    by (simp add: abs_lattice sup_commute)
+  show "a \<le> b \<Longrightarrow> - a \<le> b \<Longrightarrow> \<bar>a\<bar> \<le> b" by (fact abs_leI)
+  show "\<bar>a + b\<bar> \<le> \<bar>a\<bar> + \<bar>b\<bar>"
+  proof -
+    have g:"abs a + abs b = sup (a+b) (sup (-a-b) (sup (-a+b) (a + (-b))))" (is "_=sup ?m ?n")
+      by (simp add: abs_lattice add_sup_inf_distribs sup_ACI diff_minus)
+    have a:"a+b <= sup ?m ?n" by (simp)
+    have b:"-a-b <= ?n" by (simp) 
+    have c:"?n <= sup ?m ?n" by (simp)
+    from b c have d: "-a-b <= sup ?m ?n" by(rule order_trans)
+    have e:"-a-b = -(a+b)" by (simp add: diff_minus)
+    from a d e have "abs(a+b) <= sup ?m ?n" 
+      by (drule_tac abs_leI, auto)
+    with g[symmetric] show ?thesis by simp
+  qed
 qed
 
 end
@@ -1376,7 +1363,7 @@
         @{const_name HOL.uminus}, @{const_name HOL.minus}]
   | agrp_ord _ = ~1;
 
-fun termless_agrp (a, b) = (Term.term_lpo agrp_ord (a, b) = LESS);
+fun termless_agrp (a, b) = (TermOrd.term_lpo agrp_ord (a, b) = LESS);
 
 local
   val ac1 = mk_meta_eq @{thm add_assoc};
--- a/src/HOL/Plain.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Plain.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -1,15 +1,14 @@
-(*  Title:      HOL/Plain.thy
-    ID:         $Id$
-
-Contains fundamental HOL tools and packages.  Does not include Hilbert_Choice.
-*)
-
 header {* Plain HOL *}
 
 theory Plain
 imports Datatype FunDef Record SAT Extraction
 begin
 
+text {*
+  Plain bootstrap of fundamental HOL tools and packages; does not
+  include @{text Hilbert_Choice}.
+*}
+
 ML {* path_add "~~/src/HOL/Library" *}
 
 end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Polynomial.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -0,0 +1,1132 @@
+(*  Title:      HOL/Polynomial.thy
+    Author:     Brian Huffman
+                Based on an earlier development by Clemens Ballarin
+*)
+
+header {* Univariate Polynomials *}
+
+theory Polynomial
+imports Plain SetInterval
+begin
+
+subsection {* Definition of type @{text poly} *}
+
+typedef (Poly) 'a poly = "{f::nat \<Rightarrow> 'a::zero. \<exists>n. \<forall>i>n. f i = 0}"
+  morphisms coeff Abs_poly
+  by auto
+
+lemma expand_poly_eq: "p = q \<longleftrightarrow> (\<forall>n. coeff p n = coeff q n)"
+by (simp add: coeff_inject [symmetric] expand_fun_eq)
+
+lemma poly_ext: "(\<And>n. coeff p n = coeff q n) \<Longrightarrow> p = q"
+by (simp add: expand_poly_eq)
+
+
+subsection {* Degree of a polynomial *}
+
+definition
+  degree :: "'a::zero poly \<Rightarrow> nat" where
+  "degree p = (LEAST n. \<forall>i>n. coeff p i = 0)"
+
+lemma coeff_eq_0: "degree p < n \<Longrightarrow> coeff p n = 0"
+proof -
+  have "coeff p \<in> Poly"
+    by (rule coeff)
+  hence "\<exists>n. \<forall>i>n. coeff p i = 0"
+    unfolding Poly_def by simp
+  hence "\<forall>i>degree p. coeff p i = 0"
+    unfolding degree_def by (rule LeastI_ex)
+  moreover assume "degree p < n"
+  ultimately show ?thesis by simp
+qed
+
+lemma le_degree: "coeff p n \<noteq> 0 \<Longrightarrow> n \<le> degree p"
+  by (erule contrapos_np, rule coeff_eq_0, simp)
+
+lemma degree_le: "\<forall>i>n. coeff p i = 0 \<Longrightarrow> degree p \<le> n"
+  unfolding degree_def by (erule Least_le)
+
+lemma less_degree_imp: "n < degree p \<Longrightarrow> \<exists>i>n. coeff p i \<noteq> 0"
+  unfolding degree_def by (drule not_less_Least, simp)
+
+
+subsection {* The zero polynomial *}
+
+instantiation poly :: (zero) zero
+begin
+
+definition
+  zero_poly_def: "0 = Abs_poly (\<lambda>n. 0)"
+
+instance ..
+end
+
+lemma coeff_0 [simp]: "coeff 0 n = 0"
+  unfolding zero_poly_def
+  by (simp add: Abs_poly_inverse Poly_def)
+
+lemma degree_0 [simp]: "degree 0 = 0"
+  by (rule order_antisym [OF degree_le le0]) simp
+
+lemma leading_coeff_neq_0:
+  assumes "p \<noteq> 0" shows "coeff p (degree p) \<noteq> 0"
+proof (cases "degree p")
+  case 0
+  from `p \<noteq> 0` have "\<exists>n. coeff p n \<noteq> 0"
+    by (simp add: expand_poly_eq)
+  then obtain n where "coeff p n \<noteq> 0" ..
+  hence "n \<le> degree p" by (rule le_degree)
+  with `coeff p n \<noteq> 0` and `degree p = 0`
+  show "coeff p (degree p) \<noteq> 0" by simp
+next
+  case (Suc n)
+  from `degree p = Suc n` have "n < degree p" by simp
+  hence "\<exists>i>n. coeff p i \<noteq> 0" by (rule less_degree_imp)
+  then obtain i where "n < i" and "coeff p i \<noteq> 0" by fast
+  from `degree p = Suc n` and `n < i` have "degree p \<le> i" by simp
+  also from `coeff p i \<noteq> 0` have "i \<le> degree p" by (rule le_degree)
+  finally have "degree p = i" .
+  with `coeff p i \<noteq> 0` show "coeff p (degree p) \<noteq> 0" by simp
+qed
+
+lemma leading_coeff_0_iff [simp]: "coeff p (degree p) = 0 \<longleftrightarrow> p = 0"
+  by (cases "p = 0", simp, simp add: leading_coeff_neq_0)
+
+
+subsection {* List-style constructor for polynomials *}
+
+definition
+  pCons :: "'a::zero \<Rightarrow> 'a poly \<Rightarrow> 'a poly"
+where
+  [code del]: "pCons a p = Abs_poly (nat_case a (coeff p))"
+
+syntax
+  "_poly" :: "args \<Rightarrow> 'a poly"  ("[:(_):]")
+
+translations
+  "[:x, xs:]" == "CONST pCons x [:xs:]"
+  "[:x:]" == "CONST pCons x 0"
+
+lemma Poly_nat_case: "f \<in> Poly \<Longrightarrow> nat_case a f \<in> Poly"
+  unfolding Poly_def by (auto split: nat.split)
+
+lemma coeff_pCons:
+  "coeff (pCons a p) = nat_case a (coeff p)"
+  unfolding pCons_def
+  by (simp add: Abs_poly_inverse Poly_nat_case coeff)
+
+lemma coeff_pCons_0 [simp]: "coeff (pCons a p) 0 = a"
+  by (simp add: coeff_pCons)
+
+lemma coeff_pCons_Suc [simp]: "coeff (pCons a p) (Suc n) = coeff p n"
+  by (simp add: coeff_pCons)
+
+lemma degree_pCons_le: "degree (pCons a p) \<le> Suc (degree p)"
+by (rule degree_le, simp add: coeff_eq_0 coeff_pCons split: nat.split)
+
+lemma degree_pCons_eq:
+  "p \<noteq> 0 \<Longrightarrow> degree (pCons a p) = Suc (degree p)"
+apply (rule order_antisym [OF degree_pCons_le])
+apply (rule le_degree, simp)
+done
+
+lemma degree_pCons_0: "degree (pCons a 0) = 0"
+apply (rule order_antisym [OF _ le0])
+apply (rule degree_le, simp add: coeff_pCons split: nat.split)
+done
+
+lemma degree_pCons_eq_if [simp]:
+  "degree (pCons a p) = (if p = 0 then 0 else Suc (degree p))"
+apply (cases "p = 0", simp_all)
+apply (rule order_antisym [OF _ le0])
+apply (rule degree_le, simp add: coeff_pCons split: nat.split)
+apply (rule order_antisym [OF degree_pCons_le])
+apply (rule le_degree, simp)
+done
+
+lemma pCons_0_0 [simp]: "pCons 0 0 = 0"
+by (rule poly_ext, simp add: coeff_pCons split: nat.split)
+
+lemma pCons_eq_iff [simp]:
+  "pCons a p = pCons b q \<longleftrightarrow> a = b \<and> p = q"
+proof (safe)
+  assume "pCons a p = pCons b q"
+  then have "coeff (pCons a p) 0 = coeff (pCons b q) 0" by simp
+  then show "a = b" by simp
+next
+  assume "pCons a p = pCons b q"
+  then have "\<forall>n. coeff (pCons a p) (Suc n) =
+                 coeff (pCons b q) (Suc n)" by simp
+  then show "p = q" by (simp add: expand_poly_eq)
+qed
+
+lemma pCons_eq_0_iff [simp]: "pCons a p = 0 \<longleftrightarrow> a = 0 \<and> p = 0"
+  using pCons_eq_iff [of a p 0 0] by simp
+
+lemma Poly_Suc: "f \<in> Poly \<Longrightarrow> (\<lambda>n. f (Suc n)) \<in> Poly"
+  unfolding Poly_def
+  by (clarify, rule_tac x=n in exI, simp)
+
+lemma pCons_cases [cases type: poly]:
+  obtains (pCons) a q where "p = pCons a q"
+proof
+  show "p = pCons (coeff p 0) (Abs_poly (\<lambda>n. coeff p (Suc n)))"
+    by (rule poly_ext)
+       (simp add: Abs_poly_inverse Poly_Suc coeff coeff_pCons
+             split: nat.split)
+qed
+
+lemma pCons_induct [case_names 0 pCons, induct type: poly]:
+  assumes zero: "P 0"
+  assumes pCons: "\<And>a p. P p \<Longrightarrow> P (pCons a p)"
+  shows "P p"
+proof (induct p rule: measure_induct_rule [where f=degree])
+  case (less p)
+  obtain a q where "p = pCons a q" by (rule pCons_cases)
+  have "P q"
+  proof (cases "q = 0")
+    case True
+    then show "P q" by (simp add: zero)
+  next
+    case False
+    then have "degree (pCons a q) = Suc (degree q)"
+      by (rule degree_pCons_eq)
+    then have "degree q < degree p"
+      using `p = pCons a q` by simp
+    then show "P q"
+      by (rule less.hyps)
+  qed
+  then have "P (pCons a q)"
+    by (rule pCons)
+  then show ?case
+    using `p = pCons a q` by simp
+qed
+
+
+subsection {* Recursion combinator for polynomials *}
+
+function
+  poly_rec :: "'b \<Rightarrow> ('a::zero \<Rightarrow> 'a poly \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a poly \<Rightarrow> 'b"
+where
+  poly_rec_pCons_eq_if [simp del, code del]:
+    "poly_rec z f (pCons a p) = f a p (if p = 0 then z else poly_rec z f p)"
+by (case_tac x, rename_tac q, case_tac q, auto)
+
+termination poly_rec
+by (relation "measure (degree \<circ> snd \<circ> snd)", simp)
+   (simp add: degree_pCons_eq)
+
+lemma poly_rec_0:
+  "f 0 0 z = z \<Longrightarrow> poly_rec z f 0 = z"
+  using poly_rec_pCons_eq_if [of z f 0 0] by simp
+
+lemma poly_rec_pCons:
+  "f 0 0 z = z \<Longrightarrow> poly_rec z f (pCons a p) = f a p (poly_rec z f p)"
+  by (simp add: poly_rec_pCons_eq_if poly_rec_0)
+
+
+subsection {* Monomials *}
+
+definition
+  monom :: "'a \<Rightarrow> nat \<Rightarrow> 'a::zero poly" where
+  "monom a m = Abs_poly (\<lambda>n. if m = n then a else 0)"
+
+lemma coeff_monom [simp]: "coeff (monom a m) n = (if m=n then a else 0)"
+  unfolding monom_def
+  by (subst Abs_poly_inverse, auto simp add: Poly_def)
+
+lemma monom_0: "monom a 0 = pCons a 0"
+  by (rule poly_ext, simp add: coeff_pCons split: nat.split)
+
+lemma monom_Suc: "monom a (Suc n) = pCons 0 (monom a n)"
+  by (rule poly_ext, simp add: coeff_pCons split: nat.split)
+
+lemma monom_eq_0 [simp]: "monom 0 n = 0"
+  by (rule poly_ext) simp
+
+lemma monom_eq_0_iff [simp]: "monom a n = 0 \<longleftrightarrow> a = 0"
+  by (simp add: expand_poly_eq)
+
+lemma monom_eq_iff [simp]: "monom a n = monom b n \<longleftrightarrow> a = b"
+  by (simp add: expand_poly_eq)
+
+lemma degree_monom_le: "degree (monom a n) \<le> n"
+  by (rule degree_le, simp)
+
+lemma degree_monom_eq: "a \<noteq> 0 \<Longrightarrow> degree (monom a n) = n"
+  apply (rule order_antisym [OF degree_monom_le])
+  apply (rule le_degree, simp)
+  done
+
+
+subsection {* Addition and subtraction *}
+
+instantiation poly :: (comm_monoid_add) comm_monoid_add
+begin
+
+definition
+  plus_poly_def [code del]:
+    "p + q = Abs_poly (\<lambda>n. coeff p n + coeff q n)"
+
+lemma Poly_add:
+  fixes f g :: "nat \<Rightarrow> 'a"
+  shows "\<lbrakk>f \<in> Poly; g \<in> Poly\<rbrakk> \<Longrightarrow> (\<lambda>n. f n + g n) \<in> Poly"
+  unfolding Poly_def
+  apply (clarify, rename_tac m n)
+  apply (rule_tac x="max m n" in exI, simp)
+  done
+
+lemma coeff_add [simp]:
+  "coeff (p + q) n = coeff p n + coeff q n"
+  unfolding plus_poly_def
+  by (simp add: Abs_poly_inverse coeff Poly_add)
+
+instance proof
+  fix p q r :: "'a poly"
+  show "(p + q) + r = p + (q + r)"
+    by (simp add: expand_poly_eq add_assoc)
+  show "p + q = q + p"
+    by (simp add: expand_poly_eq add_commute)
+  show "0 + p = p"
+    by (simp add: expand_poly_eq)
+qed
+
+end
+
+instance poly ::
+  ("{cancel_ab_semigroup_add,comm_monoid_add}") cancel_ab_semigroup_add
+proof
+  fix p q r :: "'a poly"
+  assume "p + q = p + r" thus "q = r"
+    by (simp add: expand_poly_eq)
+qed
+
+instantiation poly :: (ab_group_add) ab_group_add
+begin
+
+definition
+  uminus_poly_def [code del]:
+    "- p = Abs_poly (\<lambda>n. - coeff p n)"
+
+definition
+  minus_poly_def [code del]:
+    "p - q = Abs_poly (\<lambda>n. coeff p n - coeff q n)"
+
+lemma Poly_minus:
+  fixes f :: "nat \<Rightarrow> 'a"
+  shows "f \<in> Poly \<Longrightarrow> (\<lambda>n. - f n) \<in> Poly"
+  unfolding Poly_def by simp
+
+lemma Poly_diff:
+  fixes f g :: "nat \<Rightarrow> 'a"
+  shows "\<lbrakk>f \<in> Poly; g \<in> Poly\<rbrakk> \<Longrightarrow> (\<lambda>n. f n - g n) \<in> Poly"
+  unfolding diff_minus by (simp add: Poly_add Poly_minus)
+
+lemma coeff_minus [simp]: "coeff (- p) n = - coeff p n"
+  unfolding uminus_poly_def
+  by (simp add: Abs_poly_inverse coeff Poly_minus)
+
+lemma coeff_diff [simp]:
+  "coeff (p - q) n = coeff p n - coeff q n"
+  unfolding minus_poly_def
+  by (simp add: Abs_poly_inverse coeff Poly_diff)
+
+instance proof
+  fix p q :: "'a poly"
+  show "- p + p = 0"
+    by (simp add: expand_poly_eq)
+  show "p - q = p + - q"
+    by (simp add: expand_poly_eq diff_minus)
+qed
+
+end
+
+lemma add_pCons [simp]:
+  "pCons a p + pCons b q = pCons (a + b) (p + q)"
+  by (rule poly_ext, simp add: coeff_pCons split: nat.split)
+
+lemma minus_pCons [simp]:
+  "- pCons a p = pCons (- a) (- p)"
+  by (rule poly_ext, simp add: coeff_pCons split: nat.split)
+
+lemma diff_pCons [simp]:
+  "pCons a p - pCons b q = pCons (a - b) (p - q)"
+  by (rule poly_ext, simp add: coeff_pCons split: nat.split)
+
+lemma degree_add_le_max: "degree (p + q) \<le> max (degree p) (degree q)"
+  by (rule degree_le, auto simp add: coeff_eq_0)
+
+lemma degree_add_le:
+  "\<lbrakk>degree p \<le> n; degree q \<le> n\<rbrakk> \<Longrightarrow> degree (p + q) \<le> n"
+  by (auto intro: order_trans degree_add_le_max)
+
+lemma degree_add_less:
+  "\<lbrakk>degree p < n; degree q < n\<rbrakk> \<Longrightarrow> degree (p + q) < n"
+  by (auto intro: le_less_trans degree_add_le_max)
+
+lemma degree_add_eq_right:
+  "degree p < degree q \<Longrightarrow> degree (p + q) = degree q"
+  apply (cases "q = 0", simp)
+  apply (rule order_antisym)
+  apply (simp add: degree_add_le)
+  apply (rule le_degree)
+  apply (simp add: coeff_eq_0)
+  done
+
+lemma degree_add_eq_left:
+  "degree q < degree p \<Longrightarrow> degree (p + q) = degree p"
+  using degree_add_eq_right [of q p]
+  by (simp add: add_commute)
+
+lemma degree_minus [simp]: "degree (- p) = degree p"
+  unfolding degree_def by simp
+
+lemma degree_diff_le_max: "degree (p - q) \<le> max (degree p) (degree q)"
+  using degree_add_le [where p=p and q="-q"]
+  by (simp add: diff_minus)
+
+lemma degree_diff_le:
+  "\<lbrakk>degree p \<le> n; degree q \<le> n\<rbrakk> \<Longrightarrow> degree (p - q) \<le> n"
+  by (simp add: diff_minus degree_add_le)
+
+lemma degree_diff_less:
+  "\<lbrakk>degree p < n; degree q < n\<rbrakk> \<Longrightarrow> degree (p - q) < n"
+  by (simp add: diff_minus degree_add_less)
+
+lemma add_monom: "monom a n + monom b n = monom (a + b) n"
+  by (rule poly_ext) simp
+
+lemma diff_monom: "monom a n - monom b n = monom (a - b) n"
+  by (rule poly_ext) simp
+
+lemma minus_monom: "- monom a n = monom (-a) n"
+  by (rule poly_ext) simp
+
+lemma coeff_setsum: "coeff (\<Sum>x\<in>A. p x) i = (\<Sum>x\<in>A. coeff (p x) i)"
+  by (cases "finite A", induct set: finite, simp_all)
+
+lemma monom_setsum: "monom (\<Sum>x\<in>A. a x) n = (\<Sum>x\<in>A. monom (a x) n)"
+  by (rule poly_ext) (simp add: coeff_setsum)
+
+
+subsection {* Multiplication by a constant *}
+
+definition
+  smult :: "'a::comm_semiring_0 \<Rightarrow> 'a poly \<Rightarrow> 'a poly" where
+  "smult a p = Abs_poly (\<lambda>n. a * coeff p n)"
+
+lemma Poly_smult:
+  fixes f :: "nat \<Rightarrow> 'a::comm_semiring_0"
+  shows "f \<in> Poly \<Longrightarrow> (\<lambda>n. a * f n) \<in> Poly"
+  unfolding Poly_def
+  by (clarify, rule_tac x=n in exI, simp)
+
+lemma coeff_smult [simp]: "coeff (smult a p) n = a * coeff p n"
+  unfolding smult_def
+  by (simp add: Abs_poly_inverse Poly_smult coeff)
+
+lemma degree_smult_le: "degree (smult a p) \<le> degree p"
+  by (rule degree_le, simp add: coeff_eq_0)
+
+lemma smult_smult [simp]: "smult a (smult b p) = smult (a * b) p"
+  by (rule poly_ext, simp add: mult_assoc)
+
+lemma smult_0_right [simp]: "smult a 0 = 0"
+  by (rule poly_ext, simp)
+
+lemma smult_0_left [simp]: "smult 0 p = 0"
+  by (rule poly_ext, simp)
+
+lemma smult_1_left [simp]: "smult (1::'a::comm_semiring_1) p = p"
+  by (rule poly_ext, simp)
+
+lemma smult_add_right:
+  "smult a (p + q) = smult a p + smult a q"
+  by (rule poly_ext, simp add: ring_simps)
+
+lemma smult_add_left:
+  "smult (a + b) p = smult a p + smult b p"
+  by (rule poly_ext, simp add: ring_simps)
+
+lemma smult_minus_right [simp]:
+  "smult (a::'a::comm_ring) (- p) = - smult a p"
+  by (rule poly_ext, simp)
+
+lemma smult_minus_left [simp]:
+  "smult (- a::'a::comm_ring) p = - smult a p"
+  by (rule poly_ext, simp)
+
+lemma smult_diff_right:
+  "smult (a::'a::comm_ring) (p - q) = smult a p - smult a q"
+  by (rule poly_ext, simp add: ring_simps)
+
+lemma smult_diff_left:
+  "smult (a - b::'a::comm_ring) p = smult a p - smult b p"
+  by (rule poly_ext, simp add: ring_simps)
+
+lemmas smult_distribs =
+  smult_add_left smult_add_right
+  smult_diff_left smult_diff_right
+
+lemma smult_pCons [simp]:
+  "smult a (pCons b p) = pCons (a * b) (smult a p)"
+  by (rule poly_ext, simp add: coeff_pCons split: nat.split)
+
+lemma smult_monom: "smult a (monom b n) = monom (a * b) n"
+  by (induct n, simp add: monom_0, simp add: monom_Suc)
+
+
+subsection {* Multiplication of polynomials *}
+
+text {* TODO: move to SetInterval.thy *}
+lemma setsum_atMost_Suc_shift:
+  fixes f :: "nat \<Rightarrow> 'a::comm_monoid_add"
+  shows "(\<Sum>i\<le>Suc n. f i) = f 0 + (\<Sum>i\<le>n. f (Suc i))"
+proof (induct n)
+  case 0 show ?case by simp
+next
+  case (Suc n) note IH = this
+  have "(\<Sum>i\<le>Suc (Suc n). f i) = (\<Sum>i\<le>Suc n. f i) + f (Suc (Suc n))"
+    by (rule setsum_atMost_Suc)
+  also have "(\<Sum>i\<le>Suc n. f i) = f 0 + (\<Sum>i\<le>n. f (Suc i))"
+    by (rule IH)
+  also have "f 0 + (\<Sum>i\<le>n. f (Suc i)) + f (Suc (Suc n)) =
+             f 0 + ((\<Sum>i\<le>n. f (Suc i)) + f (Suc (Suc n)))"
+    by (rule add_assoc)
+  also have "(\<Sum>i\<le>n. f (Suc i)) + f (Suc (Suc n)) = (\<Sum>i\<le>Suc n. f (Suc i))"
+    by (rule setsum_atMost_Suc [symmetric])
+  finally show ?case .
+qed
+
+instantiation poly :: (comm_semiring_0) comm_semiring_0
+begin
+
+definition
+  times_poly_def [code del]:
+    "p * q = poly_rec 0 (\<lambda>a p pq. smult a q + pCons 0 pq) p"
+
+lemma mult_poly_0_left: "(0::'a poly) * q = 0"
+  unfolding times_poly_def by (simp add: poly_rec_0)
+
+lemma mult_pCons_left [simp]:
+  "pCons a p * q = smult a q + pCons 0 (p * q)"
+  unfolding times_poly_def by (simp add: poly_rec_pCons)
+
+lemma mult_poly_0_right: "p * (0::'a poly) = 0"
+  by (induct p, simp add: mult_poly_0_left, simp)
+
+lemma mult_pCons_right [simp]:
+  "p * pCons a q = smult a p + pCons 0 (p * q)"
+  by (induct p, simp add: mult_poly_0_left, simp add: ring_simps)
+
+lemmas mult_poly_0 = mult_poly_0_left mult_poly_0_right
+
+lemma mult_smult_left [simp]: "smult a p * q = smult a (p * q)"
+  by (induct p, simp add: mult_poly_0, simp add: smult_add_right)
+
+lemma mult_smult_right [simp]: "p * smult a q = smult a (p * q)"
+  by (induct q, simp add: mult_poly_0, simp add: smult_add_right)
+
+lemma mult_poly_add_left:
+  fixes p q r :: "'a poly"
+  shows "(p + q) * r = p * r + q * r"
+  by (induct r, simp add: mult_poly_0,
+                simp add: smult_distribs group_simps)
+
+instance proof
+  fix p q r :: "'a poly"
+  show 0: "0 * p = 0"
+    by (rule mult_poly_0_left)
+  show "p * 0 = 0"
+    by (rule mult_poly_0_right)
+  show "(p + q) * r = p * r + q * r"
+    by (rule mult_poly_add_left)
+  show "(p * q) * r = p * (q * r)"
+    by (induct p, simp add: mult_poly_0, simp add: mult_poly_add_left)
+  show "p * q = q * p"
+    by (induct p, simp add: mult_poly_0, simp)
+qed
+
+end
+
+instance poly :: (comm_semiring_0_cancel) comm_semiring_0_cancel ..
+
+lemma coeff_mult:
+  "coeff (p * q) n = (\<Sum>i\<le>n. coeff p i * coeff q (n-i))"
+proof (induct p arbitrary: n)
+  case 0 show ?case by simp
+next
+  case (pCons a p n) thus ?case
+    by (cases n, simp, simp add: setsum_atMost_Suc_shift
+                            del: setsum_atMost_Suc)
+qed
+
+lemma degree_mult_le: "degree (p * q) \<le> degree p + degree q"
+apply (rule degree_le)
+apply (induct p)
+apply simp
+apply (simp add: coeff_eq_0 coeff_pCons split: nat.split)
+done
+
+lemma mult_monom: "monom a m * monom b n = monom (a * b) (m + n)"
+  by (induct m, simp add: monom_0 smult_monom, simp add: monom_Suc)
+
+
+subsection {* The unit polynomial and exponentiation *}
+
+instantiation poly :: (comm_semiring_1) comm_semiring_1
+begin
+
+definition
+  one_poly_def:
+    "1 = pCons 1 0"
+
+instance proof
+  fix p :: "'a poly" show "1 * p = p"
+    unfolding one_poly_def
+    by simp
+next
+  show "0 \<noteq> (1::'a poly)"
+    unfolding one_poly_def by simp
+qed
+
+end
+
+instance poly :: (comm_semiring_1_cancel) comm_semiring_1_cancel ..
+
+lemma coeff_1 [simp]: "coeff 1 n = (if n = 0 then 1 else 0)"
+  unfolding one_poly_def
+  by (simp add: coeff_pCons split: nat.split)
+
+lemma degree_1 [simp]: "degree 1 = 0"
+  unfolding one_poly_def
+  by (rule degree_pCons_0)
+
+instantiation poly :: (comm_semiring_1) recpower
+begin
+
+primrec power_poly where
+  power_poly_0: "(p::'a poly) ^ 0 = 1"
+| power_poly_Suc: "(p::'a poly) ^ (Suc n) = p * p ^ n"
+
+instance
+  by default simp_all
+
+end
+
+instance poly :: (comm_ring) comm_ring ..
+
+instance poly :: (comm_ring_1) comm_ring_1 ..
+
+instantiation poly :: (comm_ring_1) number_ring
+begin
+
+definition
+  "number_of k = (of_int k :: 'a poly)"
+
+instance
+  by default (rule number_of_poly_def)
+
+end
+
+
+subsection {* Polynomials form an integral domain *}
+
+lemma coeff_mult_degree_sum:
+  "coeff (p * q) (degree p + degree q) =
+   coeff p (degree p) * coeff q (degree q)"
+  by (induct p, simp, simp add: coeff_eq_0)
+
+instance poly :: (idom) idom
+proof
+  fix p q :: "'a poly"
+  assume "p \<noteq> 0" and "q \<noteq> 0"
+  have "coeff (p * q) (degree p + degree q) =
+        coeff p (degree p) * coeff q (degree q)"
+    by (rule coeff_mult_degree_sum)
+  also have "coeff p (degree p) * coeff q (degree q) \<noteq> 0"
+    using `p \<noteq> 0` and `q \<noteq> 0` by simp
+  finally have "\<exists>n. coeff (p * q) n \<noteq> 0" ..
+  thus "p * q \<noteq> 0" by (simp add: expand_poly_eq)
+qed
+
+lemma degree_mult_eq:
+  fixes p q :: "'a::idom poly"
+  shows "\<lbrakk>p \<noteq> 0; q \<noteq> 0\<rbrakk> \<Longrightarrow> degree (p * q) = degree p + degree q"
+apply (rule order_antisym [OF degree_mult_le le_degree])
+apply (simp add: coeff_mult_degree_sum)
+done
+
+lemma dvd_imp_degree_le:
+  fixes p q :: "'a::idom poly"
+  shows "\<lbrakk>p dvd q; q \<noteq> 0\<rbrakk> \<Longrightarrow> degree p \<le> degree q"
+  by (erule dvdE, simp add: degree_mult_eq)
+
+
+subsection {* Long division of polynomials *}
+
+definition
+  pdivmod_rel :: "'a::field poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<Rightarrow> bool"
+where
+  [code del]:
+  "pdivmod_rel x y q r \<longleftrightarrow>
+    x = q * y + r \<and> (if y = 0 then q = 0 else r = 0 \<or> degree r < degree y)"
+
+lemma pdivmod_rel_0:
+  "pdivmod_rel 0 y 0 0"
+  unfolding pdivmod_rel_def by simp
+
+lemma pdivmod_rel_by_0:
+  "pdivmod_rel x 0 0 x"
+  unfolding pdivmod_rel_def by simp
+
+lemma eq_zero_or_degree_less:
+  assumes "degree p \<le> n" and "coeff p n = 0"
+  shows "p = 0 \<or> degree p < n"
+proof (cases n)
+  case 0
+  with `degree p \<le> n` and `coeff p n = 0`
+  have "coeff p (degree p) = 0" by simp
+  then have "p = 0" by simp
+  then show ?thesis ..
+next
+  case (Suc m)
+  have "\<forall>i>n. coeff p i = 0"
+    using `degree p \<le> n` by (simp add: coeff_eq_0)
+  then have "\<forall>i\<ge>n. coeff p i = 0"
+    using `coeff p n = 0` by (simp add: le_less)
+  then have "\<forall>i>m. coeff p i = 0"
+    using `n = Suc m` by (simp add: less_eq_Suc_le)
+  then have "degree p \<le> m"
+    by (rule degree_le)
+  then have "degree p < n"
+    using `n = Suc m` by (simp add: less_Suc_eq_le)
+  then show ?thesis ..
+qed
+
+lemma pdivmod_rel_pCons:
+  assumes rel: "pdivmod_rel x y q r"
+  assumes y: "y \<noteq> 0"
+  assumes b: "b = coeff (pCons a r) (degree y) / coeff y (degree y)"
+  shows "pdivmod_rel (pCons a x) y (pCons b q) (pCons a r - smult b y)"
+    (is "pdivmod_rel ?x y ?q ?r")
+proof -
+  have x: "x = q * y + r" and r: "r = 0 \<or> degree r < degree y"
+    using assms unfolding pdivmod_rel_def by simp_all
+
+  have 1: "?x = ?q * y + ?r"
+    using b x by simp
+
+  have 2: "?r = 0 \<or> degree ?r < degree y"
+  proof (rule eq_zero_or_degree_less)
+    show "degree ?r \<le> degree y"
+    proof (rule degree_diff_le)
+      show "degree (pCons a r) \<le> degree y"
+        using r by auto
+      show "degree (smult b y) \<le> degree y"
+        by (rule degree_smult_le)
+    qed
+  next
+    show "coeff ?r (degree y) = 0"
+      using `y \<noteq> 0` unfolding b by simp
+  qed
+
+  from 1 2 show ?thesis
+    unfolding pdivmod_rel_def
+    using `y \<noteq> 0` by simp
+qed
+
+lemma pdivmod_rel_exists: "\<exists>q r. pdivmod_rel x y q r"
+apply (cases "y = 0")
+apply (fast intro!: pdivmod_rel_by_0)
+apply (induct x)
+apply (fast intro!: pdivmod_rel_0)
+apply (fast intro!: pdivmod_rel_pCons)
+done
+
+lemma pdivmod_rel_unique:
+  assumes 1: "pdivmod_rel x y q1 r1"
+  assumes 2: "pdivmod_rel x y q2 r2"
+  shows "q1 = q2 \<and> r1 = r2"
+proof (cases "y = 0")
+  assume "y = 0" with assms show ?thesis
+    by (simp add: pdivmod_rel_def)
+next
+  assume [simp]: "y \<noteq> 0"
+  from 1 have q1: "x = q1 * y + r1" and r1: "r1 = 0 \<or> degree r1 < degree y"
+    unfolding pdivmod_rel_def by simp_all
+  from 2 have q2: "x = q2 * y + r2" and r2: "r2 = 0 \<or> degree r2 < degree y"
+    unfolding pdivmod_rel_def by simp_all
+  from q1 q2 have q3: "(q1 - q2) * y = r2 - r1"
+    by (simp add: ring_simps)
+  from r1 r2 have r3: "(r2 - r1) = 0 \<or> degree (r2 - r1) < degree y"
+    by (auto intro: degree_diff_less)
+
+  show "q1 = q2 \<and> r1 = r2"
+  proof (rule ccontr)
+    assume "\<not> (q1 = q2 \<and> r1 = r2)"
+    with q3 have "q1 \<noteq> q2" and "r1 \<noteq> r2" by auto
+    with r3 have "degree (r2 - r1) < degree y" by simp
+    also have "degree y \<le> degree (q1 - q2) + degree y" by simp
+    also have "\<dots> = degree ((q1 - q2) * y)"
+      using `q1 \<noteq> q2` by (simp add: degree_mult_eq)
+    also have "\<dots> = degree (r2 - r1)"
+      using q3 by simp
+    finally have "degree (r2 - r1) < degree (r2 - r1)" .
+    then show "False" by simp
+  qed
+qed
+
+lemmas pdivmod_rel_unique_div =
+  pdivmod_rel_unique [THEN conjunct1, standard]
+
+lemmas pdivmod_rel_unique_mod =
+  pdivmod_rel_unique [THEN conjunct2, standard]
+
+instantiation poly :: (field) ring_div
+begin
+
+definition div_poly where
+  [code del]: "x div y = (THE q. \<exists>r. pdivmod_rel x y q r)"
+
+definition mod_poly where
+  [code del]: "x mod y = (THE r. \<exists>q. pdivmod_rel x y q r)"
+
+lemma div_poly_eq:
+  "pdivmod_rel x y q r \<Longrightarrow> x div y = q"
+unfolding div_poly_def
+by (fast elim: pdivmod_rel_unique_div)
+
+lemma mod_poly_eq:
+  "pdivmod_rel x y q r \<Longrightarrow> x mod y = r"
+unfolding mod_poly_def
+by (fast elim: pdivmod_rel_unique_mod)
+
+lemma pdivmod_rel:
+  "pdivmod_rel x y (x div y) (x mod y)"
+proof -
+  from pdivmod_rel_exists
+    obtain q r where "pdivmod_rel x y q r" by fast
+  thus ?thesis
+    by (simp add: div_poly_eq mod_poly_eq)
+qed
+
+instance proof
+  fix x y :: "'a poly"
+  show "x div y * y + x mod y = x"
+    using pdivmod_rel [of x y]
+    by (simp add: pdivmod_rel_def)
+next
+  fix x :: "'a poly"
+  have "pdivmod_rel x 0 0 x"
+    by (rule pdivmod_rel_by_0)
+  thus "x div 0 = 0"
+    by (rule div_poly_eq)
+next
+  fix y :: "'a poly"
+  have "pdivmod_rel 0 y 0 0"
+    by (rule pdivmod_rel_0)
+  thus "0 div y = 0"
+    by (rule div_poly_eq)
+next
+  fix x y z :: "'a poly"
+  assume "y \<noteq> 0"
+  hence "pdivmod_rel (x + z * y) y (z + x div y) (x mod y)"
+    using pdivmod_rel [of x y]
+    by (simp add: pdivmod_rel_def left_distrib)
+  thus "(x + z * y) div y = z + x div y"
+    by (rule div_poly_eq)
+qed
+
+end
+
+lemma degree_mod_less:
+  "y \<noteq> 0 \<Longrightarrow> x mod y = 0 \<or> degree (x mod y) < degree y"
+  using pdivmod_rel [of x y]
+  unfolding pdivmod_rel_def by simp
+
+lemma div_poly_less: "degree x < degree y \<Longrightarrow> x div y = 0"
+proof -
+  assume "degree x < degree y"
+  hence "pdivmod_rel x y 0 x"
+    by (simp add: pdivmod_rel_def)
+  thus "x div y = 0" by (rule div_poly_eq)
+qed
+
+lemma mod_poly_less: "degree x < degree y \<Longrightarrow> x mod y = x"
+proof -
+  assume "degree x < degree y"
+  hence "pdivmod_rel x y 0 x"
+    by (simp add: pdivmod_rel_def)
+  thus "x mod y = x" by (rule mod_poly_eq)
+qed
+
+lemma mod_pCons:
+  fixes a and x
+  assumes y: "y \<noteq> 0"
+  defines b: "b \<equiv> coeff (pCons a (x mod y)) (degree y) / coeff y (degree y)"
+  shows "(pCons a x) mod y = (pCons a (x mod y) - smult b y)"
+unfolding b
+apply (rule mod_poly_eq)
+apply (rule pdivmod_rel_pCons [OF pdivmod_rel y refl])
+done
+
+
+subsection {* Evaluation of polynomials *}
+
+definition
+  poly :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a" where
+  "poly = poly_rec (\<lambda>x. 0) (\<lambda>a p f x. a + x * f x)"
+
+lemma poly_0 [simp]: "poly 0 x = 0"
+  unfolding poly_def by (simp add: poly_rec_0)
+
+lemma poly_pCons [simp]: "poly (pCons a p) x = a + x * poly p x"
+  unfolding poly_def by (simp add: poly_rec_pCons)
+
+lemma poly_1 [simp]: "poly 1 x = 1"
+  unfolding one_poly_def by simp
+
+lemma poly_monom:
+  fixes a x :: "'a::{comm_semiring_1,recpower}"
+  shows "poly (monom a n) x = a * x ^ n"
+  by (induct n, simp add: monom_0, simp add: monom_Suc power_Suc mult_ac)
+
+lemma poly_add [simp]: "poly (p + q) x = poly p x + poly q x"
+  apply (induct p arbitrary: q, simp)
+  apply (case_tac q, simp, simp add: ring_simps)
+  done
+
+lemma poly_minus [simp]:
+  fixes x :: "'a::comm_ring"
+  shows "poly (- p) x = - poly p x"
+  by (induct p, simp_all)
+
+lemma poly_diff [simp]:
+  fixes x :: "'a::comm_ring"
+  shows "poly (p - q) x = poly p x - poly q x"
+  by (simp add: diff_minus)
+
+lemma poly_setsum: "poly (\<Sum>k\<in>A. p k) x = (\<Sum>k\<in>A. poly (p k) x)"
+  by (cases "finite A", induct set: finite, simp_all)
+
+lemma poly_smult [simp]: "poly (smult a p) x = a * poly p x"
+  by (induct p, simp, simp add: ring_simps)
+
+lemma poly_mult [simp]: "poly (p * q) x = poly p x * poly q x"
+  by (induct p, simp_all, simp add: ring_simps)
+
+lemma poly_power [simp]:
+  fixes p :: "'a::{comm_semiring_1,recpower} poly"
+  shows "poly (p ^ n) x = poly p x ^ n"
+  by (induct n, simp, simp add: power_Suc)
+
+
+subsection {* Synthetic division *}
+
+definition
+  synthetic_divmod :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a poly \<times> 'a"
+where [code del]:
+  "synthetic_divmod p c =
+    poly_rec (0, 0) (\<lambda>a p (q, r). (pCons r q, a + c * r)) p"
+
+definition
+  synthetic_div :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a poly"
+where
+  "synthetic_div p c = fst (synthetic_divmod p c)"
+
+lemma synthetic_divmod_0 [simp]:
+  "synthetic_divmod 0 c = (0, 0)"
+  unfolding synthetic_divmod_def
+  by (simp add: poly_rec_0)
+
+lemma synthetic_divmod_pCons [simp]:
+  "synthetic_divmod (pCons a p) c =
+    (\<lambda>(q, r). (pCons r q, a + c * r)) (synthetic_divmod p c)"
+  unfolding synthetic_divmod_def
+  by (simp add: poly_rec_pCons)
+
+lemma snd_synthetic_divmod: "snd (synthetic_divmod p c) = poly p c"
+  by (induct p, simp, simp add: split_def)
+
+lemma synthetic_div_0 [simp]: "synthetic_div 0 c = 0"
+  unfolding synthetic_div_def by simp
+
+lemma synthetic_div_pCons [simp]:
+  "synthetic_div (pCons a p) c = pCons (poly p c) (synthetic_div p c)"
+  unfolding synthetic_div_def
+  by (simp add: split_def snd_synthetic_divmod)
+
+lemma synthetic_div_eq_0_iff:
+  "synthetic_div p c = 0 \<longleftrightarrow> degree p = 0"
+  by (induct p, simp, case_tac p, simp)
+
+lemma degree_synthetic_div:
+  "degree (synthetic_div p c) = degree p - 1"
+  by (induct p, simp, simp add: synthetic_div_eq_0_iff)
+
+lemma synthetic_div_correct:
+  "p + smult c (synthetic_div p c) = pCons (poly p c) (synthetic_div p c)"
+  by (induct p) simp_all
+
+lemma synthetic_div_unique_lemma: "smult c p = pCons a p \<Longrightarrow> p = 0"
+by (induct p arbitrary: a) simp_all
+
+lemma synthetic_div_unique:
+  "p + smult c q = pCons r q \<Longrightarrow> r = poly p c \<and> q = synthetic_div p c"
+apply (induct p arbitrary: q r)
+apply (simp, frule synthetic_div_unique_lemma, simp)
+apply (case_tac q, force)
+done
+
+lemma synthetic_div_correct':
+  fixes c :: "'a::comm_ring_1"
+  shows "[:-c, 1:] * synthetic_div p c + [:poly p c:] = p"
+  using synthetic_div_correct [of p c]
+  by (simp add: group_simps)
+
+lemma poly_eq_0_iff_dvd:
+  fixes c :: "'a::idom"
+  shows "poly p c = 0 \<longleftrightarrow> [:-c, 1:] dvd p"
+proof
+  assume "poly p c = 0"
+  with synthetic_div_correct' [of c p]
+  have "p = [:-c, 1:] * synthetic_div p c" by simp
+  then show "[:-c, 1:] dvd p" ..
+next
+  assume "[:-c, 1:] dvd p"
+  then obtain k where "p = [:-c, 1:] * k" by (rule dvdE)
+  then show "poly p c = 0" by simp
+qed
+
+lemma dvd_iff_poly_eq_0:
+  fixes c :: "'a::idom"
+  shows "[:c, 1:] dvd p \<longleftrightarrow> poly p (-c) = 0"
+  by (simp add: poly_eq_0_iff_dvd)
+
+lemma poly_roots_finite:
+  fixes p :: "'a::idom poly"
+  shows "p \<noteq> 0 \<Longrightarrow> finite {x. poly p x = 0}"
+proof (induct n \<equiv> "degree p" arbitrary: p)
+  case (0 p)
+  then obtain a where "a \<noteq> 0" and "p = [:a:]"
+    by (cases p, simp split: if_splits)
+  then show "finite {x. poly p x = 0}" by simp
+next
+  case (Suc n p)
+  show "finite {x. poly p x = 0}"
+  proof (cases "\<exists>x. poly p x = 0")
+    case False
+    then show "finite {x. poly p x = 0}" by simp
+  next
+    case True
+    then obtain a where "poly p a = 0" ..
+    then have "[:-a, 1:] dvd p" by (simp only: poly_eq_0_iff_dvd)
+    then obtain k where k: "p = [:-a, 1:] * k" ..
+    with `p \<noteq> 0` have "k \<noteq> 0" by auto
+    with k have "degree p = Suc (degree k)"
+      by (simp add: degree_mult_eq del: mult_pCons_left)
+    with `Suc n = degree p` have "n = degree k" by simp
+    with `k \<noteq> 0` have "finite {x. poly k x = 0}" by (rule Suc.hyps)
+    then have "finite (insert a {x. poly k x = 0})" by simp
+    then show "finite {x. poly p x = 0}"
+      by (simp add: k uminus_add_conv_diff Collect_disj_eq
+               del: mult_pCons_left)
+  qed
+qed
+
+
+subsection {* Configuration of the code generator *}
+
+code_datatype "0::'a::zero poly" pCons
+
+declare pCons_0_0 [code post]
+
+instantiation poly :: ("{zero,eq}") eq
+begin
+
+definition [code del]:
+  "eq_class.eq (p::'a poly) q \<longleftrightarrow> p = q"
+
+instance
+  by default (rule eq_poly_def)
+
+end
+
+lemma eq_poly_code [code]:
+  "eq_class.eq (0::_ poly) (0::_ poly) \<longleftrightarrow> True"
+  "eq_class.eq (0::_ poly) (pCons b q) \<longleftrightarrow> eq_class.eq 0 b \<and> eq_class.eq 0 q"
+  "eq_class.eq (pCons a p) (0::_ poly) \<longleftrightarrow> eq_class.eq a 0 \<and> eq_class.eq p 0"
+  "eq_class.eq (pCons a p) (pCons b q) \<longleftrightarrow> eq_class.eq a b \<and> eq_class.eq p q"
+unfolding eq by simp_all
+
+lemmas coeff_code [code] =
+  coeff_0 coeff_pCons_0 coeff_pCons_Suc
+
+lemmas degree_code [code] =
+  degree_0 degree_pCons_eq_if
+
+lemmas monom_poly_code [code] =
+  monom_0 monom_Suc
+
+lemma add_poly_code [code]:
+  "0 + q = (q :: _ poly)"
+  "p + 0 = (p :: _ poly)"
+  "pCons a p + pCons b q = pCons (a + b) (p + q)"
+by simp_all
+
+lemma minus_poly_code [code]:
+  "- 0 = (0 :: _ poly)"
+  "- pCons a p = pCons (- a) (- p)"
+by simp_all
+
+lemma diff_poly_code [code]:
+  "0 - q = (- q :: _ poly)"
+  "p - 0 = (p :: _ poly)"
+  "pCons a p - pCons b q = pCons (a - b) (p - q)"
+by simp_all
+
+lemmas smult_poly_code [code] =
+  smult_0_right smult_pCons
+
+lemma mult_poly_code [code]:
+  "0 * q = (0 :: _ poly)"
+  "pCons a p * q = smult a q + pCons 0 (p * q)"
+by simp_all
+
+lemmas poly_code [code] =
+  poly_0 poly_pCons
+
+lemmas synthetic_divmod_code [code] =
+  synthetic_divmod_0 synthetic_divmod_pCons
+
+text {* code generator setup for div and mod *}
+
+definition
+  pdivmod :: "'a::field poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<times> 'a poly"
+where
+  [code del]: "pdivmod x y = (x div y, x mod y)"
+
+lemma div_poly_code [code]: "x div y = fst (pdivmod x y)"
+  unfolding pdivmod_def by simp
+
+lemma mod_poly_code [code]: "x mod y = snd (pdivmod x y)"
+  unfolding pdivmod_def by simp
+
+lemma pdivmod_0 [code]: "pdivmod 0 y = (0, 0)"
+  unfolding pdivmod_def by simp
+
+lemma pdivmod_pCons [code]:
+  "pdivmod (pCons a x) y =
+    (if y = 0 then (0, pCons a x) else
+      (let (q, r) = pdivmod x y;
+           b = coeff (pCons a r) (degree y) / coeff y (degree y)
+        in (pCons b q, pCons a r - smult b y)))"
+apply (simp add: pdivmod_def Let_def, safe)
+apply (rule div_poly_eq)
+apply (erule pdivmod_rel_pCons [OF pdivmod_rel _ refl])
+apply (rule mod_poly_eq)
+apply (erule pdivmod_rel_pCons [OF pdivmod_rel _ refl])
+done
+
+end
--- a/src/HOL/ROOT.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/ROOT.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,9 +1,6 @@
-(*  Title:      HOL/ROOT.ML
- 
-Classical Higher-order Logic -- batteries included.
-*)
+(* Classical Higher-order Logic -- batteries included *)
 
+use_thy "Main";
 use_thy "Complex_Main";
 
 val HOL_proofs = ! Proofterm.proofs;
-
--- a/src/HOL/Rational.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Rational.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -851,7 +851,7 @@
 definition Fract_norm :: "int \<Rightarrow> int \<Rightarrow> rat" where
   [simp, code del]: "Fract_norm a b = Fract a b"
 
-lemma [code]: "Fract_norm a b = (if a = 0 \<or> b = 0 then 0 else let c = zgcd a b in
+lemma Fract_norm_code [code]: "Fract_norm a b = (if a = 0 \<or> b = 0 then 0 else let c = zgcd a b in
   if b > 0 then Fract (a div c) (b div c) else Fract (- (a div c)) (- (b div c)))"
   by (simp add: eq_rat Zero_rat_def Let_def Fract_norm)
 
@@ -871,7 +871,7 @@
        then c = 0 \<or> d = 0
      else if d = 0
        then a = 0 \<or> b = 0
-     else a * d =  b * c)"
+     else a * d = b * c)"
   by (auto simp add: eq eq_rat)
 
 lemma rat_eq_refl [code nbe]:
--- a/src/HOL/Ring_and_Field.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Ring_and_Field.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -235,19 +235,21 @@
 lemma minus_mult_right: "- (a * b) = a * - b"
   by (rule equals_zero_I) (simp add: right_distrib [symmetric]) 
 
+text{*Extract signs from products*}
+lemmas mult_minus_left [simp] = minus_mult_left [symmetric]
+lemmas mult_minus_right [simp] = minus_mult_right [symmetric]
+
 lemma minus_mult_minus [simp]: "- a * - b = a * b"
-  by (simp add: minus_mult_left [symmetric] minus_mult_right [symmetric])
+  by simp
 
 lemma minus_mult_commute: "- a * b = a * - b"
-  by (simp add: minus_mult_left [symmetric] minus_mult_right [symmetric])
+  by simp
 
 lemma right_diff_distrib: "a * (b - c) = a * b - a * c"
-  by (simp add: right_distrib diff_minus 
-    minus_mult_left [symmetric] minus_mult_right [symmetric]) 
+  by (simp add: right_distrib diff_minus)
 
 lemma left_diff_distrib: "(a - b) * c = a * c - b * c"
-  by (simp add: left_distrib diff_minus 
-    minus_mult_left [symmetric] minus_mult_right [symmetric]) 
+  by (simp add: left_distrib diff_minus)
 
 lemmas ring_distribs =
   right_distrib left_distrib left_diff_distrib right_diff_distrib
@@ -293,6 +295,33 @@
 subclass ring_1 ..
 subclass comm_semiring_1_cancel ..
 
+lemma dvd_minus_iff [simp]: "x dvd - y \<longleftrightarrow> x dvd y"
+proof
+  assume "x dvd - y"
+  then have "x dvd - 1 * - y" by (rule dvd_mult)
+  then show "x dvd y" by simp
+next
+  assume "x dvd y"
+  then have "x dvd - 1 * y" by (rule dvd_mult)
+  then show "x dvd - y" by simp
+qed
+
+lemma minus_dvd_iff [simp]: "- x dvd y \<longleftrightarrow> x dvd y"
+proof
+  assume "- x dvd y"
+  then obtain k where "y = - x * k" ..
+  then have "y = x * - k" by simp
+  then show "x dvd y" ..
+next
+  assume "x dvd y"
+  then obtain k where "y = x * k" ..
+  then have "y = - x * - k" by simp
+  then show "- x dvd y" ..
+qed
+
+lemma dvd_diff: "x dvd y \<Longrightarrow> x dvd z \<Longrightarrow> x dvd (y - z)"
+  by (simp add: diff_minus dvd_add dvd_minus_iff)
+
 end
 
 class ring_no_zero_divisors = ring + no_zero_divisors
@@ -397,67 +426,47 @@
 apply auto
 done
 
-lemma nonzero_inverse_minus_eq:
-  assumes "a \<noteq> 0"
-  shows "inverse (- a) = - inverse a"
-proof -
-  have "- a * inverse (- a) = - a * - inverse a"
-    using assms by simp
-  then show ?thesis unfolding mult_cancel_left using assms by simp 
-qed
-
-lemma nonzero_inverse_inverse_eq:
-  assumes "a \<noteq> 0"
-  shows "inverse (inverse a) = a"
-proof -
-  have "(inverse (inverse a) * inverse a) * a = a" 
-    using assms by (simp add: nonzero_imp_inverse_nonzero)
-  then show ?thesis using assms by (simp add: mult_assoc)
-qed
-
-lemma nonzero_inverse_eq_imp_eq:
-  assumes inveq: "inverse a = inverse b"
-    and anz:  "a \<noteq> 0"
-    and bnz:  "b \<noteq> 0"
-  shows "a = b"
-proof -
-  have "a * inverse b = a * inverse a"
-    by (simp add: inveq)
-  hence "(a * inverse b) * b = (a * inverse a) * b"
-    by simp
-  then show "a = b"
-    by (simp add: mult_assoc anz bnz)
-qed
-
-lemma inverse_1 [simp]: "inverse 1 = 1"
-proof -
-  have "inverse 1 * 1 = 1" 
-    by (rule left_inverse) (rule one_neq_zero)
-  then show ?thesis by simp
-qed
-
 lemma inverse_unique: 
   assumes ab: "a * b = 1"
   shows "inverse a = b"
 proof -
   have "a \<noteq> 0" using ab by (cases "a = 0") simp_all
-  moreover have "inverse a * (a * b) = inverse a" by (simp add: ab) 
-  ultimately show ?thesis by (simp add: mult_assoc [symmetric]) 
+  moreover have "inverse a * (a * b) = inverse a" by (simp add: ab)
+  ultimately show ?thesis by (simp add: mult_assoc [symmetric])
 qed
 
+lemma nonzero_inverse_minus_eq:
+  "a \<noteq> 0 \<Longrightarrow> inverse (- a) = - inverse a"
+  by (rule inverse_unique) simp
+
+lemma nonzero_inverse_inverse_eq:
+  "a \<noteq> 0 \<Longrightarrow> inverse (inverse a) = a"
+  by (rule inverse_unique) simp
+
+lemma nonzero_inverse_eq_imp_eq:
+  assumes "inverse a = inverse b" and "a \<noteq> 0" and "b \<noteq> 0"
+  shows "a = b"
+proof -
+  from `inverse a = inverse b`
+  have "inverse (inverse a) = inverse (inverse b)"
+    by (rule arg_cong)
+  with `a \<noteq> 0` and `b \<noteq> 0` show "a = b"
+    by (simp add: nonzero_inverse_inverse_eq)
+qed
+
+lemma inverse_1 [simp]: "inverse 1 = 1"
+  by (rule inverse_unique) simp
+
 lemma nonzero_inverse_mult_distrib: 
-  assumes anz: "a \<noteq> 0"
-    and bnz: "b \<noteq> 0"
+  assumes "a \<noteq> 0" and "b \<noteq> 0"
   shows "inverse (a * b) = inverse b * inverse a"
 proof -
-  have "inverse (a * b) * (a * b) * inverse b = inverse b" 
-    by (simp add: anz bnz)
-  hence "inverse (a * b) * a = inverse b" 
-    by (simp add: mult_assoc bnz)
-  hence "inverse (a * b) * a * inverse a = inverse b * inverse a" 
-    by simp
+  have "a * (b * inverse b) * inverse a = 1"
+    using assms by simp
+  hence "a * b * (inverse b * inverse a) = 1"
+    by (simp only: mult_assoc)
   thus ?thesis
-    by (simp add: mult_assoc anz)
+    by (rule inverse_unique)
 qed
 
 lemma division_ring_inverse_add:
@@ -1266,19 +1275,19 @@
 subsection {* Division and Unary Minus *}
 
 lemma nonzero_minus_divide_left: "b \<noteq> 0 ==> - (a/b) = (-a) / (b::'a::field)"
-by (simp add: divide_inverse minus_mult_left)
+by (simp add: divide_inverse)
 
 lemma nonzero_minus_divide_right: "b \<noteq> 0 ==> - (a/b) = a / -(b::'a::field)"
-by (simp add: divide_inverse nonzero_inverse_minus_eq minus_mult_right)
+by (simp add: divide_inverse nonzero_inverse_minus_eq)
 
 lemma nonzero_minus_divide_divide: "b \<noteq> 0 ==> (-a)/(-b) = a / (b::'a::field)"
 by (simp add: divide_inverse nonzero_inverse_minus_eq)
 
 lemma minus_divide_left: "- (a/b) = (-a) / (b::'a::field)"
-by (simp add: divide_inverse minus_mult_left [symmetric])
+by (simp add: divide_inverse)
 
 lemma minus_divide_right: "- (a/b) = a / -(b::'a::{field,division_by_zero})"
-by (simp add: divide_inverse minus_mult_right [symmetric])
+by (simp add: divide_inverse)
 
 
 text{*The effect is to extract signs from divisions*}
@@ -1286,11 +1295,6 @@
 lemmas divide_minus_right = minus_divide_right [symmetric]
 declare divide_minus_left [simp]   divide_minus_right [simp]
 
-text{*Also, extract signs from products*}
-lemmas mult_minus_left = minus_mult_left [symmetric]
-lemmas mult_minus_right = minus_mult_right [symmetric]
-declare mult_minus_left [simp]   mult_minus_right [simp]
-
 lemma minus_divide_divide [simp]:
   "(-a)/(-b) = a / (b::'a::{field,division_by_zero})"
 apply (cases "b=0", simp) 
--- a/src/HOL/Statespace/DistinctTreeProver.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Statespace/DistinctTreeProver.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
-(*  Title:      DistinctTreeProver.thy
-    ID:         $Id$
+(*  Title:      HOL/Statespace/DistinctTreeProver.thy
     Author:     Norbert Schirmer, TU Muenchen
 *)
 
@@ -453,7 +452,7 @@
   case Tip thus ?case by simp
 next
   case (Node l x d r)
-  have sub: "subtract (Node l x d r) t\<^isub>2 = Some t".
+  have sub: "subtract (Node l x d r) t\<^isub>2 = Some t" by fact
   show ?case
   proof (cases "delete x t\<^isub>2")
     case (Some t\<^isub>2')
@@ -646,9 +645,9 @@
 val const_decls = map (fn i => Syntax.no_syn 
                                  ("const" ^ string_of_int i,Type ("nat",[]))) nums
 
-val consts = sort Term.fast_term_ord 
+val consts = sort TermOrd.fast_term_ord 
               (map (fn i => Const ("DistinctTreeProver.const"^string_of_int i,Type ("nat",[]))) nums)
-val consts' = sort Term.fast_term_ord 
+val consts' = sort TermOrd.fast_term_ord 
               (map (fn i => Const ("DistinctTreeProver.const"^string_of_int i,Type ("nat",[]))) nums')
 
 val t = DistinctTreeProver.mk_tree I (Type ("nat",[])) consts
--- a/src/HOL/Statespace/StateSpaceEx.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Statespace/StateSpaceEx.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -41,7 +41,7 @@
 projection~/ injection functions that convert from abstract values to
 @{typ "nat"} and @{text "bool"}. The logical content of the locale is: *}
 
-class_locale vars' =
+locale vars' =
   fixes n::'name and b::'name
   assumes "distinct [n, b]" 
 
--- a/src/HOL/Statespace/distinct_tree_prover.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Statespace/distinct_tree_prover.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,28 +1,24 @@
-(*  Title:      distinct_tree_prover.thy
-    ID:         $Id$
+(*  Title:      HOL/Statespace/distinct_tree_prover.ML
     Author:     Norbert Schirmer, TU Muenchen
 *)
 
 signature DISTINCT_TREE_PROVER =
 sig
   datatype Direction = Left | Right
-  val mk_tree : ('a -> Term.term) -> Term.typ -> 'a list -> Term.term
-  val dest_tree : Term.term -> Term.term list
-  val find_tree :
-       Term.term -> Term.term -> Direction list option 
+  val mk_tree : ('a -> term) -> typ -> 'a list -> term
+  val dest_tree : term -> term list
+  val find_tree : term -> term -> Direction list option 
 
-  val neq_to_eq_False : Thm.thm
-  val distinctTreeProver : Thm.thm -> Direction list -> Direction list -> Thm.thm
-  val neq_x_y :
-       Proof.context -> Term.term -> Term.term -> string -> Thm.thm option   
-  val distinctFieldSolver : string list -> MetaSimplifier.solver
-  val distinctTree_tac :
-       string list -> Proof.context -> Term.term * int -> Tactical.tactic
-  val distinct_implProver : Thm.thm -> Thm.cterm -> Thm.thm
-  val subtractProver : Term.term -> Thm.cterm -> Thm.thm -> Thm.thm
-  val distinct_simproc : string list -> MetaSimplifier.simproc
+  val neq_to_eq_False : thm
+  val distinctTreeProver : thm -> Direction list -> Direction list -> thm
+  val neq_x_y : Proof.context -> term -> term -> string -> thm option   
+  val distinctFieldSolver : string list -> solver
+  val distinctTree_tac : string list -> Proof.context -> term * int -> tactic
+  val distinct_implProver : thm -> cterm -> thm
+  val subtractProver : term -> cterm -> thm -> thm
+  val distinct_simproc : string list -> simproc
   
-  val discharge : Thm.thm list -> Thm.thm -> Thm.thm
+  val discharge : thm list -> thm -> thm
 end;
 
 structure DistinctTreeProver : DISTINCT_TREE_PROVER =
@@ -83,7 +79,7 @@
   | bin_find_tree order e t = raise TERM ("find_tree: input not a tree",[t])
 
 fun find_tree e t =
-  (case bin_find_tree Term.fast_term_ord e t of
+  (case bin_find_tree TermOrd.fast_term_ord e t of
      NONE => lin_find_tree e t
    | x => x);
 
--- a/src/HOL/Statespace/state_fun.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Statespace/state_fun.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Statespace/state_fun.ML
-    ID:         $Id$
     Author:     Norbert Schirmer, TU Muenchen
 *)
 
@@ -8,17 +7,17 @@
   val lookupN : string
   val updateN : string
 
-  val mk_constr : Context.theory -> Term.typ -> Term.term
-  val mk_destr : Context.theory -> Term.typ -> Term.term
+  val mk_constr : theory -> typ -> term
+  val mk_destr : theory -> typ -> term
 
-  val lookup_simproc : MetaSimplifier.simproc
-  val update_simproc : MetaSimplifier.simproc
-  val ex_lookup_eq_simproc : MetaSimplifier.simproc
-  val ex_lookup_ss : MetaSimplifier.simpset
-  val lazy_conj_simproc : MetaSimplifier.simproc
-  val string_eq_simp_tac : int -> Tactical.tactic
+  val lookup_simproc : simproc
+  val update_simproc : simproc
+  val ex_lookup_eq_simproc : simproc
+  val ex_lookup_ss : simpset
+  val lazy_conj_simproc : simproc
+  val string_eq_simp_tac : int -> tactic
 
-  val setup : Context.theory -> Context.theory
+  val setup : theory -> theory
 end;
 
 structure StateFun: STATE_FUN = 
--- a/src/HOL/Statespace/state_space.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Statespace/state_space.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Statespace/state_space.ML
-    ID:         $Id$
     Author:     Norbert Schirmer, TU Muenchen
 *)
 
@@ -156,15 +155,13 @@
 fun add_locale name expr elems thy =
   thy 
   |> Expression.add_locale name name expr elems
-  |> (fn ((target, notes), ctxt) => TheoryTarget.begin target ctxt |>
-           fold (fn (kind, facts) => LocalTheory.notes kind facts #> snd) notes)
+  |> snd
   |> LocalTheory.exit;
 
 fun add_locale_cmd name expr elems thy =
   thy 
-  |> Expression.add_locale_cmd name name expr elems
-  |> (fn ((target, notes), ctxt) => TheoryTarget.begin target ctxt |>
-           fold (fn (kind, facts) => LocalTheory.notes kind facts #> snd) notes)
+  |> Expression.add_locale_cmd name "" expr elems
+  |> snd
   |> LocalTheory.exit;
 
 type statespace_info =
@@ -266,7 +263,7 @@
       in EVERY [rtac rule i] st
       end
 
-    fun tac ctxt = EVERY [NewLocale.intro_locales_tac true ctxt [],
+    fun tac ctxt = EVERY [Locale.intro_locales_tac true ctxt [],
                           ALLGOALS (SUBGOAL (solve_tac ctxt))]
 
   in thy
@@ -421,13 +418,16 @@
 
         val expr = ([(suffix valuetypesN name, 
                      (("",false),Expression.Positional (map SOME pars)))],[]);
-      in prove_interpretation_in (K all_tac) (suffix valuetypesN name, expr) thy end;
+      in
+        prove_interpretation_in (ALLGOALS o solve_tac o Assumption.prems_of)
+          (suffix valuetypesN name, expr) thy
+      end;
 
     fun interprete_parent (_, pname, rs) =
       let
         val expr = ([(pname, (("",false), Expression.Positional rs))],[])
       in prove_interpretation_in
-           (fn ctxt => NewLocale.intro_locales_tac false ctxt [])
+           (fn ctxt => Locale.intro_locales_tac false ctxt [])
            (full_name, expr) end;
 
     fun declare_declinfo updates lthy phi ctxt =
--- a/src/HOL/Tools/Groebner_Basis/groebner.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/Groebner_Basis/groebner.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/Groebner_Basis/groebner.ML
-    ID:         $Id$
     Author:     Amine Chaieb, TU Muenchen
 *)
 
@@ -599,14 +598,14 @@
      if length ideal = 2 then ideal else [eq_commute, eq_commute]
   val ring_dest_neg =
     fn t => let val (l,r) = dest_comb t in
-        if could_unify(term_of l,term_of ring_neg_tm) then r else raise CTERM ("ring_dest_neg", [t])
+        if Term.could_unify(term_of l,term_of ring_neg_tm) then r else raise CTERM ("ring_dest_neg", [t])
       end
 
  val ring_mk_neg = fn tm => mk_comb (ring_neg_tm) (tm);
 (*
  fun ring_dest_inv t =
     let val (l,r) = dest_comb t in
-        if could_unify(term_of l, term_of ring_inv_tm) then r else raise CTERM "ring_dest_inv"
+        if Term.could_unify(term_of l, term_of ring_inv_tm) then r else raise CTERM "ring_dest_inv"
     end
 *)
  val ring_dest_add = dest_binary ring_add_tm;
@@ -687,7 +686,7 @@
   val cjs = conjs tm
   val  rawvars = fold_rev (fn eq => fn a =>
                                        grobvars (dest_arg1 eq) (grobvars (dest_arg eq) a)) cjs []
-  val vars = sort (fn (x, y) => Term.term_ord(term_of x,term_of y))
+  val vars = sort (fn (x, y) => TermOrd.term_ord(term_of x,term_of y))
                   (distinct (op aconvc) rawvars)
  in (vars,map (grobify_equation vars) cjs)
  end;
@@ -896,7 +895,7 @@
   val vars = filter (fn v => free_in v eq) evs
   val (qs,p) = isolate_monomials vars eq
   val rs = ideal (qs @ ps) p 
-              (fn (s,t) => Term.term_ord (term_of s, term_of t))
+              (fn (s,t) => TermOrd.term_ord (term_of s, term_of t))
  in (eq,Library.take (length qs, rs) ~~ vars)
  end;
  fun subst_in_poly i p = Thm.rhs_of (ring_normalize_conv (vsubst i p));
--- a/src/HOL/Tools/Groebner_Basis/normalizer.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/Groebner_Basis/normalizer.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/Groebner_Basis/normalizer.ML
-    ID:         $Id$
     Author:     Amine Chaieb, TU Muenchen
 *)
 
@@ -615,7 +614,7 @@
 val nat_exp_ss = HOL_basic_ss addsimps (@{thms nat_number} @ nat_arith @ @{thms arith_simps} @ @{thms rel_simps})
                               addsimps [Let_def, if_False, if_True, @{thm add_0}, @{thm add_Suc}];
 
-fun simple_cterm_ord t u = Term.term_ord (term_of t, term_of u) = LESS;
+fun simple_cterm_ord t u = TermOrd.term_ord (term_of t, term_of u) = LESS;
 
 fun semiring_normalizers_ord_wrapper ctxt ({vars, semiring, ring, idom, ideal}, 
                                      {conv, dest_const, mk_const, is_const}) ord =
--- a/src/HOL/Tools/Qelim/cooper.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/Qelim/cooper.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/Qelim/cooper.ML
-    ID:         $Id$
     Author:     Amine Chaieb, TU Muenchen
 *)
 
@@ -515,7 +514,7 @@
   let val _ = ()
   in
    Qelim.gen_qelim_conv pcv postcv pcv (cons o term_of) 
-      (term_frees (term_of p)) (linearize_conv ctxt) (integer_nnf_conv ctxt) 
+      (OldTerm.term_frees (term_of p)) (linearize_conv ctxt) (integer_nnf_conv ctxt) 
       (cooperex_conv ctxt) p 
   end
   handle  CTERM s => raise COOPER ("Cooper Failed", CTERM s)
@@ -637,7 +636,7 @@
   let
     val thy = Thm.theory_of_cterm ct;
     val t = Thm.term_of ct;
-    val (vs, ps) = pairself (map_index swap) (term_frees t, term_bools [] t);
+    val (vs, ps) = pairself (map_index swap) (OldTerm.term_frees t, term_bools [] t);
   in
     Thm.cterm_of thy (Logic.mk_equals (HOLogic.mk_Trueprop t,
       HOLogic.mk_Trueprop (term_of_qf ps vs (pa (qf_of_term ps vs t)))))
--- a/src/HOL/Tools/Qelim/langford.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/Qelim/langford.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,3 +1,7 @@
+(*  Title:      HOL/Tools/Qelim/langford.ML
+    Author:     Amine Chaieb, TU Muenchen
+*)
+
 signature LANGFORD_QE = 
 sig
   val dlo_tac : Proof.context -> int -> tactic
@@ -114,7 +118,7 @@
       let val (yes,no) = partition f xs 
       in if f x then (x::yes,no) else (yes, x::no) end;
 
-fun contains x ct = member (op aconv) (term_frees (term_of ct)) (term_of x);
+fun contains x ct = member (op aconv) (OldTerm.term_frees (term_of ct)) (term_of x);
 
 fun is_eqx x eq = case term_of eq of
    Const("op =",_)$l$r => l aconv term_of x orelse r aconv term_of x
@@ -211,7 +215,7 @@
  let 
    fun all T = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "all"}
    fun gen x t = Thm.capply (all (ctyp_of_term x)) (Thm.cabs x t)
-   val ts = sort (fn (a,b) => Term.fast_term_ord (term_of a, term_of b)) (f p)
+   val ts = sort (fn (a,b) => TermOrd.fast_term_ord (term_of a, term_of b)) (f p)
    val p' = fold_rev gen ts p
  in implies_intr p' (implies_elim st (fold forall_elim ts (assume p'))) end));
 
--- a/src/HOL/Tools/Qelim/presburger.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/Qelim/presburger.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/Qelim/presburger.ML
-    ID:         $Id$
     Author:     Amine Chaieb, TU Muenchen
 *)
 
@@ -86,8 +85,8 @@
 in 
 fun is_relevant ctxt ct = 
  gen_subset (op aconv) (term_constants (term_of ct) , snd (CooperData.get ctxt))
- andalso forall (fn Free (_,T) => T mem [@{typ "int"}, @{typ nat}]) (term_frees (term_of ct))
- andalso forall (fn Var (_,T) => T mem [@{typ "int"}, @{typ nat}]) (term_vars (term_of ct));
+ andalso forall (fn Free (_,T) => T mem [@{typ "int"}, @{typ nat}]) (OldTerm.term_frees (term_of ct))
+ andalso forall (fn Var (_,T) => T mem [@{typ "int"}, @{typ nat}]) (OldTerm.term_vars (term_of ct));
 
 fun int_nat_terms ctxt ct =
  let 
@@ -104,7 +103,7 @@
  let 
    fun all T = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "all"}
    fun gen x t = Thm.capply (all (ctyp_of_term x)) (Thm.cabs x t)
-   val ts = sort (fn (a,b) => Term.fast_term_ord (term_of a, term_of b)) (f p)
+   val ts = sort (fn (a,b) => TermOrd.fast_term_ord (term_of a, term_of b)) (f p)
    val p' = fold_rev gen ts p
  in implies_intr p' (implies_elim st (fold forall_elim ts (assume p'))) end));
 
--- a/src/HOL/Tools/TFL/casesplit.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/TFL/casesplit.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/TFL/casesplit.ML
-    ID:         $Id$
     Author:     Lucas Dixon, University of Edinburgh
 
 A structure that defines a tactic to program case splits.
@@ -104,7 +103,7 @@
     end;
 
 (*
- val ty = (snd o hd o map Term.dest_Free o Term.term_frees) t;
+ val ty = (snd o hd o map Term.dest_Free o OldTerm.term_frees) t;
 *)
 
 
@@ -122,9 +121,9 @@
       val abs_ct = ctermify abst;
       val free_ct = ctermify x;
 
-      val casethm_vars = rev (Term.term_vars (Thm.concl_of case_thm));
+      val casethm_vars = rev (OldTerm.term_vars (Thm.concl_of case_thm));
 
-      val casethm_tvars = Term.term_tvars (Thm.concl_of case_thm);
+      val casethm_tvars = OldTerm.term_tvars (Thm.concl_of case_thm);
       val (Pv, Dv, type_insts) =
           case (Thm.concl_of case_thm) of
             (_ $ ((Pv as Var(P,Pty)) $ (Dv as Var(D, Dty)))) =>
@@ -170,7 +169,7 @@
       val subgoalth' = atomize_goals subgoalth;
       val gt = Data.dest_Trueprop (Logic.get_goal (Thm.prop_of subgoalth') 1);
 
-      val freets = Term.term_frees gt;
+      val freets = OldTerm.term_frees gt;
       fun getter x =
           let val (n,ty) = Term.dest_Free x in
             (if vstr = n orelse vstr = Name.dest_skolem n
--- a/src/HOL/Tools/TFL/rules.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/TFL/rules.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,9 +1,7 @@
 (*  Title:      HOL/Tools/TFL/rules.ML
-    ID:         $Id$
     Author:     Konrad Slind, Cambridge University Computer Laboratory
-    Copyright   1997  University of Cambridge
 
-Emulation of HOL inference rules for TFL
+Emulation of HOL inference rules for TFL.
 *)
 
 signature RULES =
@@ -173,7 +171,7 @@
  *---------------------------------------------------------------------------*)
 local val thy = Thm.theory_of_thm disjI1
       val prop = Thm.prop_of disjI1
-      val [P,Q] = term_vars prop
+      val [P,Q] = OldTerm.term_vars prop
       val disj1 = Thm.forall_intr (Thm.cterm_of thy Q) disjI1
 in
 fun DISJ1 thm tm = thm RS (forall_elim (D.drop_prop tm) disj1)
@@ -182,7 +180,7 @@
 
 local val thy = Thm.theory_of_thm disjI2
       val prop = Thm.prop_of disjI2
-      val [P,Q] = term_vars prop
+      val [P,Q] = OldTerm.term_vars prop
       val disj2 = Thm.forall_intr (Thm.cterm_of thy P) disjI2
 in
 fun DISJ2 tm thm = thm RS (forall_elim (D.drop_prop tm) disj2)
@@ -278,7 +276,7 @@
 local (* this is fragile *)
       val thy = Thm.theory_of_thm spec
       val prop = Thm.prop_of spec
-      val x = hd (tl (term_vars prop))
+      val x = hd (tl (OldTerm.term_vars prop))
       val cTV = ctyp_of thy (type_of x)
       val gspec = forall_intr (cterm_of thy x) spec
 in
@@ -295,7 +293,7 @@
 (* Not optimized! Too complicated. *)
 local val thy = Thm.theory_of_thm allI
       val prop = Thm.prop_of allI
-      val [P] = add_term_vars (prop, [])
+      val [P] = OldTerm.add_term_vars (prop, [])
       fun cty_theta s = map (fn (i, (S, ty)) => (ctyp_of s (TVar (i, S)), ctyp_of s ty))
       fun ctm_theta s = map (fn (i, (_, tm2)) =>
                              let val ctm2 = cterm_of s tm2
@@ -325,7 +323,7 @@
    let val thy = Thm.theory_of_thm thm
        val prop = Thm.prop_of thm
        val tycheck = cterm_of thy
-       val vlist = map tycheck (add_term_vars (prop, []))
+       val vlist = map tycheck (OldTerm.add_term_vars (prop, []))
   in GENL vlist thm
   end;
 
@@ -365,7 +363,7 @@
 
 local val thy = Thm.theory_of_thm exI
       val prop = Thm.prop_of exI
-      val [P,x] = term_vars prop
+      val [P,x] = OldTerm.term_vars prop
 in
 fun EXISTS (template,witness) thm =
    let val thy = Thm.theory_of_thm thm
@@ -516,7 +514,7 @@
                 val (f,args) = S.strip_comb (get_lhs eq)
                 val (vstrl,_) = S.strip_abs f
                 val names  =
-                  Name.variant_list (add_term_names(body, [])) (map (#1 o dest_Free) vstrl)
+                  Name.variant_list (OldTerm.add_term_names(body, [])) (map (#1 o dest_Free) vstrl)
             in get (rst, n+1, (names,n)::L) end
             handle TERM _ => get (rst, n+1, L)
               | U.ERR _ => get (rst, n+1, L);
@@ -676,7 +674,7 @@
      fun prover used ss thm =
      let fun cong_prover ss thm =
          let val dummy = say "cong_prover:"
-             val cntxt = MetaSimplifier.prems_of_ss ss
+             val cntxt = Simplifier.prems_of_ss ss
              val dummy = print_thms "cntxt:" cntxt
              val dummy = say "cong rule:"
              val dummy = say (Display.string_of_thm thm)
@@ -689,7 +687,7 @@
                      val ants = map tych (Logic.strip_imp_prems imp)
                      val eq = Logic.strip_imp_concl imp
                      val lhs = tych(get_lhs eq)
-                     val ss' = MetaSimplifier.add_prems (map ASSUME ants) ss
+                     val ss' = Simplifier.add_prems (map ASSUME ants) ss
                      val lhs_eq_lhs1 = MetaSimplifier.rewrite_cterm (false,true,false) (prover used) ss' lhs
                        handle U.ERR _ => Thm.reflexive lhs
                      val dummy = print_thms "proven:" [lhs_eq_lhs1]
@@ -710,7 +708,7 @@
                   val Q = get_lhs eq1
                   val QeqQ1 = pbeta_reduce (tych Q)
                   val Q1 = #2(D.dest_eq(cconcl QeqQ1))
-                  val ss' = MetaSimplifier.add_prems (map ASSUME ants1) ss
+                  val ss' = Simplifier.add_prems (map ASSUME ants1) ss
                   val Q1eeqQ2 = MetaSimplifier.rewrite_cterm (false,true,false) (prover used') ss' Q1
                                 handle U.ERR _ => Thm.reflexive Q1
                   val Q2 = #2 (Logic.dest_equals (Thm.prop_of Q1eeqQ2))
@@ -760,12 +758,12 @@
 
         fun restrict_prover ss thm =
           let val dummy = say "restrict_prover:"
-              val cntxt = rev(MetaSimplifier.prems_of_ss ss)
+              val cntxt = rev(Simplifier.prems_of_ss ss)
               val dummy = print_thms "cntxt:" cntxt
               val thy = Thm.theory_of_thm thm
               val Const("==>",_) $ (Const("Trueprop",_) $ A) $ _ = Thm.prop_of thm
               fun genl tm = let val vlist = subtract (op aconv) globals
-                                           (add_term_frees(tm,[]))
+                                           (OldTerm.add_term_frees(tm,[]))
                             in fold_rev Forall vlist tm
                             end
               (*--------------------------------------------------------------
@@ -805,7 +803,7 @@
     (if (is_cong thm) then cong_prover else restrict_prover) ss thm
     end
     val ctm = cprop_of th
-    val names = add_term_names (term_of ctm, [])
+    val names = OldTerm.add_term_names (term_of ctm, [])
     val th1 = MetaSimplifier.rewrite_cterm(false,true,false)
       (prover names) (ss0 addsimps [cut_lemma'] addeqcongs congs) ctm
     val th2 = equal_elim th1 th
--- a/src/HOL/Tools/TFL/tfl.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/TFL/tfl.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,7 +1,5 @@
 (*  Title:      HOL/Tools/TFL/tfl.ML
-    ID:         $Id$
     Author:     Konrad Slind, Cambridge University Computer Laboratory
-    Copyright   1997  University of Cambridge
 
 First part of main module.
 *)
@@ -332,7 +330,7 @@
      val dummy = map (no_repeat_vars thy) pats
      val rows = ListPair.zip (map (fn x => ([]:term list,[x])) pats,
                               map (fn (t,i) => (t,(i,true))) (enumerate R))
-     val names = foldr add_term_names [] R
+     val names = foldr OldTerm.add_term_names [] R
      val atype = type_of(hd pats)
      and aname = Name.variant names "a"
      val a = Free(aname,atype)
@@ -494,7 +492,7 @@
      val tych = Thry.typecheck thy
      val WFREC_THM0 = R.ISPEC (tych functional) Thms.WFREC_COROLLARY
      val Const("All",_) $ Abs(Rname,Rtype,_) = concl WFREC_THM0
-     val R = Free (Name.variant (foldr add_term_names [] eqns) Rname,
+     val R = Free (Name.variant (foldr OldTerm.add_term_names [] eqns) Rname,
                    Rtype)
      val WFREC_THM = R.ISPECL [tych R, tych g] WFREC_THM0
      val ([proto_def, WFR],_) = S.strip_imp(concl WFREC_THM)
@@ -692,7 +690,7 @@
  let val tych = Thry.typecheck thy
      val ty_info = Thry.induct_info thy
  in fn pats =>
- let val names = foldr add_term_names [] pats
+ let val names = foldr OldTerm.add_term_names [] pats
      val T = type_of (hd pats)
      val aname = Name.variant names "a"
      val vname = Name.variant (aname::names) "v"
@@ -845,7 +843,7 @@
     val (pats,TCsl) = ListPair.unzip pat_TCs_list
     val case_thm = complete_cases thy pats
     val domain = (type_of o hd) pats
-    val Pname = Name.variant (foldr (Library.foldr add_term_names)
+    val Pname = Name.variant (foldr (Library.foldr OldTerm.add_term_names)
                               [] (pats::TCsl)) "P"
     val P = Free(Pname, domain --> HOLogic.boolT)
     val Sinduct = R.SPEC (tych P) Sinduction
@@ -856,7 +854,7 @@
     val cases = map (fn pat => Term.betapply (Sinduct_assumf, pat)) pats
     val tasks = U.zip3 cases TCl' (R.CONJUNCTS Rinduct_assum)
     val proved_cases = map (prove_case fconst thy) tasks
-    val v = Free (Name.variant (foldr add_term_names [] (map concl proved_cases))
+    val v = Free (Name.variant (foldr OldTerm.add_term_names [] (map concl proved_cases))
                     "v",
                   domain)
     val vtyped = tych v
--- a/src/HOL/Tools/TFL/usyntax.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/TFL/usyntax.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,7 +1,5 @@
 (*  Title:      HOL/Tools/TFL/usyntax.ML
-    ID:         $Id$
     Author:     Konrad Slind, Cambridge University Computer Laboratory
-    Copyright   1997  University of Cambridge
 
 Emulation of HOL's abstract syntax functions.
 *)
@@ -114,7 +112,7 @@
 
 val is_vartype = can dest_vtype;
 
-val type_vars  = map mk_prim_vartype o typ_tvars
+val type_vars  = map mk_prim_vartype o OldTerm.typ_tvars
 fun type_varsl L = distinct (op =) (fold (curry op @ o type_vars) L []);
 
 val alpha  = mk_vartype "'a"
@@ -144,7 +142,7 @@
 
 
 
-val type_vars_in_term = map mk_prim_vartype o term_tvars;
+val type_vars_in_term = map mk_prim_vartype o OldTerm.term_tvars;
 
 
 
@@ -321,7 +319,7 @@
 
 
 (* Need to reverse? *)
-fun gen_all tm = list_mk_forall(term_frees tm, tm);
+fun gen_all tm = list_mk_forall(OldTerm.term_frees tm, tm);
 
 (* Destructing a cterm to a list of Terms *)
 fun strip_comb tm =
--- a/src/HOL/Tools/arith_data.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/arith_data.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/arith_data.ML
-    ID:         $Id$
     Author:     Markus Wenzel, Stefan Berghofer, and Tobias Nipkow
 
 Basic arithmetic proof tools.
@@ -7,9 +6,8 @@
 
 signature ARITH_DATA =
 sig
-  val prove_conv: tactic -> (MetaSimplifier.simpset -> tactic)
-    -> MetaSimplifier.simpset -> term * term -> thm
-  val simp_all_tac: thm list -> MetaSimplifier.simpset -> tactic
+  val prove_conv: tactic -> (simpset -> tactic) -> simpset -> term * term -> thm
+  val simp_all_tac: thm list -> simpset -> tactic
 
   val mk_sum: term list -> term
   val mk_norm_sum: term list -> term
--- a/src/HOL/Tools/cnf_funcs.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/cnf_funcs.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,8 +1,6 @@
 (*  Title:      HOL/Tools/cnf_funcs.ML
-    ID:         $Id$
     Author:     Alwen Tiu, QSL Team, LORIA (http://qsl.loria.fr)
-    Author:     Tjark Weber
-    Copyright   2005-2006
+    Author:     Tjark Weber, TU Muenchen
 
   FIXME: major overlaps with the code in meson.ML
 
@@ -499,7 +497,7 @@
 					NONE
 		in
 			Int.max (max, getOpt (idx, 0))
-		end) (term_frees simp) 0)
+		end) (OldTerm.term_frees simp) 0)
 	(* finally, convert to definitional CNF (this should preserve the simplification) *)
 	val cnfx_thm = make_cnfx_thm_from_nnf simp
 in
--- a/src/HOL/Tools/datatype_abs_proofs.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/datatype_abs_proofs.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/datatype_abs_proofs.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Proofs and defintions independent of concrete representation
@@ -10,8 +9,7 @@
  - characteristic equations for primrec combinators
  - characteristic equations for case combinators
  - equations for splitting "P (case ...)" expressions
-  - "nchotomy" and "case_cong" theorems for TFL
-
+ - "nchotomy" and "case_cong" theorems for TFL
 *)
 
 signature DATATYPE_ABS_PROOFS =
@@ -98,7 +96,7 @@
 
     val descr' = List.concat descr;
     val recTs = get_rec_types descr' sorts;
-    val used = foldr add_typ_tfree_names [] recTs;
+    val used = foldr OldTerm.add_typ_tfree_names [] recTs;
     val newTs = Library.take (length (hd descr), recTs);
 
     val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
@@ -157,7 +155,7 @@
         InductivePackage.add_inductive_global (serial_string ())
           {quiet_mode = ! quiet_mode, verbose = false, kind = Thm.internalK,
             alt_name = Binding.name big_rec_name', coind = false, no_elim = false, no_ind = true,
-            skip_mono = true}
+            skip_mono = true, fork_mono = false}
           (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts))
           (map dest_Free rec_fns)
           (map (fn x => (Attrib.empty_binding, x)) rec_intr_ts) [] thy0;
@@ -281,7 +279,7 @@
 
     val descr' = List.concat descr;
     val recTs = get_rec_types descr' sorts;
-    val used = foldr add_typ_tfree_names [] recTs;
+    val used = foldr OldTerm.add_typ_tfree_names [] recTs;
     val newTs = Library.take (length (hd descr), recTs);
     val T' = TFree (Name.variant used "'t", HOLogic.typeS);
 
@@ -425,8 +423,8 @@
         val (Const ("Trueprop", _) $ (Const ("op =", _) $ _ $ Ma)) = tm;
         val cert = cterm_of thy;
         val nchotomy' = nchotomy RS spec;
-        val nchotomy'' = cterm_instantiate
-          [(cert (hd (add_term_vars (concl_of nchotomy', []))), cert Ma)] nchotomy'
+        val [v] = Term.add_vars (concl_of nchotomy') [];
+        val nchotomy'' = cterm_instantiate [(cert (Var v), cert Ma)] nchotomy'
       in
         SkipProof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
           (fn {prems, ...} => 
--- a/src/HOL/Tools/datatype_aux.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/datatype_aux.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/datatype_aux.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Auxiliary functions for defining datatypes.
@@ -131,7 +130,7 @@
     val flt = if null indnames then I else
       filter (fn Free (s, _) => s mem indnames | _ => false);
     fun abstr (t1, t2) = (case t1 of
-        NONE => (case flt (term_frees t2) of
+        NONE => (case flt (OldTerm.term_frees t2) of
             [Free (s, T)] => SOME (absfree (s, T, t2))
           | _ => NONE)
       | SOME (_ $ t') => SOME (Abs ("x", fastype_of t', abstract_over (t', t2))))
--- a/src/HOL/Tools/datatype_case.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/datatype_case.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,7 +1,6 @@
 (*  Title:      HOL/Tools/datatype_case.ML
-    ID:         $Id$
     Author:     Konrad Slind, Cambridge University Computer Laboratory
-                Stefan Berghofer, TU Muenchen
+    Author:     Stefan Berghofer, TU Muenchen
 
 Nested case expressions on datatypes.
 *)
@@ -55,15 +54,12 @@
  * b=false --> i = ~1
  *---------------------------------------------------------------------------*)
 
-fun pattern_map f (tm,x) = (f tm, x);
-
-fun pattern_subst theta = pattern_map (subst_free theta);
+fun pattern_subst theta (tm, x) = (subst_free theta tm, x);
 
 fun row_of_pat x = fst (snd x);
 
-fun add_row_used ((prfx, pats), (tm, tag)) used =
-  foldl add_term_free_names (foldl add_term_free_names
-    (add_term_free_names (tm, used)) pats) prfx;
+fun add_row_used ((prfx, pats), (tm, tag)) =
+  fold Term.add_free_names (tm :: pats @ prfx);
 
 (* try to preserve names given by user *)
 fun default_names names ts =
@@ -140,7 +136,7 @@
                     let
                       val Ts = map type_of rstp;
                       val xs = Name.variant_list
-                        (foldl add_term_free_names used' gvars)
+                        (fold Term.add_free_names gvars used')
                         (replicate (length rstp) "x")
                     in
                       [((prfx, gvars @ map Free (xs ~~ Ts)),
@@ -222,7 +218,7 @@
               | SOME {case_name, constructors} =>
                 let
                   val pty = body_type cT;
-                  val used' = foldl add_term_free_names used rstp;
+                  val used' = fold Term.add_free_names rstp used;
                   val nrows = maps (expand constructors used' pty) rows;
                   val subproblems = partition ty_match ty_inst type_of used'
                     constructors pty range_ty nrows;
@@ -336,7 +332,7 @@
         | prep_pat t used = case_error ("Bad pattern: " ^ Syntax.string_of_term ctxt t);
       fun dest_case1 (t as Const ("_case1", _) $ l $ r) =
             let val (l', cnstrts) = strip_constraints l
-            in ((fst (prep_pat l' (add_term_free_names (t, []))), r), cnstrts)
+            in ((fst (prep_pat l' (Term.add_free_names t [])), r), cnstrts)
             end
         | dest_case1 t = case_error "dest_case1";
       fun dest_case2 (Const ("_case2", _) $ t $ u) = t :: dest_case2 u
@@ -366,7 +362,7 @@
             val _ = if length zs < i then raise CASE_ERROR ("", 0) else ();
             val (xs, ys) = chop i zs;
             val u = list_abs (ys, strip_abs_body t);
-            val xs' = map Free (Name.variant_list (add_term_names (u, used))
+            val xs' = map Free (Name.variant_list (OldTerm.add_term_names (u, used))
               (map fst xs) ~~ map snd xs)
           in (xs', subst_bounds (rev xs', u)) end;
         fun is_dependent i t =
@@ -424,11 +420,11 @@
 (* destruct nested patterns *)
 
 fun strip_case' dest (pat, rhs) =
-  case dest (add_term_free_names (pat, [])) rhs of
+  case dest (Term.add_free_names pat []) rhs of
     SOME (exp as Free _, clauses) =>
-      if member op aconv (term_frees pat) exp andalso
+      if member op aconv (OldTerm.term_frees pat) exp andalso
         not (exists (fn (_, rhs') =>
-          member op aconv (term_frees rhs') exp) clauses)
+          member op aconv (OldTerm.term_frees rhs') exp) clauses)
       then
         maps (strip_case' dest) (map (fn (pat', rhs') =>
           (subst_free [(exp, pat')] pat, rhs')) clauses)
@@ -451,7 +447,7 @@
     val thy = ProofContext.theory_of ctxt;
     val consts = ProofContext.consts_of ctxt;
     fun mk_clause (pat, rhs) =
-      let val xs = term_frees pat
+      let val xs = Term.add_frees pat []
       in
         Syntax.const "_case1" $
           map_aterms
@@ -459,8 +455,8 @@
               | Const (s, _) => Const (Consts.extern_early consts s, dummyT)
               | t => t) pat $
           map_aterms
-            (fn x as Free (s, _) =>
-                  if member op aconv xs x then Syntax.mark_bound s else x
+            (fn x as Free (s, T) =>
+                  if member (op =) xs (s, T) then Syntax.mark_bound s else x
               | t => t) rhs
       end
   in case strip_case' (tab_of thy) true (list_comb (Syntax.const cname, ts)) of
--- a/src/HOL/Tools/datatype_codegen.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/datatype_codegen.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,6 +1,5 @@
 (*  Title:      HOL/Tools/datatype_codegen.ML
-    ID:         $Id$
-    Author:     Stefan Berghofer & Florian Haftmann, TU Muenchen
+    Author:     Stefan Berghofer and Florian Haftmann, TU Muenchen
 
 Code generator facilities for inductive datatypes.
 *)
@@ -217,8 +216,8 @@
       let
         val ts1 = Library.take (i, ts);
         val t :: ts2 = Library.drop (i, ts);
-        val names = foldr add_term_names
-          (map (fst o fst o dest_Var) (foldr add_term_vars [] ts1)) ts1;
+        val names = foldr OldTerm.add_term_names
+          (map (fst o fst o dest_Var) (foldr OldTerm.add_term_vars [] ts1)) ts1;
         val (Ts, dT) = split_last (Library.take (i+1, fst (strip_type T)));
 
         fun pcase [] [] [] gr = ([], gr)
@@ -414,7 +413,7 @@
     val inject2 = (#inject o DatatypePackage.the_datatype thy) dtco;
     val ctxt = ProofContext.init thy;
     val simpset = Simplifier.context ctxt
-      (MetaSimplifier.empty_ss addsimprocs [DatatypePackage.distinct_simproc]);
+      (Simplifier.empty_ss addsimprocs [DatatypePackage.distinct_simproc]);
     val cos = map (fn (co, tys) =>
         (Const (co, tys ---> Type (dtco, map TFree vs)), tys)) cs;
     val tac = ALLGOALS (simp_tac simpset)
--- a/src/HOL/Tools/datatype_package.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/datatype_package.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/datatype_package.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Datatype package for Isabelle/HOL.
@@ -492,7 +491,7 @@
       ^ Syntax.string_of_typ_global thy T);
     fun type_of_constr (cT as (_, T)) =
       let
-        val frees = typ_tfrees T;
+        val frees = OldTerm.typ_tfrees T;
         val (tyco, vs) = ((apsnd o map) (dest_TFree) o dest_Type o snd o strip_type) T
           handle TYPE _ => no_constr cT
         val _ = if has_duplicates (eq_fst (op =)) vs then no_constr cT else ();
@@ -583,7 +582,7 @@
         fun prep_constr (cname, cargs, mx') (constrs, constr_syntax', sorts') =
           let
             val (cargs', sorts'') = Library.foldl (prep_typ tmp_thy) (([], sorts'), cargs);
-            val _ = (case fold (curry add_typ_tfree_names) cargs' [] \\ tvs of
+            val _ = (case fold (curry OldTerm.add_typ_tfree_names) cargs' [] \\ tvs of
                 [] => ()
               | vs => error ("Extra type variables on rhs: " ^ commas vs))
           in (constrs @ [((if flat_names then Sign.full_bname tmp_thy else
--- a/src/HOL/Tools/datatype_prop.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/datatype_prop.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/datatype_prop.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Characteristic properties of datatypes.
@@ -206,7 +205,7 @@
   let
     val descr' = List.concat descr;
     val recTs = get_rec_types descr' sorts;
-    val used = foldr add_typ_tfree_names [] recTs;
+    val used = foldr OldTerm.add_typ_tfree_names [] recTs;
 
     val (rec_result_Ts, reccomb_fn_Ts) = make_primrec_Ts descr sorts used;
 
@@ -256,7 +255,7 @@
   let
     val descr' = List.concat descr;
     val recTs = get_rec_types descr' sorts;
-    val used = foldr add_typ_tfree_names [] recTs;
+    val used = foldr OldTerm.add_typ_tfree_names [] recTs;
     val newTs = Library.take (length (hd descr), recTs);
     val T' = TFree (Name.variant used "'t", HOLogic.typeS);
 
@@ -303,7 +302,7 @@
   let
     val descr' = List.concat descr;
     val recTs = get_rec_types descr' sorts;
-    val used' = foldr add_typ_tfree_names [] recTs;
+    val used' = foldr OldTerm.add_typ_tfree_names [] recTs;
     val newTs = Library.take (length (hd descr), recTs);
     val T' = TFree (Name.variant used' "'t", HOLogic.typeS);
     val P = Free ("P", T' --> HOLogic.boolT);
--- a/src/HOL/Tools/datatype_rep_proofs.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/datatype_rep_proofs.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/datatype_rep_proofs.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Definitional introduction of datatypes
@@ -8,7 +7,6 @@
  - injectivity of constructors
  - distinctness of constructors
  - induction theorem
-
 *)
 
 signature DATATYPE_REP_PROOFS =
@@ -85,7 +83,7 @@
     val branchT = if null branchTs then HOLogic.unitT
       else BalancedTree.make (fn (T, U) => Type ("+", [T, U])) branchTs;
     val arities = get_arities descr' \ 0;
-    val unneeded_vars = hd tyvars \\ foldr add_typ_tfree_names [] (leafTs' @ branchTs);
+    val unneeded_vars = hd tyvars \\ foldr OldTerm.add_typ_tfree_names [] (leafTs' @ branchTs);
     val leafTs = leafTs' @ (map (fn n => TFree (n, (the o AList.lookup (op =) sorts) n)) unneeded_vars);
     val recTs = get_rec_types descr' sorts;
     val newTs = Library.take (length (hd descr), recTs);
@@ -185,7 +183,7 @@
         InductivePackage.add_inductive_global (serial_string ())
           {quiet_mode = ! quiet_mode, verbose = false, kind = Thm.internalK,
            alt_name = Binding.name big_rec_name, coind = false, no_elim = true, no_ind = false,
-           skip_mono = true}
+           skip_mono = true, fork_mono = false}
           (map (fn s => ((Binding.name s, UnivT'), NoSyn)) rep_set_names') []
           (map (fn x => (Attrib.empty_binding, x)) intr_ts) [] thy1;
 
@@ -369,7 +367,7 @@
         val prop = Thm.prop_of thm;
         val _ $ (_ $ ((S as Const (_, Type (_, [U, _]))) $ _ )) $
           (_ $ (_ $ (r $ (a $ _)) $ _)) = Type.freeze prop;
-        val used = add_term_tfree_names (a, []);
+        val used = OldTerm.add_term_tfree_names (a, []);
 
         fun mk_thm i =
           let
--- a/src/HOL/Tools/function_package/context_tree.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/function_package/context_tree.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,12 +1,10 @@
 (*  Title:      HOL/Tools/function_package/context_tree.ML
-    ID:         $Id$
     Author:     Alexander Krauss, TU Muenchen
 
 A package for general recursive function definitions. 
 Builds and traverses trees of nested contexts along a term.
 *)
 
-
 signature FUNDEF_CTXTREE =
 sig
     type ctxt = (string * typ) list * thm list (* poor man's contexts: fixes + assumes *)
@@ -82,7 +80,7 @@
     let 
       val t' = snd (dest_all_all t)
       val (assumes, concl) = Logic.strip_horn t'
-    in (fold (curry add_term_vars) assumes [], term_vars concl)
+    in (fold Term.add_vars assumes [], Term.add_vars concl [])
     end
 
 fun cong_deps crule =
@@ -91,7 +89,9 @@
     in
       IntGraph.empty
         |> fold (fn (i,_)=> IntGraph.new_node (i,i)) num_branches
-        |> fold_product (fn (i,(c1,_)) => fn (j,(_, t2)) => if i = j orelse null (c1 inter t2) then I else IntGraph.add_edge_acyclic (i,j))
+        |> fold_product (fn (i, (c1, _)) => fn (j, (_, t2)) => 
+               if i = j orelse null (c1 inter t2) 
+               then I else IntGraph.add_edge_acyclic (i,j))
              num_branches num_branches
     end
     
--- a/src/HOL/Tools/function_package/decompose.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/function_package/decompose.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -19,7 +19,7 @@
 structure Decompose : DECOMPOSE =
 struct
 
-structure TermGraph = GraphFun(type key = term val ord = Term.fast_term_ord);
+structure TermGraph = GraphFun(type key = term val ord = TermOrd.fast_term_ord);
 
 
 fun derive_chains ctxt chain_tac cont D = Termination.CALLS (fn (cs, i) =>
--- a/src/HOL/Tools/function_package/fundef_core.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/function_package/fundef_core.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,9 +1,8 @@
 (*  Title:      HOL/Tools/function_package/fundef_core.ML
-    ID:         $Id$
     Author:     Alexander Krauss, TU Muenchen
 
-A package for general recursive function definitions.
-Main functionality
+A package for general recursive function definitions:
+Main functionality.
 *)
 
 signature FUNDEF_CORE =
@@ -437,7 +436,7 @@
                                   |> implies_intr (cterm_of thy (HOLogic.mk_Trueprop (mk_acc domT R $ x)))
                                   |> forall_intr (cterm_of thy x)
                                   |> (fn it => Drule.compose_single (it, 2, acc_induct_rule)) (* "EX! y. (?x,y):G" *)
-                                  |> (fn it => fold (forall_intr o cterm_of thy) (term_vars (prop_of it)) it)
+                                  |> (fn it => fold (forall_intr o cterm_of thy o Var) (Term.add_vars (prop_of it) []) it)
 
         val goalstate =  Conjunction.intr graph_is_function complete
                           |> Thm.close_derivation
--- a/src/HOL/Tools/function_package/fundef_datatype.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/function_package/fundef_datatype.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/function_package/fundef_datatype.ML
-    ID:         $Id$
     Author:     Alexander Krauss, TU Muenchen
 
 A package for general recursive function definitions.
@@ -64,9 +63,10 @@
 
 fun inst_case_thm thy x P thm =
     let
-        val [Pv, xv] = term_vars (prop_of thm)
+        val [Pv, xv] = Term.add_vars (prop_of thm) []
     in
-        cterm_instantiate [(cterm_of thy xv, cterm_of thy x), (cterm_of thy Pv, cterm_of thy P)] thm
+        cterm_instantiate [(cterm_of thy (Var xv), cterm_of thy x), 
+                           (cterm_of thy (Var Pv), cterm_of thy P)] thm
     end
 
 
@@ -199,14 +199,14 @@
 fun pat_completeness ctxt = Method.SIMPLE_METHOD' (pat_complete_tac ctxt)
 
 val by_pat_completeness_simp =
-    Proof.global_terminal_proof
+    Proof.global_future_terminal_proof
       (Method.Basic (pat_completeness, Position.none),
        SOME (Method.Source_i (Args.src (("HOL.auto", []), Position.none))))
 
-fun termination_by method =
+fun termination_by method int =
     FundefPackage.setup_termination_proof NONE
-    #> Proof.global_terminal_proof
-      (Method.Basic (method, Position.none), NONE)
+    #> Proof.global_future_terminal_proof
+      (Method.Basic (method, Position.none), NONE) int
 
 fun mk_catchall fixes arities =
     let
@@ -304,19 +304,19 @@
 
 local structure P = OuterParse and K = OuterKeyword in
 
-fun fun_cmd config fixes statements flags lthy =
+fun fun_cmd config fixes statements flags int lthy =
   let val group = serial_string () in
     lthy
       |> LocalTheory.set_group group
       |> FundefPackage.add_fundef fixes statements config flags
-      |> by_pat_completeness_simp
+      |> by_pat_completeness_simp int
       |> LocalTheory.restore
       |> LocalTheory.set_group group
-      |> termination_by (FundefCommon.get_termination_prover (Context.Proof lthy))
+      |> termination_by (FundefCommon.get_termination_prover (Context.Proof lthy)) int
   end;
 
 val _ =
-  OuterSyntax.local_theory "fun" "define general recursive functions (short version)" K.thy_decl
+  OuterSyntax.local_theory' "fun" "define general recursive functions (short version)" K.thy_decl
   (fundef_parser fun_config
      >> (fn ((config, fixes), (flags, statements)) => fun_cmd config fixes statements flags));
 
--- a/src/HOL/Tools/function_package/fundef_lib.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/function_package/fundef_lib.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/function_package/fundef_lib.ML
-    ID:         $Id$
     Author:     Alexander Krauss, TU Muenchen
 
 A package for general recursive function definitions. 
@@ -163,7 +162,7 @@
           else if js = []
             then mk (foldr1 mk (map (nth xs) is), Const (neu, ty))
             else mk (foldr1 mk (map (nth xs) is), foldr1 mk (map (nth xs) js)))))
-     (K (MetaSimplifier.rewrite_goals_tac ac
+     (K (rewrite_goals_tac ac
          THEN rtac Drule.reflexive_thm 1))
  end
 
--- a/src/HOL/Tools/function_package/inductive_wrap.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/function_package/inductive_wrap.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -48,7 +48,8 @@
               coind = false,
               no_elim = false,
               no_ind = false,
-              skip_mono = true}
+              skip_mono = true,
+              fork_mono = false}
             [((Binding.name R, T), NoSyn)] (* the relation *)
             [] (* no parameters *)
             (map (fn t => (Attrib.empty_binding, t)) defs) (* the intros *)
--- a/src/HOL/Tools/function_package/size.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/function_package/size.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,6 +1,5 @@
 (*  Title:      HOL/Tools/function_package/size.ML
-    ID:         $Id$
-    Author:     Stefan Berghofer, Florian Haftmann, TU Muenchen
+    Author:     Stefan Berghofer, Florian Haftmann & Alexander Krauss, TU Muenchen
 
 Size functions for datatypes.
 *)
@@ -81,7 +80,7 @@
     val param_size = AList.lookup op = param_size_fs;
 
     val extra_rewrites = descr |> map (#1 o snd) |> distinct op = |>
-      List.mapPartial (Option.map snd o lookup_size thy) |> flat;
+      map_filter (Option.map snd o lookup_size thy) |> flat;
     val extra_size = Option.map fst o lookup_size thy;
 
     val (((size_names, size_fns), def_names), def_names') =
@@ -180,7 +179,7 @@
       let
         val Ts = map (typ_of_dtyp descr sorts) cargs;
         val tnames = Name.variant_list f_names (DatatypeProp.make_tnames Ts);
-        val ts = List.mapPartial (fn (sT as (s, T), dt) =>
+        val ts = map_filter (fn (sT as (s, T), dt) =>
           Option.map (fn sz => sz $ Free sT)
             (if p dt then size_of_type size_of extra_size size_ofp T
              else NONE)) (tnames ~~ Ts ~~ cargs)
--- a/src/HOL/Tools/function_package/termination.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/function_package/termination.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,4 +1,4 @@
-(*  Title:       HOL/Tools/function_package/termination_data.ML
+(*  Title:       HOL/Tools/function_package/termination.ML
     Author:      Alexander Krauss, TU Muenchen
 
 Context data for termination proofs
@@ -50,9 +50,9 @@
 
 open FundefLib
 
-val term2_ord = prod_ord Term.fast_term_ord Term.fast_term_ord
+val term2_ord = prod_ord TermOrd.fast_term_ord TermOrd.fast_term_ord
 structure Term2tab = TableFun(type key = term * term val ord = term2_ord);
-structure Term3tab = TableFun(type key = term * (term * term) val ord = prod_ord Term.fast_term_ord term2_ord);
+structure Term3tab = TableFun(type key = term * (term * term) val ord = prod_ord TermOrd.fast_term_ord term2_ord);
 
 (** Analyzing binary trees **)
 
--- a/src/HOL/Tools/inductive_codegen.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/inductive_codegen.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/inductive_codegen.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Code generator for inductive predicates.
@@ -58,7 +57,7 @@
       | _ => (warn thm; thy))
     | SOME (Const (s, _), _) =>
         let
-          val cs = foldr add_term_consts [] (prems_of thm);
+          val cs = fold Term.add_const_names (Thm.prems_of thm) [];
           val rules = Symtab.lookup_list intros s;
           val nparms = (case optnparms of
             SOME k => k
@@ -129,7 +128,7 @@
   cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map
     string_of_mode ms)) modes));
 
-val term_vs = map (fst o fst o dest_Var) o term_vars;
+val term_vs = map (fst o fst o dest_Var) o OldTerm.term_vars;
 val terms_vs = distinct (op =) o List.concat o (map term_vs);
 
 (** collect all Vars in a term (with duplicates!) **)
@@ -491,7 +490,7 @@
       | SOME (names, thyname, nparms, intrs) =>
           mk_ind_def thy defs gr dep names (if_library thyname module)
             [] (prep_intrs intrs) nparms))
-            (gr, foldr add_term_consts [] ts)
+            (gr, fold Term.add_const_names ts [])
 
 and mk_ind_def thy defs gr dep names module modecs intrs nparms =
   add_edge_acyclic (hd names, dep) gr handle
--- a/src/HOL/Tools/inductive_package.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/inductive_package.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -46,7 +46,7 @@
     (Binding.T * string option * mixfix) list ->
     (Attrib.binding * string) list ->
     (Facts.ref * Attrib.src list) list ->
-    local_theory -> inductive_result * local_theory
+    bool -> local_theory -> inductive_result * local_theory
   val add_inductive_global: string -> inductive_flags ->
     ((Binding.T * typ) * mixfix) list -> (string * typ) list -> (Attrib.binding * term) list ->
     thm list -> theory -> inductive_result * theory
@@ -74,9 +74,9 @@
     (Binding.T * string option * mixfix) list ->
     (Binding.T * string option * mixfix) list ->
     (Attrib.binding * string) list -> (Facts.ref * Attrib.src list) list ->
-    local_theory -> inductive_result * local_theory
+    bool -> local_theory -> inductive_result * local_theory
   val gen_ind_decl: add_ind_def -> bool ->
-    OuterParse.token list -> (local_theory -> local_theory) * OuterParse.token list
+    OuterParse.token list -> (bool -> local_theory -> local_theory) * OuterParse.token list
 end;
 
 structure InductivePackage: INDUCTIVE_PACKAGE =
@@ -312,10 +312,11 @@
 
 (* prove monotonicity *)
 
-fun prove_mono quiet_mode skip_mono predT fp_fun monos ctxt =
- (message (quiet_mode orelse skip_mono andalso !quick_and_dirty)
+fun prove_mono quiet_mode skip_mono fork_mono predT fp_fun monos ctxt =
+ (message (quiet_mode orelse skip_mono andalso !quick_and_dirty orelse fork_mono)
     "  Proving monotonicity ...";
-  (if skip_mono then SkipProof.prove else Goal.prove) ctxt [] []
+  (if skip_mono then SkipProof.prove else if fork_mono then Goal.prove_future else Goal.prove) ctxt
+    [] []
     (HOLogic.mk_Trueprop
       (Const (@{const_name Orderings.mono}, (predT --> predT) --> HOLogic.boolT) $ fp_fun))
     (fn _ => EVERY [rtac @{thm monoI} 1,
@@ -582,7 +583,7 @@
 
 (** specification of (co)inductive predicates **)
 
-fun mk_ind_def quiet_mode skip_mono alt_name coind cs intr_ts monos params cnames_syn ctxt =
+fun mk_ind_def quiet_mode skip_mono fork_mono alt_name coind cs intr_ts monos params cnames_syn ctxt =
   let
     val fp_name = if coind then @{const_name Inductive.gfp} else @{const_name Inductive.lfp};
 
@@ -662,9 +663,11 @@
     val (consts_defs, ctxt'') = fold_map (LocalTheory.define Thm.internalK) specs ctxt';
     val preds = (case cs of [_] => [rec_const] | _ => map #1 consts_defs);
 
-    val mono = prove_mono quiet_mode skip_mono predT fp_fun monos ctxt''
+    val mono = prove_mono quiet_mode skip_mono fork_mono predT fp_fun monos ctxt'';
+    val ((_, [mono']), ctxt''') =
+      LocalTheory.note Thm.internalK (Attrib.empty_binding, [mono]) ctxt'';
 
-  in (ctxt'', rec_name, mono, fp_def', map (#2 o #2) consts_defs,
+  in (ctxt''', rec_name, mono', fp_def', map (#2 o #2) consts_defs,
     list_comb (rec_const, params), preds, argTs, bs, xs)
   end;
 
@@ -718,7 +721,7 @@
 
 type inductive_flags =
   {quiet_mode: bool, verbose: bool, kind: string, alt_name: Binding.T,
-   coind: bool, no_elim: bool, no_ind: bool, skip_mono: bool}
+   coind: bool, no_elim: bool, no_ind: bool, skip_mono: bool, fork_mono: bool}
 
 type add_ind_def =
   inductive_flags ->
@@ -726,7 +729,7 @@
   term list -> (Binding.T * mixfix) list ->
   local_theory -> inductive_result * local_theory
 
-fun add_ind_def {quiet_mode, verbose, kind, alt_name, coind, no_elim, no_ind, skip_mono}
+fun add_ind_def {quiet_mode, verbose, kind, alt_name, coind, no_elim, no_ind, skip_mono, fork_mono}
     cs intros monos params cnames_syn ctxt =
   let
     val _ = null cnames_syn andalso error "No inductive predicates given";
@@ -739,7 +742,7 @@
       apfst split_list (split_list (map (check_rule ctxt cs params) intros));
 
     val (ctxt1, rec_name, mono, fp_def, rec_preds_defs, rec_const, preds,
-      argTs, bs, xs) = mk_ind_def quiet_mode skip_mono alt_name coind cs intr_ts
+      argTs, bs, xs) = mk_ind_def quiet_mode skip_mono fork_mono alt_name coind cs intr_ts
         monos params cnames_syn ctxt;
 
     val (intrs, unfold) = prove_intrs quiet_mode coind mono fp_def (length bs + length xs)
@@ -780,7 +783,7 @@
 (* external interfaces *)
 
 fun gen_add_inductive_i mk_def
-    (flags as {quiet_mode, verbose, kind, alt_name, coind, no_elim, no_ind, skip_mono})
+    (flags as {quiet_mode, verbose, kind, alt_name, coind, no_elim, no_ind, skip_mono, fork_mono})
     cnames_syn pnames spec monos lthy =
   let
     val thy = ProofContext.theory_of lthy;
@@ -836,7 +839,7 @@
     ||> fold (snd oo LocalTheory.abbrev Syntax.mode_default) abbrevs
   end;
 
-fun gen_add_inductive mk_def verbose coind cnames_syn pnames_syn intro_srcs raw_monos lthy =
+fun gen_add_inductive mk_def verbose coind cnames_syn pnames_syn intro_srcs raw_monos int lthy =
   let
     val ((vars, specs), _) = lthy |> ProofContext.set_mode ProofContext.mode_abbrev
       |> Specification.read_specification
@@ -845,7 +848,8 @@
     val intrs = map (apsnd the_single) specs;
     val monos = Attrib.eval_thms lthy raw_monos;
     val flags = {quiet_mode = false, verbose = verbose, kind = Thm.theoremK,
-      alt_name = Binding.empty, coind = coind, no_elim = false, no_ind = false, skip_mono = false};
+      alt_name = Binding.empty, coind = coind, no_elim = false, no_ind = false,
+      skip_mono = false, fork_mono = not int};
   in
     lthy
     |> LocalTheory.set_group (serial_string ())
@@ -956,12 +960,12 @@
   Scan.optional (P.$$$ "where" |-- P.!!! SpecParse.specification) [] --
   Scan.optional (P.$$$ "monos" |-- P.!!! SpecParse.xthms1) []
   >> (fn (((preds, params), specs), monos) =>
-      (snd o gen_add_inductive mk_def true coind preds params (flatten_specification specs) monos));
+      (snd oo gen_add_inductive mk_def true coind preds params (flatten_specification specs) monos));
 
 val ind_decl = gen_ind_decl add_ind_def;
 
-val _ = OuterSyntax.local_theory "inductive" "define inductive predicates" K.thy_decl (ind_decl false);
-val _ = OuterSyntax.local_theory "coinductive" "define coinductive predicates" K.thy_decl (ind_decl true);
+val _ = OuterSyntax.local_theory' "inductive" "define inductive predicates" K.thy_decl (ind_decl false);
+val _ = OuterSyntax.local_theory' "coinductive" "define coinductive predicates" K.thy_decl (ind_decl true);
 
 val _ =
   OuterSyntax.local_theory "inductive_cases"
--- a/src/HOL/Tools/inductive_realizer.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/inductive_realizer.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,9 +1,8 @@
 (*  Title:      HOL/Tools/inductive_realizer.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Porgram extraction from proofs involving inductive predicates:
-Realizers for induction and elimination rules
+Realizers for induction and elimination rules.
 *)
 
 signature INDUCTIVE_REALIZER =
@@ -50,7 +49,7 @@
       t $ strip_all' used names Q
   | strip_all' _ _ t = t;
 
-fun strip_all t = strip_all' (add_term_free_names (t, [])) [] t;
+fun strip_all t = strip_all' (Term.add_free_names t []) [] t;
 
 fun strip_one name (Const ("all", _) $ Abs (s, T, Const ("==>", _) $ P $ Q)) =
       (subst_bound (Free (name, T), P), subst_bound (Free (name, T), Q))
@@ -60,11 +59,11 @@
       (Var ((a, i), T), vs) => (case strip_type T of
         (_, Type (s, _)) => if s mem ["bool"] then (a, T) :: vs else vs
       | _ => vs)
-    | (_, vs) => vs) [] (term_vars prop);
+    | (_, vs) => vs) [] (OldTerm.term_vars prop);
 
 fun dt_of_intrs thy vs nparms intrs =
   let
-    val iTs = term_tvars (prop_of (hd intrs));
+    val iTs = OldTerm.term_tvars (prop_of (hd intrs));
     val Tvs = map TVar iTs;
     val (Const (s, _), ts) = strip_comb (HOLogic.dest_Trueprop
       (Logic.strip_imp_concl (prop_of (hd intrs))));
@@ -101,7 +100,7 @@
 fun mk_realizes_eqn n vs nparms intrs =
   let
     val concl = HOLogic.dest_Trueprop (concl_of (hd intrs));
-    val iTs = term_tvars concl;
+    val iTs = OldTerm.term_tvars concl;
     val Tvs = map TVar iTs;
     val (h as Const (s, T), us) = strip_comb concl;
     val params = List.take (us, nparms);
@@ -147,7 +146,7 @@
     val conclT = Extraction.etype_of thy vs [] (Logic.strip_imp_concl rule');
     val used = map (fst o dest_Free) args;
 
-    fun is_rec t = not (null (term_consts t inter rsets));
+    val is_rec = exists_Const (fn (c, _) => member (op =) rsets c);
 
     fun is_meta (Const ("all", _) $ Abs (s, _, P)) = is_meta P
       | is_meta (Const ("==>", _) $ _ $ Q) = is_meta Q
@@ -276,7 +275,7 @@
   let
     val qualifier = NameSpace.qualifier (name_of_thm induct);
     val inducts = PureThy.get_thms thy (NameSpace.qualified qualifier "inducts");
-    val iTs = term_tvars (prop_of (hd intrs));
+    val iTs = OldTerm.term_tvars (prop_of (hd intrs));
     val ar = length vs + length iTs;
     val params = InductivePackage.params_of raw_induct;
     val arities = InductivePackage.arities_of raw_induct;
@@ -348,7 +347,7 @@
     val (ind_info, thy3') = thy2 |>
       InductivePackage.add_inductive_global (serial_string ())
         {quiet_mode = false, verbose = false, kind = Thm.theoremK, alt_name = Binding.empty,
-          coind = false, no_elim = false, no_ind = false, skip_mono = false}
+          coind = false, no_elim = false, no_ind = false, skip_mono = false, fork_mono = false}
         rlzpreds rlzparams (map (fn (rintr, intr) =>
           ((Binding.name (Sign.base_name (name_of_thm intr)), []),
            subst_atomic rlzpreds' (Logic.unvarify rintr)))
@@ -371,7 +370,7 @@
           (vs' @ Ps) rec_names rss' intrs dummies;
         val rlzs = map (fn (r, ind) => Extraction.realizes_of thy (vs' @ Ps) r
           (prop_of ind)) (rs ~~ inducts);
-        val used = foldr add_term_free_names [] rlzs;
+        val used = fold Term.add_free_names rlzs [];
         val rnames = Name.variant_list used (replicate (length inducts) "r");
         val rnames' = Name.variant_list
           (used @ rnames) (replicate (length intrs) "s");
--- a/src/HOL/Tools/inductive_set_package.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/inductive_set_package.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -20,7 +20,7 @@
     (Binding.T * string option * mixfix) list ->
     (Binding.T * string option * mixfix) list ->
     (Attrib.binding * string) list -> (Facts.ref * Attrib.src list) list ->
-    local_theory -> InductivePackage.inductive_result * local_theory
+    bool -> local_theory -> InductivePackage.inductive_result * local_theory
   val codegen_preproc: theory -> thm list -> thm list
   val setup: theory -> theory
 end;
@@ -169,7 +169,7 @@
     ({to_set_simps = to_set_simps1, to_pred_simps = to_pred_simps1,
       set_arities = set_arities1, pred_arities = pred_arities1},
      {to_set_simps = to_set_simps2, to_pred_simps = to_pred_simps2,
-      set_arities = set_arities2, pred_arities = pred_arities2}) =
+      set_arities = set_arities2, pred_arities = pred_arities2}) : T =
     {to_set_simps = Thm.merge_thms (to_set_simps1, to_set_simps2),
      to_pred_simps = Thm.merge_thms (to_pred_simps1, to_pred_simps2),
      set_arities = Symtab.merge_list op = (set_arities1, set_arities2),
@@ -403,7 +403,8 @@
 
 (**** definition of inductive sets ****)
 
-fun add_ind_set_def {quiet_mode, verbose, kind, alt_name, coind, no_elim, no_ind, skip_mono}
+fun add_ind_set_def
+    {quiet_mode, verbose, kind, alt_name, coind, no_elim, no_ind, skip_mono, fork_mono}
     cs intros monos params cnames_syn ctxt =
   let
     val thy = ProofContext.theory_of ctxt;
@@ -468,7 +469,8 @@
     val ({preds, intrs, elims, raw_induct, ...}, ctxt1) =
       InductivePackage.add_ind_def
         {quiet_mode = quiet_mode, verbose = verbose, kind = kind, alt_name = Binding.empty,
-          coind = coind, no_elim = no_elim, no_ind = no_ind, skip_mono = skip_mono}
+          coind = coind, no_elim = no_elim, no_ind = no_ind,
+          skip_mono = skip_mono, fork_mono = fork_mono}
         cs' intros' monos' params1 cnames_syn' ctxt;
 
     (* define inductive sets using previously defined predicates *)
@@ -550,10 +552,10 @@
 val ind_set_decl = InductivePackage.gen_ind_decl add_ind_set_def;
 
 val _ =
-  OuterSyntax.local_theory "inductive_set" "define inductive sets" K.thy_decl (ind_set_decl false);
+  OuterSyntax.local_theory' "inductive_set" "define inductive sets" K.thy_decl (ind_set_decl false);
 
 val _ =
-  OuterSyntax.local_theory "coinductive_set" "define coinductive sets" K.thy_decl (ind_set_decl true);
+  OuterSyntax.local_theory' "coinductive_set" "define coinductive sets" K.thy_decl (ind_set_decl true);
 
 end;
 
--- a/src/HOL/Tools/int_arith.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/int_arith.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
-(*  Title:      HOL/int_arith1.ML
-    ID:         $Id$
+(*  Title:      HOL/Tools/int_arith1.ML
     Authors:    Larry Paulson and Tobias Nipkow
 
 Simprocs and decision procedure for linear arithmetic.
@@ -66,12 +65,12 @@
     EQUAL => int_ord (Int.sign i, Int.sign j) 
   | ord => ord);
 
-(*This resembles Term.term_ord, but it puts binary numerals before other
+(*This resembles TermOrd.term_ord, but it puts binary numerals before other
   non-atomic terms.*)
 local open Term 
 in 
 fun numterm_ord (Abs (_, T, t), Abs(_, U, u)) =
-      (case numterm_ord (t, u) of EQUAL => typ_ord (T, U) | ord => ord)
+      (case numterm_ord (t, u) of EQUAL => TermOrd.typ_ord (T, U) | ord => ord)
   | numterm_ord
      (Const(@{const_name Int.number_of}, _) $ v, Const(@{const_name Int.number_of}, _) $ w) =
      num_ord (HOLogic.dest_numeral v, HOLogic.dest_numeral w)
@@ -81,7 +80,7 @@
       (case int_ord (size_of_term t, size_of_term u) of
         EQUAL =>
           let val (f, ts) = strip_comb t and (g, us) = strip_comb u in
-            (case hd_ord (f, g) of EQUAL => numterms_ord (ts, us) | ord => ord)
+            (case TermOrd.hd_ord (f, g) of EQUAL => numterms_ord (ts, us) | ord => ord)
           end
       | ord => ord)
 and numterms_ord (ts, us) = list_ord numterm_ord (ts, us)
@@ -164,7 +163,7 @@
 (*Express t as a product of (possibly) a numeral with other sorted terms*)
 fun dest_coeff sign (Const (@{const_name HOL.uminus}, _) $ t) = dest_coeff (~sign) t
   | dest_coeff sign t =
-    let val ts = sort Term.term_ord (dest_prod t)
+    let val ts = sort TermOrd.term_ord (dest_prod t)
         val (n, ts') = find_first_numeral [] ts
                           handle TERM _ => (1, ts)
     in (sign*n, mk_prod (Term.fastype_of t) ts') end;
--- a/src/HOL/Tools/lin_arith.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/lin_arith.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,6 +1,5 @@
 (*  Title:      HOL/Tools/lin_arith.ML
-    ID:         $Id$
-    Author:     Tjark Weber and Tobias Nipkow
+    Author:     Tjark Weber and Tobias Nipkow, TU Muenchen
 
 HOL setup for linear arithmetic (see Provers/Arith/fast_lin_arith.ML).
 *)
@@ -99,8 +98,9 @@
             tactics: arith_tactic list};
   val empty = {splits = [], inj_consts = [], discrete = [], tactics = []};
   val extend = I;
-  fun merge _ ({splits= splits1, inj_consts= inj_consts1, discrete= discrete1, tactics= tactics1},
-             {splits= splits2, inj_consts= inj_consts2, discrete= discrete2, tactics= tactics2}) =
+  fun merge _
+   ({splits= splits1, inj_consts= inj_consts1, discrete= discrete1, tactics= tactics1},
+    {splits= splits2, inj_consts= inj_consts2, discrete= discrete2, tactics= tactics2}) : T =
    {splits = Library.merge Thm.eq_thm_prop (splits1, splits2),
     inj_consts = Library.merge (op =) (inj_consts1, inj_consts2),
     discrete = Library.merge (op =) (discrete1, discrete2),
@@ -149,9 +149,9 @@
 
 fun add_atom (t : term) (m : Rat.rat) (p : (term * Rat.rat) list, i : Rat.rat) :
              (term * Rat.rat) list * Rat.rat =
-  case AList.lookup (op =) p t of
+  case AList.lookup Pattern.aeconv p t of
       NONE   => ((t, m) :: p, i)
-    | SOME n => (AList.update (op =) (t, Rat.add n m) p, i);
+    | SOME n => (AList.update Pattern.aeconv (t, Rat.add n m) p, i);
 
 (* decompose nested multiplications, bracketing them to the right and combining
    all their coefficients
@@ -340,7 +340,7 @@
 
 fun subst_term ([] : (term * term) list) (t : term) = t
   | subst_term pairs                     t          =
-      (case AList.lookup (op aconv) pairs t of
+      (case AList.lookup Pattern.aeconv pairs t of
         SOME new =>
           new
       | NONE     =>
@@ -682,7 +682,7 @@
 
 fun negated_term_occurs_positively (terms : term list) : bool =
   List.exists
-    (fn (Trueprop $ (Const ("Not", _) $ t)) => member (op aconv) terms (Trueprop $ t)
+    (fn (Trueprop $ (Const ("Not", _) $ t)) => member Pattern.aeconv terms (Trueprop $ t)
       | _                                   => false)
     terms;
 
--- a/src/HOL/Tools/meson.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/meson.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,11 +1,8 @@
 (*  Title:      HOL/Tools/meson.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1992  University of Cambridge
 
 The MESON resolution proof procedure for HOL.
-
-When making clauses, avoids using the rewriter -- instead uses RS recursively
+When making clauses, avoids using the rewriter -- instead uses RS recursively.
 *)
 
 signature MESON =
@@ -400,7 +397,7 @@
   Also rejects functions whose arguments are Booleans or other functions.*)
 fun is_fol_term thy t =
     Term.is_first_order ["all","All","Ex"] t andalso
-    not (exists (has_bool o fastype_of) (term_vars t)  orelse
+    not (exists_subterm (fn Var (_, T) => has_bool T | _ => false) t  orelse
          has_bool_arg_const t  orelse
          exists_Const (higher_inst_const thy) t orelse
          has_meta_conn t);
@@ -699,4 +696,3 @@
   handle Subscript => Seq.empty;
 
 end;
-
--- a/src/HOL/Tools/metis_tools.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/metis_tools.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -2,7 +2,7 @@
     Author:     Kong W. Susanto and Lawrence C. Paulson, CU Computer Laboratory
     Copyright   Cambridge University 2007
 
-HOL setup for the Metis prover (version 2.0 from 2007).
+HOL setup for the Metis prover.
 *)
 
 signature METIS_TOOLS =
@@ -280,9 +280,10 @@
 
   fun cterm_incr_types thy idx = cterm_of thy o (map_types (Logic.incr_tvar idx));
 
-  fun inst_excluded_middle th thy i_atm =
-    let val vx = hd (term_vars (prop_of th))
-        val substs = [(cterm_of thy vx, cterm_of thy i_atm)]
+  fun inst_excluded_middle thy i_atm =
+    let val th = EXCLUDED_MIDDLE
+        val [vx] = Term.add_vars (prop_of th) []
+        val substs = [(cterm_of thy (Var vx), cterm_of thy i_atm)]
     in  cterm_instantiate substs th  end;
 
   (* INFERENCE RULE: AXIOM *)
@@ -291,7 +292,7 @@
 
   (* INFERENCE RULE: ASSUME *)
   fun assume_inf ctxt atm =
-    inst_excluded_middle EXCLUDED_MIDDLE
+    inst_excluded_middle
       (ProofContext.theory_of ctxt)
       (singleton (fol_terms_to_hol ctxt) (Metis.Term.Fn atm));
 
@@ -301,12 +302,12 @@
   fun inst_inf ctxt thpairs fsubst th =    
     let val thy = ProofContext.theory_of ctxt
         val i_th   = lookth thpairs th
-        val i_th_vars = term_vars (prop_of i_th)
-        fun find_var x = valOf (List.find (fn Term.Var((a,_),_) => a=x) i_th_vars)
+        val i_th_vars = Term.add_vars (prop_of i_th) []
+        fun find_var x = valOf (List.find (fn ((a,_),_) => a=x) i_th_vars)
         fun subst_translation (x,y) =
               let val v = find_var x
                   val t = fol_term_to_hol_RAW ctxt y (*we call infer_types below*)
-              in  SOME (cterm_of thy v, t)  end
+              in  SOME (cterm_of thy (Var v), t)  end
               handle Option =>
                   (Output.debug (fn() => "List.find failed for the variable " ^ x ^
                                          " in " ^ Display.string_of_thm i_th);
@@ -370,7 +371,7 @@
     end;
 
   (* INFERENCE RULE: REFL *)
-  val refl_x = cterm_of (the_context ()) (hd (term_vars (prop_of REFL_THM)));
+  val refl_x = cterm_of (the_context ()) (Var (hd (Term.add_vars (prop_of REFL_THM) [])));
   val refl_idx = 1 + Thm.maxidx_of REFL_THM;
 
   fun refl_inf ctxt t =
@@ -429,7 +430,7 @@
         val subst' = incr_indexes (imax+1) (if pos then subst_em else ssubst_em)
         val _ = Output.debug (fn () => "subst' " ^ Display.string_of_thm subst')
         val eq_terms = map (pairself (cterm_of thy))
-                           (ListPair.zip (term_vars (prop_of subst'), [tm_abs, tm_subst, i_tm]))
+                           (ListPair.zip (OldTerm.term_vars (prop_of subst'), [tm_abs, tm_subst, i_tm]))
     in  cterm_instantiate eq_terms subst'  end;
 
   val factor = Seq.hd o distinct_subgoals_tac;
--- a/src/HOL/Tools/nat_simprocs.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/nat_simprocs.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,7 +1,5 @@
-(*  Title:      HOL/nat_simprocs.ML
-    ID:         $Id$
+(*  Title:      HOL/Tools/nat_simprocs.ML
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   2000  University of Cambridge
 
 Simprocs for nat numerals.
 *)
@@ -81,7 +79,7 @@
 
 (*Express t as a product of (possibly) a numeral with other factors, sorted*)
 fun dest_coeff t =
-    let val ts = sort Term.term_ord (dest_prod t)
+    let val ts = sort TermOrd.term_ord (dest_prod t)
         val (n, _, ts') = find_first_numeral [] ts
                           handle TERM _ => (1, one, ts)
     in (n, mk_prod ts') end;
--- a/src/HOL/Tools/numeral_syntax.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/numeral_syntax.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/numeral_syntax.ML
-    ID:         $Id$
     Authors:    Markus Wenzel, TU Muenchen
 
 Concrete syntax for generic numerals -- preserves leading zeros/ones.
@@ -19,12 +18,11 @@
 
 fun mk_bin num =
   let
-    val {leading_zeros = z, value, ...} = Syntax.read_xnum num;
     fun bit b bs = HOLogic.mk_bit b $ bs;
-    fun mk 0 = (* FIXME funpow z (bit 0) *) (Syntax.const @{const_name Int.Pls})
-      | mk ~1 = (* FIXME funpow z (bit 1) *) (Syntax.const @{const_name Int.Min})
+    fun mk 0 = Syntax.const @{const_name Int.Pls}
+      | mk ~1 = Syntax.const @{const_name Int.Min}
       | mk i = let val (q, r) = Integer.div_mod i 2 in bit r (mk q) end;
-  in mk value end;
+  in mk (#value (Syntax.read_xnum num)) end;
 
 in
 
@@ -65,15 +63,18 @@
     else sign ^ implode (replicate z "0") ^ num
   end;
 
+fun syntax_numeral t =
+  Syntax.const "_Numeral" $ (Syntax.const "_numeral" $ Syntax.free (dest_bin_str t));
+
 in
 
 fun numeral_tr' show_sorts (*"number_of"*) (Type ("fun", [_, T])) (t :: ts) =
-      let val t' = Syntax.const "_Numeral" $ Syntax.free (dest_bin_str t) in
-        if not (! show_types) andalso can Term.dest_Type T then t'
-        else Syntax.const Syntax.constrainC $ t' $ Syntax.term_of_typ show_sorts T
-      end
+      let val t' =
+        if not (! show_types) andalso can Term.dest_Type T then syntax_numeral t
+        else Syntax.const Syntax.constrainC $ syntax_numeral t $ Syntax.term_of_typ show_sorts T
+      in list_comb (t', ts) end
   | numeral_tr' _ (*"number_of"*) T (t :: ts) =
-      if T = dummyT then Syntax.const "_Numeral" $ Syntax.free (dest_bin_str t)
+      if T = dummyT then list_comb (syntax_numeral t, ts)
       else raise Match
   | numeral_tr' _ (*"number_of"*) _ _ = raise Match;
 
--- a/src/HOL/Tools/old_primrec_package.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/old_primrec_package.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,6 +1,6 @@
 (*  Title:      HOL/Tools/old_primrec_package.ML
-    ID:         $Id$
-    Author:     Stefan Berghofer, TU Muenchen and Norbert Voelker, FernUni Hagen
+    Author:     Norbert Voelker, FernUni Hagen
+    Author:     Stefan Berghofer, TU Muenchen
 
 Package for defining functions on datatypes by primitive recursion.
 *)
@@ -34,14 +34,13 @@
   same type in all introduction rules*)
 fun unify_consts thy cs intr_ts =
   (let
-    val add_term_consts_2 = fold_aterms (fn Const c => insert (op =) c | _ => I);
     fun varify (t, (i, ts)) =
       let val t' = map_types (Logic.incr_tvar (i + 1)) (snd (Type.varify [] t))
       in (maxidx_of_term t', t'::ts) end;
     val (i, cs') = foldr varify (~1, []) cs;
     val (i', intr_ts') = foldr varify (i, []) intr_ts;
-    val rec_consts = fold add_term_consts_2 cs' [];
-    val intr_consts = fold add_term_consts_2 intr_ts' [];
+    val rec_consts = fold Term.add_consts cs' [];
+    val intr_consts = fold Term.add_consts intr_ts' [];
     fun unify (cname, cT) =
       let val consts = map snd (filter (fn (c, _) => c = cname) intr_consts)
       in fold (Sign.typ_unify thy) ((replicate (length consts) cT) ~~ consts) end;
@@ -58,7 +57,7 @@
 fun process_eqn thy eq rec_fns =
   let
     val (lhs, rhs) =
-      if null (term_vars eq) then
+      if null (Term.add_vars eq []) then
         HOLogic.dest_eq (HOLogic.dest_Trueprop eq)
         handle TERM _ => raise RecError "not a proper equation"
       else raise RecError "illegal schematic variable(s)";
@@ -91,7 +90,7 @@
     else
      (check_vars "repeated variable names in pattern: " (duplicates (op =) lfrees);
       check_vars "extra variables on rhs: "
-        (map dest_Free (term_frees rhs) \\ lfrees);
+        (map dest_Free (OldTerm.term_frees rhs) \\ lfrees);
       case AList.lookup (op =) rec_fns fnameT of
         NONE =>
           (fnameT, (tname, rpos, [(cname, (ls, cargs, rs, rhs, eq))]))::rec_fns
@@ -161,7 +160,7 @@
               val recs = filter (is_rec_type o snd) (cargs' ~~ cargs);
               val rargs = map fst recs;
               val subs = map (rpair dummyT o fst)
-                (rev (rename_wrt_term rhs rargs));
+                (rev (Term.rename_wrt_term rhs rargs));
               val (rhs', (fnameTs'', fnss'')) =
                   (subst (map (fn ((x, y), z) =>
                                (Free x, (body_index y, Free z)))
--- a/src/HOL/Tools/primrec_package.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/primrec_package.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -69,7 +69,7 @@
     else
      (check_vars "repeated variable names in pattern: " (duplicates (op =) lfrees);
       check_vars "extra variables on rhs: "
-        (map dest_Free (term_frees rhs) |> subtract (op =) lfrees
+        (map dest_Free (OldTerm.term_frees rhs) |> subtract (op =) lfrees
           |> filter_out (is_fixed o fst));
       case AList.lookup (op =) rec_fns fname of
         NONE =>
@@ -141,7 +141,7 @@
               val recs = filter (is_rec_type o snd) (cargs' ~~ cargs);
               val rargs = map fst recs;
               val subs = map (rpair dummyT o fst)
-                (rev (rename_wrt_term rhs rargs));
+                (rev (Term.rename_wrt_term rhs rargs));
               val (rhs', (fnames'', fnss'')) = subst (map2 (fn (x, y) => fn z =>
                 (Free x, (body_index y, Free z))) recs subs) rhs (fnames', fnss')
                   handle PrimrecError (s, NONE) => primrec_error_eqn s eq
--- a/src/HOL/Tools/recfun_codegen.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/recfun_codegen.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/recfun_codegen.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Code generator for recursive functions.
@@ -46,7 +45,7 @@
   | expand_eta thy (thms as thm :: _) =
       let
         val (_, ty) = Code_Unit.const_typ_eqn thm;
-      in if (null o Term.typ_tvars) ty orelse (null o fst o strip_type) ty
+      in if null (Term.add_tvarsT ty []) orelse (null o fst o strip_type) ty
         then thms
         else map (Code_Unit.expand_eta thy 1) thms
       end;
--- a/src/HOL/Tools/record_package.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/record_package.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/record_package.ML
-    ID:         $Id$
     Author:     Wolfgang Naraschewski, Norbert Schirmer and Markus Wenzel, TU Muenchen
 
 Extensible records with structural subtyping in HOL.
@@ -934,7 +933,7 @@
                  then (let
                         val P=cterm_of thy (abstract_over_fun_app trm);
                         val thm = mk_fun_apply_eq trm thy;
-                        val PV = cterm_of thy (hd (term_vars (prop_of thm)));
+                        val PV = cterm_of thy (hd (OldTerm.term_vars (prop_of thm)));
                         val thm' = cterm_instantiate [(PV,P)] thm;
                        in SOME  thm' end handle TERM _ => NONE)
                 else NONE)
@@ -1305,7 +1304,7 @@
         | _ => false);
 
     val goal = nth (Thm.prems_of st) (i - 1);
-    val frees = List.filter (is_recT o type_of) (term_frees goal);
+    val frees = List.filter (is_recT o type_of) (OldTerm.term_frees goal);
 
     fun mk_split_free_tac free induct_thm i =
         let val cfree = cterm_of thy free;
@@ -1384,14 +1383,14 @@
   let
     val ctxt' = fold (Variable.declare_typ o TFree) env ctxt;
     val T = Syntax.read_typ ctxt' raw_T;
-    val env' = Term.add_typ_tfrees (T, env);
+    val env' = OldTerm.add_typ_tfrees (T, env);
   in (T, env') end;
 
 fun cert_typ ctxt raw_T env =
   let
     val thy = ProofContext.theory_of ctxt;
     val T = Type.no_tvars (Sign.certify_typ thy raw_T) handle TYPE (msg, _, _) => error msg;
-    val env' = Term.add_typ_tfrees (T, env);
+    val env' = OldTerm.add_typ_tfrees (T, env);
   in (T, env') end;
 
 
@@ -1426,7 +1425,7 @@
           (hd (rev (Logic.strip_assums_hyp (hd (prems_of rule')))))))), true)
       | [x] => (head_of x, false));
     val rule'' = cterm_instantiate (map (pairself cert) (case (rev params) of
-        [] => (case AList.lookup (op =) (map dest_Free (term_frees (prop_of st))) s of
+        [] => (case AList.lookup (op =) (map dest_Free (OldTerm.term_frees (prop_of st))) s of
           NONE => sys_error "try_param_tac: no such variable"
         | SOME T => [(P, if ca then concl else lambda (Free (s, T)) concl),
             (x, Free (s, T))])
@@ -1777,7 +1776,7 @@
     val names = map fst fields;
     val extN = full bname;
     val types = map snd fields;
-    val alphas_fields = foldr add_typ_tfree_names [] types;
+    val alphas_fields = foldr OldTerm.add_typ_tfree_names [] types;
     val alphas_ext = alphas inter alphas_fields;
     val len = length fields;
     val variants = Name.variant_list (moreN::rN::rN ^ "'"::wN::parent_variants) (map fst bfields);
@@ -2226,7 +2225,7 @@
     val init_env =
       (case parent of
         NONE => []
-      | SOME (types, _) => foldr Term.add_typ_tfrees [] types);
+      | SOME (types, _) => foldr OldTerm.add_typ_tfrees [] types);
 
 
     (* fields *)
--- a/src/HOL/Tools/refute.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/refute.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,7 +1,5 @@
 (*  Title:      HOL/Tools/refute.ML
-    ID:         $Id$
-    Author:     Tjark Weber
-    Copyright   2003-2007
+    Author:     Tjark Weber, TU Muenchen
 
 Finite model generation for HOL formulas, using a SAT solver.
 *)
@@ -391,11 +389,6 @@
 (* TRANSLATION HOL -> PROPOSITIONAL LOGIC, BOOLEAN ASSIGNMENT -> MODEL       *)
 (* ------------------------------------------------------------------------- *)
 
-  (* (''a * 'b) list -> ''a -> 'b *)
-
-  fun lookup xs key =
-    Option.valOf (AList.lookup (op =) xs key);
-
 (* ------------------------------------------------------------------------- *)
 (* typ_of_dtyp: converts a data type ('DatatypeAux.dtyp') into a type        *)
 (*              ('Term.typ'), given type parameters for the data type's type *)
@@ -407,12 +400,12 @@
 
   fun typ_of_dtyp descr typ_assoc (DatatypeAux.DtTFree a) =
     (* replace a 'DtTFree' variable by the associated type *)
-    lookup typ_assoc (DatatypeAux.DtTFree a)
+    the (AList.lookup (op =) typ_assoc (DatatypeAux.DtTFree a))
     | typ_of_dtyp descr typ_assoc (DatatypeAux.DtType (s, ds)) =
     Type (s, map (typ_of_dtyp descr typ_assoc) ds)
     | typ_of_dtyp descr typ_assoc (DatatypeAux.DtRec i) =
     let
-      val (s, ds, _) = lookup descr i
+      val (s, ds, _) = the (AList.lookup (op =) descr i)
     in
       Type (s, map (typ_of_dtyp descr typ_assoc) ds)
     end;
@@ -426,7 +419,7 @@
   fun close_form t =
   let
     (* (Term.indexname * Term.typ) list *)
-    val vars = sort_wrt (fst o fst) (map dest_Var (term_vars t))
+    val vars = sort_wrt (fst o fst) (map dest_Var (OldTerm.term_vars t))
   in
     Library.foldl (fn (t', ((x, i), T)) =>
       (Term.all T) $ Abs (x, T, abstract_over (Var ((x, i), T), t')))
@@ -792,7 +785,7 @@
       (* replace the (at most one) schematic type variable in each axiom *)
       (* by the actual type 'T'                                          *)
       val monomorphic_class_axioms = map (fn (axname, ax) =>
-        (case Term.term_tvars ax of
+        (case Term.add_tvars ax [] of
           [] =>
           (axname, ax)
         | [(idx, S)] =>
@@ -852,14 +845,14 @@
       | Const ("The", T)                =>
         let
           val ax = specialize_type thy ("The", T)
-            (lookup axioms "HOL.the_eq_trivial")
+            (the (AList.lookup (op =) axioms "HOL.the_eq_trivial"))
         in
           collect_this_axiom ("HOL.the_eq_trivial", ax) axs
         end
       | Const ("Hilbert_Choice.Eps", T) =>
         let
           val ax = specialize_type thy ("Hilbert_Choice.Eps", T)
-            (lookup axioms "Hilbert_Choice.someI")
+            (the (AList.lookup (op =) axioms "Hilbert_Choice.someI"))
         in
           collect_this_axiom ("Hilbert_Choice.someI", ax) axs
         end
@@ -944,23 +937,20 @@
 (*               and all mutually recursive IDTs are considered.             *)
 (* ------------------------------------------------------------------------- *)
 
-  (* theory -> Term.term -> Term.typ list *)
-
   fun ground_types thy t =
   let
-    (* Term.typ * Term.typ list -> Term.typ list *)
-    fun collect_types (T, acc) =
+    fun collect_types T acc =
       (case T of
-        Type ("fun", [T1, T2]) => collect_types (T1, collect_types (T2, acc))
+        Type ("fun", [T1, T2]) => collect_types T1 (collect_types T2 acc)
       | Type ("prop", [])      => acc
-      | Type ("set", [T1])     => collect_types (T1, acc)
+      | Type ("set", [T1])     => collect_types T1 acc
       | Type (s, Ts)           =>
         (case DatatypePackage.get_datatype thy s of
           SOME info =>  (* inductive datatype *)
           let
             val index        = #index info
             val descr        = #descr info
-            val (_, typs, _) = lookup descr index
+            val (_, typs, _) = the (AList.lookup (op =) descr index)
             val typ_assoc    = typs ~~ Ts
             (* sanity check: every element in 'dtyps' must be a *)
             (* 'DtTFree'                                        *)
@@ -978,15 +968,15 @@
             in
               case d of
                 DatatypeAux.DtTFree _ =>
-                collect_types (dT, acc)
+                collect_types dT acc
               | DatatypeAux.DtType (_, ds) =>
-                collect_types (dT, foldr collect_dtyp acc ds)
+                collect_types dT (foldr collect_dtyp acc ds)
               | DatatypeAux.DtRec i =>
                 if dT mem acc then
                   acc  (* prevent infinite recursion *)
                 else
                   let
-                    val (_, dtyps, dconstrs) = lookup descr i
+                    val (_, dtyps, dconstrs) = the (AList.lookup (op =) descr i)
                     (* if the current type is a recursive IDT (i.e. a depth *)
                     (* is required), add it to 'acc'                        *)
                     val acc_dT = if Library.exists (fn (_, ds) =>
@@ -1010,11 +1000,11 @@
         | NONE =>
           (* not an inductive datatype, e.g. defined via "typedef" or *)
           (* "typedecl"                                               *)
-          insert (op =) T (foldr collect_types acc Ts))
+          insert (op =) T (fold collect_types Ts acc))
       | TFree _                => insert (op =) T acc
       | TVar _                 => insert (op =) T acc)
   in
-    it_term_types collect_types (t, [])
+    fold_types collect_types t []
   end;
 
 (* ------------------------------------------------------------------------- *)
@@ -1170,7 +1160,7 @@
             let
               val index           = #index info
               val descr           = #descr info
-              val (_, _, constrs) = lookup descr index
+              val (_, _, constrs) = the (AList.lookup (op =) descr index)
             in
               (* recursive datatype? *)
               Library.exists (fn (_, ds) =>
@@ -1289,12 +1279,12 @@
     (* terms containing them (their logical meaning is that there EXISTS a *)
     (* type s.t. ...; to refute such a formula, we would have to show that *)
     (* for ALL types, not ...)                                             *)
-    val _ = null (term_tvars t) orelse
+    val _ = null (Term.add_tvars t []) orelse
       error "Term to be refuted contains schematic type variables"
 
     (* existential closure over schematic variables *)
     (* (Term.indexname * Term.typ) list *)
-    val vars = sort_wrt (fst o fst) (map dest_Var (term_vars t))
+    val vars = sort_wrt (fst o fst) (map dest_Var (OldTerm.term_vars t))
     (* Term.term *)
     val ex_closure = Library.foldl (fn (t', ((x, i), T)) =>
       (HOLogic.exists_const T) $
@@ -1662,7 +1652,7 @@
       fun interpret_groundtype () =
       let
         (* the model must specify a size for ground types *)
-        val size = (if T = Term.propT then 2 else lookup typs T)
+        val size = if T = Term.propT then 2 else the (AList.lookup (op =) typs T)
         val next = next_idx+size
         (* check if 'maxvars' is large enough *)
         val _    = (if next-1>maxvars andalso maxvars>0 then
@@ -2025,7 +2015,7 @@
             let
               val index               = #index info
               val descr               = #descr info
-              val (_, dtyps, constrs) = lookup descr index
+              val (_, dtyps, constrs) = the (AList.lookup (op =) descr index)
               val typ_assoc           = dtyps ~~ Ts
               (* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
               val _ = if Library.exists (fn d =>
@@ -2149,7 +2139,7 @@
             let
               val index               = #index info
               val descr               = #descr info
-              val (_, dtyps, constrs) = lookup descr index
+              val (_, dtyps, constrs) = the (AList.lookup (op =) descr index)
               val typ_assoc           = dtyps ~~ Ts
               (* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
               val _ = if Library.exists (fn d =>
@@ -2209,7 +2199,7 @@
             let
               val index               = #index info
               val descr               = #descr info
-              val (_, dtyps, constrs) = lookup descr index
+              val (_, dtyps, constrs) = the (AList.lookup (op =) descr index)
               val typ_assoc           = dtyps ~~ Ts'
               (* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
               val _ = if Library.exists (fn d =>
@@ -2405,7 +2395,7 @@
                   (* the index of some mutually recursive IDT               *)
                   val index         = #index info
                   val descr         = #descr info
-                  val (_, dtyps, _) = lookup descr index
+                  val (_, dtyps, _) = the (AList.lookup (op =) descr index)
                   (* sanity check: we assume that the order of constructors *)
                   (*               in 'descr' is the same as the order of   *)
                   (*               corresponding parameters, otherwise the  *)
@@ -2464,7 +2454,7 @@
                         end
                       | DatatypeAux.DtRec i =>
                         let
-                          val (_, ds, _) = lookup descr i
+                          val (_, ds, _) = the (AList.lookup (op =) descr i)
                           val (_, Ts)    = dest_Type T
                         in
                           rec_typ_assoc ((d, T)::acc) ((ds ~~ Ts) @ xs)
@@ -2478,10 +2468,10 @@
                         raise REFUTE ("IDT_recursion_interpreter",
                           "different type associations for the same dtyp"))
                   (* (DatatypeAux.dtyp * Term.typ) list *)
-                  val typ_assoc = List.filter
+                  val typ_assoc = filter
                     (fn (DatatypeAux.DtTFree _, _) => true | (_, _) => false)
                     (rec_typ_assoc []
-                      (#2 (lookup descr idt_index) ~~ (snd o dest_Type) IDT))
+                      (#2 (the (AList.lookup (op =) descr idt_index)) ~~ (snd o dest_Type) IDT))
                   (* sanity check: typ_assoc must associate types to the   *)
                   (*               elements of 'dtyps' (and only to those) *)
                   val _ = if not (Library.eq_set (dtyps, map fst typ_assoc))
@@ -2605,7 +2595,7 @@
                           SOME args => (n, args)
                         | NONE      => get_cargs_rec (n+1, xs))
                     in
-                      get_cargs_rec (0, lookup mc_intrs idx)
+                      get_cargs_rec (0, the (AList.lookup (op =) mc_intrs idx))
                     end
                   (* computes one entry in 'REC_OPERATORS', and recursively *)
                   (* all entries needed for it, where 'idx' gives the       *)
@@ -2613,7 +2603,7 @@
                   (* int -> int -> interpretation *)
                   fun compute_array_entry idx elem =
                   let
-                    val arr          = lookup REC_OPERATORS idx
+                    val arr          = the (AList.lookup (op =) REC_OPERATORS idx)
                     val (flag, intr) = Array.sub (arr, elem)
                   in
                     if flag then
@@ -2627,7 +2617,7 @@
                         val (c, args) = get_cargs idx elem
                         (* find the indices of the constructor's /recursive/ *)
                         (* arguments                                         *)
-                        val (_, _, constrs) = lookup descr idx
+                        val (_, _, constrs) = the (AList.lookup (op =) descr idx)
                         val (_, dtyps)      = List.nth (constrs, c)
                         val rec_dtyps_args  = List.filter
                           (DatatypeAux.is_rec_type o fst) (dtyps ~~ args)
@@ -2679,7 +2669,7 @@
                         result
                       end
                   end
-                  val idt_size = Array.length (lookup REC_OPERATORS idt_index)
+                  val idt_size = Array.length (the (AList.lookup (op =) REC_OPERATORS idt_index))
                   (* sanity check: the size of 'IDT' should be 'idt_size' *)
                   val _ = if idt_size <> size_of_type thy (typs, []) IDT then
                         raise REFUTE ("IDT_recursion_interpreter",
@@ -2978,8 +2968,8 @@
         (* nat -> nat -> interpretation *)
         fun append m n =
           let
-            val (len_m, off_m) = lookup lenoff_lists m
-            val (len_n, off_n) = lookup lenoff_lists n
+            val (len_m, off_m) = the (AList.lookup (op =) lenoff_lists m)
+            val (len_n, off_n) = the (AList.lookup (op =) lenoff_lists n)
             val len_elem = len_m + len_n
             val off_elem = off_m * power (size_elem, len_n) + off_n
           in
@@ -3267,7 +3257,7 @@
           val (typs, _)           = model
           val index               = #index info
           val descr               = #descr info
-          val (_, dtyps, constrs) = lookup descr index
+          val (_, dtyps, constrs) = the (AList.lookup (op =) descr index)
           val typ_assoc           = dtyps ~~ Ts
           (* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
           val _ = if Library.exists (fn d =>
--- a/src/HOL/Tools/res_atp.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/res_atp.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,7 +1,4 @@
-(*  Author: Jia Meng, Cambridge University Computer Laboratory, NICTA
-    ID: $Id$
-    Copyright 2004 University of Cambridge
-*)
+(*  Author: Jia Meng, Cambridge University Computer Laboratory, NICTA *)
 
 signature RES_ATP =
 sig
@@ -443,11 +440,11 @@
 fun delete_type cset = Symtab.delete_safe "HOL.type" cset;
 
 fun tvar_classes_of_terms ts =
-  let val sorts_list = map (map #2 o term_tvars) ts
+  let val sorts_list = map (map #2 o OldTerm.term_tvars) ts
   in  Symtab.keys (delete_type (foldl add_classes Symtab.empty sorts_list))  end;
 
 fun tfree_classes_of_terms ts =
-  let val sorts_list = map (map #2 o term_tfrees) ts
+  let val sorts_list = map (map #2 o OldTerm.term_tfrees) ts
   in  Symtab.keys (delete_type (foldl add_classes Symtab.empty sorts_list))  end;
 
 (*fold type constructors*)
@@ -507,11 +504,8 @@
 fun is_taut (Const ("Trueprop", _) $ Const ("True", _)) = true
   | is_taut _ = false;
 
-(*True if the term contains a variable whose (atomic) type is in the given list.*)
-fun has_typed_var tycons =
-  let fun var_tycon (Var (_, Type(a,_))) = a mem_string tycons
-        | var_tycon _ = false
-  in  exists var_tycon o term_vars  end;
+fun has_typed_var tycons = exists_subterm
+  (fn Var (_, Type (a, _)) => member (op =) tycons a | _ => false);
 
 (*Clauses are forbidden to contain variables of these types. The typical reason is that
   they lead to unsoundness. Note that "unit" satisfies numerous equations like ?X=().
--- a/src/HOL/Tools/res_axioms.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/res_axioms.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,6 +1,4 @@
 (*  Author: Jia Meng, Cambridge University Computer Laboratory
-    ID: $Id$
-    Copyright 2004 University of Cambridge
 
 Transformation of axiom rules (elim/intro/etc) into CNF forms.
 *)
@@ -75,7 +73,7 @@
           (*Existential: declare a Skolem function, then insert into body and continue*)
           let
             val cname = "sko_" ^ s ^ "_" ^ Int.toString (inc nref)
-            val args0 = term_frees xtp  (*get the formal parameter list*)
+            val args0 = OldTerm.term_frees xtp  (*get the formal parameter list*)
             val Ts = map type_of args0
             val extraTs = rhs_extra_types (Ts ---> T) xtp
             val argsx = map (fn T => Free (gensym "vsk", T)) extraTs
@@ -91,7 +89,7 @@
           in dec_sko (subst_bound (list_comb (c, args), p)) (ax :: axs, thy'') end
       | dec_sko (Const ("All", _) $ (xtp as Abs (a, T, p))) thx =
           (*Universal quant: insert a free variable into body and continue*)
-          let val fname = Name.variant (add_term_names (p, [])) a
+          let val fname = Name.variant (OldTerm.add_term_names (p, [])) a
           in dec_sko (subst_bound (Free (fname, T), p)) thx end
       | dec_sko (Const ("op &", _) $ p $ q) thx = dec_sko q (dec_sko p thx)
       | dec_sko (Const ("op |", _) $ p $ q) thx = dec_sko q (dec_sko p thx)
@@ -105,7 +103,7 @@
       fun dec_sko (Const ("Ex",_) $ (xtp as Abs(_,T,p))) defs =
             (*Existential: declare a Skolem function, then insert into body and continue*)
             let val skos = map (#1 o Logic.dest_equals) defs  (*existing sko fns*)
-                val args = term_frees xtp \\ skos  (*the formal parameters*)
+                val args = OldTerm.term_frees xtp \\ skos  (*the formal parameters*)
                 val Ts = map type_of args
                 val cT = Ts ---> T
                 val id = "sko_" ^ s ^ "_" ^ Int.toString (inc sko_count)
@@ -119,7 +117,7 @@
             end
         | dec_sko (Const ("All",_) $ (xtp as Abs(a,T,p))) defs =
             (*Universal quant: insert a free variable into body and continue*)
-            let val fname = Name.variant (add_term_names (p,[])) a
+            let val fname = Name.variant (OldTerm.add_term_names (p,[])) a
             in dec_sko (subst_bound (Free(fname,T), p)) defs end
         | dec_sko (Const ("op &", _) $ p $ q) defs = dec_sko q (dec_sko p defs)
         | dec_sko (Const ("op |", _) $ p $ q) defs = dec_sko q (dec_sko p defs)
@@ -158,9 +156,9 @@
 
 val monomorphic = not o Term.exists_type (Term.exists_subtype Term.is_TVar);
 
-val [f_B,g_B] = map (cterm_of (the_context ())) (term_vars (prop_of @{thm abs_B}));
-val [g_C,f_C] = map (cterm_of (the_context ())) (term_vars (prop_of @{thm abs_C}));
-val [f_S,g_S] = map (cterm_of (the_context ())) (term_vars (prop_of @{thm abs_S}));
+val [f_B,g_B] = map (cterm_of (the_context ())) (OldTerm.term_vars (prop_of @{thm abs_B}));
+val [g_C,f_C] = map (cterm_of (the_context ())) (OldTerm.term_vars (prop_of @{thm abs_C}));
+val [f_S,g_S] = map (cterm_of (the_context ())) (OldTerm.term_vars (prop_of @{thm abs_S}));
 
 (*FIXME: requires more use of cterm constructors*)
 fun abstract ct =
@@ -458,7 +456,7 @@
           |> fold (mark_seen o #1) new_facts
           |> fold_map skolem_def (sort_distinct (Thm.thm_ord o pairself snd) new_thms)
           |>> map_filter I;
-        val cache_entries = ParList.map skolem_cnfs defs;
+        val cache_entries = Par_List.map skolem_cnfs defs;
       in SOME (fold update_cache cache_entries thy') end
   end;
 
@@ -495,8 +493,8 @@
       val defs = filter (is_absko o Thm.term_of) newhyps
       val remaining_hyps = filter_out (member (op aconv) (map Thm.term_of defs))
                                       (map Thm.term_of hyps)
-      val fixed = term_frees (concl_of st) @
-                  foldl (gen_union (op aconv)) [] (map term_frees remaining_hyps)
+      val fixed = OldTerm.term_frees (concl_of st) @
+                  foldl (gen_union (op aconv)) [] (map OldTerm.term_frees remaining_hyps)
   in Seq.of_list [LocalDefs.expand (filter (is_okdef fixed o Thm.term_of) defs) st] end;
 
 
--- a/src/HOL/Tools/res_reconstruct.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/res_reconstruct.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,7 +1,4 @@
-(*  ID:         $Id$
-    Author:     L C Paulson and Claire Quigley
-    Copyright   2004  University of Cambridge
-*)
+(*  Author:     L C Paulson and Claire Quigley, Cambridge University Computer Laboratory *)
 
 (***************************************************************************)
 (*  Code to deal with the transfer of proofs from a prover process         *)
@@ -243,7 +240,7 @@
     singleton (Syntax.check_terms ctxt) (TypeInfer.constrain HOLogic.boolT (fix_sorts vt dt))
   end;
 
-fun gen_all_vars t = fold_rev Logic.all (term_vars t) t;
+fun gen_all_vars t = fold_rev Logic.all (OldTerm.term_vars t) t;
 
 fun ints_of_stree_aux (Int n, ns) = n::ns
   | ints_of_stree_aux (Br(_,ts), ns) = foldl ints_of_stree_aux ns ts;
@@ -253,7 +250,7 @@
 fun decode_tstp vt0 (name, role, ts, annots) ctxt =
   let val deps = case annots of NONE => [] | SOME (source,_) => ints_of_stree source
       val cl = clause_of_strees ctxt vt0 ts
-  in  ((name, role, cl, deps), fold Variable.declare_term (term_frees cl) ctxt)  end;
+  in  ((name, role, cl, deps), fold Variable.declare_term (OldTerm.term_frees cl) ctxt)  end;
 
 fun dest_tstp ((((name, role), ts), annots), chs) =
   case chs of
@@ -408,8 +405,8 @@
 fun add_wanted_prfline ctxt ((lno, t, []), (nlines, lines)) =
       (nlines, (lno, t, []) :: lines)   (*conjecture clauses must be kept*)
   | add_wanted_prfline ctxt ((lno, t, deps), (nlines, lines)) =
-      if eq_types t orelse not (null (term_tvars t)) orelse
-         exists bad_free (term_frees t) orelse
+      if eq_types t orelse not (null (Term.add_tvars t [])) orelse
+         exists_subterm bad_free t orelse
          (not (null lines) andalso   (*final line can't be deleted for these reasons*)
           (length deps < 2 orelse nlines mod (Config.get ctxt modulus) <> 0))   
       then (nlines+1, map (replace_deps (lno, deps)) lines) (*Delete line*)
--- a/src/HOL/Tools/sat_funcs.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/sat_funcs.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,9 +1,6 @@
 (*  Title:      HOL/Tools/sat_funcs.ML
-    ID:         $Id$
     Author:     Stephan Merz and Alwen Tiu, QSL Team, LORIA (http://qsl.loria.fr)
-    Author:     Tjark Weber
-    Copyright   2005-2006
-
+    Author:     Tjark Weber, TU Muenchen
 
 Proof reconstruction from SAT solvers.
 
@@ -323,7 +320,7 @@
 	(* Conjunction.intr_balanced and fold Thm.weaken below, is quadratic for *)
 	(* terms sorted in descending order, while only linear for terms      *)
 	(* sorted in ascending order                                          *)
-	val sorted_clauses = sort (Term.fast_term_ord o pairself Thm.term_of) nontrivial_clauses
+	val sorted_clauses = sort (TermOrd.fast_term_ord o pairself Thm.term_of) nontrivial_clauses
 	val _ = if !trace_sat then
 			tracing ("Sorted non-trivial clauses:\n" ^ cat_lines (map Display.string_of_cterm sorted_clauses))
 		else ()
--- a/src/HOL/Tools/specification_package.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/specification_package.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/specification_package.ML
-    ID:         $Id$
     Author:     Sebastian Skalberg, TU Muenchen
 
 Package for defining constants by specification.
@@ -118,7 +117,7 @@
 
         fun proc_single prop =
             let
-                val frees = Term.term_frees prop
+                val frees = OldTerm.term_frees prop
                 val _ = forall (fn v => Sign.of_sort thy (type_of v,HOLogic.typeS)) frees
                   orelse error "Specificaton: Only free variables of sort 'type' allowed"
                 val prop_closed = foldr (fn ((vname,T),prop) => HOLogic.mk_all (vname,T,prop)) prop (map dest_Free frees)
--- a/src/HOL/Tools/split_rule.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/split_rule.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/split_rule.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, David von Oheimb, and Markus Wenzel, TU Muenchen
 
 Some tools for managing tupled arguments and abstractions in rules.
@@ -105,7 +104,7 @@
 
 (*curries ALL function variables occurring in a rule's conclusion*)
 fun split_rule rl =
-  fold_rev split_rule_var' (Term.term_vars (concl_of rl)) rl
+  fold_rev split_rule_var' (OldTerm.term_vars (concl_of rl)) rl
   |> remove_internal_split
   |> Drule.standard;
 
--- a/src/HOL/Tools/string_syntax.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/string_syntax.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/string_syntax.ML
-    ID:         $Id$
     Author:     Makarius
 
 Concrete syntax for hex chars and strings.
@@ -43,8 +42,8 @@
 fun dest_char (Syntax.Appl [Syntax.Constant "Char", c1, c2]) = dest_chr c1 c2
   | dest_char _ = raise Match;
 
-fun syntax_xstr cs =
-  Syntax.Appl [Syntax.Constant "_xstr", Syntax.Variable (Syntax.implode_xstr cs)];
+fun syntax_string cs =
+  Syntax.Appl [Syntax.Constant "_inner_string", Syntax.Variable (Syntax.implode_xstr cs)];
 
 
 fun char_ast_tr [Syntax.Variable xstr] =
@@ -53,7 +52,7 @@
     | _ => error ("Single character expected: " ^ xstr))
   | char_ast_tr asts = raise AST ("char_ast_tr", asts);
 
-fun char_ast_tr' [c1, c2] = Syntax.Appl [Syntax.Constant "_Char", syntax_xstr [dest_chr c1 c2]]
+fun char_ast_tr' [c1, c2] = Syntax.Appl [Syntax.Constant "_Char", syntax_string [dest_chr c1 c2]]
   | char_ast_tr' _ = raise Match;
 
 
@@ -70,7 +69,7 @@
   | string_ast_tr asts = raise AST ("string_tr", asts);
 
 fun list_ast_tr' [args] = Syntax.Appl [Syntax.Constant "_String",
-        syntax_xstr (map dest_char (Syntax.unfold_ast "_args" args))]
+        syntax_string (map dest_char (Syntax.unfold_ast "_args" args))]
   | list_ast_tr' ts = raise Match;
 
 
--- a/src/HOL/Tools/typecopy_package.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/typecopy_package.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/typecopy_package.ML
-    ID:         $Id$
     Author:     Florian Haftmann, TU Muenchen
 
 Introducing copies of types using trivial typedefs; datatype-like abstraction.
@@ -75,7 +74,8 @@
 fun add_typecopy (raw_tyco, raw_vs) raw_ty constr_proj thy =
   let
     val ty = Sign.certify_typ thy raw_ty;
-    val vs = AList.make (the_default HOLogic.typeS o AList.lookup (op =) (typ_tfrees ty)) raw_vs;
+    val vs =
+      AList.make (the_default HOLogic.typeS o AList.lookup (op =) (Term.add_tfreesT ty [])) raw_vs;
     val tac = Tactic.rtac UNIV_witness 1;
     fun add_info tyco ( { abs_type = ty_abs, rep_type = ty_rep, Abs_name = c_abs,
       Rep_name = c_rep, Abs_inject = inject,
--- a/src/HOL/Tools/typedef_package.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Tools/typedef_package.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -176,7 +176,7 @@
     fun show_names pairs = commas_quote (map fst pairs);
 
     val illegal_vars =
-      if null (term_vars set) andalso null (term_tvars set) then []
+      if null (Term.add_vars set []) andalso null (Term.add_tvars set []) then []
       else ["Illegal schematic variable(s) on rhs"];
 
     val dup_lhs_tfrees =
@@ -188,8 +188,8 @@
       | extras => ["Extra type variables on rhs: " ^ show_names extras]);
 
     val illegal_frees =
-      (case term_frees set of [] => []
-      | xs => ["Illegal variables on rhs: " ^ show_names (map dest_Free xs)]);
+      (case Term.add_frees set [] of [] => []
+      | xs => ["Illegal variables on rhs: " ^ show_names xs]);
 
     val errs = illegal_vars @ dup_lhs_tfrees @ extra_rhs_tfrees @ illegal_frees;
     val _ = if null errs then () else error (cat_lines errs);
--- a/src/HOL/Univ_Poly.thy	Mon Jan 19 20:37:08 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1049 +0,0 @@
-(*  Title:       Univ_Poly.thy
-    Author:      Amine Chaieb
-*)
-
-header{*Univariate Polynomials*}
-
-theory Univ_Poly
-imports Plain List
-begin
-
-text{* Application of polynomial as a function. *}
-
-primrec (in semiring_0) poly :: "'a list => 'a  => 'a" where
-  poly_Nil:  "poly [] x = 0"
-| poly_Cons: "poly (h#t) x = h + x * poly t x"
-
-
-subsection{*Arithmetic Operations on Polynomials*}
-
-text{*addition*}
-
-primrec (in semiring_0) padd :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"  (infixl "+++" 65) 
-where
-  padd_Nil:  "[] +++ l2 = l2"
-| padd_Cons: "(h#t) +++ l2 = (if l2 = [] then h#t
-                            else (h + hd l2)#(t +++ tl l2))"
-
-text{*Multiplication by a constant*}
-primrec (in semiring_0) cmult :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list"  (infixl "%*" 70) where
-   cmult_Nil:  "c %* [] = []"
-|  cmult_Cons: "c %* (h#t) = (c * h)#(c %* t)"
-
-text{*Multiplication by a polynomial*}
-primrec (in semiring_0) pmult :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"  (infixl "***" 70)
-where
-   pmult_Nil:  "[] *** l2 = []"
-|  pmult_Cons: "(h#t) *** l2 = (if t = [] then h %* l2
-                              else (h %* l2) +++ ((0) # (t *** l2)))"
-
-text{*Repeated multiplication by a polynomial*}
-primrec (in semiring_0) mulexp :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a  list \<Rightarrow> 'a list" where
-   mulexp_zero:  "mulexp 0 p q = q"
-|  mulexp_Suc:   "mulexp (Suc n) p q = p *** mulexp n p q"
-
-text{*Exponential*}
-primrec (in semiring_1) pexp :: "'a list \<Rightarrow> nat \<Rightarrow> 'a list"  (infixl "%^" 80) where
-   pexp_0:   "p %^ 0 = [1]"
-|  pexp_Suc: "p %^ (Suc n) = p *** (p %^ n)"
-
-text{*Quotient related value of dividing a polynomial by x + a*}
-(* Useful for divisor properties in inductive proofs *)
-primrec (in field) "pquot" :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a list" where
-   pquot_Nil:  "pquot [] a= []"
-|  pquot_Cons: "pquot (h#t) a = (if t = [] then [h]
-                   else (inverse(a) * (h - hd( pquot t a)))#(pquot t a))"
-
-text{*normalization of polynomials (remove extra 0 coeff)*}
-primrec (in semiring_0) pnormalize :: "'a list \<Rightarrow> 'a list" where
-  pnormalize_Nil:  "pnormalize [] = []"
-| pnormalize_Cons: "pnormalize (h#p) = (if ( (pnormalize p) = [])
-                                     then (if (h = 0) then [] else [h])
-                                     else (h#(pnormalize p)))"
-
-definition (in semiring_0) "pnormal p = ((pnormalize p = p) \<and> p \<noteq> [])"
-definition (in semiring_0) "nonconstant p = (pnormal p \<and> (\<forall>x. p \<noteq> [x]))"
-text{*Other definitions*}
-
-definition (in ring_1)
-  poly_minus :: "'a list => 'a list" ("-- _" [80] 80) where
-  "-- p = (- 1) %* p"
-
-definition (in semiring_0)
-  divides :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"  (infixl "divides" 70) where
-  [code del]: "p1 divides p2 = (\<exists>q. poly p2 = poly(p1 *** q))"
-
-    --{*order of a polynomial*}
-definition (in ring_1) order :: "'a => 'a list => nat" where
-  "order a p = (SOME n. ([-a, 1] %^ n) divides p &
-                      ~ (([-a, 1] %^ (Suc n)) divides p))"
-
-     --{*degree of a polynomial*}
-definition (in semiring_0) degree :: "'a list => nat" where 
-  "degree p = length (pnormalize p) - 1"
-
-     --{*squarefree polynomials --- NB with respect to real roots only.*}
-definition (in ring_1)
-  rsquarefree :: "'a list => bool" where
-  "rsquarefree p = (poly p \<noteq> poly [] &
-                     (\<forall>a. (order a p = 0) | (order a p = 1)))"
-
-context semiring_0
-begin
-
-lemma padd_Nil2[simp]: "p +++ [] = p"
-by (induct p) auto
-
-lemma padd_Cons_Cons: "(h1 # p1) +++ (h2 # p2) = (h1 + h2) # (p1 +++ p2)"
-by auto
-
-lemma pminus_Nil[simp]: "-- [] = []"
-by (simp add: poly_minus_def)
-
-lemma pmult_singleton: "[h1] *** p1 = h1 %* p1" by simp
-end
-
-lemma (in semiring_1) poly_ident_mult[simp]: "1 %* t = t" by (induct "t", auto)
-
-lemma (in semiring_0) poly_simple_add_Cons[simp]: "[a] +++ ((0)#t) = (a#t)"
-by simp
-
-text{*Handy general properties*}
-
-lemma (in comm_semiring_0) padd_commut: "b +++ a = a +++ b"
-proof(induct b arbitrary: a)
-  case Nil thus ?case by auto
-next
-  case (Cons b bs a) thus ?case by (cases a, simp_all add: add_commute)
-qed
-
-lemma (in comm_semiring_0) padd_assoc: "\<forall>b c. (a +++ b) +++ c = a +++ (b +++ c)"
-apply (induct a arbitrary: b c)
-apply (simp, clarify)
-apply (case_tac b, simp_all add: add_ac)
-done
-
-lemma (in semiring_0) poly_cmult_distr: "a %* ( p +++ q) = (a %* p +++ a %* q)"
-apply (induct p arbitrary: q,simp)
-apply (case_tac q, simp_all add: right_distrib)
-done
-
-lemma (in ring_1) pmult_by_x[simp]: "[0, 1] *** t = ((0)#t)"
-apply (induct "t", simp)
-apply (auto simp add: mult_zero_left poly_ident_mult padd_commut)
-apply (case_tac t, auto)
-done
-
-text{*properties of evaluation of polynomials.*}
-
-lemma (in semiring_0) poly_add: "poly (p1 +++ p2) x = poly p1 x + poly p2 x"
-proof(induct p1 arbitrary: p2)
-  case Nil thus ?case by simp
-next
-  case (Cons a as p2) thus ?case 
-    by (cases p2, simp_all  add: add_ac right_distrib)
-qed
-
-lemma (in comm_semiring_0) poly_cmult: "poly (c %* p) x = c * poly p x"
-apply (induct "p") 
-apply (case_tac [2] "x=zero")
-apply (auto simp add: right_distrib mult_ac)
-done
-
-lemma (in comm_semiring_0) poly_cmult_map: "poly (map (op * c) p) x = c*poly p x"
-  by (induct p, auto simp add: right_distrib mult_ac)
-
-lemma (in comm_ring_1) poly_minus: "poly (-- p) x = - (poly p x)"
-apply (simp add: poly_minus_def)
-apply (auto simp add: poly_cmult minus_mult_left[symmetric])
-done
-
-lemma (in comm_semiring_0) poly_mult: "poly (p1 *** p2) x = poly p1 x * poly p2 x"
-proof(induct p1 arbitrary: p2)
-  case Nil thus ?case by simp
-next
-  case (Cons a as p2)
-  thus ?case by (cases as, 
-    simp_all add: poly_cmult poly_add left_distrib right_distrib mult_ac)
-qed
-
-class recpower_semiring = semiring + recpower
-class recpower_semiring_1 = semiring_1 + recpower
-class recpower_semiring_0 = semiring_0 + recpower
-class recpower_ring = ring + recpower
-class recpower_ring_1 = ring_1 + recpower
-subclass (in recpower_ring_1) recpower_ring ..
-class recpower_comm_semiring_1 = recpower + comm_semiring_1
-class recpower_comm_ring_1 = recpower + comm_ring_1
-subclass (in recpower_comm_ring_1) recpower_comm_semiring_1 ..
-class recpower_idom = recpower + idom
-subclass (in recpower_idom) recpower_comm_ring_1 ..
-class idom_char_0 = idom + ring_char_0
-class recpower_idom_char_0 = recpower + idom_char_0
-subclass (in recpower_idom_char_0) recpower_idom ..
-
-lemma (in recpower_comm_ring_1) poly_exp: "poly (p %^ n) x = (poly p x) ^ n"
-apply (induct "n")
-apply (auto simp add: poly_cmult poly_mult power_Suc)
-done
-
-text{*More Polynomial Evaluation Lemmas*}
-
-lemma  (in semiring_0) poly_add_rzero[simp]: "poly (a +++ []) x = poly a x"
-by simp
-
-lemma (in comm_semiring_0) poly_mult_assoc: "poly ((a *** b) *** c) x = poly (a *** (b *** c)) x"
-  by (simp add: poly_mult mult_assoc)
-
-lemma (in semiring_0) poly_mult_Nil2[simp]: "poly (p *** []) x = 0"
-by (induct "p", auto)
-
-lemma (in comm_semiring_1) poly_exp_add: "poly (p %^ (n + d)) x = poly( p %^ n *** p %^ d) x"
-apply (induct "n")
-apply (auto simp add: poly_mult mult_assoc)
-done
-
-subsection{*Key Property: if @{term "f(a) = 0"} then @{term "(x - a)"} divides
- @{term "p(x)"} *}
-
-lemma (in comm_ring_1) lemma_poly_linear_rem: "\<forall>h. \<exists>q r. h#t = [r] +++ [-a, 1] *** q"
-proof(induct t)
-  case Nil
-  {fix h have "[h] = [h] +++ [- a, 1] *** []" by simp}
-  thus ?case by blast
-next
-  case (Cons  x xs)
-  {fix h 
-    from Cons.hyps[rule_format, of x] 
-    obtain q r where qr: "x#xs = [r] +++ [- a, 1] *** q" by blast
-    have "h#x#xs = [a*r + h] +++ [-a, 1] *** (r#q)" 
-      using qr by(cases q, simp_all add: ring_simps diff_def[symmetric] 
-	minus_mult_left[symmetric] right_minus)
-    hence "\<exists>q r. h#x#xs = [r] +++ [-a, 1] *** q" by blast}
-  thus ?case by blast
-qed
-
-lemma (in comm_ring_1) poly_linear_rem: "\<exists>q r. h#t = [r] +++ [-a, 1] *** q"
-by (cut_tac t = t and a = a in lemma_poly_linear_rem, auto)
-
-
-lemma (in comm_ring_1) poly_linear_divides: "(poly p a = 0) = ((p = []) | (\<exists>q. p = [-a, 1] *** q))"
-proof-
-  {assume p: "p = []" hence ?thesis by simp}
-  moreover
-  {fix x xs assume p: "p = x#xs"
-    {fix q assume "p = [-a, 1] *** q" hence "poly p a = 0" 
-	by (simp add: poly_add poly_cmult minus_mult_left[symmetric])}
-    moreover
-    {assume p0: "poly p a = 0"
-      from poly_linear_rem[of x xs a] obtain q r 
-      where qr: "x#xs = [r] +++ [- a, 1] *** q" by blast
-      have "r = 0" using p0 by (simp only: p qr poly_mult poly_add) simp
-      hence "\<exists>q. p = [- a, 1] *** q" using p qr  apply - apply (rule exI[where x=q])apply auto apply (cases q) apply auto done}
-    ultimately have ?thesis using p by blast}
-  ultimately show ?thesis by (cases p, auto)
-qed
-
-lemma (in semiring_0) lemma_poly_length_mult[simp]: "\<forall>h k a. length (k %* p +++  (h # (a %* p))) = Suc (length p)"
-by (induct "p", auto)
-
-lemma (in semiring_0) lemma_poly_length_mult2[simp]: "\<forall>h k. length (k %* p +++  (h # p)) = Suc (length p)"
-by (induct "p", auto)
-
-lemma (in ring_1) poly_length_mult[simp]: "length([-a,1] *** q) = Suc (length q)"
-by auto
-
-subsection{*Polynomial length*}
-
-lemma (in semiring_0) poly_cmult_length[simp]: "length (a %* p) = length p"
-by (induct "p", auto)
-
-lemma (in semiring_0) poly_add_length: "length (p1 +++ p2) = max (length p1) (length p2)"
-apply (induct p1 arbitrary: p2, simp_all)
-apply arith
-done
-
-lemma (in semiring_0) poly_root_mult_length[simp]: "length([a,b] *** p) = Suc (length p)"
-by (simp add: poly_add_length)
-
-lemma (in idom) poly_mult_not_eq_poly_Nil[simp]: 
- "poly (p *** q) x \<noteq> poly [] x \<longleftrightarrow> poly p x \<noteq> poly [] x \<and> poly q x \<noteq> poly [] x"
-by (auto simp add: poly_mult)
-
-lemma (in idom) poly_mult_eq_zero_disj: "poly (p *** q) x = 0 \<longleftrightarrow> poly p x = 0 \<or> poly q x = 0"
-by (auto simp add: poly_mult)
-
-text{*Normalisation Properties*}
-
-lemma (in semiring_0) poly_normalized_nil: "(pnormalize p = []) --> (poly p x = 0)"
-by (induct "p", auto)
-
-text{*A nontrivial polynomial of degree n has no more than n roots*}
-lemma (in idom) poly_roots_index_lemma:
-   assumes p: "poly p x \<noteq> poly [] x" and n: "length p = n" 
-  shows "\<exists>i. \<forall>x. poly p x = 0 \<longrightarrow> (\<exists>m\<le>n. x = i m)"
-  using p n
-proof(induct n arbitrary: p x)
-  case 0 thus ?case by simp 
-next
-  case (Suc n p x)
-  {assume C: "\<And>i. \<exists>x. poly p x = 0 \<and> (\<forall>m\<le>Suc n. x \<noteq> i m)"
-    from Suc.prems have p0: "poly p x \<noteq> 0" "p\<noteq> []" by auto
-    from p0(1)[unfolded poly_linear_divides[of p x]] 
-    have "\<forall>q. p \<noteq> [- x, 1] *** q" by blast
-    from C obtain a where a: "poly p a = 0" by blast
-    from a[unfolded poly_linear_divides[of p a]] p0(2) 
-    obtain q where q: "p = [-a, 1] *** q" by blast
-    have lg: "length q = n" using q Suc.prems(2) by simp
-    from q p0 have qx: "poly q x \<noteq> poly [] x" 
-      by (auto simp add: poly_mult poly_add poly_cmult)
-    from Suc.hyps[OF qx lg] obtain i where 
-      i: "\<forall>x. poly q x = 0 \<longrightarrow> (\<exists>m\<le>n. x = i m)" by blast
-    let ?i = "\<lambda>m. if m = Suc n then a else i m"
-    from C[of ?i] obtain y where y: "poly p y = 0" "\<forall>m\<le> Suc n. y \<noteq> ?i m" 
-      by blast
-    from y have "y = a \<or> poly q y = 0" 
-      by (simp only: q poly_mult_eq_zero_disj poly_add) (simp add: ring_simps)
-    with i[rule_format, of y] y(1) y(2) have False apply auto 
-      apply (erule_tac x="m" in allE)
-      apply auto
-      done}
-  thus ?case by blast
-qed
-
-
-lemma (in idom) poly_roots_index_length: "poly p x \<noteq> poly [] x ==>
-      \<exists>i. \<forall>x. (poly p x = 0) --> (\<exists>n. n \<le> length p & x = i n)"
-by (blast intro: poly_roots_index_lemma)
-
-lemma (in idom) poly_roots_finite_lemma1: "poly p x \<noteq> poly [] x ==>
-      \<exists>N i. \<forall>x. (poly p x = 0) --> (\<exists>n. (n::nat) < N & x = i n)"
-apply (drule poly_roots_index_length, safe)
-apply (rule_tac x = "Suc (length p)" in exI)
-apply (rule_tac x = i in exI) 
-apply (simp add: less_Suc_eq_le)
-done
-
-
-lemma (in idom) idom_finite_lemma:
-  assumes P: "\<forall>x. P x --> (\<exists>n. n < length j & x = j!n)"
-  shows "finite {x. P x}"
-proof-
-  let ?M = "{x. P x}"
-  let ?N = "set j"
-  have "?M \<subseteq> ?N" using P by auto
-  thus ?thesis using finite_subset by auto
-qed
-
-
-lemma (in idom) poly_roots_finite_lemma2: "poly p x \<noteq> poly [] x ==>
-      \<exists>i. \<forall>x. (poly p x = 0) --> x \<in> set i"
-apply (drule poly_roots_index_length, safe)
-apply (rule_tac x="map (\<lambda>n. i n) [0 ..< Suc (length p)]" in exI)
-apply (auto simp add: image_iff)
-apply (erule_tac x="x" in allE, clarsimp)
-by (case_tac "n=length p", auto simp add: order_le_less)
-
-lemma UNIV_nat_infinite: "\<not> finite (UNIV :: nat set)"
-  unfolding finite_conv_nat_seg_image
-proof(auto simp add: expand_set_eq image_iff)
-  fix n::nat and f:: "nat \<Rightarrow> nat"
-  let ?N = "{i. i < n}"
-  let ?fN = "f ` ?N"
-  let ?y = "Max ?fN + 1"
-  from nat_seg_image_imp_finite[of "?fN" "f" n] 
-  have thfN: "finite ?fN" by simp
-  {assume "n =0" hence "\<exists>x. \<forall>xa<n. x \<noteq> f xa" by auto}
-  moreover
-  {assume nz: "n \<noteq> 0"
-    hence thne: "?fN \<noteq> {}" by (auto simp add: neq0_conv)
-    have "\<forall>x\<in> ?fN. Max ?fN \<ge> x" using nz Max_ge_iff[OF thfN thne] by auto
-    hence "\<forall>x\<in> ?fN. ?y > x" by auto
-    hence "?y \<notin> ?fN" by auto
-    hence "\<exists>x. \<forall>xa<n. x \<noteq> f xa" by auto }
-  ultimately show "\<exists>x. \<forall>xa<n. x \<noteq> f xa" by blast
-qed
-
-lemma (in ring_char_0) UNIV_ring_char_0_infinte: 
-  "\<not> (finite (UNIV:: 'a set))" 
-proof
-  assume F: "finite (UNIV :: 'a set)"
-  have th0: "of_nat ` UNIV \<subseteq> UNIV" by simp
-  from finite_subset[OF th0] have th: "finite (of_nat ` UNIV :: 'a set)" .
-  have th': "inj_on (of_nat::nat \<Rightarrow> 'a) (UNIV)"
-    unfolding inj_on_def by auto
-  from finite_imageD[OF th th'] UNIV_nat_infinite 
-  show False by blast
-qed
-
-lemma (in idom_char_0) poly_roots_finite: "(poly p \<noteq> poly []) = 
-  finite {x. poly p x = 0}"
-proof
-  assume H: "poly p \<noteq> poly []"
-  show "finite {x. poly p x = (0::'a)}"
-    using H
-    apply -
-    apply (erule contrapos_np, rule ext)
-    apply (rule ccontr)
-    apply (clarify dest!: poly_roots_finite_lemma2)
-    using finite_subset
-  proof-
-    fix x i
-    assume F: "\<not> finite {x. poly p x = (0\<Colon>'a)}" 
-      and P: "\<forall>x. poly p x = (0\<Colon>'a) \<longrightarrow> x \<in> set i"
-    let ?M= "{x. poly p x = (0\<Colon>'a)}"
-    from P have "?M \<subseteq> set i" by auto
-    with finite_subset F show False by auto
-  qed
-next
-  assume F: "finite {x. poly p x = (0\<Colon>'a)}"
-  show "poly p \<noteq> poly []" using F UNIV_ring_char_0_infinte by auto  
-qed
-
-text{*Entirety and Cancellation for polynomials*}
-
-lemma (in idom_char_0) poly_entire_lemma2: 
-  assumes p0: "poly p \<noteq> poly []" and q0: "poly q \<noteq> poly []"
-  shows "poly (p***q) \<noteq> poly []"
-proof-
-  let ?S = "\<lambda>p. {x. poly p x = 0}"
-  have "?S (p *** q) = ?S p \<union> ?S q" by (auto simp add: poly_mult)
-  with p0 q0 show ?thesis  unfolding poly_roots_finite by auto
-qed
-
-lemma (in idom_char_0) poly_entire: 
-  "poly (p *** q) = poly [] \<longleftrightarrow> poly p = poly [] \<or> poly q = poly []"
-using poly_entire_lemma2[of p q] 
-by auto (rule ext, simp add: poly_mult)+
-
-lemma (in idom_char_0) poly_entire_neg: "(poly (p *** q) \<noteq> poly []) = ((poly p \<noteq> poly []) & (poly q \<noteq> poly []))"
-by (simp add: poly_entire)
-
-lemma fun_eq: " (f = g) = (\<forall>x. f x = g x)"
-by (auto intro!: ext)
-
-lemma (in comm_ring_1) poly_add_minus_zero_iff: "(poly (p +++ -- q) = poly []) = (poly p = poly q)"
-by (auto simp add: ring_simps poly_add poly_minus_def fun_eq poly_cmult minus_mult_left[symmetric])
-
-lemma (in comm_ring_1) poly_add_minus_mult_eq: "poly (p *** q +++ --(p *** r)) = poly (p *** (q +++ -- r))"
-by (auto simp add: poly_add poly_minus_def fun_eq poly_mult poly_cmult right_distrib minus_mult_left[symmetric] minus_mult_right[symmetric])
-
-subclass (in idom_char_0) comm_ring_1 ..
-lemma (in idom_char_0) poly_mult_left_cancel: "(poly (p *** q) = poly (p *** r)) = (poly p = poly [] | poly q = poly r)"
-proof-
-  have "poly (p *** q) = poly (p *** r) \<longleftrightarrow> poly (p *** q +++ -- (p *** r)) = poly []" by (simp only: poly_add_minus_zero_iff)
-  also have "\<dots> \<longleftrightarrow> poly p = poly [] | poly q = poly r"
-    by (auto intro: ext simp add: poly_add_minus_mult_eq poly_entire poly_add_minus_zero_iff)
-  finally show ?thesis .
-qed
-
-lemma (in recpower_idom) poly_exp_eq_zero[simp]:
-     "(poly (p %^ n) = poly []) = (poly p = poly [] & n \<noteq> 0)"
-apply (simp only: fun_eq add: all_simps [symmetric]) 
-apply (rule arg_cong [where f = All]) 
-apply (rule ext)
-apply (induct n)
-apply (auto simp add: poly_exp poly_mult)
-done
-
-lemma (in semiring_1) one_neq_zero[simp]: "1 \<noteq> 0" using zero_neq_one by blast
-lemma (in comm_ring_1) poly_prime_eq_zero[simp]: "poly [a,1] \<noteq> poly []"
-apply (simp add: fun_eq)
-apply (rule_tac x = "minus one a" in exI)
-apply (unfold diff_minus)
-apply (subst add_commute)
-apply (subst add_assoc)
-apply simp
-done 
-
-lemma (in recpower_idom) poly_exp_prime_eq_zero: "(poly ([a, 1] %^ n) \<noteq> poly [])"
-by auto
-
-text{*A more constructive notion of polynomials being trivial*}
-
-lemma (in idom_char_0) poly_zero_lemma': "poly (h # t) = poly [] ==> h = 0 & poly t = poly []"
-apply(simp add: fun_eq)
-apply (case_tac "h = zero")
-apply (drule_tac [2] x = zero in spec, auto) 
-apply (cases "poly t = poly []", simp) 
-proof-
-  fix x
-  assume H: "\<forall>x. x = (0\<Colon>'a) \<or> poly t x = (0\<Colon>'a)"  and pnz: "poly t \<noteq> poly []"
-  let ?S = "{x. poly t x = 0}"
-  from H have "\<forall>x. x \<noteq>0 \<longrightarrow> poly t x = 0" by blast
-  hence th: "?S \<supseteq> UNIV - {0}" by auto
-  from poly_roots_finite pnz have th': "finite ?S" by blast
-  from finite_subset[OF th th'] UNIV_ring_char_0_infinte
-  show "poly t x = (0\<Colon>'a)" by simp
-  qed
-
-lemma (in idom_char_0) poly_zero: "(poly p = poly []) = list_all (%c. c = 0) p"
-apply (induct "p", simp)
-apply (rule iffI)
-apply (drule poly_zero_lemma', auto)
-done
-
-lemma (in idom_char_0) poly_0: "list_all (\<lambda>c. c = 0) p \<Longrightarrow> poly p x = 0"
-  unfolding poly_zero[symmetric] by simp
-
-
-
-text{*Basics of divisibility.*}
-
-lemma (in idom) poly_primes: "([a, 1] divides (p *** q)) = ([a, 1] divides p | [a, 1] divides q)"
-apply (auto simp add: divides_def fun_eq poly_mult poly_add poly_cmult left_distrib [symmetric])
-apply (drule_tac x = "uminus a" in spec)
-apply (simp add: poly_linear_divides poly_add poly_cmult left_distrib [symmetric])
-apply (cases "p = []")
-apply (rule exI[where x="[]"])
-apply simp
-apply (cases "q = []")
-apply (erule allE[where x="[]"], simp)
-
-apply clarsimp
-apply (cases "\<exists>q\<Colon>'a list. p = a %* q +++ ((0\<Colon>'a) # q)")
-apply (clarsimp simp add: poly_add poly_cmult)
-apply (rule_tac x="qa" in exI)
-apply (simp add: left_distrib [symmetric])
-apply clarsimp
-
-apply (auto simp add: right_minus poly_linear_divides poly_add poly_cmult left_distrib [symmetric])
-apply (rule_tac x = "pmult qa q" in exI)
-apply (rule_tac [2] x = "pmult p qa" in exI)
-apply (auto simp add: poly_add poly_mult poly_cmult mult_ac)
-done
-
-lemma (in comm_semiring_1) poly_divides_refl[simp]: "p divides p"
-apply (simp add: divides_def)
-apply (rule_tac x = "[one]" in exI)
-apply (auto simp add: poly_mult fun_eq)
-done
-
-lemma (in comm_semiring_1) poly_divides_trans: "[| p divides q; q divides r |] ==> p divides r"
-apply (simp add: divides_def, safe)
-apply (rule_tac x = "pmult qa qaa" in exI)
-apply (auto simp add: poly_mult fun_eq mult_assoc)
-done
-
-
-lemma (in recpower_comm_semiring_1) poly_divides_exp: "m \<le> n ==> (p %^ m) divides (p %^ n)"
-apply (auto simp add: le_iff_add)
-apply (induct_tac k)
-apply (rule_tac [2] poly_divides_trans)
-apply (auto simp add: divides_def)
-apply (rule_tac x = p in exI)
-apply (auto simp add: poly_mult fun_eq mult_ac)
-done
-
-lemma (in recpower_comm_semiring_1) poly_exp_divides: "[| (p %^ n) divides q;  m\<le>n |] ==> (p %^ m) divides q"
-by (blast intro: poly_divides_exp poly_divides_trans)
-
-lemma (in comm_semiring_0) poly_divides_add:
-   "[| p divides q; p divides r |] ==> p divides (q +++ r)"
-apply (simp add: divides_def, auto)
-apply (rule_tac x = "padd qa qaa" in exI)
-apply (auto simp add: poly_add fun_eq poly_mult right_distrib)
-done
-
-lemma (in comm_ring_1) poly_divides_diff:
-   "[| p divides q; p divides (q +++ r) |] ==> p divides r"
-apply (simp add: divides_def, auto)
-apply (rule_tac x = "padd qaa (poly_minus qa)" in exI)
-apply (auto simp add: poly_add fun_eq poly_mult poly_minus right_diff_distrib compare_rls add_ac)
-done
-
-lemma (in comm_ring_1) poly_divides_diff2: "[| p divides r; p divides (q +++ r) |] ==> p divides q"
-apply (erule poly_divides_diff)
-apply (auto simp add: poly_add fun_eq poly_mult divides_def add_ac)
-done
-
-lemma (in semiring_0) poly_divides_zero: "poly p = poly [] ==> q divides p"
-apply (simp add: divides_def)
-apply (rule exI[where x="[]"])
-apply (auto simp add: fun_eq poly_mult)
-done
-
-lemma (in semiring_0) poly_divides_zero2[simp]: "q divides []"
-apply (simp add: divides_def)
-apply (rule_tac x = "[]" in exI)
-apply (auto simp add: fun_eq)
-done
-
-text{*At last, we can consider the order of a root.*}
-
-lemma (in idom_char_0)  poly_order_exists_lemma:
-  assumes lp: "length p = d" and p: "poly p \<noteq> poly []"
-  shows "\<exists>n q. p = mulexp n [-a, 1] q \<and> poly q a \<noteq> 0"
-using lp p
-proof(induct d arbitrary: p)
-  case 0 thus ?case by simp
-next
-  case (Suc n p)
-  {assume p0: "poly p a = 0"
-    from Suc.prems have h: "length p = Suc n" "poly p \<noteq> poly []" by blast
-    hence pN: "p \<noteq> []" by - (rule ccontr, simp)
-    from p0[unfolded poly_linear_divides] pN  obtain q where 
-      q: "p = [-a, 1] *** q" by blast
-    from q h p0 have qh: "length q = n" "poly q \<noteq> poly []" 
-      apply -
-      apply simp
-      apply (simp only: fun_eq)
-      apply (rule ccontr)
-      apply (simp add: fun_eq poly_add poly_cmult minus_mult_left[symmetric])
-      done
-    from Suc.hyps[OF qh] obtain m r where 
-      mr: "q = mulexp m [-a,1] r" "poly r a \<noteq> 0" by blast    
-    from mr q have "p = mulexp (Suc m) [-a,1] r \<and> poly r a \<noteq> 0" by simp
-    hence ?case by blast}
-  moreover
-  {assume p0: "poly p a \<noteq> 0"
-    hence ?case using Suc.prems apply simp by (rule exI[where x="0::nat"], simp)}
-  ultimately show ?case by blast
-qed
-
-
-lemma (in recpower_comm_semiring_1) poly_mulexp: "poly (mulexp n p q) x = (poly p x) ^ n * poly q x"
-by(induct n, auto simp add: poly_mult power_Suc mult_ac)
-
-lemma (in comm_semiring_1) divides_left_mult:
-  assumes d:"(p***q) divides r" shows "p divides r \<and> q divides r"
-proof-
-  from d obtain t where r:"poly r = poly (p***q *** t)"
-    unfolding divides_def by blast
-  hence "poly r = poly (p *** (q *** t))"
-    "poly r = poly (q *** (p***t))" by(auto simp add: fun_eq poly_mult mult_ac)
-  thus ?thesis unfolding divides_def by blast
-qed
-
-
-
-(* FIXME: Tidy up *)
-
-lemma (in recpower_semiring_1) 
-  zero_power_iff: "0 ^ n = (if n = 0 then 1 else 0)"
-  by (induct n, simp_all add: power_Suc)
-
-lemma (in recpower_idom_char_0) poly_order_exists:
-  assumes lp: "length p = d" and p0: "poly p \<noteq> poly []"
-  shows "\<exists>n. ([-a, 1] %^ n) divides p & ~(([-a, 1] %^ (Suc n)) divides p)"
-proof-
-let ?poly = poly
-let ?mulexp = mulexp
-let ?pexp = pexp
-from lp p0
-show ?thesis
-apply -
-apply (drule poly_order_exists_lemma [where a=a], assumption, clarify)  
-apply (rule_tac x = n in exI, safe)
-apply (unfold divides_def)
-apply (rule_tac x = q in exI)
-apply (induct_tac "n", simp)
-apply (simp (no_asm_simp) add: poly_add poly_cmult poly_mult right_distrib mult_ac)
-apply safe
-apply (subgoal_tac "?poly (?mulexp n [uminus a, one] q) \<noteq> ?poly (pmult (?pexp [uminus a, one] (Suc n)) qa)") 
-apply simp 
-apply (induct_tac "n")
-apply (simp del: pmult_Cons pexp_Suc)
-apply (erule_tac Q = "?poly q a = zero" in contrapos_np)
-apply (simp add: poly_add poly_cmult minus_mult_left[symmetric])
-apply (rule pexp_Suc [THEN ssubst])
-apply (rule ccontr)
-apply (simp add: poly_mult_left_cancel poly_mult_assoc del: pmult_Cons pexp_Suc)
-done
-qed
-
-
-lemma (in semiring_1) poly_one_divides[simp]: "[1] divides p"
-by (simp add: divides_def, auto)
-
-lemma (in recpower_idom_char_0) poly_order: "poly p \<noteq> poly []
-      ==> EX! n. ([-a, 1] %^ n) divides p &
-                 ~(([-a, 1] %^ (Suc n)) divides p)"
-apply (auto intro: poly_order_exists simp add: less_linear simp del: pmult_Cons pexp_Suc)
-apply (cut_tac x = y and y = n in less_linear)
-apply (drule_tac m = n in poly_exp_divides)
-apply (auto dest: Suc_le_eq [THEN iffD2, THEN [2] poly_exp_divides]
-            simp del: pmult_Cons pexp_Suc)
-done
-
-text{*Order*}
-
-lemma some1_equalityD: "[| n = (@n. P n); EX! n. P n |] ==> P n"
-by (blast intro: someI2)
-
-lemma (in recpower_idom_char_0) order:
-      "(([-a, 1] %^ n) divides p &
-        ~(([-a, 1] %^ (Suc n)) divides p)) =
-        ((n = order a p) & ~(poly p = poly []))"
-apply (unfold order_def)
-apply (rule iffI)
-apply (blast dest: poly_divides_zero intro!: some1_equality [symmetric] poly_order)
-apply (blast intro!: poly_order [THEN [2] some1_equalityD])
-done
-
-lemma (in recpower_idom_char_0) order2: "[| poly p \<noteq> poly [] |]
-      ==> ([-a, 1] %^ (order a p)) divides p &
-              ~(([-a, 1] %^ (Suc(order a p))) divides p)"
-by (simp add: order del: pexp_Suc)
-
-lemma (in recpower_idom_char_0) order_unique: "[| poly p \<noteq> poly []; ([-a, 1] %^ n) divides p;
-         ~(([-a, 1] %^ (Suc n)) divides p)
-      |] ==> (n = order a p)"
-by (insert order [of a n p], auto) 
-
-lemma (in recpower_idom_char_0) order_unique_lemma: "(poly p \<noteq> poly [] & ([-a, 1] %^ n) divides p &
-         ~(([-a, 1] %^ (Suc n)) divides p))
-      ==> (n = order a p)"
-by (blast intro: order_unique)
-
-lemma (in ring_1) order_poly: "poly p = poly q ==> order a p = order a q"
-by (auto simp add: fun_eq divides_def poly_mult order_def)
-
-lemma (in semiring_1) pexp_one[simp]: "p %^ (Suc 0) = p"
-apply (induct "p")
-apply (auto simp add: numeral_1_eq_1)
-done
-
-lemma (in comm_ring_1) lemma_order_root:
-     " 0 < n & [- a, 1] %^ n divides p & ~ [- a, 1] %^ (Suc n) divides p
-             \<Longrightarrow> poly p a = 0"
-apply (induct n arbitrary: a p, blast)
-apply (auto simp add: divides_def poly_mult simp del: pmult_Cons)
-done
-
-lemma (in recpower_idom_char_0) order_root: "(poly p a = 0) = ((poly p = poly []) | order a p \<noteq> 0)"
-proof-
-  let ?poly = poly
-  show ?thesis 
-apply (case_tac "?poly p = ?poly []", auto)
-apply (simp add: poly_linear_divides del: pmult_Cons, safe)
-apply (drule_tac [!] a = a in order2)
-apply (rule ccontr)
-apply (simp add: divides_def poly_mult fun_eq del: pmult_Cons, blast)
-using neq0_conv
-apply (blast intro: lemma_order_root)
-done
-qed
-
-lemma (in recpower_idom_char_0) order_divides: "(([-a, 1] %^ n) divides p) = ((poly p = poly []) | n \<le> order a p)"
-proof-
-  let ?poly = poly
-  show ?thesis 
-apply (case_tac "?poly p = ?poly []", auto)
-apply (simp add: divides_def fun_eq poly_mult)
-apply (rule_tac x = "[]" in exI)
-apply (auto dest!: order2 [where a=a]
-	    intro: poly_exp_divides simp del: pexp_Suc)
-done
-qed
-
-lemma (in recpower_idom_char_0) order_decomp:
-     "poly p \<noteq> poly []
-      ==> \<exists>q. (poly p = poly (([-a, 1] %^ (order a p)) *** q)) &
-                ~([-a, 1] divides q)"
-apply (unfold divides_def)
-apply (drule order2 [where a = a])
-apply (simp add: divides_def del: pexp_Suc pmult_Cons, safe)
-apply (rule_tac x = q in exI, safe)
-apply (drule_tac x = qa in spec)
-apply (auto simp add: poly_mult fun_eq poly_exp mult_ac simp del: pmult_Cons)
-done
-
-text{*Important composition properties of orders.*}
-lemma order_mult: "poly (p *** q) \<noteq> poly []
-      ==> order a (p *** q) = order a p + order (a::'a::{recpower_idom_char_0}) q"
-apply (cut_tac a = a and p = "p *** q" and n = "order a p + order a q" in order)
-apply (auto simp add: poly_entire simp del: pmult_Cons)
-apply (drule_tac a = a in order2)+
-apply safe
-apply (simp add: divides_def fun_eq poly_exp_add poly_mult del: pmult_Cons, safe)
-apply (rule_tac x = "qa *** qaa" in exI)
-apply (simp add: poly_mult mult_ac del: pmult_Cons)
-apply (drule_tac a = a in order_decomp)+
-apply safe
-apply (subgoal_tac "[-a,1] divides (qa *** qaa) ")
-apply (simp add: poly_primes del: pmult_Cons)
-apply (auto simp add: divides_def simp del: pmult_Cons)
-apply (rule_tac x = qb in exI)
-apply (subgoal_tac "poly ([-a, 1] %^ (order a p) *** (qa *** qaa)) = poly ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))")
-apply (drule poly_mult_left_cancel [THEN iffD1], force)
-apply (subgoal_tac "poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** (qa *** qaa))) = poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))) ")
-apply (drule poly_mult_left_cancel [THEN iffD1], force)
-apply (simp add: fun_eq poly_exp_add poly_mult mult_ac del: pmult_Cons)
-done
-
-lemma (in recpower_idom_char_0) order_mult: 
-  assumes pq0: "poly (p *** q) \<noteq> poly []"
-  shows "order a (p *** q) = order a p + order a q"
-proof-
-  let ?order = order
-  let ?divides = "op divides"
-  let ?poly = poly
-from pq0 
-show ?thesis
-apply (cut_tac a = a and p = "pmult p q" and n = "?order a p + ?order a q" in order)
-apply (auto simp add: poly_entire simp del: pmult_Cons)
-apply (drule_tac a = a in order2)+
-apply safe
-apply (simp add: divides_def fun_eq poly_exp_add poly_mult del: pmult_Cons, safe)
-apply (rule_tac x = "pmult qa qaa" in exI)
-apply (simp add: poly_mult mult_ac del: pmult_Cons)
-apply (drule_tac a = a in order_decomp)+
-apply safe
-apply (subgoal_tac "?divides [uminus a,one ] (pmult qa qaa) ")
-apply (simp add: poly_primes del: pmult_Cons)
-apply (auto simp add: divides_def simp del: pmult_Cons)
-apply (rule_tac x = qb in exI)
-apply (subgoal_tac "?poly (pmult (pexp [uminus a, one] (?order a p)) (pmult qa qaa)) = ?poly (pmult (pexp [uminus a, one] (?order a p)) (pmult [uminus a, one] qb))")
-apply (drule poly_mult_left_cancel [THEN iffD1], force)
-apply (subgoal_tac "?poly (pmult (pexp [uminus a, one ] (order a q)) (pmult (pexp [uminus a, one] (order a p)) (pmult qa qaa))) = ?poly (pmult (pexp [uminus a, one] (order a q)) (pmult (pexp [uminus a, one] (order a p)) (pmult [uminus a, one] qb))) ")
-apply (drule poly_mult_left_cancel [THEN iffD1], force)
-apply (simp add: fun_eq poly_exp_add poly_mult mult_ac del: pmult_Cons)
-done
-qed
-
-lemma (in recpower_idom_char_0) order_root2: "poly p \<noteq> poly [] ==> (poly p a = 0) = (order a p \<noteq> 0)"
-by (rule order_root [THEN ssubst], auto)
-
-lemma (in semiring_1) pmult_one[simp]: "[1] *** p = p" by auto
-
-lemma (in semiring_0) poly_Nil_zero: "poly [] = poly [0]"
-by (simp add: fun_eq)
-
-lemma (in recpower_idom_char_0) rsquarefree_decomp:
-     "[| rsquarefree p; poly p a = 0 |]
-      ==> \<exists>q. (poly p = poly ([-a, 1] *** q)) & poly q a \<noteq> 0"
-apply (simp add: rsquarefree_def, safe)
-apply (frule_tac a = a in order_decomp)
-apply (drule_tac x = a in spec)
-apply (drule_tac a = a in order_root2 [symmetric])
-apply (auto simp del: pmult_Cons)
-apply (rule_tac x = q in exI, safe)
-apply (simp add: poly_mult fun_eq)
-apply (drule_tac p1 = q in poly_linear_divides [THEN iffD1])
-apply (simp add: divides_def del: pmult_Cons, safe)
-apply (drule_tac x = "[]" in spec)
-apply (auto simp add: fun_eq)
-done
-
-
-text{*Normalization of a polynomial.*}
-
-lemma (in semiring_0) poly_normalize[simp]: "poly (pnormalize p) = poly p"
-apply (induct "p")
-apply (auto simp add: fun_eq)
-done
-
-text{*The degree of a polynomial.*}
-
-lemma (in semiring_0) lemma_degree_zero:
-     "list_all (%c. c = 0) p \<longleftrightarrow>  pnormalize p = []"
-by (induct "p", auto)
-
-lemma (in idom_char_0) degree_zero: 
-  assumes pN: "poly p = poly []" shows"degree p = 0"
-proof-
-  let ?pn = pnormalize
-  from pN
-  show ?thesis 
-    apply (simp add: degree_def)
-    apply (case_tac "?pn p = []")
-    apply (auto simp add: poly_zero lemma_degree_zero )
-    done
-qed
-
-lemma (in semiring_0) pnormalize_sing: "(pnormalize [x] = [x]) \<longleftrightarrow> x \<noteq> 0" by simp
-lemma (in semiring_0) pnormalize_pair: "y \<noteq> 0 \<longleftrightarrow> (pnormalize [x, y] = [x, y])" by simp
-lemma (in semiring_0) pnormal_cons: "pnormal p \<Longrightarrow> pnormal (c#p)" 
-  unfolding pnormal_def by simp
-lemma (in semiring_0) pnormal_tail: "p\<noteq>[] \<Longrightarrow> pnormal (c#p) \<Longrightarrow> pnormal p"
-  unfolding pnormal_def 
-  apply (cases "pnormalize p = []", auto)
-  by (cases "c = 0", auto)
-
-
-lemma (in semiring_0) pnormal_last_nonzero: "pnormal p ==> last p \<noteq> 0"
-proof(induct p)
-  case Nil thus ?case by (simp add: pnormal_def)
-next 
-  case (Cons a as) thus ?case
-    apply (simp add: pnormal_def)
-    apply (cases "pnormalize as = []", simp_all)
-    apply (cases "as = []", simp_all)
-    apply (cases "a=0", simp_all)
-    apply (cases "a=0", simp_all)
-    done
-qed
-
-lemma (in semiring_0) pnormal_length: "pnormal p \<Longrightarrow> 0 < length p"
-  unfolding pnormal_def length_greater_0_conv by blast
-
-lemma (in semiring_0) pnormal_last_length: "\<lbrakk>0 < length p ; last p \<noteq> 0\<rbrakk> \<Longrightarrow> pnormal p"
-  apply (induct p, auto)
-  apply (case_tac "p = []", auto)
-  apply (simp add: pnormal_def)
-  by (rule pnormal_cons, auto)
-
-lemma (in semiring_0) pnormal_id: "pnormal p \<longleftrightarrow> (0 < length p \<and> last p \<noteq> 0)"
-  using pnormal_last_length pnormal_length pnormal_last_nonzero by blast
-
-lemma (in idom_char_0) poly_Cons_eq: "poly (c#cs) = poly (d#ds) \<longleftrightarrow> c=d \<and> poly cs = poly ds" (is "?lhs \<longleftrightarrow> ?rhs")
-proof
-  assume eq: ?lhs
-  hence "\<And>x. poly ((c#cs) +++ -- (d#ds)) x = 0"
-    by (simp only: poly_minus poly_add ring_simps) simp
-  hence "poly ((c#cs) +++ -- (d#ds)) = poly []" by - (rule ext, simp) 
-  hence "c = d \<and> list_all (\<lambda>x. x=0) ((cs +++ -- ds))"
-    unfolding poly_zero by (simp add: poly_minus_def ring_simps minus_mult_left[symmetric])
-  hence "c = d \<and> (\<forall>x. poly (cs +++ -- ds) x = 0)"
-    unfolding poly_zero[symmetric] by simp 
-  thus ?rhs  apply (simp add: poly_minus poly_add ring_simps) apply (rule ext, simp) done
-next
-  assume ?rhs then show ?lhs  by -  (rule ext,simp)
-qed
-  
-lemma (in idom_char_0) pnormalize_unique: "poly p = poly q \<Longrightarrow> pnormalize p = pnormalize q"
-proof(induct q arbitrary: p)
-  case Nil thus ?case by (simp only: poly_zero lemma_degree_zero) simp
-next
-  case (Cons c cs p)
-  thus ?case
-  proof(induct p)
-    case Nil
-    hence "poly [] = poly (c#cs)" by blast
-    then have "poly (c#cs) = poly [] " by simp 
-    thus ?case by (simp only: poly_zero lemma_degree_zero) simp
-  next
-    case (Cons d ds)
-    hence eq: "poly (d # ds) = poly (c # cs)" by blast
-    hence eq': "\<And>x. poly (d # ds) x = poly (c # cs) x" by simp
-    hence "poly (d # ds) 0 = poly (c # cs) 0" by blast
-    hence dc: "d = c" by auto
-    with eq have "poly ds = poly cs"
-      unfolding  poly_Cons_eq by simp
-    with Cons.prems have "pnormalize ds = pnormalize cs" by blast
-    with dc show ?case by simp
-  qed
-qed
-
-lemma (in idom_char_0) degree_unique: assumes pq: "poly p = poly q"
-  shows "degree p = degree q"
-using pnormalize_unique[OF pq] unfolding degree_def by simp
-
-lemma (in semiring_0) pnormalize_length: "length (pnormalize p) \<le> length p" by (induct p, auto)
-
-lemma (in semiring_0) last_linear_mul_lemma: 
-  "last ((a %* p) +++ (x#(b %* p))) = (if p=[] then x else b*last p)"
-
-apply (induct p arbitrary: a x b, auto)
-apply (subgoal_tac "padd (cmult aa p) (times b a # cmult b p) \<noteq> []", simp)
-apply (induct_tac p, auto)
-done
-
-lemma (in semiring_1) last_linear_mul: assumes p:"p\<noteq>[]" shows "last ([a,1] *** p) = last p"
-proof-
-  from p obtain c cs where cs: "p = c#cs" by (cases p, auto)
-  from cs have eq:"[a,1] *** p = (a %* (c#cs)) +++ (0#(1 %* (c#cs)))"
-    by (simp add: poly_cmult_distr)
-  show ?thesis using cs
-    unfolding eq last_linear_mul_lemma by simp
-qed
-
-lemma (in semiring_0) pnormalize_eq: "last p \<noteq> 0 \<Longrightarrow> pnormalize p = p"
-  apply (induct p, auto)
-  apply (case_tac p, auto)+
-  done
-
-lemma (in semiring_0) last_pnormalize: "pnormalize p \<noteq> [] \<Longrightarrow> last (pnormalize p) \<noteq> 0"
-  by (induct p, auto)
-
-lemma (in semiring_0) pnormal_degree: "last p \<noteq> 0 \<Longrightarrow> degree p = length p - 1"
-  using pnormalize_eq[of p] unfolding degree_def by simp
-
-lemma (in semiring_0) poly_Nil_ext: "poly [] = (\<lambda>x. 0)" by (rule ext) simp
-
-lemma (in idom_char_0) linear_mul_degree: assumes p: "poly p \<noteq> poly []"
-  shows "degree ([a,1] *** p) = degree p + 1"
-proof-
-  from p have pnz: "pnormalize p \<noteq> []"
-    unfolding poly_zero lemma_degree_zero .
-  
-  from last_linear_mul[OF pnz, of a] last_pnormalize[OF pnz]
-  have l0: "last ([a, 1] *** pnormalize p) \<noteq> 0" by simp
-  from last_pnormalize[OF pnz] last_linear_mul[OF pnz, of a]
-    pnormal_degree[OF l0] pnormal_degree[OF last_pnormalize[OF pnz]] pnz
- 
-
-  have th: "degree ([a,1] *** pnormalize p) = degree (pnormalize p) + 1" 
-    by (auto simp add: poly_length_mult)
-
-  have eqs: "poly ([a,1] *** pnormalize p) = poly ([a,1] *** p)"
-    by (rule ext) (simp add: poly_mult poly_add poly_cmult)
-  from degree_unique[OF eqs] th
-  show ?thesis by (simp add: degree_unique[OF poly_normalize])
-qed
-
-lemma (in idom_char_0) linear_pow_mul_degree:
-  "degree([a,1] %^n *** p) = (if poly p = poly [] then 0 else degree p + n)"
-proof(induct n arbitrary: a p)
-  case (0 a p)
-  {assume p: "poly p = poly []"
-    hence ?case using degree_unique[OF p] by (simp add: degree_def)}
-  moreover
-  {assume p: "poly p \<noteq> poly []" hence ?case by (auto simp add: poly_Nil_ext) }
-  ultimately show ?case by blast
-next
-  case (Suc n a p)
-  have eq: "poly ([a,1] %^(Suc n) *** p) = poly ([a,1] %^ n *** ([a,1] *** p))"
-    apply (rule ext, simp add: poly_mult poly_add poly_cmult)
-    by (simp add: mult_ac add_ac right_distrib)
-  note deq = degree_unique[OF eq]
-  {assume p: "poly p = poly []"
-    with eq have eq': "poly ([a,1] %^(Suc n) *** p) = poly []" 
-      by - (rule ext,simp add: poly_mult poly_cmult poly_add)
-    from degree_unique[OF eq'] p have ?case by (simp add: degree_def)}
-  moreover
-  {assume p: "poly p \<noteq> poly []"
-    from p have ap: "poly ([a,1] *** p) \<noteq> poly []"
-      using poly_mult_not_eq_poly_Nil unfolding poly_entire by auto 
-    have eq: "poly ([a,1] %^(Suc n) *** p) = poly ([a,1]%^n *** ([a,1] *** p))"
-     by (rule ext, simp add: poly_mult poly_add poly_exp poly_cmult mult_ac add_ac right_distrib)
-   from ap have ap': "(poly ([a,1] *** p) = poly []) = False" by blast
-   have  th0: "degree ([a,1]%^n *** ([a,1] *** p)) = degree ([a,1] *** p) + n"
-     apply (simp only: Suc.hyps[of a "pmult [a,one] p"] ap')
-     by simp
-    
-   from degree_unique[OF eq] ap p th0 linear_mul_degree[OF p, of a]
-   have ?case by (auto simp del: poly.simps)}
-  ultimately show ?case by blast
-qed
-
-lemma (in recpower_idom_char_0) order_degree: 
-  assumes p0: "poly p \<noteq> poly []"
-  shows "order a p \<le> degree p"
-proof-
-  from order2[OF p0, unfolded divides_def]
-  obtain q where q: "poly p = poly ([- a, 1]%^ (order a p) *** q)" by blast
-  {assume "poly q = poly []"
-    with q p0 have False by (simp add: poly_mult poly_entire)}
-  with degree_unique[OF q, unfolded linear_pow_mul_degree] 
-  show ?thesis by auto
-qed
-
-text{*Tidier versions of finiteness of roots.*}
-
-lemma (in idom_char_0) poly_roots_finite_set: "poly p \<noteq> poly [] ==> finite {x. poly p x = 0}"
-unfolding poly_roots_finite .
-
-text{*bound for polynomial.*}
-
-lemma poly_mono: "abs(x) \<le> k ==> abs(poly p (x::'a::{ordered_idom})) \<le> poly (map abs p) k"
-apply (induct "p", auto)
-apply (rule_tac y = "abs a + abs (x * poly p x)" in order_trans)
-apply (rule abs_triangle_ineq)
-apply (auto intro!: mult_mono simp add: abs_mult)
-done
-
-lemma (in semiring_0) poly_Sing: "poly [c] x = c" by simp
-
-end
--- a/src/HOL/Word/WordArith.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/Word/WordArith.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -22,7 +22,7 @@
 proof
 qed (unfold word_sle_def word_sless_def, auto)
 
-class_interpretation signed: linorder ["word_sle" "word_sless"] 
+interpretation signed!: linorder "word_sle" "word_sless"
   by (rule signed_linorder)
 
 lemmas word_arith_wis = 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/base.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -0,0 +1,2 @@
+(*side-entry for HOL-Base*)
+use_thy "Code_Setup";
--- a/src/HOL/ex/ImperativeQuicksort.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/ex/ImperativeQuicksort.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,5 @@
 theory ImperativeQuicksort
-imports Imperative_HOL Subarray Multiset Efficient_Nat
+imports "~~/src/HOL/Imperative_HOL/Imperative_HOL" Subarray Multiset Efficient_Nat
 begin
 
 text {* We prove QuickSort correct in the Relational Calculus. *}
--- a/src/HOL/ex/MIR.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/ex/MIR.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -5899,7 +5899,7 @@
   let 
     val thy = Thm.theory_of_cterm ct;
     val t = Thm.term_of ct;
-    val fs = term_frees t;
+    val fs = OldTerm.term_frees t;
     val vs = fs ~~ (0 upto (length fs - 1));
     val qe = if proofs then @{code mirlfrqe} else @{code mircfrqe};
     val t' = (term_of_fm vs o qe o fm_of_term vs) t;
--- a/src/HOL/ex/ROOT.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/ex/ROOT.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -11,12 +11,14 @@
   "Word",
   "Eval_Examples",
   "Quickcheck",
-  "Term_Of_Syntax"
+  "Term_Of_Syntax",
+  "Codegenerator",
+  "Codegenerator_Pretty",
+  "NormalForm",
+  "../NumberTheory/Factorization",
+  "../Library/BigO"
 ];
 
-no_document use_thy "Codegenerator";
-no_document use_thy "Codegenerator_Pretty";
-
 use_thys [
   "Numeral",
   "ImperativeQuicksort",
@@ -57,56 +59,43 @@
   "Meson_Test",
   "Code_Antiq",
   "Termination",
-  "Coherent"
+  "Coherent",
+  "Dense_Linear_Order_Ex",
+  "PresburgerEx",
+  "Reflected_Presburger",
+  "Reflection",
+  "ReflectionEx",
+  "BinEx",
+  "Sqrt",
+  "Sqrt_Script",
+  "BigO_Complex",
+  "Arithmetic_Series_Complex",
+  "HarmonicSeries",
+  "MIR",
+  "ReflectedFerrack",
+  "Refute_Examples",
+  "Quickcheck_Examples"
 ];
 
-setmp Proofterm.proofs 2 time_use_thy "Hilbert_Classical";
+setmp Proofterm.proofs 2 use_thy "Hilbert_Classical";
 
-time_use_thy "Dense_Linear_Order_Ex";
-time_use_thy "PresburgerEx";
-time_use_thy "Reflected_Presburger";
 
-use_thys ["Reflection", "ReflectionEx"];
+use_thy "SVC_Oracle";
 
-time_use_thy "SVC_Oracle";
-
-(*check if user has SVC installed*)
 fun svc_enabled () = getenv "SVC_HOME" <> "";
 fun if_svc_enabled f x = if svc_enabled () then f x else ();
 
-if_svc_enabled time_use_thy "svc_test";
+if_svc_enabled use_thy "svc_test";
+
 
 (* requires a proof-generating SAT solver (zChaff or MiniSAT) to be *)
 (* installed:                                                       *)
-try time_use_thy "SAT_Examples";
+try use_thy "SAT_Examples";
 
 (* requires zChaff (or some other reasonably fast SAT solver) to be *)
 (* installed:                                                       *)
 if getenv "ZCHAFF_HOME" <> "" then
-  time_use_thy "Sudoku"
+  use_thy "Sudoku"
 else ();
 
-time_use_thy "Refute_Examples";
-time_use_thy "Quickcheck_Examples";
-no_document time_use_thy "NormalForm";
-
 HTML.with_charset "utf-8" (no_document use_thys) ["Hebrew", "Chinese"];
-
-no_document use_thys [
-  "../NumberTheory/Factorization",
-  "../Library/BigO"
-];
-
-use_thys [
-  "BinEx",
-  "Sqrt",
-  "Sqrt_Script",
-  "BigO_Complex",
-
-  "Arithmetic_Series_Complex",
-  "HarmonicSeries",
-
-  "MIR",
-  "ReflectedFerrack"
-];
-
--- a/src/HOL/ex/ReflectedFerrack.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/ex/ReflectedFerrack.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -1,4 +1,4 @@
-(*  Title:      Complex/ex/ReflectedFerrack.thy
+(*  Title:      HOL/ex/ReflectedFerrack.thy
     Author:     Amine Chaieb
 *)
 
@@ -2070,7 +2070,7 @@
   let 
     val thy = Thm.theory_of_cterm ct;
     val t = Thm.term_of ct;
-    val fs = term_frees t;
+    val fs = OldTerm.term_frees t;
     val vs = fs ~~ (0 upto (length fs - 1));
     val res = HOLogic.mk_Trueprop (HOLogic.mk_eq (t, term_of_fm vs (@{code linrqe} (fm_of_term vs t))));
   in Thm.cterm_of thy res end
--- a/src/HOL/ex/Reflected_Presburger.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/ex/Reflected_Presburger.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -2039,7 +2039,7 @@
   let
     val thy = Thm.theory_of_cterm ct;
     val t = Thm.term_of ct;
-    val fs = term_frees t;
+    val fs = OldTerm.term_frees t;
     val bs = term_bools [] t;
     val vs = fs ~~ (0 upto (length fs - 1))
     val ps = bs ~~ (0 upto (length bs - 1))
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/Subarray.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -0,0 +1,66 @@
+theory Subarray
+imports Array Sublist
+begin
+
+definition subarray :: "nat \<Rightarrow> nat \<Rightarrow> ('a::heap) array \<Rightarrow> heap \<Rightarrow> 'a list"
+where
+  "subarray n m a h \<equiv> sublist' n m (get_array a h)"
+
+lemma subarray_upd: "i \<ge> m \<Longrightarrow> subarray n m a (Heap.upd a i v h) = subarray n m a h"
+apply (simp add: subarray_def Heap.upd_def)
+apply (simp add: sublist'_update1)
+done
+
+lemma subarray_upd2: " i < n  \<Longrightarrow> subarray n m a (Heap.upd a i v h) = subarray n m a h"
+apply (simp add: subarray_def Heap.upd_def)
+apply (subst sublist'_update2)
+apply fastsimp
+apply simp
+done
+
+lemma subarray_upd3: "\<lbrakk> n \<le> i; i < m\<rbrakk> \<Longrightarrow> subarray n m a (Heap.upd a i v h) = subarray n m a h[i - n := v]"
+unfolding subarray_def Heap.upd_def
+by (simp add: sublist'_update3)
+
+lemma subarray_Nil: "n \<ge> m \<Longrightarrow> subarray n m a h = []"
+by (simp add: subarray_def sublist'_Nil')
+
+lemma subarray_single: "\<lbrakk> n < Heap.length a h \<rbrakk> \<Longrightarrow> subarray n (Suc n) a h = [get_array a h ! n]" 
+by (simp add: subarray_def Heap.length_def sublist'_single)
+
+lemma length_subarray: "m \<le> Heap.length a h \<Longrightarrow> List.length (subarray n m a h) = m - n"
+by (simp add: subarray_def Heap.length_def length_sublist')
+
+lemma length_subarray_0: "m \<le> Heap.length a h \<Longrightarrow> List.length (subarray 0 m a h) = m"
+by (simp add: length_subarray)
+
+lemma subarray_nth_array_Cons: "\<lbrakk> i < Heap.length a h; i < j \<rbrakk> \<Longrightarrow> (get_array a h ! i) # subarray (Suc i) j a h = subarray i j a h"
+unfolding Heap.length_def subarray_def
+by (simp add: sublist'_front)
+
+lemma subarray_nth_array_back: "\<lbrakk> i < j; j \<le> Heap.length a h\<rbrakk> \<Longrightarrow> subarray i j a h = subarray i (j - 1) a h @ [get_array a h ! (j - 1)]"
+unfolding Heap.length_def subarray_def
+by (simp add: sublist'_back)
+
+lemma subarray_append: "\<lbrakk> i < j; j < k \<rbrakk> \<Longrightarrow> subarray i j a h @ subarray j k a h = subarray i k a h"
+unfolding subarray_def
+by (simp add: sublist'_append)
+
+lemma subarray_all: "subarray 0 (Heap.length a h) a h = get_array a h"
+unfolding Heap.length_def subarray_def
+by (simp add: sublist'_all)
+
+lemma nth_subarray: "\<lbrakk> k < j - i; j \<le> Heap.length a h \<rbrakk> \<Longrightarrow> subarray i j a h ! k = get_array a h ! (i + k)"
+unfolding Heap.length_def subarray_def
+by (simp add: nth_sublist')
+
+lemma subarray_eq_samelength_iff: "Heap.length a h = Heap.length a h' \<Longrightarrow> (subarray i j a h = subarray i j a h') = (\<forall>i'. i \<le> i' \<and> i' < j \<longrightarrow> get_array a h ! i' = get_array a h' ! i')"
+unfolding Heap.length_def subarray_def by (rule sublist'_eq_samelength_iff)
+
+lemma all_in_set_subarray_conv: "(\<forall>j. j \<in> set (subarray l r a h) \<longrightarrow> P j) = (\<forall>k. l \<le> k \<and> k < r \<and> k < Heap.length a h \<longrightarrow> P (get_array a h ! k))"
+unfolding subarray_def Heap.length_def by (rule all_in_set_sublist'_conv)
+
+lemma ball_in_set_subarray_conv: "(\<forall>j \<in> set (subarray l r a h). P j) = (\<forall>k. l \<le> k \<and> k < r \<and> k < Heap.length a h \<longrightarrow> P (get_array a h ! k))"
+unfolding subarray_def Heap.length_def by (rule ball_in_set_sublist'_conv)
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/Sublist.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -0,0 +1,507 @@
+(* $Id$ *)
+
+header {* Slices of lists *}
+
+theory Sublist
+imports Multiset
+begin
+
+
+lemma sublist_split: "i \<le> j \<and> j \<le> k \<Longrightarrow> sublist xs {i..<j} @ sublist xs {j..<k} = sublist xs {i..<k}" 
+apply (induct xs arbitrary: i j k)
+apply simp
+apply (simp only: sublist_Cons)
+apply simp
+apply safe
+apply simp
+apply (erule_tac x="0" in meta_allE)
+apply (erule_tac x="j - 1" in meta_allE)
+apply (erule_tac x="k - 1" in meta_allE)
+apply (subgoal_tac "0 \<le> j - 1 \<and> j - 1 \<le> k - 1")
+apply simp
+apply (subgoal_tac "{ja. Suc ja < j} = {0..<j - Suc 0}")
+apply (subgoal_tac "{ja. j \<le> Suc ja \<and> Suc ja < k} = {j - Suc 0..<k - Suc 0}")
+apply (subgoal_tac "{j. Suc j < k} = {0..<k - Suc 0}")
+apply simp
+apply fastsimp
+apply fastsimp
+apply fastsimp
+apply fastsimp
+apply (erule_tac x="i - 1" in meta_allE)
+apply (erule_tac x="j - 1" in meta_allE)
+apply (erule_tac x="k - 1" in meta_allE)
+apply (subgoal_tac " {ja. i \<le> Suc ja \<and> Suc ja < j} = {i - 1 ..<j - 1}")
+apply (subgoal_tac " {ja. j \<le> Suc ja \<and> Suc ja < k} = {j - 1..<k - 1}")
+apply (subgoal_tac "{j. i \<le> Suc j \<and> Suc j < k} = {i - 1..<k - 1}")
+apply (subgoal_tac " i - 1 \<le> j - 1 \<and> j - 1 \<le> k - 1")
+apply simp
+apply fastsimp
+apply fastsimp
+apply fastsimp
+apply fastsimp
+done
+
+lemma sublist_update1: "i \<notin> inds \<Longrightarrow> sublist (xs[i := v]) inds = sublist xs inds"
+apply (induct xs arbitrary: i inds)
+apply simp
+apply (case_tac i)
+apply (simp add: sublist_Cons)
+apply (simp add: sublist_Cons)
+done
+
+lemma sublist_update2: "i \<in> inds \<Longrightarrow> sublist (xs[i := v]) inds = (sublist xs inds)[(card {k \<in> inds. k < i}):= v]"
+proof (induct xs arbitrary: i inds)
+  case Nil thus ?case by simp
+next
+  case (Cons x xs)
+  thus ?case
+  proof (cases i)
+    case 0 with Cons show ?thesis by (simp add: sublist_Cons)
+  next
+    case (Suc i')
+    with Cons show ?thesis
+      apply simp
+      apply (simp add: sublist_Cons)
+      apply auto
+      apply (auto simp add: nat.split)
+      apply (simp add: card_less)
+      apply (simp add: card_less)
+      apply (simp add: card_less_Suc[symmetric])
+      apply (simp add: card_less_Suc2)
+      done
+  qed
+qed
+
+lemma sublist_update: "sublist (xs[i := v]) inds = (if i \<in> inds then (sublist xs inds)[(card {k \<in> inds. k < i}) := v] else sublist xs inds)"
+by (simp add: sublist_update1 sublist_update2)
+
+lemma sublist_take: "sublist xs {j. j < m} = take m xs"
+apply (induct xs arbitrary: m)
+apply simp
+apply (case_tac m)
+apply simp
+apply (simp add: sublist_Cons)
+done
+
+lemma sublist_take': "sublist xs {0..<m} = take m xs"
+apply (induct xs arbitrary: m)
+apply simp
+apply (case_tac m)
+apply simp
+apply (simp add: sublist_Cons sublist_take)
+done
+
+lemma sublist_all[simp]: "sublist xs {j. j < length xs} = xs"
+apply (induct xs)
+apply simp
+apply (simp add: sublist_Cons)
+done
+
+lemma sublist_all'[simp]: "sublist xs {0..<length xs} = xs"
+apply (induct xs)
+apply simp
+apply (simp add: sublist_Cons)
+done
+
+lemma sublist_single: "a < length xs \<Longrightarrow> sublist xs {a} = [xs ! a]"
+apply (induct xs arbitrary: a)
+apply simp
+apply(case_tac aa)
+apply simp
+apply (simp add: sublist_Cons)
+apply simp
+apply (simp add: sublist_Cons)
+done
+
+lemma sublist_is_Nil: "\<forall>i \<in> inds. i \<ge> length xs \<Longrightarrow> sublist xs inds = []" 
+apply (induct xs arbitrary: inds)
+apply simp
+apply (simp add: sublist_Cons)
+apply auto
+apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
+apply auto
+done
+
+lemma sublist_Nil': "sublist xs inds = [] \<Longrightarrow> \<forall>i \<in> inds. i \<ge> length xs"
+apply (induct xs arbitrary: inds)
+apply simp
+apply (simp add: sublist_Cons)
+apply (auto split: if_splits)
+apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
+apply (case_tac x, auto)
+done
+
+lemma sublist_Nil[simp]: "(sublist xs inds = []) = (\<forall>i \<in> inds. i \<ge> length xs)"
+apply (induct xs arbitrary: inds)
+apply simp
+apply (simp add: sublist_Cons)
+apply auto
+apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
+apply (case_tac x, auto)
+done
+
+lemma sublist_eq_subseteq: " \<lbrakk> inds' \<subseteq> inds; sublist xs inds = sublist ys inds \<rbrakk> \<Longrightarrow> sublist xs inds' = sublist ys inds'"
+apply (induct xs arbitrary: ys inds inds')
+apply simp
+apply (drule sym, rule sym)
+apply (simp add: sublist_Nil, fastsimp)
+apply (case_tac ys)
+apply (simp add: sublist_Nil, fastsimp)
+apply (auto simp add: sublist_Cons)
+apply (erule_tac x="list" in meta_allE)
+apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
+apply (erule_tac x="{j. Suc j \<in> inds'}" in meta_allE)
+apply fastsimp
+apply (erule_tac x="list" in meta_allE)
+apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
+apply (erule_tac x="{j. Suc j \<in> inds'}" in meta_allE)
+apply fastsimp
+done
+
+lemma sublist_eq: "\<lbrakk> \<forall>i \<in> inds. ((i < length xs) \<and> (i < length ys)) \<or> ((i \<ge> length xs ) \<and> (i \<ge> length ys)); \<forall>i \<in> inds. xs ! i = ys ! i \<rbrakk> \<Longrightarrow> sublist xs inds = sublist ys inds"
+apply (induct xs arbitrary: ys inds)
+apply simp
+apply (rule sym, simp add: sublist_Nil)
+apply (case_tac ys)
+apply (simp add: sublist_Nil)
+apply (auto simp add: sublist_Cons)
+apply (erule_tac x="list" in meta_allE)
+apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
+apply fastsimp
+apply (erule_tac x="list" in meta_allE)
+apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
+apply fastsimp
+done
+
+lemma sublist_eq_samelength: "\<lbrakk> length xs = length ys; \<forall>i \<in> inds. xs ! i = ys ! i \<rbrakk> \<Longrightarrow> sublist xs inds = sublist ys inds"
+by (rule sublist_eq, auto)
+
+lemma sublist_eq_samelength_iff: "length xs = length ys \<Longrightarrow> (sublist xs inds = sublist ys inds) = (\<forall>i \<in> inds. xs ! i = ys ! i)"
+apply (induct xs arbitrary: ys inds)
+apply simp
+apply (rule sym, simp add: sublist_Nil)
+apply (case_tac ys)
+apply (simp add: sublist_Nil)
+apply (auto simp add: sublist_Cons)
+apply (case_tac i)
+apply auto
+apply (case_tac i)
+apply auto
+done
+
+section {* Another sublist function *}
+
+function sublist' :: "nat \<Rightarrow> nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
+where
+  "sublist' n m [] = []"
+| "sublist' n 0 xs = []"
+| "sublist' 0 (Suc m) (x#xs) = (x#sublist' 0 m xs)"
+| "sublist' (Suc n) (Suc m) (x#xs) = sublist' n m xs"
+by pat_completeness auto
+termination by lexicographic_order
+
+subsection {* Proving equivalence to the other sublist command *}
+
+lemma sublist'_sublist: "sublist' n m xs = sublist xs {j. n \<le> j \<and> j < m}"
+apply (induct xs arbitrary: n m)
+apply simp
+apply (case_tac n)
+apply (case_tac m)
+apply simp
+apply (simp add: sublist_Cons)
+apply (case_tac m)
+apply simp
+apply (simp add: sublist_Cons)
+done
+
+
+lemma "sublist' n m xs = sublist xs {n..<m}"
+apply (induct xs arbitrary: n m)
+apply simp
+apply (case_tac n, case_tac m)
+apply simp
+apply simp
+apply (simp add: sublist_take')
+apply (case_tac m)
+apply simp
+apply (simp add: sublist_Cons sublist'_sublist)
+done
+
+
+subsection {* Showing equivalence to use of drop and take for definition *}
+
+lemma "sublist' n m xs = take (m - n) (drop n xs)"
+apply (induct xs arbitrary: n m)
+apply simp
+apply (case_tac m)
+apply simp
+apply (case_tac n)
+apply simp
+apply simp
+done
+
+subsection {* General lemma about sublist *}
+
+lemma sublist'_Nil[simp]: "sublist' i j [] = []"
+by simp
+
+lemma sublist'_Cons[simp]: "sublist' i (Suc j) (x#xs) = (case i of 0 \<Rightarrow> (x # sublist' 0 j xs) | Suc i' \<Rightarrow>  sublist' i' j xs)"
+by (cases i) auto
+
+lemma sublist'_Cons2[simp]: "sublist' i j (x#xs) = (if (j = 0) then [] else ((if (i = 0) then [x] else []) @ sublist' (i - 1) (j - 1) xs))"
+apply (cases j)
+apply auto
+apply (cases i)
+apply auto
+done
+
+lemma sublist_n_0: "sublist' n 0 xs = []"
+by (induct xs, auto)
+
+lemma sublist'_Nil': "n \<ge> m \<Longrightarrow> sublist' n m xs = []"
+apply (induct xs arbitrary: n m)
+apply simp
+apply (case_tac m)
+apply simp
+apply (case_tac n)
+apply simp
+apply simp
+done
+
+lemma sublist'_Nil2: "n \<ge> length xs \<Longrightarrow> sublist' n m xs = []"
+apply (induct xs arbitrary: n m)
+apply simp
+apply (case_tac m)
+apply simp
+apply (case_tac n)
+apply simp
+apply simp
+done
+
+lemma sublist'_Nil3: "(sublist' n m xs = []) = ((n \<ge> m) \<or> (n \<ge> length xs))"
+apply (induct xs arbitrary: n m)
+apply simp
+apply (case_tac m)
+apply simp
+apply (case_tac n)
+apply simp
+apply simp
+done
+
+lemma sublist'_notNil: "\<lbrakk> n < length xs; n < m \<rbrakk> \<Longrightarrow> sublist' n m xs \<noteq> []"
+apply (induct xs arbitrary: n m)
+apply simp
+apply (case_tac m)
+apply simp
+apply (case_tac n)
+apply simp
+apply simp
+done
+
+lemma sublist'_single: "n < length xs \<Longrightarrow> sublist' n (Suc n) xs = [xs ! n]"
+apply (induct xs arbitrary: n)
+apply simp
+apply simp
+apply (case_tac n)
+apply (simp add: sublist_n_0)
+apply simp
+done
+
+lemma sublist'_update1: "i \<ge> m \<Longrightarrow> sublist' n m (xs[i:=v]) = sublist' n m xs"
+apply (induct xs arbitrary: n m i)
+apply simp
+apply simp
+apply (case_tac i)
+apply simp
+apply simp
+done
+
+lemma sublist'_update2: "i < n \<Longrightarrow> sublist' n m (xs[i:=v]) = sublist' n m xs"
+apply (induct xs arbitrary: n m i)
+apply simp
+apply simp
+apply (case_tac i)
+apply simp
+apply simp
+done
+
+lemma sublist'_update3: "\<lbrakk>n \<le> i; i < m\<rbrakk> \<Longrightarrow> sublist' n m (xs[i := v]) = (sublist' n m xs)[i - n := v]"
+proof (induct xs arbitrary: n m i)
+  case Nil thus ?case by auto
+next
+  case (Cons x xs)
+  thus ?case
+    apply -
+    apply auto
+    apply (cases i)
+    apply auto
+    apply (cases i)
+    apply auto
+    done
+qed
+
+lemma "\<lbrakk> sublist' i j xs = sublist' i j ys; n \<ge> i; m \<le> j \<rbrakk> \<Longrightarrow> sublist' n m xs = sublist' n m ys"
+proof (induct xs arbitrary: i j ys n m)
+  case Nil
+  thus ?case
+    apply -
+    apply (rule sym, drule sym)
+    apply (simp add: sublist'_Nil)
+    apply (simp add: sublist'_Nil3)
+    apply arith
+    done
+next
+  case (Cons x xs i j ys n m)
+  note c = this
+  thus ?case
+  proof (cases m)
+    case 0 thus ?thesis by (simp add: sublist_n_0)
+  next
+    case (Suc m')
+    note a = this
+    thus ?thesis
+    proof (cases n)
+      case 0 note b = this
+      show ?thesis
+      proof (cases ys)
+	case Nil  with a b Cons.prems show ?thesis by (simp add: sublist'_Nil3)
+      next
+	case (Cons y ys)
+	show ?thesis
+	proof (cases j)
+	  case 0 with a b Cons.prems show ?thesis by simp
+	next
+	  case (Suc j') with a b Cons.prems Cons show ?thesis 
+	    apply -
+	    apply (simp, rule Cons.hyps [of "0" "j'" "ys" "0" "m'"], auto)
+	    done
+	qed
+      qed
+    next
+      case (Suc n')
+      show ?thesis
+      proof (cases ys)
+	case Nil with Suc a Cons.prems show ?thesis by (auto simp add: sublist'_Nil3)
+      next
+	case (Cons y ys) with Suc a Cons.prems show ?thesis
+	  apply -
+	  apply simp
+	  apply (cases j)
+	  apply simp
+	  apply (cases i)
+	  apply simp
+	  apply (rule_tac j="nat" in Cons.hyps [of "0" _ "ys" "n'" "m'"])
+	  apply simp
+	  apply simp
+	  apply simp
+	  apply simp
+	  apply (rule_tac i="nata" and j="nat" in Cons.hyps [of _ _ "ys" "n'" "m'"])
+	  apply simp
+	  apply simp
+	  apply simp
+	  done
+      qed
+    qed
+  qed
+qed
+
+lemma length_sublist': "j \<le> length xs \<Longrightarrow> length (sublist' i j xs) = j - i"
+by (induct xs arbitrary: i j, auto)
+
+lemma sublist'_front: "\<lbrakk> i < j; i < length xs \<rbrakk> \<Longrightarrow> sublist' i j xs = xs ! i # sublist' (Suc i) j xs"
+apply (induct xs arbitrary: a i j)
+apply simp
+apply (case_tac j)
+apply simp
+apply (case_tac i)
+apply simp
+apply simp
+done
+
+lemma sublist'_back: "\<lbrakk> i < j; j \<le> length xs \<rbrakk> \<Longrightarrow> sublist' i j xs = sublist' i (j - 1) xs @ [xs ! (j - 1)]"
+apply (induct xs arbitrary: a i j)
+apply simp
+apply simp
+apply (case_tac j)
+apply simp
+apply auto
+apply (case_tac nat)
+apply auto
+done
+
+(* suffices that j \<le> length xs and length ys *) 
+lemma sublist'_eq_samelength_iff: "length xs = length ys \<Longrightarrow> (sublist' i j xs  = sublist' i j ys) = (\<forall>i'. i \<le> i' \<and> i' < j \<longrightarrow> xs ! i' = ys ! i')"
+proof (induct xs arbitrary: ys i j)
+  case Nil thus ?case by simp
+next
+  case (Cons x xs)
+  thus ?case
+    apply -
+    apply (cases ys)
+    apply simp
+    apply simp
+    apply auto
+    apply (case_tac i', auto)
+    apply (erule_tac x="Suc i'" in allE, auto)
+    apply (erule_tac x="i' - 1" in allE, auto)
+    apply (case_tac i', auto)
+    apply (erule_tac x="Suc i'" in allE, auto)
+    done
+qed
+
+lemma sublist'_all[simp]: "sublist' 0 (length xs) xs = xs"
+by (induct xs, auto)
+
+lemma sublist'_sublist': "sublist' n m (sublist' i j xs) = sublist' (i + n) (min (i + m) j) xs" 
+by (induct xs arbitrary: i j n m) (auto simp add: min_diff)
+
+lemma sublist'_append: "\<lbrakk> i \<le> j; j \<le> k \<rbrakk> \<Longrightarrow>(sublist' i j xs) @ (sublist' j k xs) = sublist' i k xs"
+by (induct xs arbitrary: i j k) auto
+
+lemma nth_sublist': "\<lbrakk> k < j - i; j \<le> length xs \<rbrakk> \<Longrightarrow> (sublist' i j xs) ! k = xs ! (i + k)"
+apply (induct xs arbitrary: i j k)
+apply auto
+apply (case_tac k)
+apply auto
+apply (case_tac i)
+apply auto
+done
+
+lemma set_sublist': "set (sublist' i j xs) = {x. \<exists>k. i \<le> k \<and> k < j \<and> k < List.length xs \<and> x = xs ! k}"
+apply (simp add: sublist'_sublist)
+apply (simp add: set_sublist)
+apply auto
+done
+
+lemma all_in_set_sublist'_conv: "(\<forall>j. j \<in> set (sublist' l r xs) \<longrightarrow> P j) = (\<forall>k. l \<le> k \<and> k < r \<and> k < List.length xs \<longrightarrow> P (xs ! k))"
+unfolding set_sublist' by blast
+
+lemma ball_in_set_sublist'_conv: "(\<forall>j \<in> set (sublist' l r xs). P j) = (\<forall>k. l \<le> k \<and> k < r \<and> k < List.length xs \<longrightarrow> P (xs ! k))"
+unfolding set_sublist' by blast
+
+
+lemma multiset_of_sublist:
+assumes l_r: "l \<le> r \<and> r \<le> List.length xs"
+assumes left: "\<forall> i. i < l \<longrightarrow> (xs::'a list) ! i = ys ! i"
+assumes right: "\<forall> i. i \<ge> r \<longrightarrow> (xs::'a list) ! i = ys ! i"
+assumes multiset: "multiset_of xs = multiset_of ys"
+  shows "multiset_of (sublist' l r xs) = multiset_of (sublist' l r ys)"
+proof -
+  from l_r have xs_def: "xs = (sublist' 0 l xs) @ (sublist' l r xs) @ (sublist' r (List.length xs) xs)" (is "_ = ?xs_long") 
+    by (simp add: sublist'_append)
+  from multiset have length_eq: "List.length xs = List.length ys" by (rule multiset_of_eq_length)
+  with l_r have ys_def: "ys = (sublist' 0 l ys) @ (sublist' l r ys) @ (sublist' r (List.length ys) ys)" (is "_ = ?ys_long") 
+    by (simp add: sublist'_append)
+  from xs_def ys_def multiset have "multiset_of ?xs_long = multiset_of ?ys_long" by simp
+  moreover
+  from left l_r length_eq have "sublist' 0 l xs = sublist' 0 l ys"
+    by (auto simp add: length_sublist' nth_sublist' intro!: nth_equalityI)
+  moreover
+  from right l_r length_eq have "sublist' r (List.length xs) xs = sublist' r (List.length ys) ys"
+    by (auto simp add: length_sublist' nth_sublist' intro!: nth_equalityI)
+  moreover
+  ultimately show ?thesis by (simp add: multiset_of_append)
+qed
+
+
+end
--- a/src/HOL/ex/coopertac.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/ex/coopertac.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -47,7 +47,7 @@
     val (fm',np) =  foldr (fn ((x, T), (fm,n)) => mk_all ((x, T), (fm,n)))
       (foldr HOLogic.mk_imp c rhs, np) ps
     val (vs, _) = List.partition (fn t => q orelse (type_of t) = nT)
-      (term_frees fm' @ term_vars fm');
+      (OldTerm.term_frees fm' @ OldTerm.term_vars fm');
     val fm2 = foldr mk_all2 fm' vs
   in (fm2, np + length vs, length rhs) end;
 
--- a/src/HOL/ex/linrtac.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/ex/linrtac.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -54,7 +54,7 @@
     val (fm',np) =  foldr (fn ((x, T), (fm,n)) => mk_all ((x, T), (fm,n)))
       (foldr HOLogic.mk_imp c rhs, np) ps
     val (vs, _) = List.partition (fn t => q orelse (type_of t) = HOLogic.natT)
-      (term_frees fm' @ term_vars fm');
+      (OldTerm.term_frees fm' @ OldTerm.term_vars fm');
     val fm2 = foldr mk_all2 fm' vs
   in (fm2, np + length vs, length rhs) end;
 
--- a/src/HOL/ex/mirtac.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/ex/mirtac.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
-(*  Title:      HOL/Complex/ex/mirtac.ML
-    ID:         $Id$
+(*  Title:      HOL/ex/mirtac.ML
     Author:     Amine Chaieb, TU Muenchen
 *)
 
@@ -74,7 +73,7 @@
     val (fm',np) =  foldr (fn ((x, T), (fm,n)) => mk_all ((x, T), (fm,n)))
       (foldr HOLogic.mk_imp c rhs, np) ps
     val (vs, _) = List.partition (fn t => q orelse (type_of t) = nT)
-      (term_frees fm' @ term_vars fm');
+      (OldTerm.term_frees fm' @ OldTerm.term_vars fm');
     val fm2 = foldr mk_all2 fm' vs
   in (fm2, np + length vs, length rhs) end;
 
--- a/src/HOL/ex/reflection.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/ex/reflection.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,19 +1,18 @@
-(*
-    ID:         $Id$
-    Author:     Amine Chaieb, TU Muenchen
+(*  Author:     Amine Chaieb, TU Muenchen
 
 A trial for automatical reification.
 *)
 
-signature REFLECTION = sig
+signature REFLECTION =
+sig
   val genreify_tac: Proof.context -> thm list -> term option -> int -> tactic
   val reflection_tac: Proof.context -> thm list -> thm list -> term option -> int -> tactic
   val gen_reflection_tac: Proof.context -> (cterm -> thm)
     -> thm list -> thm list -> term option -> int -> tactic
 end;
 
-structure Reflection : REFLECTION
-= struct
+structure Reflection : REFLECTION =
+struct
 
 val ext2 = thm "ext2";
 val nth_Cons_0 = thm "nth_Cons_0";
@@ -38,10 +37,10 @@
    val (((_,_),[th']), ctxt') = Variable.import_thms true [th] ctxt
    val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of th'))
    fun add_fterms (t as t1 $ t2) = 
-       if exists (fn f => could_unify (t |> strip_comb |> fst, f)) fs then insert (op aconv) t
+       if exists (fn f => Term.could_unify (t |> strip_comb |> fst, f)) fs then insert (op aconv) t
        else add_fterms t1 #> add_fterms t2
      | add_fterms (t as Abs(xn,xT,t')) = 
-       if (fN mem (term_consts t)) then (fn _ => [t]) else (fn _ => [])
+       if exists_Const (fn (c, _) => c = fN) t then (fn _ => [t]) else (fn _ => [])
      | add_fterms _ = I
    val fterms = add_fterms rhs []
    val (xs, ctxt'') = Variable.variant_fixes (replicate (length fterms) "x") ctxt'
@@ -139,7 +138,7 @@
         val (tyenv, tmenv) =
         Pattern.match thy
         ((fst o HOLogic.dest_eq o HOLogic.dest_Trueprop) (concl_of cong), t)
-        (Envir.type_env (Envir.empty 0),Term.Vartab.empty)
+        (Envir.type_env (Envir.empty 0), Vartab.empty)
         val (fnvs,invs) = List.partition (fn ((vn,_),_) => vn mem vns) (Vartab.dest tmenv)
         val (fts,its) = 
 	    (map (snd o snd) fnvs,
@@ -191,7 +190,7 @@
    val rhs_P = subst_free subst rhs
    val (tyenv, tmenv) = Pattern.match 
 	                    thy (rhs_P, t)
-	                    (Envir.type_env (Envir.empty 0),Term.Vartab.empty)
+	                    (Envir.type_env (Envir.empty 0), Vartab.empty)
    val sbst = Envir.subst_vars (tyenv, tmenv)
    val sbsT = Envir.typ_subst_TVars tyenv
    val subst_ty = map (fn (n,(s,t)) => (certT (TVar (n, s)), certT t)) 
--- a/src/HOL/ex/svc_funcs.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/ex/svc_funcs.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,9 +1,8 @@
 (*  Title:      HOL/ex/svc_funcs.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson
     Copyright   1999  University of Cambridge
 
-Translation functions for the interface to SVC
+Translation functions for the interface to SVC.
 
 Based upon the work of Soren T. Heilmann
 
@@ -126,7 +125,7 @@
    pos ["positive"]: true if an assumption, false if a goal*)
  fun expr_of pos t =
   let
-    val params = rev (rename_wrt_term t (Term.strip_all_vars t))
+    val params = rev (Term.rename_wrt_term t (Term.strip_all_vars t))
     and body   = Term.strip_all_body t
     val nat_vars = ref ([] : string list)
     (*translation of a variable: record all natural numbers*)
--- a/src/HOL/main.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/main.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,7 +1,2 @@
-(*  Title:      HOL/main.ML
-    ID:         $Id$
- 
-Classical Higher-order Logic -- only "Main".
-*)
-
+(*side-entry for HOL-Main*)
 use_thy "Main";
--- a/src/HOL/plain.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOL/plain.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,6 +1,2 @@
-(*  Title:      HOL/plain.ML
- 
-Classical Higher-order Logic -- plain Tool bootstrap.
-*)
-
+(*side-entry for HOL-Plain*)
 use_thy "Plain";
--- a/src/HOLCF/Cfun.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOLCF/Cfun.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -7,8 +7,7 @@
 header {* The type of continuous functions *}
 
 theory Cfun
-imports Pcpodef Ffun
-uses ("Tools/cont_proc.ML")
+imports Pcpodef Ffun Product_Cpo
 begin
 
 defaultsort cpo
@@ -301,7 +300,7 @@
 
 text {* cont2cont lemma for @{term Rep_CFun} *}
 
-lemma cont2cont_Rep_CFun:
+lemma cont2cont_Rep_CFun [cont2cont]:
   assumes f: "cont (\<lambda>x. f x)"
   assumes t: "cont (\<lambda>x. t x)"
   shows "cont (\<lambda>x. (f x)\<cdot>(t x))"
@@ -321,6 +320,11 @@
 
 text {* cont2cont Lemma for @{term "%x. LAM y. f x y"} *}
 
+text {*
+  Not suitable as a cont2cont rule, because on nested lambdas
+  it causes exponential blow-up in the number of subgoals.
+*}
+
 lemma cont2cont_LAM:
   assumes f1: "\<And>x. cont (\<lambda>y. f x y)"
   assumes f2: "\<And>y. cont (\<lambda>x. f x y)"
@@ -331,17 +335,40 @@
   from f2 show "cont f" by (rule cont2cont_lambda)
 qed
 
-text {* continuity simplification procedure *}
+text {*
+  This version does work as a cont2cont rule, since it
+  has only a single subgoal.
+*}
+
+lemma cont2cont_LAM' [cont2cont]:
+  fixes f :: "'a::cpo \<Rightarrow> 'b::cpo \<Rightarrow> 'c::cpo"
+  assumes f: "cont (\<lambda>p. f (fst p) (snd p))"
+  shows "cont (\<lambda>x. \<Lambda> y. f x y)"
+proof (rule cont2cont_LAM)
+  fix x :: 'a
+  have "cont (\<lambda>y. (x, y))"
+    by (rule cont_pair2)
+  with f have "cont (\<lambda>y. f (fst (x, y)) (snd (x, y)))"
+    by (rule cont2cont_app3)
+  thus "cont (\<lambda>y. f x y)"
+    by (simp only: fst_conv snd_conv)
+next
+  fix y :: 'b
+  have "cont (\<lambda>x. (x, y))"
+    by (rule cont_pair1)
+  with f have "cont (\<lambda>x. f (fst (x, y)) (snd (x, y)))"
+    by (rule cont2cont_app3)
+  thus "cont (\<lambda>x. f x y)"
+    by (simp only: fst_conv snd_conv)
+qed
+
+lemma cont2cont_LAM_discrete [cont2cont]:
+  "(\<And>y::'a::discrete_cpo. cont (\<lambda>x. f x y)) \<Longrightarrow> cont (\<lambda>x. \<Lambda> y. f x y)"
+by (simp add: cont2cont_LAM)
 
 lemmas cont_lemmas1 =
   cont_const cont_id cont_Rep_CFun2 cont2cont_Rep_CFun cont2cont_LAM
 
-use "Tools/cont_proc.ML";
-setup ContProc.setup;
-
-(*val cont_tac = (fn i => (resolve_tac cont_lemmas i));*)
-(*val cont_tacR = (fn i => (REPEAT (cont_tac i)));*)
-
 subsection {* Miscellaneous *}
 
 text {* Monotonicity of @{term Abs_CFun} *}
@@ -530,7 +557,8 @@
   monocontlub2cont [OF monofun_strictify2 contlub_strictify2, standard]
 
 lemma strictify_conv_if: "strictify\<cdot>f\<cdot>x = (if x = \<bottom> then \<bottom> else f\<cdot>x)"
-by (unfold strictify_def, simp add: cont_strictify1 cont_strictify2)
+  unfolding strictify_def
+  by (simp add: cont_strictify1 cont_strictify2 cont2cont_LAM)
 
 lemma strictify1 [simp]: "strictify\<cdot>f\<cdot>\<bottom> = \<bottom>"
 by (simp add: strictify_conv_if)
--- a/src/HOLCF/CompactBasis.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOLCF/CompactBasis.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -244,7 +244,7 @@
   assumes "ab_semigroup_idem_mult f"
   shows "fold_pd g f (PDPlus t u) = f (fold_pd g f t) (fold_pd g f u)"
 proof -
-  class_interpret ab_semigroup_idem_mult [f] by fact
+  interpret ab_semigroup_idem_mult f by fact
   show ?thesis unfolding fold_pd_def Rep_PDPlus by (simp add: image_Un fold1_Un2)
 qed
 
--- a/src/HOLCF/Cont.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOLCF/Cont.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -104,6 +104,8 @@
 apply simp
 done
 
+lemmas cont2monofunE = cont2mono [THEN monofunE]
+
 lemmas ch2ch_cont = cont2mono [THEN ch2ch_monofun]
 
 text {* right to left: @{prop "cont f \<Longrightarrow> monofun f \<and> contlub f"} *}
@@ -135,22 +137,82 @@
 apply (erule cpo_lubI)
 done
 
+subsection {* Continuity simproc *}
+
+ML {*
+structure Cont2ContData = NamedThmsFun
+  ( val name = "cont2cont" val description = "continuity intro rule" )
+*}
+
+setup {* Cont2ContData.setup *}
+
+text {*
+  Given the term @{term "cont f"}, the procedure tries to construct the
+  theorem @{term "cont f == True"}. If this theorem cannot be completely
+  solved by the introduction rules, then the procedure returns a
+  conditional rewrite rule with the unsolved subgoals as premises.
+*}
+
+setup {*
+let
+  fun solve_cont thy ss t =
+    let
+      val tr = instantiate' [] [SOME (cterm_of thy t)] Eq_TrueI;
+      val rules = Cont2ContData.get (Simplifier.the_context ss);
+      val tac = REPEAT_ALL_NEW (match_tac rules);
+    in Option.map fst (Seq.pull (tac 1 tr)) end
+
+  val proc =
+    Simplifier.simproc @{theory} "cont_proc" ["cont f"] solve_cont;
+in
+  Simplifier.map_simpset (fn ss => ss addsimprocs [proc])
+end
+*}
+
 subsection {* Continuity of basic functions *}
 
 text {* The identity function is continuous *}
 
-lemma cont_id: "cont (\<lambda>x. x)"
+lemma cont_id [cont2cont]: "cont (\<lambda>x. x)"
 apply (rule contI)
 apply (erule cpo_lubI)
 done
 
 text {* constant functions are continuous *}
 
-lemma cont_const: "cont (\<lambda>x. c)"
+lemma cont_const [cont2cont]: "cont (\<lambda>x. c)"
 apply (rule contI)
 apply (rule lub_const)
 done
 
+text {* application of functions is continuous *}
+
+lemma cont2cont_apply:
+  fixes f :: "'a::cpo \<Rightarrow> 'b::cpo \<Rightarrow> 'c::cpo" and t :: "'a \<Rightarrow> 'b"
+  assumes f1: "\<And>y. cont (\<lambda>x. f x y)"
+  assumes f2: "\<And>x. cont (\<lambda>y. f x y)"
+  assumes t: "cont (\<lambda>x. t x)"
+  shows "cont (\<lambda>x. (f x) (t x))"
+proof (rule monocontlub2cont [OF monofunI contlubI])
+  fix x y :: "'a" assume "x \<sqsubseteq> y"
+  then show "f x (t x) \<sqsubseteq> f y (t y)"
+    by (auto intro: cont2monofunE [OF f1]
+                    cont2monofunE [OF f2]
+                    cont2monofunE [OF t]
+                    trans_less)
+next
+  fix Y :: "nat \<Rightarrow> 'a" assume "chain Y"
+  then show "f (\<Squnion>i. Y i) (t (\<Squnion>i. Y i)) = (\<Squnion>i. f (Y i) (t (Y i)))"
+    by (simp only: cont2contlubE [OF t]  ch2ch_cont [OF t]
+                   cont2contlubE [OF f1] ch2ch_cont [OF f1]
+                   cont2contlubE [OF f2] ch2ch_cont [OF f2]
+                   diag_lub)
+qed
+
+lemma cont2cont_compose:
+  "\<lbrakk>cont c; cont (\<lambda>x. f x)\<rbrakk> \<Longrightarrow> cont (\<lambda>x. c (f x))"
+by (rule cont2cont_apply [OF cont_const])
+
 text {* if-then-else is continuous *}
 
 lemma cont_if [simp]:
--- a/src/HOLCF/ConvexPD.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOLCF/ConvexPD.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -296,9 +296,8 @@
 apply (simp add: PDPlus_absorb)
 done
 
-class_interpretation aci_convex_plus: ab_semigroup_idem_mult ["op +\<natural>"]
-  by unfold_locales
-    (rule convex_plus_assoc convex_plus_commute convex_plus_absorb)+
+interpretation aci_convex_plus!: ab_semigroup_idem_mult "op +\<natural>"
+  proof qed (rule convex_plus_assoc convex_plus_commute convex_plus_absorb)+
 
 lemma convex_plus_left_commute: "xs +\<natural> (ys +\<natural> zs) = ys +\<natural> (xs +\<natural> zs)"
 by (rule aci_convex_plus.mult_left_commute)
--- a/src/HOLCF/Cprod.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOLCF/Cprod.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -12,23 +12,6 @@
 
 subsection {* Type @{typ unit} is a pcpo *}
 
-instantiation unit :: sq_ord
-begin
-
-definition
-  less_unit_def [simp]: "x \<sqsubseteq> (y::unit) \<equiv> True"
-
-instance ..
-end
-
-instance unit :: discrete_cpo
-by intro_classes simp
-
-instance unit :: finite_po ..
-
-instance unit :: pcpo
-by intro_classes simp
-
 definition
   unit_when :: "'a \<rightarrow> unit \<rightarrow> 'a" where
   "unit_when = (\<Lambda> a _. a)"
@@ -39,165 +22,6 @@
 lemma unit_when [simp]: "unit_when\<cdot>a\<cdot>u = a"
 by (simp add: unit_when_def)
 
-
-subsection {* Product type is a partial order *}
-
-instantiation "*" :: (sq_ord, sq_ord) sq_ord
-begin
-
-definition
-  less_cprod_def: "(op \<sqsubseteq>) \<equiv> \<lambda>p1 p2. (fst p1 \<sqsubseteq> fst p2 \<and> snd p1 \<sqsubseteq> snd p2)"
-
-instance ..
-end
-
-instance "*" :: (po, po) po
-proof
-  fix x :: "'a \<times> 'b"
-  show "x \<sqsubseteq> x"
-    unfolding less_cprod_def by simp
-next
-  fix x y :: "'a \<times> 'b"
-  assume "x \<sqsubseteq> y" "y \<sqsubseteq> x" thus "x = y"
-    unfolding less_cprod_def Pair_fst_snd_eq
-    by (fast intro: antisym_less)
-next
-  fix x y z :: "'a \<times> 'b"
-  assume "x \<sqsubseteq> y" "y \<sqsubseteq> z" thus "x \<sqsubseteq> z"
-    unfolding less_cprod_def
-    by (fast intro: trans_less)
-qed
-
-subsection {* Monotonicity of @{text "(_,_)"}, @{term fst}, @{term snd} *}
-
-lemma prod_lessI: "\<lbrakk>fst p \<sqsubseteq> fst q; snd p \<sqsubseteq> snd q\<rbrakk> \<Longrightarrow> p \<sqsubseteq> q"
-unfolding less_cprod_def by simp
-
-lemma Pair_less_iff [simp]: "(a, b) \<sqsubseteq> (c, d) = (a \<sqsubseteq> c \<and> b \<sqsubseteq> d)"
-unfolding less_cprod_def by simp
-
-text {* Pair @{text "(_,_)"}  is monotone in both arguments *}
-
-lemma monofun_pair1: "monofun (\<lambda>x. (x, y))"
-by (simp add: monofun_def)
-
-lemma monofun_pair2: "monofun (\<lambda>y. (x, y))"
-by (simp add: monofun_def)
-
-lemma monofun_pair:
-  "\<lbrakk>x1 \<sqsubseteq> x2; y1 \<sqsubseteq> y2\<rbrakk> \<Longrightarrow> (x1, y1) \<sqsubseteq> (x2, y2)"
-by simp
-
-text {* @{term fst} and @{term snd} are monotone *}
-
-lemma monofun_fst: "monofun fst"
-by (simp add: monofun_def less_cprod_def)
-
-lemma monofun_snd: "monofun snd"
-by (simp add: monofun_def less_cprod_def)
-
-subsection {* Product type is a cpo *}
-
-lemma is_lub_Pair:
-  "\<lbrakk>range X <<| x; range Y <<| y\<rbrakk> \<Longrightarrow> range (\<lambda>i. (X i, Y i)) <<| (x, y)"
-apply (rule is_lubI [OF ub_rangeI])
-apply (simp add: less_cprod_def is_ub_lub)
-apply (frule ub2ub_monofun [OF monofun_fst])
-apply (drule ub2ub_monofun [OF monofun_snd])
-apply (simp add: less_cprod_def is_lub_lub)
-done
-
-lemma lub_cprod:
-  fixes S :: "nat \<Rightarrow> ('a::cpo \<times> 'b::cpo)"
-  assumes S: "chain S"
-  shows "range S <<| (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))"
-proof -
-  have "chain (\<lambda>i. fst (S i))"
-    using monofun_fst S by (rule ch2ch_monofun)
-  hence 1: "range (\<lambda>i. fst (S i)) <<| (\<Squnion>i. fst (S i))"
-    by (rule cpo_lubI)
-  have "chain (\<lambda>i. snd (S i))"
-    using monofun_snd S by (rule ch2ch_monofun)
-  hence 2: "range (\<lambda>i. snd (S i)) <<| (\<Squnion>i. snd (S i))"
-    by (rule cpo_lubI)
-  show "range S <<| (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))"
-    using is_lub_Pair [OF 1 2] by simp
-qed
-
-lemma thelub_cprod:
-  "chain (S::nat \<Rightarrow> 'a::cpo \<times> 'b::cpo)
-    \<Longrightarrow> (\<Squnion>i. S i) = (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))"
-by (rule lub_cprod [THEN thelubI])
-
-instance "*" :: (cpo, cpo) cpo
-proof
-  fix S :: "nat \<Rightarrow> ('a \<times> 'b)"
-  assume "chain S"
-  hence "range S <<| (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))"
-    by (rule lub_cprod)
-  thus "\<exists>x. range S <<| x" ..
-qed
-
-instance "*" :: (finite_po, finite_po) finite_po ..
-
-instance "*" :: (discrete_cpo, discrete_cpo) discrete_cpo
-proof
-  fix x y :: "'a \<times> 'b"
-  show "x \<sqsubseteq> y \<longleftrightarrow> x = y"
-    unfolding less_cprod_def Pair_fst_snd_eq
-    by simp
-qed
-
-subsection {* Product type is pointed *}
-
-lemma minimal_cprod: "(\<bottom>, \<bottom>) \<sqsubseteq> p"
-by (simp add: less_cprod_def)
-
-instance "*" :: (pcpo, pcpo) pcpo
-by intro_classes (fast intro: minimal_cprod)
-
-lemma inst_cprod_pcpo: "\<bottom> = (\<bottom>, \<bottom>)"
-by (rule minimal_cprod [THEN UU_I, symmetric])
-
-
-subsection {* Continuity of @{text "(_,_)"}, @{term fst}, @{term snd} *}
-
-lemma cont_pair1: "cont (\<lambda>x. (x, y))"
-apply (rule contI)
-apply (rule is_lub_Pair)
-apply (erule cpo_lubI)
-apply (rule lub_const)
-done
-
-lemma cont_pair2: "cont (\<lambda>y. (x, y))"
-apply (rule contI)
-apply (rule is_lub_Pair)
-apply (rule lub_const)
-apply (erule cpo_lubI)
-done
-
-lemma contlub_fst: "contlub fst"
-apply (rule contlubI)
-apply (simp add: thelub_cprod)
-done
-
-lemma contlub_snd: "contlub snd"
-apply (rule contlubI)
-apply (simp add: thelub_cprod)
-done
-
-lemma cont_fst: "cont fst"
-apply (rule monocontlub2cont)
-apply (rule monofun_fst)
-apply (rule contlub_fst)
-done
-
-lemma cont_snd: "cont snd"
-apply (rule monocontlub2cont)
-apply (rule monofun_snd)
-apply (rule contlub_snd)
-done
-
 subsection {* Continuous versions of constants *}
 
 definition
@@ -245,10 +69,10 @@
 by (simp add: cpair_eq_pair)
 
 lemma cpair_less [iff]: "(<a, b> \<sqsubseteq> <a', b'>) = (a \<sqsubseteq> a' \<and> b \<sqsubseteq> b')"
-by (simp add: cpair_eq_pair less_cprod_def)
+by (simp add: cpair_eq_pair)
 
 lemma cpair_defined_iff [iff]: "(<x, y> = \<bottom>) = (x = \<bottom> \<and> y = \<bottom>)"
-by (simp add: inst_cprod_pcpo cpair_eq_pair)
+by (simp add: cpair_eq_pair)
 
 lemma cpair_strict [simp]: "\<langle>\<bottom>, \<bottom>\<rangle> = \<bottom>"
 by simp
@@ -273,10 +97,10 @@
 by (simp add: cpair_eq_pair csnd_def cont_snd)
 
 lemma cfst_strict [simp]: "cfst\<cdot>\<bottom> = \<bottom>"
-unfolding inst_cprod_pcpo2 by (rule cfst_cpair)
+by (simp add: cfst_def)
 
 lemma csnd_strict [simp]: "csnd\<cdot>\<bottom> = \<bottom>"
-unfolding inst_cprod_pcpo2 by (rule csnd_cpair)
+by (simp add: csnd_def)
 
 lemma cpair_cfst_csnd: "\<langle>cfst\<cdot>p, csnd\<cdot>p\<rangle> = p"
 by (cases p rule: cprodE, simp)
@@ -302,19 +126,10 @@
 by (rule compactI, simp add: csnd_less_iff)
 
 lemma compact_cpair: "\<lbrakk>compact x; compact y\<rbrakk> \<Longrightarrow> compact <x, y>"
-by (rule compactI, simp add: less_cprod)
+by (simp add: cpair_eq_pair)
 
 lemma compact_cpair_iff [simp]: "compact <x, y> = (compact x \<and> compact y)"
-apply (safe intro!: compact_cpair)
-apply (drule compact_cfst, simp)
-apply (drule compact_csnd, simp)
-done
-
-instance "*" :: (chfin, chfin) chfin
-apply intro_classes
-apply (erule compact_imp_max_in_chain)
-apply (rule_tac p="\<Squnion>i. Y i" in cprodE, simp)
-done
+by (simp add: cpair_eq_pair)
 
 lemma lub_cprod2: 
   "chain S \<Longrightarrow> range S <<| <\<Squnion>i. cfst\<cdot>(S i), \<Squnion>i. csnd\<cdot>(S i)>"
--- a/src/HOLCF/Dsum.thy	Mon Jan 19 20:37:08 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,251 +0,0 @@
-(*  Title:      HOLCF/Dsum.thy
-    Author:     Brian Huffman
-*)
-
-header {* The cpo of disjoint sums *}
-
-theory Dsum
-imports Bifinite
-begin
-
-subsection {* Ordering on type @{typ "'a + 'b"} *}
-
-instantiation "+" :: (sq_ord, sq_ord) sq_ord
-begin
-
-definition
-  less_sum_def: "x \<sqsubseteq> y \<equiv> case x of
-         Inl a \<Rightarrow> (case y of Inl b \<Rightarrow> a \<sqsubseteq> b | Inr b \<Rightarrow> False) |
-         Inr a \<Rightarrow> (case y of Inl b \<Rightarrow> False | Inr b \<Rightarrow> a \<sqsubseteq> b)"
-
-instance ..
-end
-
-lemma Inl_less_iff [simp]: "Inl x \<sqsubseteq> Inl y = x \<sqsubseteq> y"
-unfolding less_sum_def by simp
-
-lemma Inr_less_iff [simp]: "Inr x \<sqsubseteq> Inr y = x \<sqsubseteq> y"
-unfolding less_sum_def by simp
-
-lemma Inl_less_Inr [simp]: "\<not> Inl x \<sqsubseteq> Inr y"
-unfolding less_sum_def by simp
-
-lemma Inr_less_Inl [simp]: "\<not> Inr x \<sqsubseteq> Inl y"
-unfolding less_sum_def by simp
-
-lemma Inl_mono: "x \<sqsubseteq> y \<Longrightarrow> Inl x \<sqsubseteq> Inl y"
-by simp
-
-lemma Inr_mono: "x \<sqsubseteq> y \<Longrightarrow> Inr x \<sqsubseteq> Inr y"
-by simp
-
-lemma Inl_lessE: "\<lbrakk>Inl a \<sqsubseteq> x; \<And>b. \<lbrakk>x = Inl b; a \<sqsubseteq> b\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
-by (cases x, simp_all)
-
-lemma Inr_lessE: "\<lbrakk>Inr a \<sqsubseteq> x; \<And>b. \<lbrakk>x = Inr b; a \<sqsubseteq> b\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
-by (cases x, simp_all)
-
-lemmas sum_less_elims = Inl_lessE Inr_lessE
-
-lemma sum_less_cases:
-  "\<lbrakk>x \<sqsubseteq> y;
-    \<And>a b. \<lbrakk>x = Inl a; y = Inl b; a \<sqsubseteq> b\<rbrakk> \<Longrightarrow> R;
-    \<And>a b. \<lbrakk>x = Inr a; y = Inr b; a \<sqsubseteq> b\<rbrakk> \<Longrightarrow> R\<rbrakk>
-      \<Longrightarrow> R"
-by (cases x, safe elim!: sum_less_elims, auto)
-
-subsection {* Sum type is a complete partial order *}
-
-instance "+" :: (po, po) po
-proof
-  fix x :: "'a + 'b"
-  show "x \<sqsubseteq> x"
-    by (induct x, simp_all)
-next
-  fix x y :: "'a + 'b"
-  assume "x \<sqsubseteq> y" and "y \<sqsubseteq> x" thus "x = y"
-    by (induct x, auto elim!: sum_less_elims intro: antisym_less)
-next
-  fix x y z :: "'a + 'b"
-  assume "x \<sqsubseteq> y" and "y \<sqsubseteq> z" thus "x \<sqsubseteq> z"
-    by (induct x, auto elim!: sum_less_elims intro: trans_less)
-qed
-
-lemma monofun_inv_Inl: "monofun (\<lambda>p. THE a. p = Inl a)"
-by (rule monofunI, erule sum_less_cases, simp_all)
-
-lemma monofun_inv_Inr: "monofun (\<lambda>p. THE b. p = Inr b)"
-by (rule monofunI, erule sum_less_cases, simp_all)
-
-lemma sum_chain_cases:
-  assumes Y: "chain Y"
-  assumes A: "\<And>A. \<lbrakk>chain A; Y = (\<lambda>i. Inl (A i))\<rbrakk> \<Longrightarrow> R"
-  assumes B: "\<And>B. \<lbrakk>chain B; Y = (\<lambda>i. Inr (B i))\<rbrakk> \<Longrightarrow> R"
-  shows "R"
- apply (cases "Y 0")
-  apply (rule A)
-   apply (rule ch2ch_monofun [OF monofun_inv_Inl Y])
-  apply (rule ext)
-  apply (cut_tac j=i in chain_mono [OF Y le0], simp)
-  apply (erule Inl_lessE, simp)
- apply (rule B)
-  apply (rule ch2ch_monofun [OF monofun_inv_Inr Y])
- apply (rule ext)
- apply (cut_tac j=i in chain_mono [OF Y le0], simp)
- apply (erule Inr_lessE, simp)
-done
-
-lemma is_lub_Inl: "range S <<| x \<Longrightarrow> range (\<lambda>i. Inl (S i)) <<| Inl x"
- apply (rule is_lubI)
-  apply (rule ub_rangeI)
-  apply (simp add: is_ub_lub)
- apply (frule ub_rangeD [where i=arbitrary])
- apply (erule Inl_lessE, simp)
- apply (erule is_lub_lub)
- apply (rule ub_rangeI)
- apply (drule ub_rangeD, simp)
-done
-
-lemma is_lub_Inr: "range S <<| x \<Longrightarrow> range (\<lambda>i. Inr (S i)) <<| Inr x"
- apply (rule is_lubI)
-  apply (rule ub_rangeI)
-  apply (simp add: is_ub_lub)
- apply (frule ub_rangeD [where i=arbitrary])
- apply (erule Inr_lessE, simp)
- apply (erule is_lub_lub)
- apply (rule ub_rangeI)
- apply (drule ub_rangeD, simp)
-done
-
-instance "+" :: (cpo, cpo) cpo
- apply intro_classes
- apply (erule sum_chain_cases, safe)
-  apply (rule exI)
-  apply (rule is_lub_Inl)
-  apply (erule cpo_lubI)
- apply (rule exI)
- apply (rule is_lub_Inr)
- apply (erule cpo_lubI)
-done
-
-subsection {* Continuity of @{term Inl}, @{term Inr}, @{term sum_case} *}
-
-lemma cont2cont_Inl [simp]: "cont f \<Longrightarrow> cont (\<lambda>x. Inl (f x))"
-by (fast intro: contI is_lub_Inl elim: contE)
-
-lemma cont2cont_Inr [simp]: "cont f \<Longrightarrow> cont (\<lambda>x. Inr (f x))"
-by (fast intro: contI is_lub_Inr elim: contE)
-
-lemma cont_Inl: "cont Inl"
-by (rule cont2cont_Inl [OF cont_id])
-
-lemma cont_Inr: "cont Inr"
-by (rule cont2cont_Inr [OF cont_id])
-
-lemmas ch2ch_Inl [simp] = ch2ch_cont [OF cont_Inl]
-lemmas ch2ch_Inr [simp] = ch2ch_cont [OF cont_Inr]
-
-lemmas lub_Inl = cont2contlubE [OF cont_Inl, symmetric]
-lemmas lub_Inr = cont2contlubE [OF cont_Inr, symmetric]
-
-lemma cont_sum_case1:
-  assumes f: "\<And>a. cont (\<lambda>x. f x a)"
-  assumes g: "\<And>b. cont (\<lambda>x. g x b)"
-  shows "cont (\<lambda>x. case y of Inl a \<Rightarrow> f x a | Inr b \<Rightarrow> g x b)"
-by (induct y, simp add: f, simp add: g)
-
-lemma cont_sum_case2: "\<lbrakk>cont f; cont g\<rbrakk> \<Longrightarrow> cont (sum_case f g)"
-apply (rule contI)
-apply (erule sum_chain_cases)
-apply (simp add: cont2contlubE [OF cont_Inl, symmetric] contE)
-apply (simp add: cont2contlubE [OF cont_Inr, symmetric] contE)
-done
-
-lemma cont2cont_sum_case [simp]:
-  assumes f1: "\<And>a. cont (\<lambda>x. f x a)" and f2: "\<And>x. cont (\<lambda>a. f x a)"
-  assumes g1: "\<And>b. cont (\<lambda>x. g x b)" and g2: "\<And>x. cont (\<lambda>b. g x b)"
-  assumes h: "cont (\<lambda>x. h x)"
-  shows "cont (\<lambda>x. case h x of Inl a \<Rightarrow> f x a | Inr b \<Rightarrow> g x b)"
-apply (rule cont2cont_app2 [OF cont2cont_lambda _ h])
-apply (rule cont_sum_case1 [OF f1 g1])
-apply (rule cont_sum_case2 [OF f2 g2])
-done
-
-subsection {* Compactness and chain-finiteness *}
-
-lemma compact_Inl: "compact a \<Longrightarrow> compact (Inl a)"
-apply (rule compactI2)
-apply (erule sum_chain_cases, safe)
-apply (simp add: lub_Inl)
-apply (erule (2) compactD2)
-apply (simp add: lub_Inr)
-done
-
-lemma compact_Inr: "compact a \<Longrightarrow> compact (Inr a)"
-apply (rule compactI2)
-apply (erule sum_chain_cases, safe)
-apply (simp add: lub_Inl)
-apply (simp add: lub_Inr)
-apply (erule (2) compactD2)
-done
-
-lemma compact_Inl_rev: "compact (Inl a) \<Longrightarrow> compact a"
-unfolding compact_def
-by (drule adm_subst [OF cont_Inl], simp)
-
-lemma compact_Inr_rev: "compact (Inr a) \<Longrightarrow> compact a"
-unfolding compact_def
-by (drule adm_subst [OF cont_Inr], simp)
-
-lemma compact_Inl_iff [simp]: "compact (Inl a) = compact a"
-by (safe elim!: compact_Inl compact_Inl_rev)
-
-lemma compact_Inr_iff [simp]: "compact (Inr a) = compact a"
-by (safe elim!: compact_Inr compact_Inr_rev)
-
-instance "+" :: (chfin, chfin) chfin
-apply intro_classes
-apply (erule compact_imp_max_in_chain)
-apply (case_tac "\<Squnion>i. Y i", simp_all)
-done
-
-instance "+" :: (finite_po, finite_po) finite_po ..
-
-instance "+" :: (discrete_cpo, discrete_cpo) discrete_cpo
-by intro_classes (simp add: less_sum_def split: sum.split)
-
-subsection {* Sum type is a bifinite domain *}
-
-instantiation "+" :: (profinite, profinite) profinite
-begin
-
-definition
-  approx_sum_def: "approx =
-    (\<lambda>n. \<Lambda> x. case x of Inl a \<Rightarrow> Inl (approx n\<cdot>a) | Inr b \<Rightarrow> Inr (approx n\<cdot>b))"
-
-lemma approx_Inl [simp]: "approx n\<cdot>(Inl x) = Inl (approx n\<cdot>x)"
-  unfolding approx_sum_def by simp
-
-lemma approx_Inr [simp]: "approx n\<cdot>(Inr x) = Inr (approx n\<cdot>x)"
-  unfolding approx_sum_def by simp
-
-instance proof
-  fix i :: nat and x :: "'a + 'b"
-  show "chain (approx :: nat \<Rightarrow> 'a + 'b \<rightarrow> 'a + 'b)"
-    unfolding approx_sum_def
-    by (rule ch2ch_LAM, case_tac x, simp_all)
-  show "(\<Squnion>i. approx i\<cdot>x) = x"
-    by (induct x, simp_all add: lub_Inl lub_Inr)
-  show "approx i\<cdot>(approx i\<cdot>x) = approx i\<cdot>x"
-    by (induct x, simp_all)
-  have "{x::'a + 'b. approx i\<cdot>x = x} \<subseteq>
-        {x::'a. approx i\<cdot>x = x} <+> {x::'b. approx i\<cdot>x = x}"
-    by (rule subsetI, case_tac x, simp_all add: InlI InrI)
-  thus "finite {x::'a + 'b. approx i\<cdot>x = x}"
-    by (rule finite_subset,
-        intro finite_Plus finite_fixes_approx)
-qed
-
-end
-
-end
--- a/src/HOLCF/Fixrec.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOLCF/Fixrec.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -55,6 +55,7 @@
   "maybe_when\<cdot>f\<cdot>r\<cdot>fail = f"
   "maybe_when\<cdot>f\<cdot>r\<cdot>(return\<cdot>x) = r\<cdot>x"
 by (simp_all add: return_def fail_def maybe_when_def cont_Rep_maybe
+                  cont2cont_LAM
                   cont_Abs_maybe Abs_maybe_inverse Rep_maybe_strict)
 
 translations
@@ -216,19 +217,19 @@
 
 syntax
   "_pat" :: "'a"
-  "_var" :: "'a"
+  "_variable" :: "'a"
   "_noargs" :: "'a"
 
 translations
-  "_Case1 p r" => "CONST branch (_pat p)\<cdot>(_var p r)"
-  "_var (_args x y) r" => "CONST csplit\<cdot>(_var x (_var y r))"
-  "_var _noargs r" => "CONST unit_when\<cdot>r"
+  "_Case1 p r" => "CONST branch (_pat p)\<cdot>(_variable p r)"
+  "_variable (_args x y) r" => "CONST csplit\<cdot>(_variable x (_variable y r))"
+  "_variable _noargs r" => "CONST unit_when\<cdot>r"
 
 parse_translation {*
 (* rewrites (_pat x) => (return) *)
-(* rewrites (_var x t) => (Abs_CFun (%x. t)) *)
+(* rewrites (_variable x t) => (Abs_CFun (%x. t)) *)
   [("_pat", K (Syntax.const "Fixrec.return")),
-   mk_binder_tr ("_var", "Abs_CFun")];
+   mk_binder_tr ("_variable", "Abs_CFun")];
 *}
 
 text {* Printing Case expressions *}
@@ -250,7 +251,7 @@
             val abs = case t of Abs abs => abs
                 | _ => ("x", dummyT, incr_boundvars 1 t $ Bound 0);
             val (x, t') = atomic_abs_tr' abs;
-          in (Syntax.const "_var" $ x, t') end
+          in (Syntax.const "_variable" $ x, t') end
     |   dest_LAM _ = raise Match; (* too few vars: abort translation *)
 
     fun Case1_tr' [Const(@{const_syntax branch},_) $ p, r] =
@@ -261,7 +262,7 @@
 *}
 
 translations
-  "x" <= "_match Fixrec.return (_var x)"
+  "x" <= "_match Fixrec.return (_variable x)"
 
 
 subsection {* Pattern combinators for data constructors *}
@@ -320,18 +321,18 @@
 
 text {* Parse translations (variables) *}
 translations
-  "_var (XCONST cpair\<cdot>x\<cdot>y) r" => "_var (_args x y) r"
-  "_var (XCONST spair\<cdot>x\<cdot>y) r" => "_var (_args x y) r"
-  "_var (XCONST sinl\<cdot>x) r" => "_var x r"
-  "_var (XCONST sinr\<cdot>x) r" => "_var x r"
-  "_var (XCONST up\<cdot>x) r" => "_var x r"
-  "_var (XCONST TT) r" => "_var _noargs r"
-  "_var (XCONST FF) r" => "_var _noargs r"
-  "_var (XCONST ONE) r" => "_var _noargs r"
+  "_variable (XCONST cpair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r"
+  "_variable (XCONST spair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r"
+  "_variable (XCONST sinl\<cdot>x) r" => "_variable x r"
+  "_variable (XCONST sinr\<cdot>x) r" => "_variable x r"
+  "_variable (XCONST up\<cdot>x) r" => "_variable x r"
+  "_variable (XCONST TT) r" => "_variable _noargs r"
+  "_variable (XCONST FF) r" => "_variable _noargs r"
+  "_variable (XCONST ONE) r" => "_variable _noargs r"
 
 translations
-  "_var (CONST cpair\<cdot>x\<cdot>y) r" => "_var (_args x y) r"
-  "_var (CONST spair\<cdot>x\<cdot>y) r" => "_var (_args x y) r"
+  "_variable (CONST cpair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r"
+  "_variable (CONST spair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r"
 
 text {* Print translations *}
 translations
@@ -437,14 +438,14 @@
 
 text {* Parse translations (variables) *}
 translations
-  "_var _ r" => "_var _noargs r"
-  "_var (_as_pat x y) r" => "_var (_args x y) r"
-  "_var (_lazy_pat x) r" => "_var x r"
+  "_variable _ r" => "_variable _noargs r"
+  "_variable (_as_pat x y) r" => "_variable (_args x y) r"
+  "_variable (_lazy_pat x) r" => "_variable x r"
 
 text {* Print translations *}
 translations
   "_" <= "_match (CONST wild_pat) _noargs"
-  "_as_pat x (_match p v)" <= "_match (CONST as_pat p) (_args (_var x) v)"
+  "_as_pat x (_match p v)" <= "_match (CONST as_pat p) (_args (_variable x) v)"
   "_lazy_pat (_match p v)" <= "_match (CONST lazy_pat p) v"
 
 text {* Lazy patterns in lambda abstractions *}
--- a/src/HOLCF/HOLCF.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOLCF/HOLCF.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -6,17 +6,17 @@
 
 theory HOLCF
 imports
-  Domain ConvexPD Algebraic Universal Dsum Main
+  Domain ConvexPD Algebraic Universal Sum_Cpo Main
 uses
   "holcf_logic.ML"
   "Tools/cont_consts.ML"
+  "Tools/cont_proc.ML"
   "Tools/domain/domain_library.ML"
   "Tools/domain/domain_syntax.ML"
   "Tools/domain/domain_axioms.ML"
   "Tools/domain/domain_theorems.ML"
   "Tools/domain/domain_extender.ML"
   "Tools/adm_tac.ML"
-
 begin
 
 defaultsort pcpo
@@ -24,7 +24,7 @@
 declaration {* fn _ =>
   Simplifier.map_ss (fn simpset => simpset addSolver
     (mk_solver' "adm_tac" (fn ss =>
-      adm_tac (cut_facts_tac (Simplifier.prems_of_ss ss) THEN' cont_tacRs ss))));
+      Adm.adm_tac (cut_facts_tac (Simplifier.prems_of_ss ss) THEN' cont_tacRs ss))));
 *}
 
 end
--- a/src/HOLCF/IsaMakefile	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOLCF/IsaMakefile	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,3 @@
-#
-# $Id$
 #
 # IsaMakefile for HOLCF
 #
@@ -41,7 +39,6 @@
   Discrete.thy \
   Deflation.thy \
   Domain.thy \
-  Dsum.thy \
   Eventual.thy \
   Ffun.thy \
   Fixrec.thy \
@@ -54,8 +51,10 @@
   Pcpodef.thy \
   Pcpo.thy \
   Porder.thy \
+  Product_Cpo.thy \
   Sprod.thy \
   Ssum.thy \
+  Sum_Cpo.thy \
   Tr.thy \
   Universal.thy \
   UpperPD.thy \
--- a/src/HOLCF/LowerPD.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOLCF/LowerPD.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -250,9 +250,8 @@
 apply (simp add: PDPlus_absorb)
 done
 
-class_interpretation aci_lower_plus: ab_semigroup_idem_mult ["op +\<flat>"]
-  by unfold_locales
-    (rule lower_plus_assoc lower_plus_commute lower_plus_absorb)+
+interpretation aci_lower_plus!: ab_semigroup_idem_mult "op +\<flat>"
+  proof qed (rule lower_plus_assoc lower_plus_commute lower_plus_absorb)+
 
 lemma lower_plus_left_commute: "xs +\<flat> (ys +\<flat> zs) = ys +\<flat> (xs +\<flat> zs)"
 by (rule aci_lower_plus.mult_left_commute)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOLCF/Product_Cpo.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -0,0 +1,247 @@
+(*  Title:      HOLCF/Product_Cpo.thy
+    Author:     Franz Regensburger
+*)
+
+header {* The cpo of cartesian products *}
+
+theory Product_Cpo
+imports Adm
+begin
+
+defaultsort cpo
+
+subsection {* Type @{typ unit} is a pcpo *}
+
+instantiation unit :: sq_ord
+begin
+
+definition
+  less_unit_def [simp]: "x \<sqsubseteq> (y::unit) \<equiv> True"
+
+instance ..
+end
+
+instance unit :: discrete_cpo
+by intro_classes simp
+
+instance unit :: finite_po ..
+
+instance unit :: pcpo
+by intro_classes simp
+
+
+subsection {* Product type is a partial order *}
+
+instantiation "*" :: (sq_ord, sq_ord) sq_ord
+begin
+
+definition
+  less_cprod_def: "(op \<sqsubseteq>) \<equiv> \<lambda>p1 p2. (fst p1 \<sqsubseteq> fst p2 \<and> snd p1 \<sqsubseteq> snd p2)"
+
+instance ..
+end
+
+instance "*" :: (po, po) po
+proof
+  fix x :: "'a \<times> 'b"
+  show "x \<sqsubseteq> x"
+    unfolding less_cprod_def by simp
+next
+  fix x y :: "'a \<times> 'b"
+  assume "x \<sqsubseteq> y" "y \<sqsubseteq> x" thus "x = y"
+    unfolding less_cprod_def Pair_fst_snd_eq
+    by (fast intro: antisym_less)
+next
+  fix x y z :: "'a \<times> 'b"
+  assume "x \<sqsubseteq> y" "y \<sqsubseteq> z" thus "x \<sqsubseteq> z"
+    unfolding less_cprod_def
+    by (fast intro: trans_less)
+qed
+
+subsection {* Monotonicity of @{text "(_,_)"}, @{term fst}, @{term snd} *}
+
+lemma prod_lessI: "\<lbrakk>fst p \<sqsubseteq> fst q; snd p \<sqsubseteq> snd q\<rbrakk> \<Longrightarrow> p \<sqsubseteq> q"
+unfolding less_cprod_def by simp
+
+lemma Pair_less_iff [simp]: "(a, b) \<sqsubseteq> (c, d) \<longleftrightarrow> a \<sqsubseteq> c \<and> b \<sqsubseteq> d"
+unfolding less_cprod_def by simp
+
+text {* Pair @{text "(_,_)"}  is monotone in both arguments *}
+
+lemma monofun_pair1: "monofun (\<lambda>x. (x, y))"
+by (simp add: monofun_def)
+
+lemma monofun_pair2: "monofun (\<lambda>y. (x, y))"
+by (simp add: monofun_def)
+
+lemma monofun_pair:
+  "\<lbrakk>x1 \<sqsubseteq> x2; y1 \<sqsubseteq> y2\<rbrakk> \<Longrightarrow> (x1, y1) \<sqsubseteq> (x2, y2)"
+by simp
+
+text {* @{term fst} and @{term snd} are monotone *}
+
+lemma monofun_fst: "monofun fst"
+by (simp add: monofun_def less_cprod_def)
+
+lemma monofun_snd: "monofun snd"
+by (simp add: monofun_def less_cprod_def)
+
+subsection {* Product type is a cpo *}
+
+lemma is_lub_Pair:
+  "\<lbrakk>range X <<| x; range Y <<| y\<rbrakk> \<Longrightarrow> range (\<lambda>i. (X i, Y i)) <<| (x, y)"
+apply (rule is_lubI [OF ub_rangeI])
+apply (simp add: less_cprod_def is_ub_lub)
+apply (frule ub2ub_monofun [OF monofun_fst])
+apply (drule ub2ub_monofun [OF monofun_snd])
+apply (simp add: less_cprod_def is_lub_lub)
+done
+
+lemma lub_cprod:
+  fixes S :: "nat \<Rightarrow> ('a::cpo \<times> 'b::cpo)"
+  assumes S: "chain S"
+  shows "range S <<| (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))"
+proof -
+  have "chain (\<lambda>i. fst (S i))"
+    using monofun_fst S by (rule ch2ch_monofun)
+  hence 1: "range (\<lambda>i. fst (S i)) <<| (\<Squnion>i. fst (S i))"
+    by (rule cpo_lubI)
+  have "chain (\<lambda>i. snd (S i))"
+    using monofun_snd S by (rule ch2ch_monofun)
+  hence 2: "range (\<lambda>i. snd (S i)) <<| (\<Squnion>i. snd (S i))"
+    by (rule cpo_lubI)
+  show "range S <<| (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))"
+    using is_lub_Pair [OF 1 2] by simp
+qed
+
+lemma thelub_cprod:
+  "chain (S::nat \<Rightarrow> 'a::cpo \<times> 'b::cpo)
+    \<Longrightarrow> (\<Squnion>i. S i) = (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))"
+by (rule lub_cprod [THEN thelubI])
+
+instance "*" :: (cpo, cpo) cpo
+proof
+  fix S :: "nat \<Rightarrow> ('a \<times> 'b)"
+  assume "chain S"
+  hence "range S <<| (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))"
+    by (rule lub_cprod)
+  thus "\<exists>x. range S <<| x" ..
+qed
+
+instance "*" :: (finite_po, finite_po) finite_po ..
+
+instance "*" :: (discrete_cpo, discrete_cpo) discrete_cpo
+proof
+  fix x y :: "'a \<times> 'b"
+  show "x \<sqsubseteq> y \<longleftrightarrow> x = y"
+    unfolding less_cprod_def Pair_fst_snd_eq
+    by simp
+qed
+
+subsection {* Product type is pointed *}
+
+lemma minimal_cprod: "(\<bottom>, \<bottom>) \<sqsubseteq> p"
+by (simp add: less_cprod_def)
+
+instance "*" :: (pcpo, pcpo) pcpo
+by intro_classes (fast intro: minimal_cprod)
+
+lemma inst_cprod_pcpo: "\<bottom> = (\<bottom>, \<bottom>)"
+by (rule minimal_cprod [THEN UU_I, symmetric])
+
+lemma Pair_defined_iff [simp]: "(x, y) = \<bottom> \<longleftrightarrow> x = \<bottom> \<and> y = \<bottom>"
+unfolding inst_cprod_pcpo by simp
+
+lemma fst_strict [simp]: "fst \<bottom> = \<bottom>"
+unfolding inst_cprod_pcpo by (rule fst_conv)
+
+lemma csnd_strict [simp]: "snd \<bottom> = \<bottom>"
+unfolding inst_cprod_pcpo by (rule snd_conv)
+
+lemma Pair_strict [simp]: "(\<bottom>, \<bottom>) = \<bottom>"
+by simp
+
+lemma split_strict [simp]: "split f \<bottom> = f \<bottom> \<bottom>"
+unfolding split_def by simp
+
+subsection {* Continuity of @{text "(_,_)"}, @{term fst}, @{term snd} *}
+
+lemma cont_pair1: "cont (\<lambda>x. (x, y))"
+apply (rule contI)
+apply (rule is_lub_Pair)
+apply (erule cpo_lubI)
+apply (rule lub_const)
+done
+
+lemma cont_pair2: "cont (\<lambda>y. (x, y))"
+apply (rule contI)
+apply (rule is_lub_Pair)
+apply (rule lub_const)
+apply (erule cpo_lubI)
+done
+
+lemma contlub_fst: "contlub fst"
+apply (rule contlubI)
+apply (simp add: thelub_cprod)
+done
+
+lemma contlub_snd: "contlub snd"
+apply (rule contlubI)
+apply (simp add: thelub_cprod)
+done
+
+lemma cont_fst: "cont fst"
+apply (rule monocontlub2cont)
+apply (rule monofun_fst)
+apply (rule contlub_fst)
+done
+
+lemma cont_snd: "cont snd"
+apply (rule monocontlub2cont)
+apply (rule monofun_snd)
+apply (rule contlub_snd)
+done
+
+lemma cont2cont_Pair [cont2cont]:
+  assumes f: "cont (\<lambda>x. f x)"
+  assumes g: "cont (\<lambda>x. g x)"
+  shows "cont (\<lambda>x. (f x, g x))"
+apply (rule cont2cont_apply [OF _ cont_pair1 f])
+apply (rule cont2cont_apply [OF _ cont_pair2 g])
+apply (rule cont_const)
+done
+
+lemmas cont2cont_fst [cont2cont] = cont2cont_compose [OF cont_fst]
+
+lemmas cont2cont_snd [cont2cont] = cont2cont_compose [OF cont_snd]
+
+subsection {* Compactness and chain-finiteness *}
+
+lemma fst_less_iff: "fst (x::'a \<times> 'b) \<sqsubseteq> y \<longleftrightarrow> x \<sqsubseteq> (y, snd x)"
+unfolding less_cprod_def by simp
+
+lemma snd_less_iff: "snd (x::'a \<times> 'b) \<sqsubseteq> y = x \<sqsubseteq> (fst x, y)"
+unfolding less_cprod_def by simp
+
+lemma compact_fst: "compact x \<Longrightarrow> compact (fst x)"
+by (rule compactI, simp add: fst_less_iff)
+
+lemma compact_snd: "compact x \<Longrightarrow> compact (snd x)"
+by (rule compactI, simp add: snd_less_iff)
+
+lemma compact_Pair: "\<lbrakk>compact x; compact y\<rbrakk> \<Longrightarrow> compact (x, y)"
+by (rule compactI, simp add: less_cprod_def)
+
+lemma compact_Pair_iff [simp]: "compact (x, y) \<longleftrightarrow> compact x \<and> compact y"
+apply (safe intro!: compact_Pair)
+apply (drule compact_fst, simp)
+apply (drule compact_snd, simp)
+done
+
+instance "*" :: (chfin, chfin) chfin
+apply intro_classes
+apply (erule compact_imp_max_in_chain)
+apply (case_tac "\<Squnion>i. Y i", simp)
+done
+
+end
--- a/src/HOLCF/Ssum.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOLCF/Ssum.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -188,7 +188,7 @@
 
 lemma beta_sscase:
   "sscase\<cdot>f\<cdot>g\<cdot>s = (\<Lambda><t, x, y>. If t then f\<cdot>x else g\<cdot>y fi)\<cdot>(Rep_Ssum s)"
-unfolding sscase_def by (simp add: cont_Rep_Ssum)
+unfolding sscase_def by (simp add: cont_Rep_Ssum cont2cont_LAM)
 
 lemma sscase1 [simp]: "sscase\<cdot>f\<cdot>g\<cdot>\<bottom> = \<bottom>"
 unfolding beta_sscase by (simp add: Rep_Ssum_strict)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOLCF/Sum_Cpo.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -0,0 +1,251 @@
+(*  Title:      HOLCF/Sum_Cpo.thy
+    Author:     Brian Huffman
+*)
+
+header {* The cpo of disjoint sums *}
+
+theory Sum_Cpo
+imports Bifinite
+begin
+
+subsection {* Ordering on type @{typ "'a + 'b"} *}
+
+instantiation "+" :: (sq_ord, sq_ord) sq_ord
+begin
+
+definition
+  less_sum_def: "x \<sqsubseteq> y \<equiv> case x of
+         Inl a \<Rightarrow> (case y of Inl b \<Rightarrow> a \<sqsubseteq> b | Inr b \<Rightarrow> False) |
+         Inr a \<Rightarrow> (case y of Inl b \<Rightarrow> False | Inr b \<Rightarrow> a \<sqsubseteq> b)"
+
+instance ..
+end
+
+lemma Inl_less_iff [simp]: "Inl x \<sqsubseteq> Inl y = x \<sqsubseteq> y"
+unfolding less_sum_def by simp
+
+lemma Inr_less_iff [simp]: "Inr x \<sqsubseteq> Inr y = x \<sqsubseteq> y"
+unfolding less_sum_def by simp
+
+lemma Inl_less_Inr [simp]: "\<not> Inl x \<sqsubseteq> Inr y"
+unfolding less_sum_def by simp
+
+lemma Inr_less_Inl [simp]: "\<not> Inr x \<sqsubseteq> Inl y"
+unfolding less_sum_def by simp
+
+lemma Inl_mono: "x \<sqsubseteq> y \<Longrightarrow> Inl x \<sqsubseteq> Inl y"
+by simp
+
+lemma Inr_mono: "x \<sqsubseteq> y \<Longrightarrow> Inr x \<sqsubseteq> Inr y"
+by simp
+
+lemma Inl_lessE: "\<lbrakk>Inl a \<sqsubseteq> x; \<And>b. \<lbrakk>x = Inl b; a \<sqsubseteq> b\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
+by (cases x, simp_all)
+
+lemma Inr_lessE: "\<lbrakk>Inr a \<sqsubseteq> x; \<And>b. \<lbrakk>x = Inr b; a \<sqsubseteq> b\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
+by (cases x, simp_all)
+
+lemmas sum_less_elims = Inl_lessE Inr_lessE
+
+lemma sum_less_cases:
+  "\<lbrakk>x \<sqsubseteq> y;
+    \<And>a b. \<lbrakk>x = Inl a; y = Inl b; a \<sqsubseteq> b\<rbrakk> \<Longrightarrow> R;
+    \<And>a b. \<lbrakk>x = Inr a; y = Inr b; a \<sqsubseteq> b\<rbrakk> \<Longrightarrow> R\<rbrakk>
+      \<Longrightarrow> R"
+by (cases x, safe elim!: sum_less_elims, auto)
+
+subsection {* Sum type is a complete partial order *}
+
+instance "+" :: (po, po) po
+proof
+  fix x :: "'a + 'b"
+  show "x \<sqsubseteq> x"
+    by (induct x, simp_all)
+next
+  fix x y :: "'a + 'b"
+  assume "x \<sqsubseteq> y" and "y \<sqsubseteq> x" thus "x = y"
+    by (induct x, auto elim!: sum_less_elims intro: antisym_less)
+next
+  fix x y z :: "'a + 'b"
+  assume "x \<sqsubseteq> y" and "y \<sqsubseteq> z" thus "x \<sqsubseteq> z"
+    by (induct x, auto elim!: sum_less_elims intro: trans_less)
+qed
+
+lemma monofun_inv_Inl: "monofun (\<lambda>p. THE a. p = Inl a)"
+by (rule monofunI, erule sum_less_cases, simp_all)
+
+lemma monofun_inv_Inr: "monofun (\<lambda>p. THE b. p = Inr b)"
+by (rule monofunI, erule sum_less_cases, simp_all)
+
+lemma sum_chain_cases:
+  assumes Y: "chain Y"
+  assumes A: "\<And>A. \<lbrakk>chain A; Y = (\<lambda>i. Inl (A i))\<rbrakk> \<Longrightarrow> R"
+  assumes B: "\<And>B. \<lbrakk>chain B; Y = (\<lambda>i. Inr (B i))\<rbrakk> \<Longrightarrow> R"
+  shows "R"
+ apply (cases "Y 0")
+  apply (rule A)
+   apply (rule ch2ch_monofun [OF monofun_inv_Inl Y])
+  apply (rule ext)
+  apply (cut_tac j=i in chain_mono [OF Y le0], simp)
+  apply (erule Inl_lessE, simp)
+ apply (rule B)
+  apply (rule ch2ch_monofun [OF monofun_inv_Inr Y])
+ apply (rule ext)
+ apply (cut_tac j=i in chain_mono [OF Y le0], simp)
+ apply (erule Inr_lessE, simp)
+done
+
+lemma is_lub_Inl: "range S <<| x \<Longrightarrow> range (\<lambda>i. Inl (S i)) <<| Inl x"
+ apply (rule is_lubI)
+  apply (rule ub_rangeI)
+  apply (simp add: is_ub_lub)
+ apply (frule ub_rangeD [where i=arbitrary])
+ apply (erule Inl_lessE, simp)
+ apply (erule is_lub_lub)
+ apply (rule ub_rangeI)
+ apply (drule ub_rangeD, simp)
+done
+
+lemma is_lub_Inr: "range S <<| x \<Longrightarrow> range (\<lambda>i. Inr (S i)) <<| Inr x"
+ apply (rule is_lubI)
+  apply (rule ub_rangeI)
+  apply (simp add: is_ub_lub)
+ apply (frule ub_rangeD [where i=arbitrary])
+ apply (erule Inr_lessE, simp)
+ apply (erule is_lub_lub)
+ apply (rule ub_rangeI)
+ apply (drule ub_rangeD, simp)
+done
+
+instance "+" :: (cpo, cpo) cpo
+ apply intro_classes
+ apply (erule sum_chain_cases, safe)
+  apply (rule exI)
+  apply (rule is_lub_Inl)
+  apply (erule cpo_lubI)
+ apply (rule exI)
+ apply (rule is_lub_Inr)
+ apply (erule cpo_lubI)
+done
+
+subsection {* Continuity of @{term Inl}, @{term Inr}, @{term sum_case} *}
+
+lemma cont2cont_Inl [simp]: "cont f \<Longrightarrow> cont (\<lambda>x. Inl (f x))"
+by (fast intro: contI is_lub_Inl elim: contE)
+
+lemma cont2cont_Inr [simp]: "cont f \<Longrightarrow> cont (\<lambda>x. Inr (f x))"
+by (fast intro: contI is_lub_Inr elim: contE)
+
+lemma cont_Inl: "cont Inl"
+by (rule cont2cont_Inl [OF cont_id])
+
+lemma cont_Inr: "cont Inr"
+by (rule cont2cont_Inr [OF cont_id])
+
+lemmas ch2ch_Inl [simp] = ch2ch_cont [OF cont_Inl]
+lemmas ch2ch_Inr [simp] = ch2ch_cont [OF cont_Inr]
+
+lemmas lub_Inl = cont2contlubE [OF cont_Inl, symmetric]
+lemmas lub_Inr = cont2contlubE [OF cont_Inr, symmetric]
+
+lemma cont_sum_case1:
+  assumes f: "\<And>a. cont (\<lambda>x. f x a)"
+  assumes g: "\<And>b. cont (\<lambda>x. g x b)"
+  shows "cont (\<lambda>x. case y of Inl a \<Rightarrow> f x a | Inr b \<Rightarrow> g x b)"
+by (induct y, simp add: f, simp add: g)
+
+lemma cont_sum_case2: "\<lbrakk>cont f; cont g\<rbrakk> \<Longrightarrow> cont (sum_case f g)"
+apply (rule contI)
+apply (erule sum_chain_cases)
+apply (simp add: cont2contlubE [OF cont_Inl, symmetric] contE)
+apply (simp add: cont2contlubE [OF cont_Inr, symmetric] contE)
+done
+
+lemma cont2cont_sum_case [simp]:
+  assumes f1: "\<And>a. cont (\<lambda>x. f x a)" and f2: "\<And>x. cont (\<lambda>a. f x a)"
+  assumes g1: "\<And>b. cont (\<lambda>x. g x b)" and g2: "\<And>x. cont (\<lambda>b. g x b)"
+  assumes h: "cont (\<lambda>x. h x)"
+  shows "cont (\<lambda>x. case h x of Inl a \<Rightarrow> f x a | Inr b \<Rightarrow> g x b)"
+apply (rule cont2cont_app2 [OF cont2cont_lambda _ h])
+apply (rule cont_sum_case1 [OF f1 g1])
+apply (rule cont_sum_case2 [OF f2 g2])
+done
+
+subsection {* Compactness and chain-finiteness *}
+
+lemma compact_Inl: "compact a \<Longrightarrow> compact (Inl a)"
+apply (rule compactI2)
+apply (erule sum_chain_cases, safe)
+apply (simp add: lub_Inl)
+apply (erule (2) compactD2)
+apply (simp add: lub_Inr)
+done
+
+lemma compact_Inr: "compact a \<Longrightarrow> compact (Inr a)"
+apply (rule compactI2)
+apply (erule sum_chain_cases, safe)
+apply (simp add: lub_Inl)
+apply (simp add: lub_Inr)
+apply (erule (2) compactD2)
+done
+
+lemma compact_Inl_rev: "compact (Inl a) \<Longrightarrow> compact a"
+unfolding compact_def
+by (drule adm_subst [OF cont_Inl], simp)
+
+lemma compact_Inr_rev: "compact (Inr a) \<Longrightarrow> compact a"
+unfolding compact_def
+by (drule adm_subst [OF cont_Inr], simp)
+
+lemma compact_Inl_iff [simp]: "compact (Inl a) = compact a"
+by (safe elim!: compact_Inl compact_Inl_rev)
+
+lemma compact_Inr_iff [simp]: "compact (Inr a) = compact a"
+by (safe elim!: compact_Inr compact_Inr_rev)
+
+instance "+" :: (chfin, chfin) chfin
+apply intro_classes
+apply (erule compact_imp_max_in_chain)
+apply (case_tac "\<Squnion>i. Y i", simp_all)
+done
+
+instance "+" :: (finite_po, finite_po) finite_po ..
+
+instance "+" :: (discrete_cpo, discrete_cpo) discrete_cpo
+by intro_classes (simp add: less_sum_def split: sum.split)
+
+subsection {* Sum type is a bifinite domain *}
+
+instantiation "+" :: (profinite, profinite) profinite
+begin
+
+definition
+  approx_sum_def: "approx =
+    (\<lambda>n. \<Lambda> x. case x of Inl a \<Rightarrow> Inl (approx n\<cdot>a) | Inr b \<Rightarrow> Inr (approx n\<cdot>b))"
+
+lemma approx_Inl [simp]: "approx n\<cdot>(Inl x) = Inl (approx n\<cdot>x)"
+  unfolding approx_sum_def by simp
+
+lemma approx_Inr [simp]: "approx n\<cdot>(Inr x) = Inr (approx n\<cdot>x)"
+  unfolding approx_sum_def by simp
+
+instance proof
+  fix i :: nat and x :: "'a + 'b"
+  show "chain (approx :: nat \<Rightarrow> 'a + 'b \<rightarrow> 'a + 'b)"
+    unfolding approx_sum_def
+    by (rule ch2ch_LAM, case_tac x, simp_all)
+  show "(\<Squnion>i. approx i\<cdot>x) = x"
+    by (induct x, simp_all add: lub_Inl lub_Inr)
+  show "approx i\<cdot>(approx i\<cdot>x) = approx i\<cdot>x"
+    by (induct x, simp_all)
+  have "{x::'a + 'b. approx i\<cdot>x = x} \<subseteq>
+        {x::'a. approx i\<cdot>x = x} <+> {x::'b. approx i\<cdot>x = x}"
+    by (rule subsetI, case_tac x, simp_all add: InlI InrI)
+  thus "finite {x::'a + 'b. approx i\<cdot>x = x}"
+    by (rule finite_subset,
+        intro finite_Plus finite_fixes_approx)
+qed
+
+end
+
+end
--- a/src/HOLCF/Tools/adm_tac.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOLCF/Tools/adm_tac.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,18 +1,16 @@
-(*  ID:         $Id$
-    Author:     Stefan Berghofer, TU Muenchen
+(*  Author:     Stefan Berghofer, TU Muenchen
 
 Admissibility tactic.
 
 Checks whether adm_subst theorem is applicable to the current proof
 state:
 
-  [| cont t; adm P |] ==> adm (%x. P (t x))
+  cont t ==> adm P ==> adm (%x. P (t x))
 
 "t" is instantiated with a term of chain-finite type, so that
 adm_chfin can be applied:
 
   adm (P::'a::{chfin,pcpo} => bool)
-
 *)
 
 signature ADM =
@@ -39,21 +37,19 @@
       if i = lev then [[(Bound 0, path)]]
       else []
   | find_subterms (t as (Abs (_, _, t2))) lev path =
-      if List.filter (fn x => x<=lev)
-           (add_loose_bnos (t, 0, [])) = [lev] then
-        [(incr_bv (~lev, 0, t), path)]::
+      if filter (fn x => x <= lev) (add_loose_bnos (t, 0, [])) = [lev]
+      then
+        [(incr_bv (~lev, 0, t), path)] ::
         (find_subterms t2 (lev+1) (0::path))
       else find_subterms t2 (lev+1) (0::path)
   | find_subterms (t as (t1 $ t2)) lev path =
       let val ts1 = find_subterms t1 lev (0::path);
           val ts2 = find_subterms t2 lev (1::path);
           fun combine [] y = []
-            | combine (x::xs) ys =
-                (map (fn z => x @ z) ys) @ (combine xs ys)
+            | combine (x::xs) ys = map (fn z => x @ z) ys @ combine xs ys
       in
-        (if List.filter (fn x => x<=lev)
-              (add_loose_bnos (t, 0, [])) = [lev] then
-           [[(incr_bv (~lev, 0, t), path)]]
+        (if filter (fn x => x <= lev) (add_loose_bnos (t, 0, [])) = [lev]
+         then [[(incr_bv (~lev, 0, t), path)]]
          else []) @
         (if ts1 = [] then ts2
          else if ts2 = [] then ts1
@@ -65,7 +61,7 @@
 (*** make term for instantiation of predicate "P" in adm_subst theorem ***)
 
 fun make_term t path paths lev =
-  if path mem paths then Bound lev
+  if member (op =) paths path then Bound lev
   else case t of
       (Abs (s, T, t1)) => Abs (s, T, make_term t1 (0::path) paths (lev+1))
     | (t1 $ t2) => (make_term t1 (0::path) paths lev) $
@@ -79,30 +75,24 @@
   | eq_terms (ts as (t, _) :: _) = forall (fn (t2, _) => t2 aconv t) ts;
 
 
-(*figure out internal names*)
-val chfin_pcpoS = Sign.intern_sort (the_context ()) ["chfin", "pcpo"];
-val cont_name = Sign.intern_const (the_context ()) "cont";
-val adm_name = Sign.intern_const (the_context ()) "adm";
-
-
 (*** check whether type of terms in list is chain finite ***)
 
-fun is_chfin sign T params ((t, _)::_) =
+fun is_chfin thy T params ((t, _)::_) =
   let val parTs = map snd (rev params)
-  in Sign.of_sort sign (fastype_of1 (T::parTs, t), chfin_pcpoS) end;
+  in Sign.of_sort thy (fastype_of1 (T::parTs, t), @{sort "{chfin,pcpo}"}) end;
 
 
 (*** try to prove that terms in list are continuous
      if successful, add continuity theorem to list l ***)
 
-fun prove_cont tac sign s T prems params (l, ts as ((t, _)::_)) =
+fun prove_cont tac thy s T prems params (ts as ((t, _)::_)) l =
   let val parTs = map snd (rev params);
        val contT = (T --> (fastype_of1 (T::parTs, t))) --> HOLogic.boolT;
        fun mk_all [] t = t
          | mk_all ((a,T)::Ts) t = Term.all T $ (Abs (a, T, mk_all Ts t));
-       val t = HOLogic.mk_Trueprop((Const (cont_name, contT)) $ (Abs(s, T, t)));
+       val t = HOLogic.mk_Trueprop (Const (@{const_name cont}, contT) $ Abs (s, T, t));
        val t' = mk_all params (Logic.list_implies (prems, t));
-       val thm = Goal.prove (ProofContext.init sign) [] [] t' (K (tac 1));
+       val thm = Goal.prove (ProofContext.init thy) [] [] t' (K (tac 1));
   in (ts, thm)::l end
   handle ERROR _ => l;
 
@@ -111,71 +101,59 @@
 
 fun inst_adm_subst_thm state i params s T subt t paths =
   let
-      val sign = Thm.theory_of_thm state;
-      val j = Thm.maxidx_of state + 1;
-      val parTs = map snd (rev params);
-      val rule = Thm.lift_rule (Thm.cprem_of state i) @{thm adm_subst};
-      val types = valOf o (fst (Drule.types_sorts rule));
-      val tT = types ("t", j);
-      val PT = types ("P", j);
-      fun mk_abs [] t = t
-        | mk_abs ((a,T)::Ts) t = Abs (a, T, mk_abs Ts t);
-      val tt = cterm_of sign (mk_abs (params @ [(s, T)]) subt);
-      val Pt = cterm_of sign (mk_abs (params @ [(s, fastype_of1 (T::parTs, subt))])
-                     (make_term t [] paths 0));
-      val tye = Sign.typ_match sign (tT, #T (rep_cterm tt)) Vartab.empty;
-      val tye' = Sign.typ_match sign (PT, #T (rep_cterm Pt)) tye;
-      val ctye = map (fn (ixn, (S, T)) =>
-        (ctyp_of sign (TVar (ixn, S)), ctyp_of sign T)) (Vartab.dest tye');
-      val tv = cterm_of sign (Var (("t", j), Envir.typ_subst_TVars tye' tT));
-      val Pv = cterm_of sign (Var (("P", j), Envir.typ_subst_TVars tye' PT));
-      val rule' = instantiate (ctye, [(tv, tt), (Pv, Pt)]) rule
+    val thy = Thm.theory_of_thm state;
+    val j = Thm.maxidx_of state + 1;
+    val parTs = map snd (rev params);
+    val rule = Thm.lift_rule (Thm.cprem_of state i) @{thm adm_subst};
+    val types = the o fst (Drule.types_sorts rule);
+    val tT = types ("t", j);
+    val PT = types ("P", j);
+    fun mk_abs [] t = t
+      | mk_abs ((a,T)::Ts) t = Abs (a, T, mk_abs Ts t);
+    val tt = cterm_of thy (mk_abs (params @ [(s, T)]) subt);
+    val Pt = cterm_of thy (mk_abs (params @ [(s, fastype_of1 (T::parTs, subt))])
+                   (make_term t [] paths 0));
+    val tye = Sign.typ_match thy (tT, #T (rep_cterm tt)) Vartab.empty;
+    val tye' = Sign.typ_match thy (PT, #T (rep_cterm Pt)) tye;
+    val ctye = map (fn (ixn, (S, T)) =>
+      (ctyp_of thy (TVar (ixn, S)), ctyp_of thy T)) (Vartab.dest tye');
+    val tv = cterm_of thy (Var (("t", j), Envir.typ_subst_TVars tye' tT));
+    val Pv = cterm_of thy (Var (("P", j), Envir.typ_subst_TVars tye' PT));
+    val rule' = instantiate (ctye, [(tv, tt), (Pv, Pt)]) rule
   in rule' end;
 
 
-(*** extract subgoal i from proof state ***)
-
-fun nth_subgoal i thm = List.nth (prems_of thm, i-1);
-
-
 (*** the admissibility tactic ***)
 
-fun try_dest_adm (Const _ $ (Const (name, _) $ Abs abs)) =
-      if name = adm_name then SOME abs else NONE
+fun try_dest_adm (Const _ $ (Const (@{const_name adm}, _) $ Abs abs)) = SOME abs
   | try_dest_adm _ = NONE;
 
-fun adm_tac tac i state =
-  state |>
-  let val goali = nth_subgoal i state in
-    (case try_dest_adm (Logic.strip_assums_concl goali) of
-      NONE => no_tac
-    | SOME (s, T, t) =>
-        let
-          val sign = Thm.theory_of_thm state;
-          val prems = Logic.strip_assums_hyp goali;
-          val params = Logic.strip_params goali;
-          val ts = find_subterms t 0 [];
-          val ts' = List.filter eq_terms ts;
-          val ts'' = List.filter (is_chfin sign T params) ts';
-          val thms = Library.foldl (prove_cont tac sign s T prems params) ([], ts'');
-        in
-          (case thms of
-            ((ts as ((t', _)::_), cont_thm)::_) =>
-              let
-                val paths = map snd ts;
-                val rule = inst_adm_subst_thm state i params s T t' t paths;
-              in
-                compose_tac (false, rule, 2) i THEN
-                rtac cont_thm i THEN
-                REPEAT (assume_tac i) THEN
-                rtac @{thm adm_chfin} i
-              end 
-          | [] => no_tac)
-        end)
-    end;
-
+fun adm_tac tac i state = (i, state) |-> SUBGOAL (fn (goali, _) =>
+  (case try_dest_adm (Logic.strip_assums_concl goali) of
+    NONE => no_tac
+  | SOME (s, T, t) =>
+      let
+        val thy = Thm.theory_of_thm state;
+        val prems = Logic.strip_assums_hyp goali;
+        val params = Logic.strip_params goali;
+        val ts = find_subterms t 0 [];
+        val ts' = filter eq_terms ts;
+        val ts'' = filter (is_chfin thy T params) ts';
+        val thms = fold (prove_cont tac thy s T prems params) ts'' [];
+      in
+        (case thms of
+          ((ts as ((t', _)::_), cont_thm) :: _) =>
+            let
+              val paths = map snd ts;
+              val rule = inst_adm_subst_thm state i params s T t' t paths;
+            in
+              compose_tac (false, rule, 2) i THEN
+              resolve_tac [cont_thm] i THEN
+              REPEAT (assume_tac i) THEN
+              resolve_tac [@{thm adm_chfin}] i
+            end
+        | [] => no_tac)
+      end));
 
 end;
 
-
-open Adm;
--- a/src/HOLCF/Tools/cont_proc.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOLCF/Tools/cont_proc.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -33,21 +33,15 @@
 val cont_R = @{thm cont_Rep_CFun2};
 
 (* checks whether a term contains no dangling bound variables *)
-val is_closed_term =
-  let
-    fun bound_less i (t $ u) =
-          bound_less i t andalso bound_less i u
-      | bound_less i (Abs (_, _, t)) = bound_less (i+1) t
-      | bound_less i (Bound n) = n < i
-      | bound_less i _ = true; (* Const, Free, and Var are OK *)
-  in bound_less 0 end;
+fun is_closed_term t = not (Term.loose_bvar (t, 0));
 
 (* checks whether a term is written entirely in the LCF sublanguage *)
 fun is_lcf_term (Const (@{const_name Rep_CFun}, _) $ t $ u) =
       is_lcf_term t andalso is_lcf_term u
   | is_lcf_term (Const (@{const_name Abs_CFun}, _) $ Abs (_, _, t)) =
       is_lcf_term t
-  | is_lcf_term (Const (@{const_name Abs_CFun}, _) $ _) = false
+  | is_lcf_term (Const (@{const_name Abs_CFun}, _) $ t) =
+      is_lcf_term (Term.incr_boundvars 1 t $ Bound 0)
   | is_lcf_term (Bound _) = true
   | is_lcf_term t = is_closed_term t;
 
@@ -73,9 +67,9 @@
   fun lam [] = ([], cont_K)
     | lam (x::ys) =
     let
-      (* should use "standard" for thms that are used multiple times *)
+      (* should use "close_derivation" for thms that are used multiple times *)
       (* it seems to allow for sharing in explicit proof objects *)
-      val x' = standard (k x);
+      val x' = Thm.close_derivation (k x);
       val Lx = x' RS cont_L;
     in (map (fn y => SOME (k y RS Lx)) ys, x') end;
 
@@ -92,6 +86,12 @@
       val (cs, ls) = cont_thms1 b t;
       val (cs', l) = lam cs;
     in (cs', l::ls) end
+    | cont_thms1 b (Const (@{const_name Abs_CFun}, _) $ t) =
+    let
+      val t' = Term.incr_boundvars 1 t $ Bound 0;
+      val (cs, ls) = cont_thms1 b t';
+      val (cs', l) = lam cs;
+    in (cs', l::ls) end
     | cont_thms1 _ (Bound n) = (var n, [])
     | cont_thms1 _ _ = ([], []);
 in
--- a/src/HOLCF/Tools/domain/domain_syntax.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOLCF/Tools/domain/domain_syntax.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -89,7 +89,7 @@
     fun arg1 n (con,_,args) = foldr cabs (expvar n) (argvars n args);
     fun when1 n m = if n = m then arg1 n else K (Constant "UU");
 
-    fun app_var x = mk_appl (Constant "_var") [x, Variable "rhs"];
+    fun app_var x = mk_appl (Constant "_variable") [x, Variable "rhs"];
     fun app_pat x = mk_appl (Constant "_pat") [x];
     fun args_list [] = Constant "_noargs"
     |   args_list xs = foldr1 (app "_args") xs;
--- a/src/HOLCF/Tools/domain/domain_theorems.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOLCF/Tools/domain/domain_theorems.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -10,6 +10,12 @@
 
 structure Domain_Theorems = struct
 
+val quiet_mode = ref false;
+val trace_domain = ref false;
+
+fun message s = if !quiet_mode then () else writeln s;
+fun trace s = if !trace_domain then tracing s else ();
+
 local
 
 val adm_impl_admw = @{thm adm_impl_admw};
@@ -115,7 +121,7 @@
 fun theorems (((dname, _), cons) : eq, eqs : eq list) thy =
 let
 
-val dummy = writeln ("Proving isomorphism properties of domain "^dname^" ...");
+val _ = message ("Proving isomorphism properties of domain "^dname^" ...");
 val pg = pg' thy;
 
 (* ----- getting the axioms and definitions --------------------------------- *)
@@ -158,6 +164,8 @@
 
 (* ----- generating beta reduction rules from definitions-------------------- *)
 
+val _ = trace " Proving beta reduction rules...";
+
 local
   fun arglist (Const _ $ Abs (s, _, t)) =
     let
@@ -179,7 +187,9 @@
     in pg (def::betas) appl (K [rtac reflexive_thm 1]) end;
 end;
 
+val _ = trace "Proving when_appl...";
 val when_appl = appl_of_def ax_when_def;
+val _ = trace "Proving con_appls...";
 val con_appls = map appl_of_def axs_con_def;
 
 local
@@ -229,7 +239,9 @@
     rewrite_goals_tac [mk_meta_eq iso_swap],
     rtac thm3 1];
 in
+  val _ = trace " Proving exhaust...";
   val exhaust = pg con_appls (mk_trp exh) (K tacs);
+  val _ = trace " Proving casedist...";
   val casedist =
     standard (rewrite_rule exh_casedists (exhaust RS exh_casedist0));
 end;
@@ -239,6 +251,7 @@
   fun bound_fun i _ = Bound (length cons - i);
   val when_app = list_ccomb (%%:(dname^"_when"), mapn bound_fun 1 cons);
 in
+  val _ = trace " Proving when_strict...";
   val when_strict =
     let
       val axs = [when_appl, mk_meta_eq rep_strict];
@@ -246,6 +259,7 @@
       val tacs = [resolve_tac [sscase1, ssplit1, strictify1] 1];
     in pg axs goal (K tacs) end;
 
+  val _ = trace " Proving when_apps...";
   val when_apps =
     let
       fun one_when n (con,args) =
@@ -276,6 +290,7 @@
       val tacs = [asm_simp_tac (HOLCF_ss addsimps when_rews) 1];
     in pg axs_dis_def goal (K tacs) end;
 
+  val _ = trace " Proving dis_apps...";
   val dis_apps = maps (fn (c,_) => map (dis_app c) cons) cons;
 
   fun dis_defin (con, args) =
@@ -288,7 +303,9 @@
           (asm_simp_tac (HOLCF_ss addsimps dis_apps) 1))];
     in pg [] goal (K tacs) end;
 
+  val _ = trace " Proving dis_stricts...";
   val dis_stricts = map dis_strict cons;
+  val _ = trace " Proving dis_defins...";
   val dis_defins = map dis_defin cons;
 in
   val dis_rews = dis_stricts @ dis_defins @ dis_apps;
@@ -301,6 +318,7 @@
       val tacs = [rtac when_strict 1];
     in pg axs_mat_def goal (K tacs) end;
 
+  val _ = trace " Proving mat_stricts...";
   val mat_stricts = map mat_strict cons;
 
   fun one_mat c (con, args) =
@@ -314,6 +332,7 @@
       val tacs = [asm_simp_tac (HOLCF_ss addsimps when_rews) 1];
     in pg axs_mat_def goal (K tacs) end;
 
+  val _ = trace " Proving mat_apps...";
   val mat_apps =
     maps (fn (c,_) => map (one_mat c) cons) cons;
 in
@@ -346,7 +365,9 @@
       val tacs = [asm_simp_tac (HOLCF_ss addsimps when_rews) 1];
     in pg axs goal (K tacs) end;
 
+  val _ = trace " Proving pat_stricts...";
   val pat_stricts = map pat_strict cons;
+  val _ = trace " Proving pat_apps...";
   val pat_apps = maps (fn c => map (pat_app c) cons) cons;
 in
   val pat_rews = pat_stricts @ pat_apps;
@@ -374,7 +395,9 @@
         asm_simp_tac (HOLCF_ss addsimps dis_rews) 1];
     in pg [] goal tacs end;
 in
+  val _ = trace " Proving con_stricts...";
   val con_stricts = maps con_strict cons;
+  val _ = trace " Proving pat_defins...";
   val con_defins = map con_defin cons;
   val con_rews = con_stricts @ con_defins;
 end;
@@ -391,6 +414,7 @@
         REPEAT (resolve_tac rules 1 ORELSE atac 1)];
     in pg con_appls goal (K tacs) end;
 in
+  val _ = trace " Proving con_compacts...";
   val con_compacts = map con_compact cons;
 end;
 
@@ -402,6 +426,7 @@
   fun sel_strict (_, args) =
     List.mapPartial (Option.map one_sel o sel_of) args;
 in
+  val _ = trace " Proving sel_stricts...";
   val sel_stricts = maps sel_strict cons;
 end;
 
@@ -439,6 +464,7 @@
   fun one_con (c, args) =
     flat (map_filter I (mapn (one_sel' c) 0 args));
 in
+  val _ = trace " Proving sel_apps...";
   val sel_apps = maps one_con cons;
 end;
 
@@ -453,6 +479,7 @@
           (asm_simp_tac (HOLCF_ss addsimps sel_apps) 1))];
     in pg [] goal (K tacs) end;
 in
+  val _ = trace " Proving sel_defins...";
   val sel_defins =
     if length cons = 1
     then List.mapPartial (fn arg => Option.map sel_defin (sel_of arg))
@@ -463,6 +490,7 @@
 val sel_rews = sel_stricts @ sel_defins @ sel_apps;
 val rev_contrapos = @{thm rev_contrapos};
 
+val _ = trace " Proving dist_les...";
 val distincts_le =
   let
     fun dist (con1, args1) (con2, args2) =
@@ -487,6 +515,8 @@
       | distincts (c::cs) = (map (distinct c) cs) :: distincts cs;
   in distincts cons end;
 val dist_les = flat (flat distincts_le);
+
+val _ = trace " Proving dist_eqs...";
 val dist_eqs =
   let
     fun distinct (_,args1) ((_,args2), leqs) =
@@ -517,6 +547,7 @@
     in pg con_appls prop end;
   val cons' = List.filter (fn (_,args) => args<>[]) cons;
 in
+  val _ = trace " Proving inverts...";
   val inverts =
     let
       val abs_less = ax_abs_iso RS (allI RS injection_less);
@@ -524,6 +555,7 @@
         [asm_full_simp_tac (HOLCF_ss addsimps [abs_less, spair_less]) 1];
     in map (fn (con, args) => pgterm (op <<) con args (K tacs)) cons' end;
 
+  val _ = trace " Proving injects...";
   val injects =
     let
       val abs_eq = ax_abs_iso RS (allI RS injection_eq);
@@ -551,6 +583,7 @@
       val tacs2 = [asm_simp_tac (HOLCF_ss addsimps when_apps) 1];
     in pg [ax_copy_def] goal (fn ctxt => (tacs1 ctxt @ tacs2)) end;
 in
+  val _ = trace " Proving copy_apps...";
   val copy_apps = map copy_app cons;
 end;
 
@@ -565,6 +598,7 @@
 
   fun has_nonlazy_rec (_, args) = exists is_nonlazy_rec args;
 in
+  val _ = trace " Proving copy_stricts...";
   val copy_stricts = map one_strict (List.filter has_nonlazy_rec cons);
 end;
 
@@ -603,7 +637,7 @@
 val conss  = map  snd        eqs;
 val comp_dname = Sign.full_bname thy comp_dnam;
 
-val d = writeln("Proving induction properties of domain "^comp_dname^" ...");
+val _ = message ("Proving induction properties of domain "^comp_dname^" ...");
 val pg = pg' thy;
 
 (* ----- getting the composite axiom and definitions ------------------------ *)
@@ -639,6 +673,7 @@
   val copy_con_rews  = copy_rews @ con_rews;
   val copy_take_defs =
     (if n_eqs = 1 then [] else [ax_copy2_def]) @ axs_take_def;
+  val _ = trace " Proving take_stricts...";
   val take_stricts =
     let
       fun one_eq ((dn, args), _) = strict (dc_take dn $ %:"n");
@@ -656,6 +691,7 @@
     in pg axs_take_def goal (K [simp_tac iterate_Cprod_ss 1]) end;
   val take_0s = mapn take_0 1 dnames;
   fun c_UU_tac ctxt = case_UU_tac ctxt (take_stricts'::copy_con_rews) 1;
+  val _ = trace " Proving take_apps...";
   val take_apps =
     let
       fun mk_eqn dn (con, args) =
@@ -737,6 +773,7 @@
     val is_finite = forall (not o lazy_rec_to [] false) n__eqs;
   end;
 in (* local *)
+  val _ = trace " Proving finite_ind...";
   val finite_ind =
     let
       fun concf n dn = %:(P_name n) $ (dc_take dn $ %:"n" `%(x_name n));
@@ -768,6 +805,7 @@
         end;
     in pg'' thy [] goal tacf end;
 
+  val _ = trace " Proving take_lemmas...";
   val take_lemmas =
     let
       fun take_lemma n (dn, ax_reach) =
@@ -793,6 +831,7 @@
 
 (* ----- theorems concerning finiteness and induction ----------------------- *)
 
+  val _ = trace " Proving finites, ind...";
   val (finites, ind) =
     if is_finite
     then (* finite case *)
@@ -914,6 +953,7 @@
   fun bnd_arg n i = Bound(2*(n_eqs - n)-i-1);
   val take_ss = HOL_ss addsimps take_rews;
   val sproj = prj (fn s => K("fst("^s^")")) (fn s => K("snd("^s^")"));
+  val _ = trace " Proving coind_lemma...";
   val coind_lemma =
     let
       fun mk_prj n _ = proj (%:"R") eqs n $ bnd_arg n 0 $ bnd_arg n 1;
@@ -940,6 +980,7 @@
         flat (mapn (x_tacs ctxt) 0 xs);
     in pg [ax_bisim_def] goal tacs end;
 in
+  val _ = trace " Proving coind...";
   val coind = 
     let
       fun mk_prj n x = mk_trp (proj (%:"R") eqs n $ %:x $ %:(x^"'"));
--- a/src/HOLCF/Up.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOLCF/Up.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -282,10 +282,10 @@
 text {* properties of fup *}
 
 lemma fup1 [simp]: "fup\<cdot>f\<cdot>\<bottom> = \<bottom>"
-by (simp add: fup_def cont_Ifup1 cont_Ifup2 inst_up_pcpo)
+by (simp add: fup_def cont_Ifup1 cont_Ifup2 inst_up_pcpo cont2cont_LAM)
 
 lemma fup2 [simp]: "fup\<cdot>f\<cdot>(up\<cdot>x) = f\<cdot>x"
-by (simp add: up_def fup_def cont_Iup cont_Ifup1 cont_Ifup2)
+by (simp add: up_def fup_def cont_Iup cont_Ifup1 cont_Ifup2 cont2cont_LAM)
 
 lemma fup3 [simp]: "fup\<cdot>up\<cdot>x = x"
 by (cases x, simp_all)
--- a/src/HOLCF/UpperPD.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOLCF/UpperPD.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -248,9 +248,8 @@
 apply (simp add: PDPlus_absorb)
 done
 
-class_interpretation aci_upper_plus: ab_semigroup_idem_mult ["op +\<sharp>"]
-  by unfold_locales
-    (rule upper_plus_assoc upper_plus_commute upper_plus_absorb)+
+interpretation aci_upper_plus!: ab_semigroup_idem_mult "op +\<sharp>"
+  proof qed (rule upper_plus_assoc upper_plus_commute upper_plus_absorb)+
 
 lemma upper_plus_left_commute: "xs +\<sharp> (ys +\<sharp> zs) = ys +\<sharp> (xs +\<sharp> zs)"
 by (rule aci_upper_plus.mult_left_commute)
--- a/src/HOLCF/ex/Stream.thy	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/HOLCF/ex/Stream.thy	Mon Jan 19 21:20:18 2009 +0100
@@ -521,7 +521,7 @@
 section "smap"
 
 lemma smap_unfold: "smap = (\<Lambda> f t. case t of x&&xs \<Rightarrow> f$x && smap$f$xs)"
-by (insert smap_def [THEN eq_reflection, THEN fix_eq2], auto)
+by (insert smap_def [where 'a='a and 'b='b, THEN eq_reflection, THEN fix_eq2], auto)
 
 lemma smap_empty [simp]: "smap\<cdot>f\<cdot>\<bottom> = \<bottom>"
 by (subst smap_unfold, simp)
@@ -540,7 +540,7 @@
 lemma sfilter_unfold:
  "sfilter = (\<Lambda> p s. case s of x && xs \<Rightarrow>
   If p\<cdot>x then x && sfilter\<cdot>p\<cdot>xs else sfilter\<cdot>p\<cdot>xs fi)"
-by (insert sfilter_def [THEN eq_reflection, THEN fix_eq2], auto)
+by (insert sfilter_def [where 'a='a, THEN eq_reflection, THEN fix_eq2], auto)
 
 lemma strict_sfilter: "sfilter\<cdot>\<bottom> = \<bottom>"
 apply (rule ext_cfun)
--- a/src/Provers/Arith/cancel_factor.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Provers/Arith/cancel_factor.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Provers/Arith/cancel_factor.ML
-    ID:         $Id$
     Author:     Markus Wenzel and Stefan Berghofer, TU Muenchen
 
 Cancel common constant factor from balanced exression (e.g. =, <=, < of sums).
@@ -36,7 +35,7 @@
       if t aconv u then (u, k + 1) :: uks
       else (t, 1) :: (u, k) :: uks;
 
-fun count_terms ts = foldr ins_term (sort Term.term_ord ts, []);
+fun count_terms ts = foldr ins_term (sort TermOrd.term_ord ts, []);
 
 
 (* divide sum *)
--- a/src/Provers/Arith/cancel_sums.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Provers/Arith/cancel_sums.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Provers/Arith/cancel_sums.ML
-    ID:         $Id$
     Author:     Markus Wenzel and Stefan Berghofer, TU Muenchen
 
 Cancel common summands of balanced expressions:
@@ -38,11 +37,11 @@
 fun cons2 y (x, ys, z) = (x, y :: ys, z);
 fun cons12 x y (xs, ys, z) = (x :: xs, y :: ys, z);
 
-(*symmetric difference of multisets -- assumed to be sorted wrt. Logic.term_ord*)
+(*symmetric difference of multisets -- assumed to be sorted wrt. TermOrd.term_ord*)
 fun cancel ts [] vs = (ts, [], vs)
   | cancel [] us vs = ([], us, vs)
   | cancel (t :: ts) (u :: us) vs =
-      (case Term.term_ord (t, u) of
+      (case TermOrd.term_ord (t, u) of
         EQUAL => cancel ts us (t :: vs)
       | LESS => cons1 t (cancel ts (u :: us) vs)
       | GREATER => cons2 u (cancel (t :: ts) us vs));
@@ -64,7 +63,7 @@
   | SOME bal =>
       let
         val thy = ProofContext.theory_of (Simplifier.the_context ss);
-        val (ts, us) = pairself (sort Term.term_ord o Data.dest_sum) bal;
+        val (ts, us) = pairself (sort TermOrd.term_ord o Data.dest_sum) bal;
         val (ts', us', vs) = cancel ts us [];
       in
         if null vs then NONE
--- a/src/Provers/classical.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Provers/classical.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,7 +1,5 @@
 (*  Title:      Provers/classical.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1992  University of Cambridge
 
 Theorem prover for classical reasoning, including predicate calculus, set
 theory, etc.
@@ -810,9 +808,7 @@
     (fn (prem,i) =>
       let val deti =
           (*No Vars in the goal?  No need to backtrack between goals.*)
-          case term_vars prem of
-              []        => DETERM
-            | _::_      => I
+          if exists_subterm (fn Var _ => true | _ => false) prem then DETERM else I
       in  SELECT_GOAL (TRY (safe_tac cs) THEN
                        DEPTH_SOLVE (deti (depth_tac cs m 1))) i
       end);
--- a/src/Provers/coherent.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Provers/coherent.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,7 +1,6 @@
 (*  Title:      Provers/coherent.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
-                Marc Bezem, Institutt for Informatikk, Universitetet i Bergen 
+    Author:     Marc Bezem, Institutt for Informatikk, Universitetet i Bergen 
 
 Prover for coherent logic, see e.g.
 
@@ -39,7 +38,7 @@
   ClPrf of thm * (Type.tyenv * Envir.tenv) * ((indexname * typ) * term) list *
   int list * (term list * cl_prf) list;
 
-fun is_atomic t = null (term_consts t inter Data.operator_names);
+val is_atomic = not o exists_Const (member (op =) Data.operator_names o #1);
 
 local open Conv in
 
--- a/src/Provers/eqsubst.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Provers/eqsubst.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,13 +1,12 @@
 (*  Title:      Provers/eqsubst.ML
-    ID:         $Id$
-    Author:     Lucas Dixon, University of Edinburgh, lucas.dixon@ed.ac.uk
+    Author:     Lucas Dixon, University of Edinburgh
 
 A proof method to perform a substiution using an equation.
 *)
 
 signature EQSUBST =
 sig
-  (* a type abriviation for match information *)
+  (* a type abbreviation for match information *)
   type match =
        ((indexname * (sort * typ)) list (* type instantiations *)
         * (indexname * (typ * term)) list) (* term instantiations *)
@@ -224,7 +223,7 @@
       (* is it OK to ignore the type instantiation info?
          or should I be using it? *)
       val typs_unify =
-          SOME (Sign.typ_unify thry (pat_ty, tgt_ty) (Term.Vartab.empty, ix))
+          SOME (Sign.typ_unify thry (pat_ty, tgt_ty) (Vartab.empty, ix))
             handle Type.TUNIFY => NONE;
     in
       case typs_unify of
--- a/src/Provers/order.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Provers/order.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,8 +1,6 @@
-(*
-  Title:	Transitivity reasoner for partial and linear orders
-  Id:		$Id$
-  Author:	Oliver Kutter
-  Copyright:	TU Muenchen
+(*  Author:     Oliver Kutter, TU Muenchen
+
+Transitivity reasoner for partial and linear orders.
 *)
 
 (* TODO: reduce number of input thms *)
@@ -1234,7 +1232,7 @@
 
     SUBGOAL (fn (A, n, sign) =>
      let
-      val rfrees = map Free (rename_wrt_term A (Logic.strip_params A))
+      val rfrees = map Free (Term.rename_wrt_term A (Logic.strip_params A))
       val Hs = map prop_of prems @ map (fn H => subst_bounds (rfrees, H)) (Logic.strip_assums_hyp A)
       val C = subst_bounds (rfrees, Logic.strip_assums_concl A)
       val lesss = List.concat (ListPair.map (mkasm decomp less_thms sign) (Hs, 0 upto (length Hs - 1)))
--- a/src/Provers/quasi.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Provers/quasi.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,9 +1,7 @@
-(* 
-   Title:	Reasoner for simple transitivity and quasi orders. 
-   Id:		$Id$
-   Author:	Oliver Kutter
-   Copyright:	TU Muenchen
- *)
+(*  Author:     Oliver Kutter, TU Muenchen
+
+Reasoner for simple transitivity and quasi orders.
+*)
 
 (* 
  
@@ -557,7 +555,7 @@
 (* trans_tac - solves transitivity chains over <= *)
 val trans_tac  =  SUBGOAL (fn (A, n, sign) =>
  let
-  val rfrees = map Free (rename_wrt_term A (Logic.strip_params A))
+  val rfrees = map Free (Term.rename_wrt_term A (Logic.strip_params A))
   val Hs = map (fn H => subst_bounds (rfrees, H)) (Logic.strip_assums_hyp A)
   val C = subst_bounds (rfrees, Logic.strip_assums_concl A)
   val lesss = List.concat (ListPair.map (mkasm_trans  sign) (Hs, 0 upto (length Hs - 1)))
@@ -578,7 +576,7 @@
 (* quasi_tac - solves quasi orders *)
 val quasi_tac = SUBGOAL (fn (A, n, sign) =>
  let
-  val rfrees = map Free (rename_wrt_term A (Logic.strip_params A))
+  val rfrees = map Free (Term.rename_wrt_term A (Logic.strip_params A))
   val Hs = map (fn H => subst_bounds (rfrees, H)) (Logic.strip_assums_hyp A)
   val C = subst_bounds (rfrees, Logic.strip_assums_concl A)
   val lesss = List.concat (ListPair.map (mkasm_quasi sign) (Hs, 0 upto (length Hs - 1)))
--- a/src/Provers/splitter.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Provers/splitter.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -146,7 +146,7 @@
 
 fun mk_cntxt_splitthm t tt T =
   let fun repl lev t =
-    if incr_boundvars lev tt aconv t then Bound lev
+    if Pattern.aeconv(incr_boundvars lev tt, t) then Bound lev
     else case t of
         (Abs (v, T2, t)) => Abs (v, T2, repl (lev+1) t)
       | (Bound i) => Bound (if i>=lev then i+1 else i)
--- a/src/Pure/Concurrent/future.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Concurrent/future.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -45,6 +45,7 @@
   val join: 'a future -> 'a
   val map: ('a -> 'b) -> 'a future -> 'b future
   val interrupt_task: string -> unit
+  val cancel_group: group -> unit
   val cancel: 'a future -> unit
   val shutdown: unit -> unit
 end;
@@ -116,7 +117,7 @@
   ConditionVar.wait (cond, lock);
 
 fun wait_timeout timeout = (*requires SYNCHRONIZED*)
-  ConditionVar.waitUntil (cond, lock, Time.+ (Time.now (), timeout));
+  ignore (ConditionVar.waitUntil (cond, lock, Time.+ (Time.now (), timeout)));
 
 fun notify_all () = (*requires SYNCHRONIZED*)
   ConditionVar.broadcast cond;
@@ -137,17 +138,22 @@
   change workers (AList.update Thread.equal (Thread.self (), active));
 
 
-(* execute *)
+(* execute jobs *)
 
-fun execute name (task, group, run) =
+fun do_cancel group = (*requires SYNCHRONIZED*)
+  change canceled (insert Task_Queue.eq_group group);
+
+fun execute name (task, group, jobs) =
   let
     val _ = trace_active ();
-    val ok = setmp_thread_data (name, task) run ();
+    val valid = Task_Queue.is_valid group;
+    val ok = setmp_thread_data (name, task) (fn () =>
+      fold (fn job => fn ok => job valid andalso ok) jobs true) ();
     val _ = SYNCHRONIZED "execute" (fn () =>
      (change queue (Task_Queue.finish task);
       if ok then ()
       else if Task_Queue.cancel (! queue) group then ()
-      else change canceled (cons group);
+      else do_cancel group;
       notify_all ()));
   in () end;
 
@@ -205,7 +211,8 @@
     val _ = if continue then () else scheduler := NONE;
 
     val _ = notify_all ();
-    val _ = wait_timeout (Time.fromSeconds 3);
+    val _ = interruptible (fn () => wait_timeout (Time.fromSeconds 1)) ()
+      handle Exn.Interrupt => List.app do_cancel (Task_Queue.cancel_all (! queue));
   in continue end;
 
 fun scheduler_loop () =
@@ -221,16 +228,15 @@
   else ());
 
 
-(* future values: fork independent computation *)
+
+(** futures **)
 
-fun future opt_group deps pri (e: unit -> 'a) =
+(* future job: fill result *)
+
+fun future_job group (e: unit -> 'a) =
   let
-    val _ = scheduler_check "future check";
-
-    val group = (case opt_group of SOME group => group | NONE => Task_Queue.new_group ());
-
     val result = ref (NONE: 'a Exn.result option);
-    val run = Multithreading.with_attributes (Thread.getAttributes ())
+    val job = Multithreading.with_attributes (Thread.getAttributes ())
       (fn _ => fn ok =>
         let
           val res = if ok then Exn.capture e () else Exn.Exn Exn.Interrupt;
@@ -241,75 +247,123 @@
             | Exn.Exn Exn.Interrupt => (Task_Queue.invalidate_group group; true)
             | _ => false);
         in res_ok end);
+  in (result, job) end;
 
+
+(* fork *)
+
+fun fork_future opt_group deps pri e =
+  let
+    val _ = scheduler_check "future check";
+
+    val group = (case opt_group of SOME group => group | NONE => Task_Queue.new_group ());
+    val (result, job) = future_job group e;
     val task = SYNCHRONIZED "future" (fn () =>
-      change_result queue (Task_Queue.enqueue group deps pri run) before notify_all ());
+      change_result queue (Task_Queue.enqueue group deps pri job) before notify_all ());
   in Future {task = task, group = group, result = result} end;
 
-fun fork e = future NONE [] 0 e;
-fun fork_group group e = future (SOME group) [] 0 e;
-fun fork_deps deps e = future NONE (map task_of deps) 0 e;
-fun fork_pri pri e = future NONE [] pri e;
+fun fork e = fork_future NONE [] 0 e;
+fun fork_group group e = fork_future (SOME group) [] 0 e;
+fun fork_deps deps e = fork_future NONE (map task_of deps) 0 e;
+fun fork_pri pri e = fork_future NONE [] pri e;
 
 
-(* join: retrieve results *)
+(* join *)
+
+local
+
+fun get_result x = the_default (Exn.Exn (SYS_ERROR "unfinished future")) (peek x);
+
+fun join_next pending = (*requires SYNCHRONIZED*)
+  if forall is_finished pending then NONE
+  else
+    (case change_result queue Task_Queue.dequeue of
+      NONE => (worker_wait (); join_next pending)
+    | some => some);
 
-fun join_results [] = []
-  | join_results xs = uninterruptible (fn _ => fn () =>
-      let
-        val _ = scheduler_check "join check";
-        val _ = Multithreading.self_critical () andalso
-          exists (not o is_finished) xs andalso
-          error "Cannot join future values within critical section";
+fun join_loop name pending =
+  (case SYNCHRONIZED name (fn () => join_next pending) of
+    NONE => ()
+  | SOME work => (execute name work; join_loop name pending));
+
+in
+
+fun join_results xs =
+  if forall is_finished xs then map get_result xs
+  else uninterruptible (fn _ => fn () =>
+    let
+      val _ = scheduler_check "join check";
+      val _ = Multithreading.self_critical () andalso
+        error "Cannot join future values within critical section";
 
-        fun join_loop _ [] = ()
-          | join_loop name tasks =
-              (case SYNCHRONIZED name (fn () =>
-                  change_result queue (Task_Queue.dequeue_towards tasks)) of
-                NONE => ()
-              | SOME (work, tasks') => (execute name work; join_loop name tasks'));
-        val _ =
-          (case thread_data () of
-            NONE =>
-              (*alien thread -- refrain from contending for resources*)
-              while exists (not o is_finished) xs
-              do SYNCHRONIZED "join_thread" (fn () => wait ())
-          | SOME (name, task) =>
-              (*proper task -- actively work towards results*)
-              let
-                val unfinished = xs |> map_filter
-                  (fn Future {task, result = ref NONE, ...} => SOME task | _ => NONE);
-                val _ = SYNCHRONIZED "join" (fn () =>
-                  (change queue (Task_Queue.depend unfinished task); notify_all ()));
-                val _ = join_loop ("join_loop: " ^ name) unfinished;
-                val _ =
-                  while exists (not o is_finished) xs
-                  do SYNCHRONIZED "join_task" (fn () => worker_wait ());
-              in () end);
+      fun join_deps _ [] = ()
+        | join_deps name deps =
+            (case SYNCHRONIZED name (fn () =>
+                change_result queue (Task_Queue.dequeue_towards deps)) of
+              NONE => ()
+            | SOME (work, deps') => (execute name work; join_deps name deps'));
 
-      in xs |> map (fn Future {result = ref (SOME res), ...} => res) end) ();
+      val _ =
+        (case thread_data () of
+          NONE =>
+            (*alien thread -- refrain from contending for resources*)
+            while not (forall is_finished xs)
+            do SYNCHRONIZED "join_thread" (fn () => wait ())
+        | SOME (name, task) =>
+            (*proper task -- actively work towards results*)
+            let
+              val pending = filter_out is_finished xs;
+              val deps = map task_of pending;
+              val _ = SYNCHRONIZED "join" (fn () =>
+                (change queue (Task_Queue.depend deps task); notify_all ()));
+              val _ = join_deps ("join_deps: " ^ name) deps;
+              val _ = join_loop ("join_loop: " ^ name) (filter_out is_finished pending);
+            in () end);
+
+    in map get_result xs end) ();
+
+end;
 
 fun join_result x = singleton join_results x;
 fun join x = Exn.release (join_result x);
 
-fun map f x =
-  let val task = task_of x
-  in future NONE [task] (Task_Queue.pri_of_task task) (fn () => f (join x)) end;
+
+(* map *)
+
+fun map_future f x =
+  let
+    val _ = scheduler_check "map_future check";
+
+    val task = task_of x;
+    val group = Task_Queue.new_group ();
+    val (result, job) = future_job group (fn () => f (join x));
+
+    val extended = SYNCHRONIZED "map_future" (fn () =>
+      (case Task_Queue.extend task job (! queue) of
+        SOME queue' => (queue := queue'; true)
+      | NONE => false));
+  in
+    if extended then Future {task = task, group = group, result = result}
+    else fork_future NONE [task] (Task_Queue.pri_of_task task) (fn () => f (join x))
+  end;
 
 
-(* misc operations *)
+(* cancellation *)
 
 (*interrupt: permissive signal, may get ignored*)
 fun interrupt_task id = SYNCHRONIZED "interrupt"
   (fn () => Task_Queue.interrupt_external (! queue) id);
 
 (*cancel: present and future group members will be interrupted eventually*)
-fun cancel x =
+fun cancel_group group =
  (scheduler_check "cancel check";
-  SYNCHRONIZED "cancel" (fn () => (change canceled (cons (group_of x)); notify_all ())));
+  SYNCHRONIZED "cancel" (fn () => (do_cancel group; notify_all ())));
+
+fun cancel x = cancel_group (group_of x);
 
 
-(*global join and shutdown*)
+(** global join and shutdown **)
+
 fun shutdown () =
   if Multithreading.available then
    (scheduler_check "shutdown check";
@@ -323,6 +377,10 @@
       OS.Process.sleep (Time.fromMilliseconds 300))))
   else ();
 
+
+(*final declarations of this structure!*)
+val map = map_future;
+
 end;
 
 type 'a future = 'a Future.future;
--- a/src/Pure/Concurrent/mailbox.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Concurrent/mailbox.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Concurrent/mailbox.ML
-    ID:         $Id$
     Author:     Makarius
 
 Message exchange via mailbox, with non-blocking send (due to unbounded
--- a/src/Pure/Concurrent/par_list.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Concurrent/par_list.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Concurrent/par_list.ML
-    ID:         $Id$
     Author:     Makarius
 
 Parallel list combinators.
@@ -24,7 +23,7 @@
   val forall: ('a -> bool) -> 'a list -> bool
 end;
 
-structure ParList: PAR_LIST =
+structure Par_List: PAR_LIST =
 struct
 
 fun raw_map f xs =
--- a/src/Pure/Concurrent/par_list_dummy.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Concurrent/par_list_dummy.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,11 +1,10 @@
 (*  Title:      Pure/Concurrent/par_list_dummy.ML
-    ID:         $Id$
     Author:     Makarius
 
 Dummy version of parallel list combinators -- plain sequential evaluation.
 *)
 
-structure ParList: PAR_LIST =
+structure Par_List: PAR_LIST =
 struct
 
 val map = map;
--- a/src/Pure/Concurrent/simple_thread.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Concurrent/simple_thread.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Concurrent/simple_thread.ML
-    ID:         $Id$
     Author:     Makarius
 
 Simplified thread operations.
--- a/src/Pure/Concurrent/synchronized.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Concurrent/synchronized.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Concurrent/synchronized.ML
-    ID:         $Id$
     Author:     Fabian Immler and Makarius
 
 State variables with synchronized access.
--- a/src/Pure/Concurrent/task_queue.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Concurrent/task_queue.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -6,29 +6,33 @@
 
 signature TASK_QUEUE =
 sig
-  eqtype task
+  type task
   val new_task: int -> task
   val pri_of_task: task -> int
   val str_of_task: task -> string
-  eqtype group
+  type group
+  val eq_group: group * group -> bool
   val new_group: unit -> group
+  val is_valid: group -> bool
   val invalidate_group: group -> unit
   val str_of_group: group -> string
   type queue
   val empty: queue
   val is_empty: queue -> bool
   val enqueue: group -> task list -> int -> (bool -> bool) -> queue -> task * queue
+  val extend: task -> (bool -> bool) -> queue -> queue option
   val depend: task list -> task -> queue -> queue
-  val dequeue: queue -> (task * group * (unit -> bool)) option * queue
+  val dequeue: queue -> (task * group * (bool -> bool) list) option * queue
   val dequeue_towards: task list -> queue ->
-    (((task * group * (unit -> bool)) * task list) option * queue)
+    (((task * group * (bool -> bool) list) * task list) option * queue)
   val interrupt: queue -> task -> unit
   val interrupt_external: queue -> string -> unit
+  val cancel: queue -> group -> bool
+  val cancel_all: queue -> group list
   val finish: task -> queue -> queue
-  val cancel: queue -> group -> bool
 end;
 
-structure Task_Queue: TASK_QUEUE =
+structure Task_Queue:> TASK_QUEUE =
 struct
 
 (* tasks *)
@@ -46,9 +50,11 @@
 (* groups *)
 
 datatype group = Group of serial * bool ref;
+fun eq_group (Group (gid1, _), Group (gid2, _)) = gid1 = gid2;
 
 fun new_group () = Group (serial (), ref true);
 
+fun is_valid (Group (_, ref ok)) = ok;
 fun invalidate_group (Group (_, ok)) = ok := false;
 
 fun str_of_group (Group (i, ref ok)) =
@@ -58,14 +64,14 @@
 (* jobs *)
 
 datatype job =
-  Job of bool -> bool |
+  Job of (bool -> bool) list |
   Running of Thread.thread;
 
 type jobs = (group * job) Task_Graph.T;
 
 fun get_group (jobs: jobs) task = #1 (Task_Graph.get_node jobs task);
 fun get_job (jobs: jobs) task = #2 (Task_Graph.get_node jobs task);
-fun map_job task f (jobs: jobs) = Task_Graph.map_node task (apsnd f) jobs;
+fun set_job task job (jobs: jobs) = Task_Graph.map_node task (fn (group, _) => (group, job)) jobs;
 
 fun add_job task dep (jobs: jobs) =
   Task_Graph.add_edge (dep, task) jobs handle Task_Graph.UNDEF _ => jobs;
@@ -93,9 +99,14 @@
     val task = new_task pri;
     val groups' = Inttab.cons_list (gid, task) groups;
     val jobs' = jobs
-      |> Task_Graph.new_node (task, (group, Job job)) |> fold (add_job task) deps;
+      |> Task_Graph.new_node (task, (group, Job [job])) |> fold (add_job task) deps;
   in (task, make_queue groups' jobs') end;
 
+fun extend task job (Queue {groups, jobs}) =
+  (case try (get_job jobs) task of
+    SOME (Job list) => SOME (make_queue groups (set_job task (Job (job :: list)) jobs))
+  | _ => NONE);
+
 fun depend deps task (Queue {groups, jobs}) =
   make_queue groups (fold (add_job_acyclic task) deps jobs);
 
@@ -106,14 +117,13 @@
 
 fun dequeue_result NONE queue = (NONE, queue)
   | dequeue_result (SOME (result as (task, _, _))) (Queue {groups, jobs}) =
-      (SOME result, make_queue groups (map_job task (K (Running (Thread.self ()))) jobs));
+      (SOME result, make_queue groups (set_job task (Running (Thread.self ())) jobs));
 
 in
 
 fun dequeue (queue as Queue {jobs, ...}) =
   let
-    fun ready (task, ((group as Group (_, ref ok), Job job), ([], _))) =
-          SOME (task, group, (fn () => job ok))
+    fun ready (task, ((group, Job list), ([], _))) = SOME (task, group, rev list)
       | ready _ = NONE;
   in dequeue_result (Task_Graph.get_first ready jobs) queue end;
 
@@ -123,8 +133,8 @@
 
     fun ready task =
       (case Task_Graph.get_node jobs task of
-        (group as Group (_, ref ok), Job job) =>
-          if null (Task_Graph.imm_preds jobs task) then SOME (task, group, (fn () => job ok))
+        (group, Job list) =>
+          if null (Task_Graph.imm_preds jobs task) then SOME (task, group, rev list)
           else NONE
       | _ => NONE);
   in
@@ -150,16 +160,26 @@
   | NONE => ());
 
 
-(* misc operations *)
+(* termination *)
 
 fun cancel (Queue {groups, jobs, ...}) (group as Group (gid, _)) =
   let
     val _ = invalidate_group group;
     val tasks = Inttab.lookup_list groups gid;
-    val running = fold (get_job jobs #> (fn Running thread => cons thread | _ => I)) tasks [];
+    val running = fold (get_job jobs #> (fn Running t => insert Thread.equal t | _ => I)) tasks [];
     val _ = List.app SimpleThread.interrupt running;
   in null running end;
 
+fun cancel_all (Queue {jobs, ...}) =
+  let
+    fun cancel_job (group, job) (groups, running) =
+      (invalidate_group group;
+        (case job of Running t => (insert eq_group group groups, insert Thread.equal t running)
+        | _ => (groups, running)));
+    val (groups, running) = Task_Graph.fold (cancel_job o #1 o #2) jobs ([], []);
+    val _ = List.app SimpleThread.interrupt running;
+  in groups end;
+
 fun finish task (Queue {groups, jobs}) =
   let
     val Group (gid, _) = get_group jobs task;
--- a/src/Pure/General/binding.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/General/binding.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -26,6 +26,8 @@
   val base_name: T -> string
   val pos_of: T -> Position.T
   val dest: T -> (string * bool) list * string
+  val separator: string
+  val is_qualified: string -> bool
   val display: T -> string
 end
 
@@ -39,6 +41,17 @@
 val unique_names = ref true;
 
 
+(** qualification **)
+
+val separator = ".";
+val is_qualified = exists_string (fn s => s = separator);
+
+fun reject_qualified kind s =
+  if is_qualified s then
+    error ("Attempt to declare qualified " ^ kind ^ " " ^ quote s)
+  else s;
+
+
 (** binding representation **)
 
 datatype T = Binding of ((string * bool) list * string) * Position.T;
@@ -54,13 +67,14 @@
 
 fun qualify_base path name =
   if path = "" orelse name = "" then name
-  else path ^ "." ^ name;
+  else path ^ separator ^ name;
 
 val qualify = map_base o qualify_base;
   (*FIXME should all operations on bare names move here from name_space.ML ?*)
 
 fun add_prefix sticky "" b = b
-  | add_prefix sticky prfx b = (map_binding o apfst) (cons (prfx, sticky)) b;
+  | add_prefix sticky prfx b = (map_binding o apfst)
+      (cons ((*reject_qualified "prefix"*) prfx, sticky)) b;
 
 fun map_prefix f (Binding ((prefix, name), pos)) =
   f prefix (name_pos (name, pos));
--- a/src/Pure/General/buffer.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/General/buffer.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/General/buffer.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Efficient text buffers.
@@ -10,7 +9,6 @@
   type T
   val empty: T
   val add: string -> T -> T
-  val add_substring: substring -> T -> T
   val markup: Markup.T -> (T -> T) -> T -> T
   val content: T -> string
   val output: T -> TextIO.outstream -> unit
@@ -19,46 +17,18 @@
 structure Buffer: BUFFER =
 struct
 
-(* datatype *)
-
-datatype T =
-    EmptyBuffer
-  | String of string * T
-  | Substring of substring * T;
+datatype T = Buffer of string list;
 
-val empty = EmptyBuffer;
-
+val empty = Buffer [];
 
-(* add content *)
-
-fun add s buf = if s = "" then buf else String (s, buf);
-fun add_substring s buf = if Substring.isEmpty s then buf else Substring (s, buf);
+fun add "" buf = buf
+  | add x (Buffer xs) = Buffer (x :: xs);
 
 fun markup m body =
   let val (bg, en) = Markup.output m
   in add bg #> body #> add en end;
 
-
-(* cumulative content *)
-
-fun rev_content EmptyBuffer acc = acc
-  | rev_content (String (s, buf)) acc = rev_content buf (s :: acc)
-  | rev_content (Substring (s, buf)) acc = rev_content buf (Substring.string s :: acc);
-
-fun content buf = implode (rev_content buf []);
-
-
-(* file output *)
-
-fun rev_buffer EmptyBuffer acc = acc
-  | rev_buffer (String (s, buf)) acc = rev_buffer buf (String (s, acc))
-  | rev_buffer (Substring (s, buf)) acc = rev_buffer buf (Substring (s, acc));
-
-fun output buffer stream =
-  let
-    fun rev_output EmptyBuffer = ()
-      | rev_output (String (s, buf)) = (TextIO.output (stream, s); rev_output buf)
-      | rev_output (Substring (s, buf)) = (TextIO.outputSubstr (stream, s); rev_output buf);
-  in rev_output (rev_buffer buffer empty) end;
+fun content (Buffer xs) = implode (rev xs);
+fun output (Buffer xs) stream = List.app (fn s => TextIO.output (stream, s)) (rev xs);
 
 end;
--- a/src/Pure/General/lazy.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/General/lazy.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/General/lazy.ML
-    ID:         $Id$
     Author:     Florian Haftmann and Makarius, TU Muenchen
 
 Lazy evaluation with memoing.  Concurrency may lead to multiple
@@ -13,6 +12,7 @@
   val lazy: (unit -> 'a) -> 'a lazy
   val value: 'a -> 'a lazy
   val peek: 'a lazy -> 'a Exn.result option
+  val force_result: 'a lazy -> 'a Exn.result
   val force: 'a lazy -> 'a
   val map_force: ('a -> 'a) -> 'a lazy -> 'a lazy
 end
@@ -41,7 +41,7 @@
 
 (* force result *)
 
-fun force r =
+fun force_result r =
   let
     (*potentially concurrent evaluation*)
     val res =
@@ -53,7 +53,9 @@
       (case ! r of
         Lazy _ => (r := Result res; res)
       | Result res' => res'));
-  in Exn.release res' end;
+  in res' end;
+
+fun force r = Exn.release (force_result r);
 
 fun map_force f = value o f o force;
 
--- a/src/Pure/General/markup.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/General/markup.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/General/markup.ML
-    ID:         $Id$
     Author:     Makarius
 
 Common markup elements.
@@ -52,11 +51,9 @@
   val skolemN: string val skolem: T
   val boundN: string val bound: T
   val varN: string val var: T
-  val numN: string val num: T
-  val floatN: string val float: T
-  val xnumN: string val xnum: T
-  val xstrN: string val xstr: T
+  val numeralN: string val numeral: T
   val literalN: string val literal: T
+  val inner_stringN: string val inner_string: T
   val inner_commentN: string val inner_comment: T
   val sortN: string val sort: T
   val typN: string val typ: T
@@ -88,24 +85,18 @@
   val subgoalN: string val subgoal: T
   val sendbackN: string val sendback: T
   val hiliteN: string val hilite: T
+  val taskN: string
   val unprocessedN: string val unprocessed: T
-  val runningN: string val running: T
+  val runningN: string val running: string -> T
   val failedN: string val failed: T
   val finishedN: string val finished: T
   val disposedN: string val disposed: T
+  val editsN: string val edits: string -> T
+  val editN: string val edit: string -> string -> T
   val pidN: string
   val sessionN: string
-  val classN: string
-  val messageN: string val message: string -> T
   val promptN: string val prompt: T
-  val writelnN: string
-  val priorityN: string
-  val tracingN: string
-  val warningN: string
-  val errorN: string
-  val debugN: string
-  val initN: string
-  val statusN: string
+  val no_output: output * output
   val default_output: T -> output * output
   val add_mode: string -> (T -> output * output) -> unit
   val output: T -> output * output
@@ -203,11 +194,9 @@
 val (skolemN, skolem) = markup_elem "skolem";
 val (boundN, bound) = markup_elem "bound";
 val (varN, var) = markup_elem "var";
-val (numN, num) = markup_elem "num";
-val (floatN, float) = markup_elem "float";
-val (xnumN, xnum) = markup_elem "xnum";
-val (xstrN, xstr) = markup_elem "xstr";
+val (numeralN, numeral) = markup_elem "numeral";
 val (literalN, literal) = markup_elem "literal";
+val (inner_stringN, inner_string) = markup_elem "inner_string";
 val (inner_commentN, inner_comment) = markup_elem "inner_comment";
 
 val (sortN, sort) = markup_elem "sort";
@@ -263,36 +252,36 @@
 
 (* command status *)
 
+val taskN = "task";
+
 val (unprocessedN, unprocessed) = markup_elem "unprocessed";
-val (runningN, running) = markup_elem "running";
+val (runningN, running) = markup_string "running" taskN;
 val (failedN, failed) = markup_elem "failed";
 val (finishedN, finished) = markup_elem "finished";
 val (disposedN, disposed) = markup_elem "disposed";
 
 
+(* interactive documents *)
+
+val (editsN, edits) = markup_string "edits" idN;
+
+val editN = "edit";
+fun edit id state_id : T = (editN, [(idN, id), (stateN, state_id)]);
+
+
 (* messages *)
 
 val pidN = "pid";
 val sessionN = "session";
 
-val classN = "class";
-val (messageN, message) = markup_string "message" classN;
-
 val (promptN, prompt) = markup_elem "prompt";
 
-val writelnN = "writeln";
-val priorityN = "priority";
-val tracingN = "tracing";
-val warningN = "warning";
-val errorN = "error";
-val debugN = "debug";
-val initN = "init";
-val statusN = "status";
 
 
 (* print mode operations *)
 
-fun default_output (_: T) = ("", "");
+val no_output = ("", "");
+fun default_output (_: T) = no_output;
 
 local
   val default = {output = default_output};
@@ -304,7 +293,7 @@
     the_default default (Library.get_first (Symtab.lookup (! modes)) (print_mode_value ()));
 end;
 
-fun output m = if is_none m then ("", "") else #output (get_mode ()) m;
+fun output m = if is_none m then no_output else #output (get_mode ()) m;
 
 val enclose = output #-> Library.enclose;
 
--- a/src/Pure/General/markup.scala	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/General/markup.scala	Mon Jan 19 21:20:18 2009 +0100
@@ -110,6 +110,8 @@
 
   /* command status */
 
+  val TASK = "task"
+
   val UNPROCESSED = "unprocessed"
   val RUNNING = "running"
   val FAILED = "failed"
@@ -117,12 +119,33 @@
   val DISPOSED = "disposed"
 
 
+  /* interactive documents */
+
+  val EDITS = "edits"
+  val EDIT = "edit"
+
+
   /* messages */
 
   val PID = "pid"
   val SESSION = "session"
 
   val MESSAGE = "message"
+  val CLASS = "class"
+
+  val INIT = "init"
+  val STATUS = "status"
+  val WRITELN = "writeln"
+  val PRIORITY = "priority"
+  val TRACING = "tracing"
+  val WARNING = "warning"
+  val ERROR = "error"
+  val DEBUG = "debug"
+  val SYSTEM = "system"
+  val STDIN = "stdin"
+  val STDOUT = "stdout"
+  val SIGNAL = "signal"
+  val EXIT = "exit"
 
 
   /* content */
--- a/src/Pure/General/name_space.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/General/name_space.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -59,10 +59,10 @@
 (** long identifiers **)
 
 fun hidden name = "??." ^ name;
-val is_hidden = String.isPrefix "??."
+val is_hidden = String.isPrefix "??.";
 
-val separator = ".";
-val is_qualified = exists_string (fn s => s = separator);
+val separator = Binding.separator;
+val is_qualified = Binding.is_qualified;
 
 val implode_name = space_implode separator;
 val explode_name = space_explode separator;
--- a/src/Pure/General/position.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/General/position.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,8 +1,7 @@
 (*  Title:      Pure/General/position.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
-Source positions: counting Isabelle symbols -- starting from 1.
+Source positions: counting Isabelle symbols, starting from 1.
 *)
 
 signature POSITION =
@@ -19,6 +18,7 @@
   val file: string -> T
   val line: int -> T
   val line_file: int -> string -> T
+  val id: string -> T
   val get_id: T -> string option
   val put_id: string -> T -> T
   val of_properties: Properties.T -> T
@@ -86,6 +86,7 @@
 val none = Pos ((0, 0, 0), []);
 val start = Pos ((1, 1, 1), []);
 
+
 fun file_name "" = []
   | file_name name = [(Markup.fileN, name)];
 
@@ -95,11 +96,14 @@
 fun line i = line_file i "";
 
 
-(* markup properties *)
+fun id id = Pos ((0, 0, 1), [(Markup.idN, id)]);
 
 fun get_id (Pos (_, props)) = Properties.get props Markup.idN;
 fun put_id id (Pos (count, props)) = Pos (count, Properties.put (Markup.idN, id) props);
 
+
+(* markup properties *)
+
 fun of_properties props =
   let
     fun get name =
--- a/src/Pure/General/symbol.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/General/symbol.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -65,6 +65,7 @@
   val bump_string: string -> string
   val length: symbol list -> int
   val xsymbolsN: string
+  val output: string -> output * int
 end;
 
 structure Symbol: SYMBOL =
@@ -175,21 +176,22 @@
   ord space <= ord c andalso ord c <= ord "~" andalso c <> "." andalso c <> ">"
   orelse ord c >= 128;
 
-fun encode_raw str =
-  let
-    val raw0 = enclose "\\<^raw:" ">";
-    val raw1 = raw0 o implode;
-    val raw2 = enclose "\\<^raw" ">" o string_of_int o ord;
-
-    fun encode cs = enc (Library.take_prefix raw_chr cs)
-    and enc ([], []) = []
-      | enc (cs, []) = [raw1 cs]
-      | enc ([], d :: ds) = raw2 d :: encode ds
-      | enc (cs, d :: ds) = raw1 cs :: raw2 d :: encode ds;
-  in
-    if exists_string (not o raw_chr) str then implode (encode (explode str))
-    else raw0 str
-  end;
+fun encode_raw "" = ""
+  | encode_raw str =
+      let
+        val raw0 = enclose "\\<^raw:" ">";
+        val raw1 = raw0 o implode;
+        val raw2 = enclose "\\<^raw" ">" o string_of_int o ord;
+    
+        fun encode cs = enc (Library.take_prefix raw_chr cs)
+        and enc ([], []) = []
+          | enc (cs, []) = [raw1 cs]
+          | enc ([], d :: ds) = raw2 d :: encode ds
+          | enc (cs, d :: ds) = raw1 cs :: raw2 d :: encode ds;
+      in
+        if exists_string (not o raw_chr) str then implode (encode (explode str))
+        else raw0 str
+      end;
 
 
 (* diagnostics *)
@@ -519,9 +521,9 @@
 
 
 
-(** xsymbols **)
+(** symbol output **)
 
-val xsymbolsN = "xsymbols";
+(* length *)
 
 fun sym_len s =
   if not (is_printable s) then (0: int)
@@ -532,8 +534,16 @@
 
 fun sym_length ss = fold (fn s => fn n => sym_len s + n) ss 0;
 
+
+(* print mode *)
+
+val xsymbolsN = "xsymbols";
+
+fun output s = (s, sym_length (sym_explode s));
+
+
 (*final declarations of this structure!*)
+val explode = sym_explode;
 val length = sym_length;
-val explode = sym_explode;
 
 end;
--- a/src/Pure/General/xml.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/General/xml.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -79,7 +79,7 @@
   end;
 
 fun output_markup (markup as (name, atts)) =
-  if Markup.is_none markup then ("", "")
+  if Markup.is_none markup then Markup.no_output
   else (enclose "<" ">" (elem name atts), enclose "</" ">" name);
 
 
--- a/src/Pure/General/yxml.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/General/yxml.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -42,7 +42,7 @@
 (* output *)
 
 fun output_markup (markup as (name, atts)) =
-  if Markup.is_none markup then ("", "")
+  if Markup.is_none markup then Markup.no_output
   else (XY ^ name ^ implode (map (fn (a, x) => Y ^ a ^ "=" ^ x) atts) ^ X, XYX);
 
 fun element name atts body =
--- a/src/Pure/General/yxml.scala	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/General/yxml.scala	Mon Jan 19 21:20:18 2009 +0100
@@ -59,7 +59,7 @@
     val s = source.toString
     val i = s.indexOf('=')
     if (i <= 0) err_attribute()
-    (s.substring(0, i), s.substring(i + 1))
+    (s.substring(0, i).intern, s.substring(i + 1))
   }
 
 
@@ -91,7 +91,7 @@
       if (chunk == Y_string) pop()
       else {
         chunks(Y, chunk).toList match {
-          case "" :: name :: atts => push(name.toString, atts.map(parse_attrib))
+          case "" :: name :: atts => push(name.toString.intern, atts.map(parse_attrib))
           case txts => for (txt <- txts) add(XML.Text(txt.toString))
         }
       }
@@ -109,12 +109,21 @@
       case _ => err("multiple results")
     }
 
+
+  /* failsafe parsing */
+
+  private def markup_failsafe(source: CharSequence) =
+    XML.Elem (Markup.BAD, Nil,
+      List(XML.Text(source.toString.replace(X_string, "\\<^X>").replace(Y_string, "\\<^Y>"))))
+
+  def parse_body_failsafe(source: CharSequence) = {
+    try { parse_body(source) }
+    catch { case _: RuntimeException => List(markup_failsafe(source)) }
+  }
+
   def parse_failsafe(source: CharSequence) = {
     try { parse(source) }
-    catch {
-      case _: RuntimeException => XML.Elem (Markup.BAD, Nil,
-        List(XML.Text(source.toString.replace(X_string, "\\<^X>").replace(Y_string, "\\<^Y>"))))
-    }
+    catch { case _: RuntimeException => markup_failsafe(source) }
   }
 
 }
--- a/src/Pure/IsaMakefile	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/IsaMakefile	Mon Jan 19 21:20:18 2009 +0100
@@ -26,28 +26,29 @@
   Concurrent/par_list_dummy.ML Concurrent/simple_thread.ML		\
   Concurrent/synchronized.ML Concurrent/task_queue.ML General/ROOT.ML	\
   General/alist.ML General/balanced_tree.ML General/basics.ML		\
-  General/buffer.ML General/file.ML General/graph.ML General/heap.ML	\
-  General/integer.ML General/lazy.ML General/markup.ML			\
-  General/name_space.ML General/ord_list.ML General/output.ML		\
-  General/path.ML General/position.ML General/pretty.ML			\
-  General/print_mode.ML General/properties.ML General/queue.ML		\
-  General/scan.ML General/secure.ML General/seq.ML General/source.ML	\
-  General/stack.ML General/symbol.ML General/symbol_pos.ML		\
-  General/table.ML General/url.ML General/xml.ML General/yxml.ML	\
-  Isar/ROOT.ML Isar/antiquote.ML Isar/args.ML Isar/attrib.ML		\
-  Isar/auto_bind.ML Isar/calculation.ML Isar/class.ML Isar/code.ML	\
-  Isar/code_unit.ML Isar/constdefs.ML Isar/context_rules.ML		\
-  Isar/element.ML Isar/expression.ML Isar/find_theorems.ML		\
-  Isar/instance.ML Isar/isar.ML Isar/isar_cmd.ML Isar/isar_syn.ML	\
+  General/binding.ML General/buffer.ML General/file.ML			\
+  General/graph.ML General/heap.ML General/integer.ML General/lazy.ML	\
+  General/markup.ML General/name_space.ML General/ord_list.ML		\
+  General/output.ML General/path.ML General/position.ML			\
+  General/pretty.ML General/print_mode.ML General/properties.ML		\
+  General/queue.ML General/scan.ML General/secure.ML General/seq.ML	\
+  General/source.ML General/stack.ML General/symbol.ML			\
+  General/symbol_pos.ML General/table.ML General/url.ML General/xml.ML	\
+  General/yxml.ML Isar/ROOT.ML Isar/antiquote.ML Isar/args.ML		\
+  Isar/attrib.ML Isar/auto_bind.ML Isar/calculation.ML Isar/class.ML	\
+  Isar/class_target.ML Isar/code.ML Isar/code_unit.ML			\
+  Isar/constdefs.ML Isar/context_rules.ML Isar/element.ML		\
+  Isar/expression.ML Isar/find_theorems.ML Isar/isar.ML			\
+  Isar/isar_document.ML Isar/isar_cmd.ML Isar/isar_syn.ML		\
   Isar/local_defs.ML Isar/local_syntax.ML Isar/local_theory.ML		\
-  Isar/locale.ML Isar/method.ML Isar/net_rules.ML Isar/new_locale.ML	\
+  Isar/locale.ML Isar/method.ML Isar/net_rules.ML Isar/old_locale.ML	\
   Isar/object_logic.ML Isar/obtain.ML Isar/outer_keyword.ML		\
   Isar/outer_lex.ML Isar/outer_parse.ML Isar/outer_syntax.ML		\
   Isar/overloading.ML Isar/proof.ML Isar/proof_context.ML		\
   Isar/proof_display.ML Isar/proof_node.ML Isar/rule_cases.ML		\
   Isar/rule_insts.ML Isar/session.ML Isar/skip_proof.ML			\
-  Isar/spec_parse.ML Isar/specification.ML Isar/subclass.ML		\
-  Isar/theory_target.ML Isar/toplevel.ML ML-Systems/alice.ML		\
+  Isar/spec_parse.ML Isar/specification.ML Isar/theory_target.ML	\
+  Isar/toplevel.ML Isar/value_parse.ML ML-Systems/alice.ML		\
   ML-Systems/exn.ML ML-Systems/install_pp_polyml.ML			\
   ML-Systems/ml_name_space.ML ML-Systems/multithreading.ML		\
   ML-Systems/mosml.ML ML-Systems/multithreading_polyml.ML		\
@@ -73,17 +74,18 @@
   Syntax/printer.ML Syntax/simple_syntax.ML Syntax/syn_ext.ML		\
   Syntax/syn_trans.ML Syntax/syntax.ML Syntax/type_ext.ML Thy/html.ML	\
   Thy/latex.ML Thy/present.ML Thy/term_style.ML Thy/thm_deps.ML		\
-  Thy/thy_edit.ML Thy/thy_header.ML Thy/thy_info.ML Thy/thy_load.ML	\
-  Thy/thy_output.ML Tools/ROOT.ML Tools/invoke.ML			\
+  Thy/thy_header.ML Thy/thy_info.ML Thy/thy_load.ML Thy/thy_output.ML	\
+  Thy/thy_syntax.ML Tools/ROOT.ML Tools/invoke.ML			\
   Tools/isabelle_process.ML Tools/named_thms.ML Tools/xml_syntax.ML	\
   assumption.ML axclass.ML codegen.ML config.ML conjunction.ML		\
   consts.ML context.ML context_position.ML conv.ML defs.ML display.ML	\
   drule.ML envir.ML facts.ML goal.ML interpretation.ML library.ML	\
   logic.ML meta_simplifier.ML more_thm.ML morphism.ML name.ML net.ML	\
-  old_goals.ML pattern.ML primitive_defs.ML proofterm.ML pure_setup.ML	\
-  pure_thy.ML search.ML sign.ML simplifier.ML sorts.ML subgoal.ML	\
-  tactic.ML tctical.ML term.ML term_subst.ML theory.ML thm.ML type.ML	\
-  type_infer.ML unify.ML variable.ML ../Tools/quickcheck.ML
+  old_goals.ML old_term.ML pattern.ML primitive_defs.ML proofterm.ML	\
+  pure_setup.ML pure_thy.ML search.ML sign.ML simplifier.ML sorts.ML	\
+  subgoal.ML tactic.ML tctical.ML term.ML term_ord.ML term_subst.ML	\
+  theory.ML thm.ML type.ML type_infer.ML unify.ML variable.ML		\
+  ../Tools/quickcheck.ML
 	@./mk
 
 
@@ -124,6 +126,7 @@
 SCALA_FILES = General/event_bus.scala General/markup.scala		\
   General/position.scala General/swing.scala General/symbol.scala	\
   General/xml.scala General/yxml.scala Isar/isar.scala			\
+  Isar/isar_document.scala Isar/outer_keyword.scala			\
   Thy/thy_header.scala Tools/isabelle_process.scala			\
   Tools/isabelle_syntax.scala Tools/isabelle_system.scala
 
--- a/src/Pure/Isar/ROOT.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Isar/ROOT.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/ROOT.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Isar -- Intelligible Semi-Automated Reasoning for Isabelle.
@@ -21,8 +20,9 @@
 
 (*outer syntax*)
 use "outer_lex.ML";
+use "outer_keyword.ML";
 use "outer_parse.ML";
-use "outer_keyword.ML";
+use "value_parse.ML";
 use "args.ML";
 use "antiquote.ML";
 use "../ML/ml_context.ML";
@@ -53,13 +53,12 @@
 (*local theories and targets*)
 use "local_theory.ML";
 use "overloading.ML";
+use "old_locale.ML";
 use "locale.ML";
-use "new_locale.ML";
+use "class_target.ML";
+use "theory_target.ML";
 use "expression.ML";
 use "class.ML";
-use "theory_target.ML";
-use "instance.ML";
-use "subclass.ML";
 
 (*complex proof machineries*)
 use "../simplifier.ML";
@@ -80,12 +79,13 @@
 (*theory syntax*)
 use "../Thy/term_style.ML";
 use "../Thy/thy_output.ML";
-use "../Thy/thy_edit.ML";
+use "../Thy/thy_syntax.ML";
 use "../old_goals.ML";
 use "outer_syntax.ML";
 use "../Thy/thy_info.ML";
 use "session.ML";
 use "isar.ML";
+use "isar_document.ML";
 
 (*theory and proof operations*)
 use "rule_insts.ML";
--- a/src/Pure/Isar/class.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Isar/class.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,633 +1,144 @@
-(*  Title:      Pure/Isar/class.ML
-    ID:         $Id$
+(*  Title:      Pure/Isar/ML
     Author:     Florian Haftmann, TU Muenchen
 
-Type classes derived from primitive axclasses and locales.
+Type classes derived from primitive axclasses and locales - interfaces
 *)
 
 signature CLASS =
 sig
-  (*classes*)
-  val class: bstring -> class list -> Element.context_i list
-    -> theory -> string * Proof.context
-  val class_cmd: bstring -> xstring list -> Element.context list
-    -> theory -> string * Proof.context
-
-  val init: class -> theory -> Proof.context
-  val declare: class -> Properties.T
-    -> (string * mixfix) * term -> theory -> theory
-  val abbrev: class -> Syntax.mode -> Properties.T
-    -> (string * mixfix) * term -> theory -> theory
-  val refresh_syntax: class -> Proof.context -> Proof.context
-
-  val intro_classes_tac: thm list -> tactic
-  val default_intro_tac: Proof.context -> thm list -> tactic
-  val prove_subclass: class * class -> thm option -> theory -> theory
-
-  val class_prefix: string -> string
-  val is_class: theory -> class -> bool
-  val these_params: theory -> sort -> (string * (class * (string * typ))) list
-  val print_classes: theory -> unit
+  include CLASS_TARGET
+    (*FIXME the split into class_target.ML, theory_target.ML and
+      class.ML is artificial*)
 
-  (*instances*)
-  val init_instantiation: string list * (string * sort) list * sort
-    -> theory -> local_theory
-  val instantiation_instance: (local_theory -> local_theory)
-    -> local_theory -> Proof.state
-  val prove_instantiation_instance: (Proof.context -> tactic)
-    -> local_theory -> local_theory
-  val prove_instantiation_exit: (Proof.context -> tactic)
-    -> local_theory -> theory
-  val prove_instantiation_exit_result: (morphism -> 'a -> 'b)
-    -> (Proof.context -> 'b -> tactic) -> 'a -> local_theory -> 'b * theory
-  val conclude_instantiation: local_theory -> local_theory
-  val instantiation_param: local_theory -> string -> string option
-  val confirm_declaration: string -> local_theory -> local_theory
-  val pretty_instantiation: local_theory -> Pretty.T
-  val type_name: string -> string
-
-  (*old axclass layer*)
-  val axclass_cmd: bstring * xstring list
-    -> (Attrib.binding * string list) list
-    -> theory -> class * theory
-  val classrel_cmd: xstring * xstring -> theory -> Proof.state
-
-  (*old instance layer*)
-  val instance_arity: (theory -> theory) -> arity -> theory -> Proof.state
-  val instance_arity_cmd: bstring * xstring list * xstring -> theory -> Proof.state
+  val class: bstring -> class list -> Element.context_i list
+    -> theory -> string * local_theory
+  val class_cmd: bstring -> xstring list -> Element.context list
+    -> theory -> string * local_theory
+  val prove_subclass: tactic -> class -> local_theory -> local_theory
+  val subclass: class -> local_theory -> Proof.state
+  val subclass_cmd: xstring -> local_theory -> Proof.state
 end;
 
 structure Class : CLASS =
 struct
 
-(*temporary adaption code to mediate between old and new locale code*)
-
-structure Old_Locale =
-struct
-
-val intro_locales_tac = Locale.intro_locales_tac; (*already forked!*)
-
-val interpretation = Locale.interpretation;
-val interpretation_in_locale = Locale.interpretation_in_locale;
-val get_interpret_morph = Locale.get_interpret_morph;
-val Locale = Locale.Locale;
-val extern = Locale.extern;
-val intros = Locale.intros;
-val dests = Locale.dests;
-val init = Locale.init;
-val Merge = Locale.Merge;
-val parameters_of_expr = Locale.parameters_of_expr;
-val empty = Locale.empty;
-val cert_expr = Locale.cert_expr;
-val read_expr = Locale.read_expr;
-val parameters_of = Locale.parameters_of;
-val add_locale = Locale.add_locale;
-
-end;
-
-structure New_Locale =
-struct
-
-val intro_locales_tac = Locale.intro_locales_tac; (*already forked!*)
-
-val interpretation = Locale.interpretation; (*!*)
-val interpretation_in_locale = Locale.interpretation_in_locale; (*!*)
-val get_interpret_morph = Locale.get_interpret_morph; (*!*)
-fun Locale loc = ([(loc, ("", Expression.Positional []))], []);
-val extern = NewLocale.extern;
-val intros = Locale.intros; (*!*)
-val dests = Locale.dests; (*!*)
-val init = NewLocale.init;
-fun Merge locs = (map (fn loc => (loc, ("", Expression.Positional []))) locs, []);
-val parameters_of_expr = Locale.parameters_of_expr; (*!*)
-val empty = ([], []);
-val cert_expr = Locale.cert_expr; (*!"*)
-val read_expr = Locale.read_expr; (*!"*)
-val parameters_of = NewLocale.params_of; (*why typ option?*)
-val add_locale = Expression.add_locale;
-
-end;
-
-structure Locale = Old_Locale;
-
-(*proper code again*)
-
-
-(** auxiliary **)
+open Class_Target;
 
-fun prove_interpretation tac prfx_atts expr inst =
-  Locale.interpretation I prfx_atts expr inst
-  ##> Proof.global_terminal_proof
-      (Method.Basic (fn ctxt => Method.SIMPLE_METHOD (tac ctxt), Position.none), NONE)
-  ##> ProofContext.theory_of;
-
-fun prove_interpretation_in tac after_qed (name, expr) =
-  Locale.interpretation_in_locale
-      (ProofContext.theory after_qed) (name, expr)
-  #> Proof.global_terminal_proof
-      (Method.Basic (K (Method.SIMPLE_METHOD tac), Position.none), NONE)
-  #> ProofContext.theory_of;
-
-val class_prefix = Logic.const_of_class o Sign.base_name;
-
-fun class_name_morphism class =
-  Binding.map_prefix (K (Binding.add_prefix false (class_prefix class)));
-
-fun activate_class_morphism thy class inst morphism =
-  Locale.get_interpret_morph thy (class_name_morphism class) (class, "") morphism class inst;
-
-fun prove_class_interpretation class inst consts hyp_facts def_facts thy =
-  let
-    val constraints = map (fn (class, c) => map_atyps (K (TFree (Name.aT,
-      [class]))) (Sign.the_const_type thy c)) consts;
-    val no_constraints = map (map_atyps (K (TFree (Name.aT, [])))) constraints;
-    fun add_constraint c T = Sign.add_const_constraint (c, SOME T);
-    fun tac ctxt = ALLGOALS (ProofContext.fact_tac (hyp_facts @ def_facts)
-      ORELSE' (fn n => SELECT_GOAL (Locale.intro_locales_tac false ctxt []) n));
-    val prfx = class_prefix class;
-  in
-    thy
-    |> fold2 add_constraint (map snd consts) no_constraints
-    |> prove_interpretation tac (class_name_morphism class) (Locale.Locale class)
-          (map SOME inst, map (pair (Attrib.empty_binding) o Thm.prop_of) def_facts)
-    ||> fold2 add_constraint (map snd consts) constraints
-  end;
-
-
-(** primitive axclass and instance commands **)
-
-fun axclass_cmd (class, raw_superclasses) raw_specs thy =
-  let
-    val ctxt = ProofContext.init thy;
-    val superclasses = map (Sign.read_class thy) raw_superclasses;
-    val name_atts = map ((apsnd o map) (Attrib.attribute thy) o fst)
-      raw_specs;
-    val axiomss = ProofContext.read_propp (ctxt, map (map (rpair []) o snd)
-          raw_specs)
-      |> snd
-      |> (map o map) fst;
-  in
-    AxClass.define_class (class, superclasses) []
-      (name_atts ~~ axiomss) thy
-  end;
+(** define classes **)
 
 local
 
-fun gen_instance mk_prop add_thm after_qed insts thy =
-  let
-    fun after_qed' results =
-      ProofContext.theory ((fold o fold) add_thm results #> after_qed);
-  in
-    thy
-    |> ProofContext.init
-    |> Proof.theorem_i NONE after_qed' ((map (fn t => [(t, [])])
-        o mk_prop thy) insts)
-  end;
-
-in
-
-val instance_arity =
-  gen_instance (Logic.mk_arities oo Sign.cert_arity) AxClass.add_arity;
-val instance_arity_cmd =
-  gen_instance (Logic.mk_arities oo Sign.read_arity) AxClass.add_arity I;
-val classrel =
-  gen_instance (single oo (Logic.mk_classrel oo AxClass.cert_classrel)) AxClass.add_classrel I;
-val classrel_cmd =
-  gen_instance (single oo (Logic.mk_classrel oo AxClass.read_classrel)) AxClass.add_classrel I;
-
-end; (*local*)
-
-
-(** class data **)
-
-datatype class_data = ClassData of {
-
-  (* static part *)
-  consts: (string * string) list
-    (*locale parameter ~> constant name*),
-  base_sort: sort,
-  inst: term list
-    (*canonical interpretation*),
-  morphism: Morphism.morphism * ((typ Vartab.table * typ list) * (term Vartab.table * term list))
-    (*morphism cookie of canonical interpretation*),
-  assm_intro: thm option,
-  of_class: thm,
-  axiom: thm option,
-  
-  (* dynamic part *)
-  defs: thm list,
-  operations: (string * (class * (typ * term))) list
-
-};
-
-fun rep_class_data (ClassData d) = d;
-fun mk_class_data ((consts, base_sort, inst, morphism, assm_intro, of_class, axiom),
-    (defs, operations)) =
-  ClassData { consts = consts, base_sort = base_sort, inst = inst,
-    morphism = morphism, assm_intro = assm_intro, of_class = of_class, axiom = axiom,
-    defs = defs, operations = operations };
-fun map_class_data f (ClassData { consts, base_sort, inst, morphism, assm_intro,
-    of_class, axiom, defs, operations }) =
-  mk_class_data (f ((consts, base_sort, inst, morphism, assm_intro, of_class, axiom),
-    (defs, operations)));
-fun merge_class_data _ (ClassData { consts = consts,
-    base_sort = base_sort, inst = inst, morphism = morphism, assm_intro = assm_intro,
-    of_class = of_class, axiom = axiom, defs = defs1, operations = operations1 },
-  ClassData { consts = _, base_sort = _, inst = _, morphism = _, assm_intro = _,
-    of_class = _, axiom = _, defs = defs2, operations = operations2 }) =
-  mk_class_data ((consts, base_sort, inst, morphism, assm_intro, of_class, axiom),
-    (Thm.merge_thms (defs1, defs2),
-      AList.merge (op =) (K true) (operations1, operations2)));
-
-structure ClassData = TheoryDataFun
-(
-  type T = class_data Graph.T
-  val empty = Graph.empty;
-  val copy = I;
-  val extend = I;
-  fun merge _ = Graph.join merge_class_data;
-);
-
-
-(* queries *)
-
-val lookup_class_data = Option.map rep_class_data oo try o Graph.get_node o ClassData.get;
-
-fun the_class_data thy class = case lookup_class_data thy class
- of NONE => error ("Undeclared class " ^ quote class)
-  | SOME data => data;
-
-val is_class = is_some oo lookup_class_data;
-
-val ancestry = Graph.all_succs o ClassData.get;
-
-fun the_inst thy = #inst o the_class_data thy;
-
-fun these_params thy =
+fun calculate thy class sups base_sort param_map assm_axiom =
   let
-    fun params class =
-      let
-        val const_typs = (#params o AxClass.get_info thy) class;
-        val const_names = (#consts o the_class_data thy) class;
-      in
-        (map o apsnd)
-          (fn c => (class, (c, (the o AList.lookup (op =) const_typs) c))) const_names
-      end;
-  in maps params o ancestry thy end;
+    val empty_ctxt = ProofContext.init thy;
 
-fun these_assm_intros thy =
-  Graph.fold (fn (_, (data, _)) => fold (insert Thm.eq_thm)
-    ((the_list o #assm_intro o rep_class_data) data)) (ClassData.get thy) [];
-
-fun these_defs thy = maps (these o Option.map #defs o lookup_class_data thy) o ancestry thy;
-
-fun these_operations thy =
-  maps (#operations o the_class_data thy) o ancestry thy;
-
-fun morphism thy class =
-  let
-    val { inst, morphism, ... } = the_class_data thy class;
-  in activate_class_morphism thy class inst morphism end;
+    (* instantiation of canonical interpretation *)
+    (*FIXME inst_morph should be calculated manually and not instantiate constraint*)
+    val aT = TFree ("'a", base_sort);
+    val (([props], [(_, inst_morph)], export_morph), _) = empty_ctxt
+      |> Expression.cert_goal_expression ([(class, (("", false),
+           Expression.Named ((map o apsnd) Const param_map)))], []);
 
-fun print_classes thy =
-  let
-    val ctxt = ProofContext.init thy;
-    val algebra = Sign.classes_of thy;
-    val arities =
-      Symtab.empty
-      |> Symtab.fold (fn (tyco, arities) => fold (fn (class, _) =>
-           Symtab.map_default (class, []) (insert (op =) tyco)) arities)
-             ((#arities o Sorts.rep_algebra) algebra);
-    val the_arities = these o Symtab.lookup arities;
-    fun mk_arity class tyco =
-      let
-        val Ss = Sorts.mg_domain algebra tyco [class];
-      in Syntax.pretty_arity ctxt (tyco, Ss, [class]) end;
-    fun mk_param (c, ty) = Pretty.str (Sign.extern_const thy c ^ " :: "
-      ^ setmp show_sorts false (Syntax.string_of_typ ctxt o Type.strip_sorts) ty);
-    fun mk_entry class = (Pretty.block o Pretty.fbreaks o map_filter I) [
-      (SOME o Pretty.str) ("class " ^ Sign.extern_class thy class ^ ":"),
-      (SOME o Pretty.block) [Pretty.str "supersort: ",
-        (Syntax.pretty_sort ctxt o Sign.minimize_sort thy o Sign.super_classes thy) class],
-      if is_class thy class then (SOME o Pretty.str)
-        ("locale: " ^ Locale.extern thy class) else NONE,
-      ((fn [] => NONE | ps => (SOME o Pretty.block o Pretty.fbreaks)
-          (Pretty.str "parameters:" :: ps)) o map mk_param
-        o these o Option.map #params o try (AxClass.get_info thy)) class,
-      (SOME o Pretty.block o Pretty.breaks) [
-        Pretty.str "instances:",
-        Pretty.list "" "" (map (mk_arity class) (the_arities class))
-      ]
-    ]
-  in
-    (Pretty.writeln o Pretty.chunks o separate (Pretty.str "")
-      o map mk_entry o Sorts.all_classes) algebra
-  end;
-
-
-(* updaters *)
+    (* witness for canonical interpretation *)
+    val prop = try the_single props;
+    val wit = Option.map (fn prop => let
+        val sup_axioms = map_filter (fst o rules thy) sups;
+        val loc_intro_tac = case Locale.intros_of thy class
+          of (_, NONE) => all_tac
+           | (_, SOME intro) => ALLGOALS (Tactic.rtac intro);
+        val tac = loc_intro_tac
+          THEN ALLGOALS (ProofContext.fact_tac (sup_axioms @ the_list assm_axiom))
+      in Element.prove_witness empty_ctxt prop tac end) prop;
+    val axiom = Option.map Element.conclude_witness wit;
 
-fun add_class_data ((class, superclasses),
-    (params, base_sort, inst, phi, axiom, assm_intro, of_class)) thy =
-  let
-    val operations = map (fn (v_ty as (_, ty), (c, _)) =>
-      (c, (class, (ty, Free v_ty)))) params;
-    val add_class = Graph.new_node (class,
-        mk_class_data (((map o pairself) fst params, base_sort,
-          inst, phi, assm_intro, of_class, axiom), ([], operations)))
-      #> fold (curry Graph.add_edge class) superclasses;
-  in ClassData.map add_class thy end;
-
-fun register_operation class (c, (t, some_def)) thy =
-  let
-    val base_sort = (#base_sort o the_class_data thy) class;
-    val prep_typ = map_type_tvar
-      (fn (vi as (v, _), sort) => if Name.aT = v
-        then TFree (v, base_sort) else TVar (vi, sort));
-    val t' = map_types prep_typ t;
-    val ty' = Term.fastype_of t';
-  in
-    thy
-    |> (ClassData.map o Graph.map_node class o map_class_data o apsnd)
-      (fn (defs, operations) =>
-        (fold cons (the_list some_def) defs,
-          (c, (class, (ty', t'))) :: operations))
-  end;
-
-
-(** rule calculation, tactics and methods **)
-
-fun calculate_axiom thy sups base_sort assm_axiom param_map class =
-  case Locale.intros thy class
-   of (_, []) => assm_axiom
-    | (_, [intro]) =>
-      let
-        fun instantiate thy sort = Thm.instantiate ([pairself (Thm.ctyp_of thy o TVar o pair (Name.aT, 0))
-          (base_sort, sort)], map (fn (v, (c, ty)) => pairself (Thm.cterm_of thy)
-            (Var ((v, 0), map_atyps (fn _ => TVar ((Name.aT, 0), sort)) ty),
-              Const (c, map_atyps (fn _ => TVar ((Name.aT, 0), sort)) ty))) param_map);
-        val axiom_premises = map_filter (#axiom o the_class_data thy) sups
-          @ the_list assm_axiom;
-      in intro
-        |> instantiate thy [class]
-        |> (fn thm => thm OF axiom_premises)
-        |> Drule.standard'
-        |> Thm.close_derivation
-        |> SOME
-      end;
+    (* canonical interpretation *)
+    val base_morph = inst_morph
+      $> Morphism.binding_morphism
+           (Binding.add_prefix false (class_prefix class))
+      $> Element.satisfy_morphism (the_list wit);
+    val defs = these_defs thy sups;
+    val eq_morph = Element.eq_morphism thy defs;
+    val morph = base_morph $> eq_morph;
 
-fun calculate_rules thy phi sups base_sort param_map axiom class =
-  let
-    fun instantiate thy sort = Thm.instantiate ([pairself (Thm.ctyp_of thy o TVar o pair (Name.aT, 0))
-      (base_sort, sort)], map (fn (v, (c, ty)) => pairself (Thm.cterm_of thy)
-        (Var ((v, 0), map_atyps (fn _ => TVar ((Name.aT, 0), sort)) ty),
-          Const (c, map_atyps (fn _ => TVar ((Name.aT, 0), sort)) ty))) param_map);
-    val defs = these_defs thy sups;
-    val assm_intro = Locale.intros thy class
-      |> fst
-      |> map (instantiate thy base_sort)
-      |> map (MetaSimplifier.rewrite_rule defs)
-      |> map Thm.close_derivation
-      |> try the_single;
-    val fixate = Thm.instantiate
-      (map (pairself (Thm.ctyp_of thy)) [(TVar ((Name.aT, 0), []), TFree (Name.aT, base_sort)),
-        (TVar ((Name.aT, 0), base_sort), TFree (Name.aT, base_sort))], [])
-    val of_class_sups = if null sups
-      then map (fixate o Thm.class_triv thy) base_sort
-      else map (fixate o #of_class o the_class_data thy) sups;
-    val locale_dests = map Drule.standard' (Locale.dests thy class);
-    val num_trivs = case length locale_dests
-     of 0 => if is_none axiom then 0 else 1
-      | n => n;
-    val pred_trivs = if num_trivs = 0 then []
-      else the axiom
-        |> Thm.prop_of
-        |> (map_types o map_atyps o K) (TFree (Name.aT, base_sort))
-        |> (Thm.assume o Thm.cterm_of thy)
-        |> replicate num_trivs;
-    val axclass_intro = (#intro o AxClass.get_info thy) class;
-    val of_class = (fixate axclass_intro OF of_class_sups OF locale_dests OF pred_trivs)
-      |> Drule.standard'
-      |> Thm.close_derivation;
-  in (assm_intro, of_class) end;
+    (* assm_intro *)
+    fun prove_assm_intro thm = 
+      let
+        val prop = thm |> Thm.prop_of |> Logic.unvarify
+          |> Morphism.term (inst_morph $> eq_morph) 
+          |> (map_types o map_atyps) (K aT);
+        fun tac ctxt = LocalDefs.unfold_tac ctxt (map Thm.symmetric defs) (*FIXME*)
+          THEN ALLGOALS (ProofContext.fact_tac [thm]);
+      in Goal.prove_global thy [] [] prop (tac o #context) end;
+    val assm_intro = Option.map prove_assm_intro
+      (fst (Locale.intros_of thy class));
 
-fun prove_subclass (sub, sup) some_thm thy =
-  let
-    val of_class = (#of_class o the_class_data thy) sup;
-    val intro = case some_thm
-     of SOME thm => Drule.standard' (of_class OF [Drule.standard' thm])
-      | NONE => Thm.instantiate ([pairself (Thm.ctyp_of thy o TVar o pair (Name.aT, 0))
-          ([], [sub])], []) of_class;
-    val classrel = (intro OF (the_list o #axiom o the_class_data thy) sub)
-      |> Thm.close_derivation;
-  in
-    thy
-    |> AxClass.add_classrel classrel
-    |> prove_interpretation_in (ALLGOALS (ProofContext.fact_tac (the_list some_thm)))
-         I (sub, Locale.Locale sup)
-    |> ClassData.map (Graph.add_edge (sub, sup))
-  end;
-
-fun note_assm_intro class assm_intro thy =
-  thy
-  |> Sign.sticky_prefix (class_prefix class ^ "_" ^ AxClass.axiomsN)
-  |> PureThy.store_thm (AxClass.introN, assm_intro)
-  |> snd
-  |> Sign.restore_naming thy;
+    (* of_class *)
+    val of_class_prop_concl = Logic.mk_inclass (aT, class);
+    val of_class_prop = case prop of NONE => of_class_prop_concl
+      | SOME prop => Logic.mk_implies (Morphism.term inst_morph prop,
+          of_class_prop_concl) |> (map_types o map_atyps) (K aT)
+    val sup_of_classes = map (snd o rules thy) sups;
+    val loc_axiom_intros = map Drule.standard' (Locale.axioms_of thy class);
+    val axclass_intro = #intro (AxClass.get_info thy class);
+    val base_sort_trivs = Drule.sort_triv thy (aT, base_sort);
+    val tac = REPEAT (SOMEGOAL
+      (Tactic.match_tac (axclass_intro :: sup_of_classes
+         @ loc_axiom_intros @ base_sort_trivs)
+           ORELSE' Tactic.assume_tac));
+    val of_class = Goal.prove_global thy [] [] of_class_prop (K tac);
 
-fun intro_classes_tac facts st =
-  let
-    val thy = Thm.theory_of_thm st;
-    val classes = Sign.all_classes thy;
-    val class_trivs = map (Thm.class_triv thy) classes;
-    val class_intros = map_filter (try (#intro o AxClass.get_info thy)) classes;
-    val assm_intros = these_assm_intros thy;
-  in
-    Method.intros_tac (class_trivs @ class_intros @ assm_intros) facts st
-  end;
-
-fun default_intro_tac ctxt [] =
-      intro_classes_tac [] ORELSE NewLocale.intro_locales_tac true ctxt [] ORELSE
-      Locale.intro_locales_tac true ctxt []
-  | default_intro_tac _ _ = no_tac;
+  in (base_morph, morph, export_morph, axiom, assm_intro, of_class) end;
 
-fun default_tac rules ctxt facts =
-  HEADGOAL (Method.some_rule_tac rules ctxt facts) ORELSE
-    default_intro_tac ctxt facts;
-
-val _ = Context.>> (Context.map_theory
-  (Method.add_methods
-   [("intro_classes", Method.no_args (Method.METHOD intro_classes_tac),
-      "back-chain introduction rules of classes"),
-    ("default", Method.thms_ctxt_args (Method.METHOD oo default_tac),
-      "apply some intro/elim rule")]));
-
-
-(** classes and class target **)
-
-(* class context syntax *)
-
-fun synchronize_class_syntax sups base_sort ctxt =
+fun prep_class_spec prep_class process_decl thy raw_supclasses raw_elems =
   let
-    val thy = ProofContext.theory_of ctxt;
-    val algebra = Sign.classes_of thy;
-    val operations = these_operations thy sups;
-    fun subst_class_typ sort = map_type_tfree (K (TVar ((Name.aT, 0), sort)));
-    val primary_constraints =
-      (map o apsnd) (subst_class_typ base_sort o fst o snd) operations;
-    val secondary_constraints =
-      (map o apsnd) (fn (class, (ty, _)) => subst_class_typ [class] ty) operations;
-    fun declare_const (c, _) =
-      let val b = Sign.base_name c
-      in Sign.intern_const thy b = c ? Variable.declare_const (b, c) end;
-    fun improve (c, ty) = (case AList.lookup (op =) primary_constraints c
-     of SOME ty' => (case try (Type.raw_match (ty', ty)) Vartab.empty
-         of SOME tyenv => (case Vartab.lookup tyenv (Name.aT, 0)
-             of SOME (_, ty' as TVar (tvar as (vi, sort))) =>
-                  if TypeInfer.is_param vi
-                    andalso Sorts.sort_le algebra (base_sort, sort)
-                      then SOME (ty', TFree (Name.aT, base_sort))
-                      else NONE
-              | _ => NONE)
-          | NONE => NONE)
-      | NONE => NONE)
-    fun subst (c, ty) = Option.map snd (AList.lookup (op =) operations c);
-    val unchecks = map (fn (c, (_, (ty, t))) => (t, Const (c, ty))) operations;
-  in
-    ctxt
-    |> fold declare_const primary_constraints
-    |> Overloading.map_improvable_syntax (K (((primary_constraints, secondary_constraints),
-        (((improve, subst), true), unchecks)), false))
-    |> Overloading.set_primary_constraints
-  end;
-
-fun refresh_syntax class ctxt =
-  let
-    val thy = ProofContext.theory_of ctxt;
-    val base_sort = (#base_sort o the_class_data thy) class;
-  in synchronize_class_syntax [class] base_sort ctxt end;
-
-fun init_ctxt sups base_sort ctxt =
-  ctxt
-  |> Variable.declare_term
-      (Logic.mk_type (TFree (Name.aT, base_sort)))
-  |> synchronize_class_syntax sups base_sort
-  |> Overloading.add_improvable_syntax;
-
-fun init class thy =
-  thy
-  |> Locale.init class
-  |> init_ctxt [class] ((#base_sort o the_class_data thy) class);
-
-
-(* class target *)
-
-fun declare class pos ((c, mx), dict) thy =
-  let
-    val prfx = class_prefix class;
-    val thy' = thy |> Sign.add_path prfx;
-    val phi = morphism thy' class;
-    val inst = the_inst thy' class;
-    val params = map (apsnd fst o snd) (these_params thy' [class]);
-
-    val c' = Sign.full_bname thy' c;
-    val dict' = Morphism.term phi dict;
-    val dict_def = map_types Logic.unvarifyT dict';
-    val ty' = Term.fastype_of dict_def;
-    val ty'' = Type.strip_sorts ty';
-    val def_eq = Logic.mk_equals (Const (c', ty'), dict_def);
-    fun get_axiom thy = ((Thm.varifyT o Thm.axiom thy) c', thy);
-  in
-    thy'
-    |> Sign.declare_const pos ((Binding.name c, ty''), mx) |> snd
-    |> Thm.add_def false false (c, def_eq)
-    |>> Thm.symmetric
-    ||>> get_axiom
-    |-> (fn (def, def') => prove_class_interpretation class inst params [] [def]
-      #> snd
-        (*assumption: interpretation cookie does not change
-          by adding equations to interpretation*)
-      #> register_operation class (c', (dict', SOME (Thm.symmetric def')))
-      #> PureThy.store_thm (c ^ "_raw", def')
-      #> snd)
-    |> Sign.restore_naming thy
-    |> Sign.add_const_constraint (c', SOME ty')
-  end;
-
-fun abbrev class prmode pos ((c, mx), rhs) thy =
-  let
-    val prfx = class_prefix class;
-    val thy' = thy |> Sign.add_path prfx;
-
-    val unchecks = map (fn (c, (_, (ty, t))) => (t, Const (c, ty)))
-      (these_operations thy [class]);
-    val c' = Sign.full_bname thy' c;
-    val rhs' = Pattern.rewrite_term thy unchecks [] rhs;
-    val rhs'' = map_types Logic.varifyT rhs';
-    val ty' = Term.fastype_of rhs';
-  in
-    thy'
-    |> Sign.add_abbrev (#1 prmode) pos (Binding.name c, map_types Type.strip_sorts rhs'') |> snd
-    |> Sign.add_const_constraint (c', SOME ty')
-    |> Sign.notation true prmode [(Const (c', ty'), mx)]
-    |> not (#1 prmode = PrintMode.input) ? register_operation class (c', (rhs', NONE))
-    |> Sign.restore_naming thy
-  end;
-
-
-(* class definition *)
-
-local
-
-fun gen_class_spec prep_class process_expr thy raw_supclasses raw_elems =
-  let
+    (*FIXME 2009 simplify*)
     val supclasses = map (prep_class thy) raw_supclasses;
     val supsort = Sign.minimize_sort thy supclasses;
-    val sups = filter (is_class thy) supsort;
-    val supparam_names = map fst (these_params thy sups);
+    val (sups, bases) = List.partition (is_class thy) supsort;
+    val base_sort = if null sups then supsort else
+      Library.foldr (Sorts.inter_sort (Sign.classes_of thy))
+        (map (base_sort thy) sups, bases);
+    val supparams = (map o apsnd) (snd o snd) (these_params thy sups);
+    val supparam_names = map fst supparams;
     val _ = if has_duplicates (op =) supparam_names
       then error ("Duplicate parameter(s) in superclasses: "
         ^ (commas o map quote o duplicates (op =)) supparam_names)
       else ();
-    val base_sort = if null sups then supsort else
-      foldr1 (Sorts.inter_sort (Sign.classes_of thy))
-        (map (#base_sort o the_class_data thy) sups);
-    val suplocales = map Locale.Locale sups;
-    val supexpr = Locale.Merge suplocales;
-    val supparams = (map fst o Locale.parameters_of_expr thy) supexpr;
-    val mergeexpr = Locale.Merge suplocales;
+
+    val supexpr = (map (fn sup => (sup, (("", false), Expression.Positional [])))
+      sups, []);
     val constrain = Element.Constrains ((map o apsnd o map_atyps)
       (K (TFree (Name.aT, base_sort))) supparams);
+      (*FIXME 2009 perhaps better: control type variable by explicit
+      parameter instantiation of import expression*)
+    val begin_ctxt = begin sups base_sort
+      #> fold (Variable.declare_constraints o Free) ((map o apsnd o map_atyps)
+          (K (TFree (Name.aT, base_sort))) supparams) (*FIXME
+            should constraints be issued in begin?*)
+    val ((_, _, syntax_elems), _) = ProofContext.init thy
+      |> begin_ctxt
+      |> process_decl supexpr raw_elems;
     fun fork_syn (Element.Fixes xs) =
           fold_map (fn (c, ty, syn) => cons (Binding.base_name c, syn) #> pair (c, ty, NoSyn)) xs
           #>> Element.Fixes
       | fork_syn x = pair x;
-    fun fork_syntax elems =
-      let
-        val (elems', global_syntax) = fold_map fork_syn elems [];
-      in (constrain :: elems', global_syntax) end;
-    val (elems, global_syntax) =
-      ProofContext.init thy
-      |> Locale.cert_expr supexpr [constrain]
-      |> snd
-      |> init_ctxt sups base_sort
-      |> process_expr Locale.empty raw_elems
-      |> fst
-      |> fork_syntax
-  in (((sups, supparams), (supsort, base_sort, mergeexpr)), (elems, global_syntax)) end;
+    val (elems, global_syntax) = fold_map fork_syn syntax_elems [];
+  in (((sups, supparam_names), (supsort, base_sort, supexpr)), (constrain :: elems, global_syntax)) end;
 
-val read_class_spec = gen_class_spec Sign.intern_class Locale.read_expr;
-val check_class_spec = gen_class_spec (K I) Locale.cert_expr;
+val cert_class_spec = prep_class_spec (K I) Expression.cert_declaration;
+val read_class_spec = prep_class_spec Sign.intern_class Expression.cert_read_declaration;
 
 fun add_consts bname class base_sort sups supparams global_syntax thy =
   let
-    val supconsts = map fst supparams
+    (*FIXME 2009 simplify*)
+    val supconsts = supparams
       |> AList.make (snd o the o AList.lookup (op =) (these_params thy sups))
       |> (map o apsnd o apsnd o map_atyps o K o TFree) (Name.aT, [class]);
-    val all_params = map fst (Locale.parameters_of thy class);
+    val all_params = Locale.params_of thy class;
     val raw_params = (snd o chop (length supparams)) all_params;
-    fun add_const (v, raw_ty) thy =
+    fun add_const (b, SOME raw_ty, _) thy =
       let
+        val v = Binding.base_name b;
         val c = Sign.full_bname thy v;
         val ty = map_atyps (K (TFree (Name.aT, base_sort))) raw_ty;
         val ty0 = Type.strip_sorts ty;
@@ -641,7 +152,7 @@
       end;
   in
     thy
-    |> Sign.add_path (Logic.const_of_class bname)
+    |> Sign.add_path (class_prefix class)
     |> fold_map add_const raw_params
     ||> Sign.restore_naming thy
     |-> (fn params => pair (supconsts @ (map o apfst) fst params, params))
@@ -649,12 +160,13 @@
 
 fun adjungate_axclass bname class base_sort sups supsort supparams global_syntax thy =
   let
+    (*FIXME 2009 simplify*)
     fun globalize param_map = map_aterms
       (fn Free (v, ty) => Const ((fst o the o AList.lookup (op =) param_map) v, ty)
         | t => t);
-    val raw_pred = Locale.intros thy class
+    val raw_pred = Locale.intros_of thy class
       |> fst
-      |> map (Logic.unvarify o Logic.strip_imp_concl o Thm.prop_of);
+      |> Option.map (Logic.unvarify o Logic.strip_imp_concl o Thm.prop_of);
     fun get_axiom thy = case (#axioms o AxClass.get_info thy) class
      of [] => NONE
       | [thm] => SOME thm;
@@ -663,240 +175,86 @@
     |> add_consts bname class base_sort sups supparams global_syntax
     |-> (fn (param_map, params) => AxClass.define_class (bname, supsort)
           (map (fst o snd) params)
-          [((Binding.name (bname ^ "_" ^ AxClass.axiomsN), []), map (globalize param_map) raw_pred)]
+          [((Binding.empty, []),
+            Option.map (globalize param_map) raw_pred |> the_list)]
     #> snd
     #> `get_axiom
     #-> (fn assm_axiom => fold (Sign.add_const_constraint o apsnd SOME o snd) params
-    #> pair (map (Const o snd) param_map, param_map, params, assm_axiom)))
+    #> pair (param_map, params, assm_axiom)))
   end;
 
 fun gen_class prep_spec bname raw_supclasses raw_elems thy =
   let
     val class = Sign.full_bname thy bname;
-    val (((sups, supparams), (supsort, base_sort, mergeexpr)), (elems, global_syntax)) =
+    val (((sups, supparams), (supsort, base_sort, supexpr)), (elems, global_syntax)) =
       prep_spec thy raw_supclasses raw_elems;
-    val supconsts = map (apsnd fst o snd) (these_params thy sups);
   in
     thy
-    |> Locale.add_locale "" bname mergeexpr elems
-    |> snd
-    |> ProofContext.theory_of
+    |> Expression.add_locale bname "" supexpr elems
+    |> snd |> LocalTheory.exit_global
     |> adjungate_axclass bname class base_sort sups supsort supparams global_syntax
-    |-> (fn (inst, param_map, params, assm_axiom) =>
-        `(fn thy => calculate_axiom thy sups base_sort assm_axiom param_map class)
-    #-> (fn axiom =>
-        prove_class_interpretation class inst
-          (supconsts @ map (pair class o fst o snd) params) (the_list axiom) []
-    #-> (fn morphism =>
-        `(fn thy => activate_class_morphism thy class inst morphism)
-    #-> (fn phi =>
-        `(fn thy => calculate_rules thy phi sups base_sort param_map axiom class)
-    #-> (fn (assm_intro, of_class) =>
-        add_class_data ((class, sups), (params, base_sort, inst,
-          morphism, axiom, assm_intro, of_class))
-    #> fold (note_assm_intro class) (the_list assm_intro))))))
-    |> init class
+    |-> (fn (param_map, params, assm_axiom) =>
+       `(fn thy => calculate thy class sups base_sort param_map assm_axiom)
+    #-> (fn (base_morph, morph, export_morph, axiom, assm_intro, of_class) =>
+       Locale.add_registration (class, (morph, export_morph))
+    #> Locale.activate_global_facts (class, morph $> export_morph)
+    #> register class sups params base_sort base_morph axiom assm_intro of_class))
+    |> TheoryTarget.init (SOME class)
     |> pair class
   end;
 
 in
 
+val class = gen_class cert_class_spec;
 val class_cmd = gen_class read_class_spec;
-val class = gen_class check_class_spec;
 
 end; (*local*)
 
 
-
-(** instantiation target **)
-
-(* bookkeeping *)
-
-datatype instantiation = Instantiation of {
-  arities: string list * (string * sort) list * sort,
-  params: ((string * string) * (string * typ)) list
-    (*(instantiation parameter, type constructor), (local instantiation parameter, typ)*)
-}
-
-structure Instantiation = ProofDataFun
-(
-  type T = instantiation
-  fun init _ = Instantiation { arities = ([], [], []), params = [] };
-);
-
-fun mk_instantiation (arities, params) =
-  Instantiation { arities = arities, params = params };
-fun get_instantiation lthy = case Instantiation.get (LocalTheory.target_of lthy)
- of Instantiation data => data;
-fun map_instantiation f = (LocalTheory.target o Instantiation.map)
-  (fn Instantiation { arities, params } => mk_instantiation (f (arities, params)));
-
-fun the_instantiation lthy = case get_instantiation lthy
- of { arities = ([], [], []), ... } => error "No instantiation target"
-  | data => data;
-
-val instantiation_params = #params o get_instantiation;
-
-fun instantiation_param lthy v = instantiation_params lthy
-  |> find_first (fn (_, (v', _)) => v = v')
-  |> Option.map (fst o fst);
-
-
-(* syntax *)
+(** subclass relations **)
 
-fun synchronize_inst_syntax ctxt =
-  let
-    val Instantiation { arities = (_, _, sort), params = params } = Instantiation.get ctxt;
-    val thy = ProofContext.theory_of ctxt;
-    fun subst (c, ty) = case AxClass.inst_tyco_of thy (c, ty)
-         of SOME tyco => (case AList.lookup (op =) params (c, tyco)
-             of SOME (v_ty as (_, ty)) => SOME (ty, Free v_ty)
-              | NONE => NONE)
-          | NONE => NONE;
-    val unchecks =
-      map (fn ((c, _), v_ty as (_, ty)) => (Free v_ty, Const (c, ty))) params;
-  in
-    ctxt
-    |> Overloading.map_improvable_syntax
-         (fn (((primary_constraints, _), (((improve, _), _), _)), _) =>
-            (((primary_constraints, []), (((improve, subst), false), unchecks)), false))
-  end;
-
-
-(* target *)
+local
 
-val sanatize_name = (*FIXME*)
-  let
-    fun is_valid s = Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s
-      orelse s = "'" orelse s = "_";
-    val is_junk = not o is_valid andf Symbol.is_regular;
-    val junk = Scan.many is_junk;
-    val scan_valids = Symbol.scanner "Malformed input"
-      ((junk |--
-        (Scan.optional (Scan.one Symbol.is_ascii_letter) "x" ^^ (Scan.many is_valid >> implode)
-        --| junk))
-      ::: Scan.repeat ((Scan.many1 is_valid >> implode) --| junk));
-  in
-    explode #> scan_valids #> implode
-  end;
-
-fun type_name "*" = "prod"
-  | type_name "+" = "sum"
-  | type_name s = sanatize_name (NameSpace.base s); (*FIXME*)
-
-fun resort_terms pp algebra consts constraints ts =
-  let
-    fun matchings (Const (c_ty as (c, _))) = (case constraints c
-         of NONE => I
-          | SOME sorts => fold2 (curry (Sorts.meet_sort algebra))
-              (Consts.typargs consts c_ty) sorts)
-      | matchings _ = I
-    val tvartab = (fold o fold_aterms) matchings ts Vartab.empty
-      handle Sorts.CLASS_ERROR e => error (Sorts.class_error pp e);
-    val inst = map_type_tvar
-      (fn (vi, sort) => TVar (vi, the_default sort (Vartab.lookup tvartab vi)));
-  in if Vartab.is_empty tvartab then NONE else SOME ((map o map_types) inst ts) end;
-
-fun init_instantiation (tycos, vs, sort) thy =
+fun gen_subclass prep_class do_proof raw_sup lthy =
   let
-    val _ = if null tycos then error "At least one arity must be given" else ();
-    val params = these_params thy sort;
-    fun get_param tyco (param, (_, (c, ty))) = if can (AxClass.param_of_inst thy) (c, tyco)
-      then NONE else SOME ((c, tyco),
-        (param ^ "_" ^ type_name tyco, map_atyps (K (Type (tyco, map TFree vs))) ty));
-    val inst_params = map_product get_param tycos params |> map_filter I;
-    val primary_constraints = map (apsnd
-      (map_atyps (K (TVar ((Name.aT, 0), [])))) o snd o snd) params;
-    val pp = Syntax.pp_global thy;
-    val algebra = Sign.classes_of thy
-      |> fold (fn tyco => Sorts.add_arities pp
-            (tyco, map (fn class => (class, map snd vs)) sort)) tycos;
-    val consts = Sign.consts_of thy;
-    val improve_constraints = AList.lookup (op =)
-      (map (fn (_, (class, (c, _))) => (c, [[class]])) params);
-    fun resort_check ts ctxt = case resort_terms pp algebra consts improve_constraints ts
-     of NONE => NONE
-      | SOME ts' => SOME (ts', ctxt);
-    fun improve (c, ty) = case AxClass.inst_tyco_of thy (c, ty)
-     of SOME tyco => (case AList.lookup (op =) inst_params (c, tyco)
-         of SOME (_, ty') => if Type.raw_instance (ty', ty) then SOME (ty, ty') else NONE
-          | NONE => NONE)
-      | NONE => NONE;
-  in
-    thy
-    |> ProofContext.init
-    |> Instantiation.put (mk_instantiation ((tycos, vs, sort), inst_params))
-    |> fold (Variable.declare_typ o TFree) vs
-    |> fold (Variable.declare_names o Free o snd) inst_params
-    |> (Overloading.map_improvable_syntax o apfst)
-         (fn ((_, _), ((_, subst), unchecks)) =>
-            ((primary_constraints, []), (((improve, K NONE), false), [])))
-    |> Overloading.add_improvable_syntax
-    |> Context.proof_map (Syntax.add_term_check 0 "resorting" resort_check)
-    |> synchronize_inst_syntax
-  end;
+    val thy = ProofContext.theory_of lthy;
+    val proto_sup = prep_class thy raw_sup;
+    val proto_sub = case TheoryTarget.peek lthy
+     of {is_class = false, ...} => error "Not in a class context"
+      | {target, ...} => target;
+    val (sub, sup) = AxClass.cert_classrel thy (proto_sub, proto_sup)
 
-fun confirm_declaration c = (map_instantiation o apsnd)
-  (filter_out (fn (_, (c', _)) => c' = c))
-  #> LocalTheory.target synchronize_inst_syntax
+    val expr = ([(sup, (("", false), Expression.Positional []))], []);
+    val (([props], deps, export), goal_ctxt) =
+      Expression.cert_goal_expression expr lthy;
+    val some_prop = try the_single props;
+    val some_dep_morph = try the_single (map snd deps);
+    fun after_qed some_wit =
+      ProofContext.theory (register_subclass (sub, sup)
+        some_dep_morph some_wit export)
+      #> ProofContext.theory_of #> TheoryTarget.init (SOME sub);
+  in do_proof after_qed some_prop goal_ctxt end;
 
-fun gen_instantiation_instance do_proof after_qed lthy =
-  let
-    val (tycos, vs, sort) = (#arities o the_instantiation) lthy;
-    val arities_proof = maps (fn tyco => Logic.mk_arities (tyco, map snd vs, sort)) tycos;
-    fun after_qed' results =
-      LocalTheory.theory (fold (AxClass.add_arity o Thm.varifyT) results)
-      #> after_qed;
-  in
-    lthy
-    |> do_proof after_qed' arities_proof
-  end;
-
-val instantiation_instance = gen_instantiation_instance (fn after_qed => fn ts =>
-  Proof.theorem_i NONE (after_qed o map the_single) (map (fn t => [(t, [])]) ts));
-
-fun prove_instantiation_instance tac = gen_instantiation_instance (fn after_qed =>
-  fn ts => fn lthy => after_qed (map (fn t => Goal.prove lthy [] [] t
-    (fn {context, ...} => tac context)) ts) lthy) I;
-
-fun prove_instantiation_exit tac = prove_instantiation_instance tac
-  #> LocalTheory.exit_global;
+fun user_proof after_qed NONE =
+      Proof.theorem_i NONE (K (after_qed NONE)) [[]]
+      #> Element.refine_witness #> Seq.hd
+  | user_proof after_qed (SOME prop) =
+      Proof.theorem_i NONE (after_qed o SOME o Element.make_witness prop
+        o Thm.close_derivation o the_single o the_single)
+          [[(Element.mark_witness prop, [])]]
+      #> Element.refine_witness #> Seq.hd;
 
-fun prove_instantiation_exit_result f tac x lthy =
-  let
-    val phi = ProofContext.export_morphism lthy
-      (ProofContext.init (ProofContext.theory_of lthy));
-    val y = f phi x;
-  in
-    lthy
-    |> prove_instantiation_exit (fn ctxt => tac ctxt y)
-    |> pair y
-  end;
+fun tactic_proof tac after_qed NONE ctxt =
+      after_qed NONE ctxt
+  | tactic_proof tac after_qed (SOME prop) ctxt =
+      after_qed (SOME (Element.prove_witness ctxt prop tac)) ctxt;
 
-fun conclude_instantiation lthy =
-  let
-    val { arities, params } = the_instantiation lthy;
-    val (tycos, vs, sort) = arities;
-    val thy = ProofContext.theory_of lthy;
-    val _ = map (fn tyco => if Sign.of_sort thy
-        (Type (tyco, map TFree vs), sort)
-      then () else error ("Missing instance proof for type " ^ quote (Sign.extern_type thy tyco)))
-        tycos;
-  in lthy end;
+in
 
-fun pretty_instantiation lthy =
-  let
-    val { arities, params } = the_instantiation lthy;
-    val (tycos, vs, sort) = arities;
-    val thy = ProofContext.theory_of lthy;
-    fun pr_arity tyco = Syntax.pretty_arity lthy (tyco, map snd vs, sort);
-    fun pr_param ((c, _), (v, ty)) =
-      (Pretty.block o Pretty.breaks) [Pretty.str v, Pretty.str "==",
-        (Pretty.str o Sign.extern_const thy) c, Pretty.str "::", Syntax.pretty_typ_global thy ty];
-  in
-    (Pretty.block o Pretty.fbreaks)
-      (Pretty.str "instantiation" :: map pr_arity tycos @ map pr_param params)
-  end;
+val subclass = gen_subclass (K I) user_proof;
+fun prove_subclass tac = gen_subclass (K I) (tactic_proof tac);
+val subclass_cmd = gen_subclass Sign.read_class user_proof;
+
+end; (*local*)
 
 end;
-
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Isar/class_target.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -0,0 +1,629 @@
+(*  Title:      Pure/Isar/class_target.ML
+    Author:     Florian Haftmann, TU Muenchen
+
+Type classes derived from primitive axclasses and locales - mechanisms.
+*)
+
+signature CLASS_TARGET =
+sig
+  (*classes*)
+  val register: class -> class list -> ((string * typ) * (string * typ)) list
+    -> sort -> morphism -> thm option -> thm option -> thm
+    -> theory -> theory
+  val register_subclass: class * class -> morphism option -> Element.witness option
+    -> morphism -> theory -> theory
+
+  val is_class: theory -> class -> bool
+  val base_sort: theory -> class -> sort
+  val rules: theory -> class -> thm option * thm
+  val these_params: theory -> sort -> (string * (class * (string * typ))) list
+  val these_defs: theory -> sort -> thm list
+  val print_classes: theory -> unit
+
+  val begin: class list -> sort -> Proof.context -> Proof.context
+  val init: class -> theory -> Proof.context
+  val declare: class -> Properties.T
+    -> (string * mixfix) * term -> theory -> theory
+  val abbrev: class -> Syntax.mode -> Properties.T
+    -> (string * mixfix) * term -> theory -> theory
+  val class_prefix: string -> string
+  val refresh_syntax: class -> Proof.context -> Proof.context
+
+  (*instances*)
+  val init_instantiation: string list * (string * sort) list * sort
+    -> theory -> local_theory
+  val instantiation_instance: (local_theory -> local_theory)
+    -> local_theory -> Proof.state
+  val prove_instantiation_instance: (Proof.context -> tactic)
+    -> local_theory -> local_theory
+  val prove_instantiation_exit: (Proof.context -> tactic)
+    -> local_theory -> theory
+  val prove_instantiation_exit_result: (morphism -> 'a -> 'b)
+    -> (Proof.context -> 'b -> tactic) -> 'a -> local_theory -> 'b * theory
+  val conclude_instantiation: local_theory -> local_theory
+  val instantiation_param: local_theory -> string -> string option
+  val confirm_declaration: string -> local_theory -> local_theory
+  val pretty_instantiation: local_theory -> Pretty.T
+  val type_name: string -> string
+
+  val intro_classes_tac: thm list -> tactic
+  val default_intro_tac: Proof.context -> thm list -> tactic
+
+  (*old axclass layer*)
+  val axclass_cmd: bstring * xstring list
+    -> (Attrib.binding * string list) list
+    -> theory -> class * theory
+  val classrel_cmd: xstring * xstring -> theory -> Proof.state
+
+  (*old instance layer*)
+  val instance_arity: (theory -> theory) -> arity -> theory -> Proof.state
+  val instance_arity_cmd: bstring * xstring list * xstring -> theory -> Proof.state
+end;
+
+structure Class_Target : CLASS_TARGET =
+struct
+
+(** primitive axclass and instance commands **)
+
+fun axclass_cmd (class, raw_superclasses) raw_specs thy =
+  let
+    val ctxt = ProofContext.init thy;
+    val superclasses = map (Sign.read_class thy) raw_superclasses;
+    val name_atts = map ((apsnd o map) (Attrib.attribute thy) o fst)
+      raw_specs;
+    val axiomss = ProofContext.read_propp (ctxt, map (map (rpair []) o snd)
+          raw_specs)
+      |> snd
+      |> (map o map) fst;
+  in
+    AxClass.define_class (class, superclasses) []
+      (name_atts ~~ axiomss) thy
+  end;
+
+local
+
+fun gen_instance mk_prop add_thm after_qed insts thy =
+  let
+    fun after_qed' results =
+      ProofContext.theory ((fold o fold) add_thm results #> after_qed);
+  in
+    thy
+    |> ProofContext.init
+    |> Proof.theorem_i NONE after_qed' ((map (fn t => [(t, [])])
+        o mk_prop thy) insts)
+  end;
+
+in
+
+val instance_arity =
+  gen_instance (Logic.mk_arities oo Sign.cert_arity) AxClass.add_arity;
+val instance_arity_cmd =
+  gen_instance (Logic.mk_arities oo Sign.read_arity) AxClass.add_arity I;
+val classrel =
+  gen_instance (single oo (Logic.mk_classrel oo AxClass.cert_classrel)) AxClass.add_classrel I;
+val classrel_cmd =
+  gen_instance (single oo (Logic.mk_classrel oo AxClass.read_classrel)) AxClass.add_classrel I;
+
+end; (*local*)
+
+
+(** class data **)
+
+datatype class_data = ClassData of {
+
+  (* static part *)
+  consts: (string * string) list
+    (*locale parameter ~> constant name*),
+  base_sort: sort,
+  base_morph: morphism
+    (*static part of canonical morphism*),
+  assm_intro: thm option,
+  of_class: thm,
+  axiom: thm option,
+  
+  (* dynamic part *)
+  defs: thm list,
+  operations: (string * (class * (typ * term))) list
+
+};
+
+fun rep_class_data (ClassData data) = data;
+fun mk_class_data ((consts, base_sort, base_morph, assm_intro, of_class, axiom),
+    (defs, operations)) =
+  ClassData { consts = consts, base_sort = base_sort,
+    base_morph = base_morph, assm_intro = assm_intro, of_class = of_class, axiom = axiom,
+    defs = defs, operations = operations };
+fun map_class_data f (ClassData { consts, base_sort, base_morph, assm_intro,
+    of_class, axiom, defs, operations }) =
+  mk_class_data (f ((consts, base_sort, base_morph, assm_intro, of_class, axiom),
+    (defs, operations)));
+fun merge_class_data _ (ClassData { consts = consts,
+    base_sort = base_sort, base_morph = base_morph, assm_intro = assm_intro,
+    of_class = of_class, axiom = axiom, defs = defs1, operations = operations1 },
+  ClassData { consts = _, base_sort = _, base_morph = _, assm_intro = _,
+    of_class = _, axiom = _, defs = defs2, operations = operations2 }) =
+  mk_class_data ((consts, base_sort, base_morph, assm_intro, of_class, axiom),
+    (Thm.merge_thms (defs1, defs2),
+      AList.merge (op =) (K true) (operations1, operations2)));
+
+structure ClassData = TheoryDataFun
+(
+  type T = class_data Graph.T
+  val empty = Graph.empty;
+  val copy = I;
+  val extend = I;
+  fun merge _ = Graph.join merge_class_data;
+);
+
+
+(* queries *)
+
+val lookup_class_data = Option.map rep_class_data oo try o Graph.get_node o ClassData.get;
+
+fun the_class_data thy class = case lookup_class_data thy class
+ of NONE => error ("Undeclared class " ^ quote class)
+  | SOME data => data;
+
+val is_class = is_some oo lookup_class_data;
+
+val ancestry = Graph.all_succs o ClassData.get;
+val heritage = Graph.all_preds o ClassData.get;
+
+fun these_params thy =
+  let
+    fun params class =
+      let
+        val const_typs = (#params o AxClass.get_info thy) class;
+        val const_names = (#consts o the_class_data thy) class;
+      in
+        (map o apsnd)
+          (fn c => (class, (c, (the o AList.lookup (op =) const_typs) c))) const_names
+      end;
+  in maps params o ancestry thy end;
+
+val base_sort = #base_sort oo the_class_data;
+
+fun rules thy class =
+  let val { axiom, of_class, ... } = the_class_data thy class
+  in (axiom, of_class) end;
+
+fun all_assm_intros thy =
+  Graph.fold (fn (_, (data, _)) => fold (insert Thm.eq_thm)
+    ((the_list o #assm_intro o rep_class_data) data)) (ClassData.get thy) [];
+
+fun these_defs thy = maps (#defs o the_class_data thy) o ancestry thy;
+fun these_operations thy = maps (#operations o the_class_data thy) o ancestry thy;
+
+val base_morphism = #base_morph oo the_class_data;
+fun morphism thy class = base_morphism thy class
+  $> Element.eq_morphism thy (these_defs thy [class]);
+
+fun print_classes thy =
+  let
+    val ctxt = ProofContext.init thy;
+    val algebra = Sign.classes_of thy;
+    val arities =
+      Symtab.empty
+      |> Symtab.fold (fn (tyco, arities) => fold (fn (class, _) =>
+           Symtab.map_default (class, []) (insert (op =) tyco)) arities)
+             ((#arities o Sorts.rep_algebra) algebra);
+    val the_arities = these o Symtab.lookup arities;
+    fun mk_arity class tyco =
+      let
+        val Ss = Sorts.mg_domain algebra tyco [class];
+      in Syntax.pretty_arity ctxt (tyco, Ss, [class]) end;
+    fun mk_param (c, ty) = Pretty.str (Sign.extern_const thy c ^ " :: "
+      ^ setmp show_sorts false (Syntax.string_of_typ ctxt o Type.strip_sorts) ty);
+    fun mk_entry class = (Pretty.block o Pretty.fbreaks o map_filter I) [
+      (SOME o Pretty.str) ("class " ^ Sign.extern_class thy class ^ ":"),
+      (SOME o Pretty.block) [Pretty.str "supersort: ",
+        (Syntax.pretty_sort ctxt o Sign.minimize_sort thy o Sign.super_classes thy) class],
+      ((fn [] => NONE | ps => (SOME o Pretty.block o Pretty.fbreaks)
+          (Pretty.str "parameters:" :: ps)) o map mk_param
+        o these o Option.map #params o try (AxClass.get_info thy)) class,
+      (SOME o Pretty.block o Pretty.breaks) [
+        Pretty.str "instances:",
+        Pretty.list "" "" (map (mk_arity class) (the_arities class))
+      ]
+    ]
+  in
+    (Pretty.writeln o Pretty.chunks o separate (Pretty.str "")
+      o map mk_entry o Sorts.all_classes) algebra
+  end;
+
+
+(* updaters *)
+
+fun register class sups params base_sort morph
+    axiom assm_intro of_class thy =
+  let
+    val operations = map (fn (v_ty as (_, ty), (c, _)) =>
+      (c, (class, (ty, Free v_ty)))) params;
+    val add_class = Graph.new_node (class,
+        mk_class_data (((map o pairself) fst params, base_sort,
+          morph, assm_intro, of_class, axiom), ([], operations)))
+      #> fold (curry Graph.add_edge class) sups;
+  in ClassData.map add_class thy end;
+
+fun activate_defs class thms thy =
+  let
+    val eq_morph = Element.eq_morphism thy thms;
+    fun amend cls thy = Locale.amend_registration eq_morph
+      (cls, morphism thy cls) thy;
+  in fold amend (heritage thy [class]) thy end;
+
+fun register_operation class (c, (t, some_def)) thy =
+  (*FIXME 2009 does this still also work for abbrevs?*)
+  let
+    val base_sort = base_sort thy class;
+    val prep_typ = map_type_tfree
+      (fn (v, sort) => if Name.aT = v
+        then TFree (v, base_sort) else TVar ((v, 0), sort));
+    val t' = map_types prep_typ t;
+    val ty' = Term.fastype_of t';
+  in
+    thy
+    |> (ClassData.map o Graph.map_node class o map_class_data o apsnd)
+      (fn (defs, operations) =>
+        (fold cons (the_list some_def) defs,
+          (c, (class, (ty', t'))) :: operations))
+    |> is_some some_def ? activate_defs class (the_list some_def)
+  end;
+
+fun register_subclass (sub, sup) some_dep_morph some_wit export thy =
+  let
+    val intros = (snd o rules thy) sup :: map_filter I
+      [Option.map (Drule.standard' o Element.conclude_witness) some_wit,
+        (fst o rules thy) sub];
+    val tac = ALLGOALS (EVERY' (map Tactic.rtac intros));
+    val classrel = Goal.prove_global thy [] [] (Logic.mk_classrel (sub, sup))
+      (K tac);
+    val diff_sort = Sign.complete_sort thy [sup]
+      |> subtract (op =) (Sign.complete_sort thy [sub]);
+  in
+    thy
+    |> AxClass.add_classrel classrel
+    |> ClassData.map (Graph.add_edge (sub, sup))
+    |> activate_defs sub (these_defs thy diff_sort)
+    |> fold2 (fn dep_morph => fn wit => Locale.add_dependency sub (sup,
+        dep_morph $> Element.satisfy_morphism [wit] $> export))
+          (the_list some_dep_morph) (the_list some_wit)
+    |> (fn thy => fold_rev Locale.activate_global_facts
+      (Locale.get_registrations thy) thy)
+  end;
+
+
+(** classes and class target **)
+
+(* class context syntax *)
+
+fun synchronize_class_syntax sups base_sort ctxt =
+  let
+    val thy = ProofContext.theory_of ctxt;
+    val algebra = Sign.classes_of thy;
+    val operations = these_operations thy sups;
+    fun subst_class_typ sort = map_type_tfree (K (TVar ((Name.aT, 0), sort)));
+    val primary_constraints =
+      (map o apsnd) (subst_class_typ base_sort o fst o snd) operations;
+    val secondary_constraints =
+      (map o apsnd) (fn (class, (ty, _)) => subst_class_typ [class] ty) operations;
+    fun declare_const (c, _) =
+      let val b = Sign.base_name c
+      in Sign.intern_const thy b = c ? Variable.declare_const (b, c) end;
+    fun improve (c, ty) = (case AList.lookup (op =) primary_constraints c
+     of SOME ty' => (case try (Type.raw_match (ty', ty)) Vartab.empty
+         of SOME tyenv => (case Vartab.lookup tyenv (Name.aT, 0)
+             of SOME (_, ty' as TVar (tvar as (vi, sort))) =>
+                  if TypeInfer.is_param vi
+                    andalso Sorts.sort_le algebra (base_sort, sort)
+                      then SOME (ty', TFree (Name.aT, base_sort))
+                      else NONE
+              | _ => NONE)
+          | NONE => NONE)
+      | NONE => NONE)
+    fun subst (c, ty) = Option.map snd (AList.lookup (op =) operations c);
+    val unchecks = map (fn (c, (_, (ty, t))) => (t, Const (c, ty))) operations;
+  in
+    ctxt
+    |> fold declare_const primary_constraints
+    |> Overloading.map_improvable_syntax (K (((primary_constraints, secondary_constraints),
+        (((improve, subst), true), unchecks)), false))
+    |> Overloading.set_primary_constraints
+  end;
+
+fun refresh_syntax class ctxt =
+  let
+    val thy = ProofContext.theory_of ctxt;
+    val base_sort = base_sort thy class;
+  in synchronize_class_syntax [class] base_sort ctxt end;
+
+fun begin sups base_sort ctxt =
+  ctxt
+  |> Variable.declare_term
+      (Logic.mk_type (TFree (Name.aT, base_sort)))
+  |> synchronize_class_syntax sups base_sort
+  |> Overloading.add_improvable_syntax;
+
+fun init class thy =
+  thy
+  |> Locale.init class
+  |> begin [class] (base_sort thy class);
+
+
+(* class target *)
+
+val class_prefix = Logic.const_of_class o Sign.base_name;
+
+fun declare class pos ((c, mx), dict) thy =
+  let
+    val prfx = class_prefix class;
+    val thy' = thy |> Sign.add_path prfx;
+      (*FIXME 2009 use proper name morphism*)
+    val morph = morphism thy' class;
+    val params = map (apsnd fst o snd) (these_params thy' [class]);
+
+    val c' = Sign.full_bname thy' c;
+    val dict' = Morphism.term morph dict;
+    val ty' = Term.fastype_of dict';
+    val ty'' = Type.strip_sorts ty';
+      (*FIXME 2009 the tinkering with theorems here is a mess*)
+    val def_eq = Logic.mk_equals (Const (c', ty'), dict');
+    fun get_axiom thy = ((Thm.varifyT o Thm.axiom thy) c', thy);
+  in
+    thy'
+    |> Sign.declare_const pos ((Binding.name c, ty''), mx) |> snd
+    |> Thm.add_def false false (c, def_eq) (*FIXME 2009 name of theorem*)
+      (*FIXME 2009 add_def should accept binding*)
+    |>> Thm.symmetric
+    ||>> get_axiom
+    |-> (fn (def, def') => register_operation class (c', (dict', SOME (Thm.symmetric def')))
+      #> PureThy.store_thm (c ^ "_raw", def') (*FIXME 2009 name of theorem*)
+         (*FIXME 2009 store_thm etc. should accept binding*)
+      #> snd)
+    |> Sign.restore_naming thy
+    |> Sign.add_const_constraint (c', SOME ty')
+  end;
+
+fun abbrev class prmode pos ((c, mx), rhs) thy =
+  let
+    val prfx = class_prefix class;
+    val thy' = thy |> Sign.add_path prfx;
+
+    val unchecks = map (fn (c, (_, (ty, t))) => (t, Const (c, ty)))
+      (these_operations thy [class]);
+    val c' = Sign.full_bname thy' c;
+    val rhs' = Pattern.rewrite_term thy unchecks [] rhs;
+    val rhs'' = map_types Logic.varifyT rhs';
+    val ty' = Term.fastype_of rhs';
+  in
+    thy'
+    |> Sign.add_abbrev (#1 prmode) pos (Binding.name c, map_types Type.strip_sorts rhs'') |> snd
+    |> Sign.add_const_constraint (c', SOME ty')
+    |> Sign.notation true prmode [(Const (c', ty'), mx)]
+    |> not (#1 prmode = PrintMode.input) ? register_operation class (c', (rhs', NONE))
+    |> Sign.restore_naming thy
+  end;
+
+
+(** instantiation target **)
+
+(* bookkeeping *)
+
+datatype instantiation = Instantiation of {
+  arities: string list * (string * sort) list * sort,
+  params: ((string * string) * (string * typ)) list
+    (*(instantiation parameter, type constructor), (local instantiation parameter, typ)*)
+}
+
+structure Instantiation = ProofDataFun
+(
+  type T = instantiation
+  fun init _ = Instantiation { arities = ([], [], []), params = [] };
+);
+
+fun mk_instantiation (arities, params) =
+  Instantiation { arities = arities, params = params };
+fun get_instantiation lthy = case Instantiation.get (LocalTheory.target_of lthy)
+ of Instantiation data => data;
+fun map_instantiation f = (LocalTheory.target o Instantiation.map)
+  (fn Instantiation { arities, params } => mk_instantiation (f (arities, params)));
+
+fun the_instantiation lthy = case get_instantiation lthy
+ of { arities = ([], [], []), ... } => error "No instantiation target"
+  | data => data;
+
+val instantiation_params = #params o get_instantiation;
+
+fun instantiation_param lthy v = instantiation_params lthy
+  |> find_first (fn (_, (v', _)) => v = v')
+  |> Option.map (fst o fst);
+
+
+(* syntax *)
+
+fun synchronize_inst_syntax ctxt =
+  let
+    val Instantiation { arities = (_, _, sort), params = params } = Instantiation.get ctxt;
+    val thy = ProofContext.theory_of ctxt;
+    fun subst (c, ty) = case AxClass.inst_tyco_of thy (c, ty)
+         of SOME tyco => (case AList.lookup (op =) params (c, tyco)
+             of SOME (v_ty as (_, ty)) => SOME (ty, Free v_ty)
+              | NONE => NONE)
+          | NONE => NONE;
+    val unchecks =
+      map (fn ((c, _), v_ty as (_, ty)) => (Free v_ty, Const (c, ty))) params;
+  in
+    ctxt
+    |> Overloading.map_improvable_syntax
+         (fn (((primary_constraints, _), (((improve, _), _), _)), _) =>
+            (((primary_constraints, []), (((improve, subst), false), unchecks)), false))
+  end;
+
+
+(* target *)
+
+val sanatize_name = (*FIXME*)
+  let
+    fun is_valid s = Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s
+      orelse s = "'" orelse s = "_";
+    val is_junk = not o is_valid andf Symbol.is_regular;
+    val junk = Scan.many is_junk;
+    val scan_valids = Symbol.scanner "Malformed input"
+      ((junk |--
+        (Scan.optional (Scan.one Symbol.is_ascii_letter) "x" ^^ (Scan.many is_valid >> implode)
+        --| junk))
+      ::: Scan.repeat ((Scan.many1 is_valid >> implode) --| junk));
+  in
+    explode #> scan_valids #> implode
+  end;
+
+fun type_name "*" = "prod"
+  | type_name "+" = "sum"
+  | type_name s = sanatize_name (NameSpace.base s);
+
+fun resort_terms pp algebra consts constraints ts =
+  let
+    fun matchings (Const (c_ty as (c, _))) = (case constraints c
+         of NONE => I
+          | SOME sorts => fold2 (curry (Sorts.meet_sort algebra))
+              (Consts.typargs consts c_ty) sorts)
+      | matchings _ = I
+    val tvartab = (fold o fold_aterms) matchings ts Vartab.empty
+      handle Sorts.CLASS_ERROR e => error (Sorts.class_error pp e);
+    val inst = map_type_tvar
+      (fn (vi, sort) => TVar (vi, the_default sort (Vartab.lookup tvartab vi)));
+  in if Vartab.is_empty tvartab then NONE else SOME ((map o map_types) inst ts) end;
+
+fun init_instantiation (tycos, vs, sort) thy =
+  let
+    val _ = if null tycos then error "At least one arity must be given" else ();
+    val params = these_params thy sort;
+    fun get_param tyco (param, (_, (c, ty))) = if can (AxClass.param_of_inst thy) (c, tyco)
+      then NONE else SOME ((c, tyco),
+        (param ^ "_" ^ type_name tyco, map_atyps (K (Type (tyco, map TFree vs))) ty));
+    val inst_params = map_product get_param tycos params |> map_filter I;
+    val primary_constraints = map (apsnd
+      (map_atyps (K (TVar ((Name.aT, 0), [])))) o snd o snd) params;
+    val pp = Syntax.pp_global thy;
+    val algebra = Sign.classes_of thy
+      |> fold (fn tyco => Sorts.add_arities pp
+            (tyco, map (fn class => (class, map snd vs)) sort)) tycos;
+    val consts = Sign.consts_of thy;
+    val improve_constraints = AList.lookup (op =)
+      (map (fn (_, (class, (c, _))) => (c, [[class]])) params);
+    fun resort_check ts ctxt = case resort_terms pp algebra consts improve_constraints ts
+     of NONE => NONE
+      | SOME ts' => SOME (ts', ctxt);
+    fun improve (c, ty) = case AxClass.inst_tyco_of thy (c, ty)
+     of SOME tyco => (case AList.lookup (op =) inst_params (c, tyco)
+         of SOME (_, ty') => if Type.raw_instance (ty', ty) then SOME (ty, ty') else NONE
+          | NONE => NONE)
+      | NONE => NONE;
+  in
+    thy
+    |> ProofContext.init
+    |> Instantiation.put (mk_instantiation ((tycos, vs, sort), inst_params))
+    |> fold (Variable.declare_typ o TFree) vs
+    |> fold (Variable.declare_names o Free o snd) inst_params
+    |> (Overloading.map_improvable_syntax o apfst)
+         (fn ((_, _), ((_, subst), unchecks)) =>
+            ((primary_constraints, []), (((improve, K NONE), false), [])))
+    |> Overloading.add_improvable_syntax
+    |> Context.proof_map (Syntax.add_term_check 0 "resorting" resort_check)
+    |> synchronize_inst_syntax
+  end;
+
+fun confirm_declaration c = (map_instantiation o apsnd)
+  (filter_out (fn (_, (c', _)) => c' = c))
+  #> LocalTheory.target synchronize_inst_syntax
+
+fun gen_instantiation_instance do_proof after_qed lthy =
+  let
+    val (tycos, vs, sort) = (#arities o the_instantiation) lthy;
+    val arities_proof = maps (fn tyco => Logic.mk_arities (tyco, map snd vs, sort)) tycos;
+    fun after_qed' results =
+      LocalTheory.theory (fold (AxClass.add_arity o Thm.varifyT) results)
+      #> after_qed;
+  in
+    lthy
+    |> do_proof after_qed' arities_proof
+  end;
+
+val instantiation_instance = gen_instantiation_instance (fn after_qed => fn ts =>
+  Proof.theorem_i NONE (after_qed o map the_single) (map (fn t => [(t, [])]) ts));
+
+fun prove_instantiation_instance tac = gen_instantiation_instance (fn after_qed =>
+  fn ts => fn lthy => after_qed (map (fn t => Goal.prove lthy [] [] t
+    (fn {context, ...} => tac context)) ts) lthy) I;
+
+fun prove_instantiation_exit tac = prove_instantiation_instance tac
+  #> LocalTheory.exit_global;
+
+fun prove_instantiation_exit_result f tac x lthy =
+  let
+    val morph = ProofContext.export_morphism lthy
+      (ProofContext.init (ProofContext.theory_of lthy));
+    val y = f morph x;
+  in
+    lthy
+    |> prove_instantiation_exit (fn ctxt => tac ctxt y)
+    |> pair y
+  end;
+
+fun conclude_instantiation lthy =
+  let
+    val { arities, params } = the_instantiation lthy;
+    val (tycos, vs, sort) = arities;
+    val thy = ProofContext.theory_of lthy;
+    val _ = map (fn tyco => if Sign.of_sort thy
+        (Type (tyco, map TFree vs), sort)
+      then () else error ("Missing instance proof for type " ^ quote (Sign.extern_type thy tyco)))
+        tycos;
+  in lthy end;
+
+fun pretty_instantiation lthy =
+  let
+    val { arities, params } = the_instantiation lthy;
+    val (tycos, vs, sort) = arities;
+    val thy = ProofContext.theory_of lthy;
+    fun pr_arity tyco = Syntax.pretty_arity lthy (tyco, map snd vs, sort);
+    fun pr_param ((c, _), (v, ty)) =
+      (Pretty.block o Pretty.breaks) [Pretty.str v, Pretty.str "==",
+        (Pretty.str o Sign.extern_const thy) c, Pretty.str "::", Syntax.pretty_typ_global thy ty];
+  in
+    (Pretty.block o Pretty.fbreaks)
+      (Pretty.str "instantiation" :: map pr_arity tycos @ map pr_param params)
+  end;
+
+
+(** tactics and methods **)
+
+fun intro_classes_tac facts st =
+  let
+    val thy = Thm.theory_of_thm st;
+    val classes = Sign.all_classes thy;
+    val class_trivs = map (Thm.class_triv thy) classes;
+    val class_intros = map_filter (try (#intro o AxClass.get_info thy)) classes;
+    val assm_intros = all_assm_intros thy;
+  in
+    Method.intros_tac (class_trivs @ class_intros @ assm_intros) facts st
+  end;
+
+fun default_intro_tac ctxt [] =
+      intro_classes_tac [] ORELSE Old_Locale.intro_locales_tac true ctxt [] ORELSE
+      Locale.intro_locales_tac true ctxt []
+  | default_intro_tac _ _ = no_tac;
+
+fun default_tac rules ctxt facts =
+  HEADGOAL (Method.some_rule_tac rules ctxt facts) ORELSE
+    default_intro_tac ctxt facts;
+
+val _ = Context.>> (Context.map_theory
+  (Method.add_methods
+   [("intro_classes", Method.no_args (Method.METHOD intro_classes_tac),
+      "back-chain introduction rules of classes"),
+    ("default", Method.thms_ctxt_args (Method.METHOD oo default_tac),
+      "apply some intro/elim rule")]));
+
+end;
+
--- a/src/Pure/Isar/code.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Isar/code.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/code.ML
-    ID:         $Id$
     Author:     Florian Haftmann, TU Muenchen
 
 Abstract executable content of theory.  Management of data dependent on
@@ -16,8 +15,8 @@
   val del_eqn: thm -> theory -> theory
   val del_eqns: string -> theory -> theory
   val add_eqnl: string * (thm * bool) list lazy -> theory -> theory
-  val map_pre: (MetaSimplifier.simpset -> MetaSimplifier.simpset) -> theory -> theory
-  val map_post: (MetaSimplifier.simpset -> MetaSimplifier.simpset) -> theory -> theory
+  val map_pre: (simpset -> simpset) -> theory -> theory
+  val map_post: (simpset -> simpset) -> theory -> theory
   val add_inline: thm -> theory -> theory
   val add_functrans: string * (theory -> (thm * bool) list -> (thm * bool) list option) -> theory -> theory
   val del_functrans: string -> theory -> theory
@@ -186,8 +185,8 @@
 (* pre- and postprocessor *)
 
 datatype thmproc = Thmproc of {
-  pre: MetaSimplifier.simpset,
-  post: MetaSimplifier.simpset,
+  pre: simpset,
+  post: simpset,
   functrans: (string * (serial * (theory -> (thm * bool) list -> (thm * bool) list option))) list
 };
 
@@ -198,8 +197,8 @@
 fun merge_thmproc (Thmproc { pre = pre1, post = post1, functrans = functrans1 },
   Thmproc { pre = pre2, post = post2, functrans = functrans2 }) =
     let
-      val pre = MetaSimplifier.merge_ss (pre1, pre2);
-      val post = MetaSimplifier.merge_ss (post1, post2);
+      val pre = Simplifier.merge_ss (pre1, pre2);
+      val post = Simplifier.merge_ss (post1, post2);
       val functrans = AList.merge (op =) (eq_fst (op =)) (functrans1, functrans2);
     in mk_thmproc ((pre, post), functrans) end;
 
@@ -221,7 +220,7 @@
     val thmproc = merge_thmproc (thmproc1, thmproc2);
     val spec = merge_spec (spec1, spec2);
   in mk_exec (thmproc, spec) end;
-val empty_exec = mk_exec (mk_thmproc ((MetaSimplifier.empty_ss, MetaSimplifier.empty_ss), []),
+val empty_exec = mk_exec (mk_thmproc ((Simplifier.empty_ss, Simplifier.empty_ss), []),
   mk_spec ((false, Symtab.empty), (Symtab.empty, (Symtab.empty, Symtab.empty))));
 
 fun the_thmproc (Exec { thmproc = Thmproc x, ...}) = x;
@@ -417,12 +416,12 @@
       Pretty.block [
         Pretty.str "preprocessing simpset:",
         Pretty.fbrk,
-        MetaSimplifier.pretty_ss pre
+        Simplifier.pretty_ss pre
       ],
       Pretty.block [
         Pretty.str "postprocessing simpset:",
         Pretty.fbrk,
-        MetaSimplifier.pretty_ss post
+        Simplifier.pretty_ss post
       ],
       Pretty.block (
         Pretty.str "function transformers:"
--- a/src/Pure/Isar/code_unit.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Isar/code_unit.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/code_unit.ML
-    ID:         $Id$
     Author:     Florian Haftmann, TU Muenchen
 
 Basic notions of code generation.  Auxiliary.
@@ -230,7 +229,7 @@
   val empty = ([], []);
   val copy = I;
   val extend = I;
-  fun merge _ ((alias1, classes1), (alias2, classes2)) =
+  fun merge _ ((alias1, classes1), (alias2, classes2)) : T =
     (Library.merge (eq_snd Thm.eq_thm_prop) (alias1, alias2),
       Library.merge (op =) (classes1, classes2));
 );
@@ -286,7 +285,7 @@
       ^ " :: " ^ string_of_typ thy ty);
     fun last_typ c_ty ty =
       let
-        val frees = typ_tfrees ty;
+        val frees = OldTerm.typ_tfrees ty;
         val (tyco, vs) = ((apsnd o map) (dest_TFree) o dest_Type o snd o strip_type) ty
           handle TYPE _ => no_constr c_ty
         val _ = if has_duplicates (eq_fst (op =)) vs then no_constr c_ty else ();
--- a/src/Pure/Isar/element.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Isar/element.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -78,6 +78,7 @@
   val generalize_facts: Proof.context -> Proof.context ->
     (Attrib.binding * (thm list * Attrib.src list) list) list ->
     (Attrib.binding * (thm list * Attrib.src list) list) list
+  val eq_morphism: theory -> thm list -> morphism
   val transfer_morphism: theory -> morphism
   val activate: (typ, term, Facts.ref) ctxt list -> Proof.context -> context_i list * Proof.context
   val activate_i: context_i list -> Proof.context -> context_i list * Proof.context
@@ -521,6 +522,14 @@
 fun satisfy_facts witns = facts_map (morph_ctxt (satisfy_morphism witns));
 
 
+(* rewriting with equalities *)
+
+fun eq_morphism thy thms = Morphism.morphism
+ {binding = I, var = I, typ = I,
+  term = MetaSimplifier.rewrite_term thy thms [],
+  fact = map (MetaSimplifier.rewrite_rule thms)};
+
+
 (* generalize type/term parameters *)
 
 val maxidx_atts = fold Args.maxidx_values;
--- a/src/Pure/Isar/expression.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Isar/expression.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,37 +1,52 @@
 (*  Title:      Pure/Isar/expression.ML
     Author:     Clemens Ballarin, TU Muenchen
 
-New locale development --- experimental.
+Locale expressions.
 *)
 
 signature EXPRESSION =
 sig
-  datatype 'term map = Positional of 'term option list | Named of (string * 'term) list;
-  type 'term expr = (string * ((string * bool) * 'term map)) list;
-  type expression = string expr * (Binding.T * string option * mixfix) list;
-  type expression_i = term expr * (Binding.T * typ option * mixfix) list;
+  (* Locale expressions *)
+  datatype 'term map = Positional of 'term option list | Named of (string * 'term) list
+  type 'term expr = (string * ((string * bool) * 'term map)) list
+  type expression_i = term expr * (Binding.T * typ option * mixfix) list
+  type expression = string expr * (Binding.T * string option * mixfix) list
 
   (* Processing of context statements *)
+  val cert_statement: Element.context_i list -> (term * term list) list list ->
+    Proof.context -> (term * term list) list list * Proof.context
   val read_statement: Element.context list -> (string * string list) list list ->
-    Proof.context ->  (term * term list) list list * Proof.context;
-  val cert_statement: Element.context_i list -> (term * term list) list list ->
-    Proof.context -> (term * term list) list list * Proof.context;
+    Proof.context -> (term * term list) list list * Proof.context
 
   (* Declaring locales *)
-  val add_locale_cmd: string -> bstring -> expression -> Element.context list -> theory ->
-    (string * (string * (Attrib.binding * (thm list * Attrib.src list) list) list) list) *
-    Proof.context
-  val add_locale: string -> bstring -> expression_i -> Element.context_i list -> theory ->
-    (string * (string * (Attrib.binding * (thm list * Attrib.src list) list) list) list) *
-    Proof.context
+  val cert_declaration: expression_i -> Element.context_i list -> Proof.context ->
+    ((Binding.T * typ option * mixfix) list * (string * morphism) list
+      * Element.context_i list) * ((string * typ) list * Proof.context)
+  val cert_read_declaration: expression_i -> Element.context list -> Proof.context ->
+    ((Binding.T * typ option * mixfix) list * (string * morphism) list
+      * Element.context_i list) * ((string * typ) list * Proof.context)
+      (*FIXME*)
+  val read_declaration: expression -> Element.context list -> Proof.context ->
+    ((Binding.T * typ option * mixfix) list * (string * morphism) list
+      * Element.context_i list) * ((string * typ) list * Proof.context)
+  val add_locale: bstring -> bstring -> expression_i -> Element.context_i list ->
+    theory -> string * local_theory
+  val add_locale_cmd: bstring -> bstring -> expression -> Element.context list ->
+    theory -> string * local_theory
 
   (* Interpretation *)
-  val sublocale_cmd: string -> expression -> theory -> Proof.state;
-  val sublocale: string -> expression_i -> theory -> Proof.state;
-  val interpretation_cmd: expression -> (Attrib.binding * string) list -> theory -> Proof.state;
-  val interpretation: expression_i -> (Attrib.binding * term) list -> theory -> Proof.state;
-  val interpret_cmd: expression -> bool -> Proof.state -> Proof.state;
-  val interpret: expression_i -> bool -> Proof.state -> Proof.state;
+  val cert_goal_expression: expression_i -> Proof.context ->
+    (term list list * (string * morphism) list * morphism) * Proof.context
+  val read_goal_expression: expression -> Proof.context ->
+    (term list list * (string * morphism) list * morphism) * Proof.context
+  val sublocale: string -> expression_i -> theory -> Proof.state
+  val sublocale_cmd: string -> expression -> theory -> Proof.state
+  val interpretation: expression_i -> (Attrib.binding * term) list ->
+    theory -> Proof.state
+  val interpretation_cmd: expression -> (Attrib.binding * string) list ->
+    theory -> Proof.state
+  val interpret: expression_i -> bool -> Proof.state -> Proof.state
+  val interpret_cmd: expression -> bool -> Proof.state -> Proof.state
 end;
 
 
@@ -55,7 +70,7 @@
 
 (** Internalise locale names in expr **)
 
-fun intern thy instances =  map (apfst (NewLocale.intern thy)) instances;
+fun intern thy instances =  map (apfst (Locale.intern thy)) instances;
 
 
 (** Parameters of expression.
@@ -74,7 +89,7 @@
       end;
 
     fun match_bind (n, b) = (n = Binding.base_name b);
-    fun parm_eq ((b1, mx1), (b2, mx2)) =
+    fun parm_eq ((b1, mx1: mixfix), (b2, mx2)) =
       (* FIXME: cannot compare bindings for equality, instead check for equal name and syntax *)
       (Binding.base_name b1 = Binding.base_name b2) andalso
       (if mx1 = mx2 then true
@@ -85,15 +100,15 @@
       (* FIXME: cannot compare bindings for equality. *)
 
     fun params_loc loc =
-          (NewLocale.params_of thy loc |> map (fn (p, _, mx) => (p, mx)), loc);
+          (Locale.params_of thy loc |> map (fn (p, _, mx) => (p, mx)), loc);
     fun params_inst (expr as (loc, (prfx, Positional insts))) =
           let
             val (ps, loc') = params_loc loc;
-	    val d = length ps - length insts;
-	    val insts' =
-	      if d < 0 then error ("More arguments than parameters in instantiation of locale " ^
-                quote (NewLocale.extern thy loc))
-	      else insts @ replicate d NONE;
+            val d = length ps - length insts;
+            val insts' =
+              if d < 0 then error ("More arguments than parameters in instantiation of locale " ^
+                quote (Locale.extern thy loc))
+              else insts @ replicate d NONE;
             val ps' = (ps ~~ insts') |>
               map_filter (fn (p, NONE) => SOME p | (_, SOME _) => NONE);
           in (ps', (loc', (prfx, Positional insts'))) end
@@ -135,14 +150,14 @@
 
 local
 
-fun prep_inst parse_term parms (Positional insts) ctxt =
+fun prep_inst parse_term ctxt parms (Positional insts) =
       (insts ~~ parms) |> map (fn
-        (NONE, p) => Syntax.parse_term ctxt p |
+        (NONE, p) => Free (p, the (Variable.default_type ctxt p)) |
         (SOME t, _) => parse_term ctxt t)
-  | prep_inst parse_term parms (Named insts) ctxt =
+  | prep_inst parse_term ctxt parms (Named insts) =
       parms |> map (fn p => case AList.lookup (op =) insts p of
         SOME t => parse_term ctxt t |
-        NONE => Syntax.parse_term ctxt p);
+        NONE => Free (p, the (Variable.default_type ctxt p)));
 
 in
 
@@ -309,8 +324,8 @@
 fun finish_inst ctxt parms do_close (loc, (prfx, inst)) =
   let
     val thy = ProofContext.theory_of ctxt;
-    val (parm_names, parm_types) = NewLocale.params_of thy loc |>
-      map (fn (b, SOME T, _) => (Binding.base_name b, T)) |> split_list;
+    val (parm_names, parm_types) = Locale.params_of thy loc |>
+      map_split (fn (b, SOME T, _) => (Binding.base_name b, T));
     val (morph, _) = inst_morph (parm_names, parm_types) (prfx, inst) ctxt;
   in (loc, morph) end;
 
@@ -332,18 +347,19 @@
 
 local
 
-fun prep_full_context_statement parse_typ parse_prop parse_inst prep_vars prep_expr
-    strict do_close context raw_import raw_elems raw_concl =
+fun prep_full_context_statement parse_typ parse_prop prep_vars_elem parse_inst prep_vars_inst prep_expr
+    strict do_close raw_import raw_elems raw_concl ctxt1 =
   let
-    val thy = ProofContext.theory_of context;
+    val thy = ProofContext.theory_of ctxt1;
 
     val (raw_insts, fixed) = parameters_of thy strict (apfst (prep_expr thy) raw_import);
 
     fun prep_inst (loc, (prfx, inst)) (i, insts, ctxt) =
       let
-        val (parm_names, parm_types) = NewLocale.params_of thy loc |>
-          map (fn (b, SOME T, _) => (Binding.base_name b, T)) |> split_list;
-        val inst' = parse_inst parm_names inst ctxt;
+        val (parm_names, parm_types) = Locale.params_of thy loc |>
+          map_split (fn (b, SOME T, _) => (Binding.base_name b, T))
+            (*FIXME return value of Locale.params_of has strange type*)
+        val inst' = parse_inst ctxt parm_names inst;
         val parm_types' = map (TypeInfer.paramify_vars o
           Term.map_type_tvar (fn ((x, _), S) => TVar ((x, i), S)) o Logic.varifyT) parm_types;
         val inst'' = map2 TypeInfer.constrain parm_types' inst';
@@ -351,45 +367,50 @@
         val (insts'', _, _, ctxt' (* FIXME not used *) ) = check_autofix insts' [] [] ctxt;
         val inst''' = insts'' |> List.last |> snd |> snd;
         val (morph, _) = inst_morph (parm_names, parm_types) (prfx, inst''') ctxt;
-        val ctxt'' = NewLocale.activate_declarations thy (loc, morph) ctxt;
+        val ctxt'' = Locale.activate_declarations thy (loc, morph) ctxt;
       in (i+1, insts', ctxt'') end;
   
-    fun prep_elem raw_elem (insts, elems, ctxt) =
+    fun prep_elem insts raw_elem (elems, ctxt) =
       let
-        val ctxt' = declare_elem prep_vars raw_elem ctxt;
+        val ctxt' = declare_elem prep_vars_elem raw_elem ctxt;
         val elems' = elems @ [parse_elem parse_typ parse_prop ctxt' raw_elem];
         val (_, _, _, ctxt'') = check_autofix insts elems' [] ctxt';
-      in (insts, elems', ctxt') end;
+      in (elems', ctxt') end;
 
     fun prep_concl raw_concl (insts, elems, ctxt) =
       let
         val concl = parse_concl parse_prop ctxt raw_concl;
       in check_autofix insts elems concl ctxt end;
 
-    val fors = prep_vars fixed context |> fst;
-    val ctxt = context |> ProofContext.add_fixes_i fors |> snd;
-    val (_, insts', ctxt') = fold prep_inst raw_insts (0, [], ctxt);
-    val (_, elems'', ctxt'') = fold prep_elem raw_elems (insts', [], ctxt');
-    val (insts, elems, concl, ctxt) = prep_concl raw_concl (insts', elems'', ctxt'');
+    val fors = prep_vars_inst fixed ctxt1 |> fst;
+    val ctxt2 = ctxt1 |> ProofContext.add_fixes_i fors |> snd;
+    val (_, insts', ctxt3) = fold prep_inst raw_insts (0, [], ctxt2);
+    val (elems, ctxt4) = fold (prep_elem insts') raw_elems ([], ctxt3);
+    val (insts, elems', concl, ctxt5) =
+      prep_concl raw_concl (insts', elems, ctxt4);
 
     (* Retrieve parameter types *)
-    val xs = fold (fn Fixes fixes => (fn ps => ps @ map (Binding.base_name o #1) fixes) |
-      _ => fn ps => ps) (Fixes fors :: elems) [];
-    val (Ts, ctxt') = fold_map ProofContext.inferred_param xs ctxt; 
+    val xs = fold (fn Fixes fixes => (fn ps => ps @ map (Binding.base_name o #1) fixes)
+      | _ => fn ps => ps) (Fixes fors :: elems') [];
+    val (Ts, ctxt6) = fold_map ProofContext.inferred_param xs ctxt5; 
     val parms = xs ~~ Ts;  (* params from expression and elements *)
 
     val Fixes fors' = finish_primitive parms I (Fixes fors);
-    val (deps, elems') = finish ctxt' parms do_close insts elems;
+    val (deps, elems'') = finish ctxt6 parms do_close insts elems';
 
-  in ((fors', deps, elems', concl), (parms, ctxt')) end
+  in ((fors', deps, elems'', concl), (parms, ctxt6)) end
 
 in
 
+fun cert_full_context_statement x =
+  prep_full_context_statement (K I) (K I) ProofContext.cert_vars
+  make_inst ProofContext.cert_vars (K I) x;
+fun cert_read_full_context_statement x =
+  prep_full_context_statement Syntax.parse_typ Syntax.parse_prop ProofContext.read_vars
+  make_inst ProofContext.cert_vars (K I) x;
 fun read_full_context_statement x =
-  prep_full_context_statement Syntax.parse_typ Syntax.parse_prop parse_inst
-  ProofContext.read_vars intern x;
-fun cert_full_context_statement x =
-  prep_full_context_statement (K I) (K I) make_inst ProofContext.cert_vars (K I) x;
+  prep_full_context_statement Syntax.parse_typ Syntax.parse_prop ProofContext.read_vars
+  parse_inst ProofContext.read_vars intern x;
 
 end;
 
@@ -400,14 +421,17 @@
 
 fun prep_statement prep activate raw_elems raw_concl context =
   let
-     val ((_, _, elems, concl), _) = prep true false context ([], []) raw_elems raw_concl;
-     val (_, context') = activate elems (ProofContext.set_stmt true context);
+     val ((_, _, elems, concl), _) =
+       prep true false ([], []) raw_elems raw_concl context;
+     val (_, context') = context |>
+       ProofContext.set_stmt true |>
+       activate elems;
   in (concl, context') end;
 
 in
 
+fun cert_statement x = prep_statement cert_full_context_statement Element.activate_i x;
 fun read_statement x = prep_statement read_full_context_statement Element.activate x;
-fun cert_statement x = prep_statement cert_full_context_statement Element.activate_i x;
 
 end;
 
@@ -418,18 +442,22 @@
 
 fun prep_declaration prep activate raw_import raw_elems context =
   let
-    val ((fixed, deps, elems, _), (parms, ctxt')) = prep false true context raw_import raw_elems [];
+    val ((fixed, deps, elems, _), (parms, ctxt')) =
+      prep false true raw_import raw_elems [] context ;
     (* Declare parameters and imported facts *)
     val context' = context |>
       ProofContext.add_fixes_i fixed |> snd |>
-      fold NewLocale.activate_local_facts deps;
-    val (elems', _) = activate elems (ProofContext.set_stmt true context');
+      fold Locale.activate_local_facts deps;
+    val (elems', _) = context' |>
+      ProofContext.set_stmt true |>
+      activate elems;
   in ((fixed, deps, elems'), (parms, ctxt')) end;
 
 in
 
+fun cert_declaration x = prep_declaration cert_full_context_statement Element.activate_i x;
+fun cert_read_declaration x = prep_declaration cert_read_full_context_statement Element.activate x;
 fun read_declaration x = prep_declaration read_full_context_statement Element.activate x;
-fun cert_declaration x = prep_declaration cert_full_context_statement Element.activate_i x;
 
 end;
 
@@ -440,7 +468,7 @@
 
 fun props_of thy (name, morph) =
   let
-    val (asm, defs) = NewLocale.specification_of thy name;
+    val (asm, defs) = Locale.specification_of thy name;
   in
     (case asm of NONE => defs | SOME asm => asm :: defs) |> map (Morphism.term morph)
   end;
@@ -449,7 +477,8 @@
   let
     val thy = ProofContext.theory_of context;
 
-    val ((fixed, deps, _, _), _) = prep true true context expression [] [];
+    val ((fixed, deps, _, _), _) =
+      prep true true expression [] [] context;
     (* proof obligations *)
     val propss = map (props_of thy) deps;
 
@@ -467,8 +496,8 @@
     
 in
 
+fun cert_goal_expression x = prep_goal_expression cert_full_context_statement x;
 fun read_goal_expression x = prep_goal_expression read_full_context_statement x;
-fun cert_goal_expression x = prep_goal_expression cert_full_context_statement x;
 
 end;
 
@@ -526,7 +555,7 @@
 fun eval_inst ctxt (loc, morph) text =
   let
     val thy = ProofContext.theory_of ctxt;
-    val (asm, defs) = NewLocale.specification_of thy loc;
+    val (asm, defs) = Locale.specification_of thy loc;
     val asm' = Option.map (Morphism.term morph) asm;
     val defs' = map (Morphism.term morph) defs;
     val text' = text |>
@@ -536,7 +565,7 @@
       (if not (null defs)
         then eval_text ctxt false (Defines (map (fn def => (Attrib.empty_binding, (def, []))) defs'))
         else I)
-(* FIXME clone from new_locale.ML *)
+(* FIXME clone from locale.ML *)
   in text' end;
 
 fun eval_elem ctxt elem text =
@@ -594,10 +623,11 @@
     val name = Sign.full_bname thy bname;
 
     val (body, bodyT, body_eq) = atomize_spec thy norm_ts;
-    val env = Term.add_term_free_names (body, []);
+    val env = Term.add_free_names body [];
     val xs = filter (member (op =) env o #1) parms;
     val Ts = map #2 xs;
-    val extraTs = (Term.term_tfrees body \\ fold Term.add_tfreesT Ts [])
+    val extraTs =
+      (Term.add_tfrees body [] \\ fold Term.add_tfreesT Ts [])
       |> Library.sort_wrt #1 |> map TFree;
     val predT = map Term.itselfT extraTs ---> Ts ---> bodyT;
 
@@ -652,7 +682,7 @@
             |> Sign.add_path aname
             |> Sign.no_base_names
             |> PureThy.note_thmss Thm.internalK
-              [((Binding.name introN, []), [([intro], [NewLocale.unfold_attrib])])]
+              [((Binding.name introN, []), [([intro], [Locale.unfold_attrib])])]
             ||> Sign.restore_naming thy';
           in (SOME statement, SOME intro, axioms, thy'') end;
     val (b_pred, b_intro, b_axioms, thy'''') =
@@ -667,7 +697,7 @@
             |> Sign.add_path pname
             |> Sign.no_base_names
             |> PureThy.note_thmss Thm.internalK
-                 [((Binding.name introN, []), [([intro], [NewLocale.intro_attrib])]),
+                 [((Binding.name introN, []), [([intro], [Locale.intro_attrib])]),
                   ((Binding.name axiomsN, []),
                     [(map (Drule.standard o Element.conclude_witness) axioms, [])])]
             ||> Sign.restore_naming thy''';
@@ -689,22 +719,24 @@
 fun defines_to_notes thy (Defines defs) =
       Notes (Thm.definitionK, map (fn (a, (def, _)) =>
         (a, [([Assumption.assume (cterm_of thy def)],
-          [(Attrib.internal o K) NewLocale.witness_attrib])])) defs)
+          [(Attrib.internal o K) Locale.witness_attrib])])) defs)
   | defines_to_notes _ e = e;
 
 fun gen_add_locale prep_decl
-    bname predicate_name raw_imprt raw_body thy =
+    bname raw_predicate_bname raw_imprt raw_body thy =
   let
     val name = Sign.full_bname thy bname;
-    val _ = NewLocale.test_locale thy name andalso
+    val _ = Locale.defined thy name andalso
       error ("Duplicate definition of locale " ^ quote name);
 
     val ((fixed, deps, body_elems), (parms, ctxt')) =
       prep_decl raw_imprt raw_body (ProofContext.init thy);
     val text as (((_, exts'), _), defs) = eval ctxt' deps body_elems;
 
+    val predicate_bname = if raw_predicate_bname = "" then bname
+      else raw_predicate_bname;
     val ((a_statement, a_intro, a_axioms), (b_statement, b_intro, b_axioms), thy') =
-      define_preds predicate_name parms text thy;
+      define_preds predicate_bname parms text thy;
 
     val extraTs = fold Term.add_tfrees exts' [] \\ fold Term.add_tfreesT (map snd parms) [];
     val _ = if null extraTs then ()
@@ -717,15 +749,13 @@
       (body_elems |> map_filter (fn Fixes fixes => SOME fixes | _ => NONE) |> flat);
     val asm = if is_some b_statement then b_statement else a_statement;
 
-    (* These are added immediately. *)
     val notes =
         if is_some asm
         then [(Thm.internalK, [((Binding.name (bname ^ "_" ^ axiomsN), []),
           [([Assumption.assume (cterm_of thy' (the asm))],
-            [(Attrib.internal o K) NewLocale.witness_attrib])])])]
+            [(Attrib.internal o K) Locale.witness_attrib])])])]
         else [];
 
-    (* These will be added in the local theory. *)
     val notes' = body_elems |>
       map (defines_to_notes thy') |>
       map (Element.morph_ctxt a_satisfy) |>
@@ -735,19 +765,21 @@
       map_filter (fn Notes notes => SOME notes | _ => NONE);
 
     val deps' = map (fn (l, morph) => (l, morph $> b_satisfy)) deps;
+    val axioms = map Element.conclude_witness b_axioms;
 
-    val loc_ctxt = thy' |>
-      NewLocale.register_locale bname (extraTs, params)
-        (asm, rev defs) ([], [])
-        (map (fn n => (n, stamp ())) notes |> rev) (map (fn d => (d, stamp ())) deps' |> rev) |>
-      NewLocale.init name;
+    val loc_ctxt = thy'
+      |> Locale.register_locale bname (extraTs, params)
+          (asm, rev defs) (a_intro, b_intro) axioms ([], []) 
+          (map (fn n => (n, stamp ())) notes |> rev) (map (fn d => (d, stamp ())) deps' |> rev)
+      |> TheoryTarget.init (SOME name)
+      |> fold (fn (kind, facts) => LocalTheory.notes kind facts #> snd) notes';
 
-  in ((name, notes'), loc_ctxt) end;
+  in (name, loc_ctxt) end;
 
 in
 
+val add_locale = gen_add_locale cert_declaration;
 val add_locale_cmd = gen_add_locale read_declaration;
-val add_locale = gen_add_locale cert_declaration;
 
 end;
 
@@ -758,8 +790,8 @@
 
 fun prep_propp propss = propss |> map (map (rpair [] o Element.mark_witness));
 
-fun prep_result propps thmss =
-  ListPair.map (fn (props, thms) => map2 Element.make_witness props thms) (propps, thmss);
+val prep_result = map2 (fn props => fn thms =>
+  map2 Element.make_witness props (map Thm.close_derivation thms));
 
 
 (** Interpretation between locales: declaring sublocale relationships **)
@@ -770,20 +802,20 @@
     raw_target expression thy =
   let
     val target = intern thy raw_target;
-    val target_ctxt = NewLocale.init target thy;
+    val target_ctxt = Locale.init target thy;
 
     val ((propss, deps, export), goal_ctxt) = prep_expr expression target_ctxt;
     
-    fun store_dep ((name, morph), thms) =
-      NewLocale.add_dependency target (name, morph $> Element.satisfy_morphism thms $> export);
+    fun store_dep (name, morph) thms =
+      Locale.add_dependency target (name, morph $> Element.satisfy_morphism thms $> export);
 
     fun after_qed results =
       ProofContext.theory (
         (* store dependencies *)
-        fold store_dep (deps ~~ prep_result propss results) #>
+        fold2 store_dep deps (prep_result propss results) #>
         (* propagate registrations *)
-        (fn thy => fold_rev (fn reg => NewLocale.activate_global_facts reg)
-          (NewLocale.get_global_registrations thy) thy));
+        (fn thy => fold_rev Locale.activate_global_facts
+          (Locale.get_registrations thy) thy));
   in
     goal_ctxt |>
       Proof.theorem_i NONE after_qed (prep_propp propss) |>
@@ -792,8 +824,8 @@
 
 in
 
-fun sublocale_cmd x = gen_sublocale read_goal_expression NewLocale.intern x;
 fun sublocale x = gen_sublocale cert_goal_expression (K I) x;
+fun sublocale_cmd x = gen_sublocale read_goal_expression Locale.intern x;
 
 end;
 
@@ -802,52 +834,55 @@
 
 local
 
-datatype goal = Reg of string * Morphism.morphism | Eqns of Attrib.binding list;
-
 fun gen_interpretation prep_expr parse_prop prep_attr
-    expression equations thy =
+    expression equations theory =
   let
-    val ctxt = ProofContext.init thy;
+    val ((propss, regs, export), expr_ctxt) = ProofContext.init theory
+      |> prep_expr expression;
 
-    val ((propss, regs, export), expr_ctxt) = prep_expr expression ctxt;
-    
     val eqns = map (parse_prop expr_ctxt o snd) equations |> Syntax.check_terms expr_ctxt;
-    val eq_attns = map ((apsnd o map) (prep_attr thy) o fst) equations;
+    val eq_attns = map ((apsnd o map) (prep_attr theory) o fst) equations;
     val goal_ctxt = fold Variable.auto_fixes eqns expr_ctxt;
     val export' = Variable.export_morphism goal_ctxt expr_ctxt;
 
-    fun store (Reg (name, morph), thms) (regs, thy) =
-        let
-          val thms' = map (Element.morph_witness export') thms;
-          val morph' = morph $> Element.satisfy_morphism thms';
-          val add = NewLocale.add_global_registration (name, (morph', export));
-        in ((name, morph') :: regs, add thy) end
-      | store (Eqns [], []) (regs, thy) =
-        let val add = fold_rev (fn (name, morph) =>
-              NewLocale.activate_global_facts (name, morph $> export)) regs;
-        in (regs, add thy) end
-      | store (Eqns attns, thms) (regs, thy) =
-        let
-          val thms' = thms |> map (Element.conclude_witness #>
-            Morphism.thm (export' $> export) #>
-            LocalDefs.meta_rewrite_rule (ProofContext.init thy) #>
-            Drule.abs_def);
-          val eq_morph =
-            Morphism.term_morphism (MetaSimplifier.rewrite_term thy thms' []) $>
-            Morphism.thm_morphism (MetaSimplifier.rewrite_rule thms');
-          val attns' = map ((apsnd o map) (Attrib.attribute_i thy)) attns;
-          val add =
-            fold_rev (fn (name, morph) =>
-              NewLocale.amend_global_registration eq_morph (name, morph) #>
-              NewLocale.activate_global_facts (name, morph $> eq_morph $> export)) regs #>
-            PureThy.note_thmss Thm.lemmaK (attns' ~~ map (fn th => [([th], [])]) thms') #>
-            snd
-        in (regs, add thy) end;
+    fun store_reg ((name, morph), thms) thy =
+      let
+        val thms' = map (Element.morph_witness export') thms;
+        val morph' = morph $> Element.satisfy_morphism thms';
+      in
+        thy
+        |> Locale.add_registration (name, (morph', export))
+        |> pair (name, morph')
+      end;
+
+    fun store_eqns_activate regs [] thy =
+          thy
+          |> fold (fn (name, morph) =>
+               Locale.activate_global_facts (name, morph $> export)) regs
+      | store_eqns_activate regs thms thy =
+          let
+            val thms' = thms |> map (Element.conclude_witness #>
+              Morphism.thm (export' $> export) #>
+              LocalDefs.meta_rewrite_rule (ProofContext.init thy) #>
+              Drule.abs_def);
+            val eq_morph = Element.eq_morphism thy thms';
+            val eq_attns' = map ((apsnd o map) (Attrib.attribute_i thy)) eq_attns;
+          in
+            thy
+            |> fold (fn (name, morph) =>
+                Locale.amend_registration eq_morph (name, morph) #>
+                Locale.activate_global_facts (name, morph $> eq_morph $> export)) regs
+            |> PureThy.note_thmss Thm.lemmaK (eq_attns' ~~ map (fn th => [([th], [])]) thms')
+            |> snd
+          end;
 
     fun after_qed results =
-      ProofContext.theory (fn thy =>
-        fold store (map Reg regs @ [Eqns eq_attns] ~~
-          prep_result (propss @ [eqns]) results) ([], thy) |> snd);
+      let
+        val (wits_reg, wits_eq) = split_last (prep_result (propss @ [eqns]) results);
+      in ProofContext.theory (fold_map store_reg (regs ~~ wits_reg)
+        #-> (fn regs => store_eqns_activate regs wits_eq))
+      end;
+
   in
     goal_ctxt |>
       Proof.theorem_i NONE after_qed (prep_propp (propss @ [eqns])) |>
@@ -856,9 +891,9 @@
 
 in
 
+fun interpretation x = gen_interpretation cert_goal_expression (K I) (K I) x;
 fun interpretation_cmd x = gen_interpretation read_goal_expression
   Syntax.parse_prop Attrib.intern_src x;
-fun interpretation x = gen_interpretation cert_goal_expression (K I) (K I) x;
 
 end;
 
@@ -879,11 +914,11 @@
       let
         val morph' = morph $> Element.satisfy_morphism thms $> export;
       in
-        NewLocale.activate_local_facts (name, morph')
+        Locale.activate_local_facts (name, morph')
       end;
 
     fun after_qed results =
-      Proof.map_context (fold store_reg (regs ~~ prep_result propss results)) #> Seq.single;
+      Proof.map_context (fold store_reg (regs ~~ prep_result propss results));
   in
     state |> Proof.map_context (K goal_ctxt) |>
       Proof.local_goal (ProofDisplay.print_results int) (K I) ProofContext.bind_propp_i
@@ -893,8 +928,8 @@
 
 in
 
+fun interpret x = gen_interpret cert_goal_expression x;
 fun interpret_cmd x = gen_interpret read_goal_expression x;
-fun interpret x = gen_interpret cert_goal_expression x;
 
 end;
 
--- a/src/Pure/Isar/find_theorems.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Isar/find_theorems.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/find_theorems.ML
-    ID:         $Id$
     Author:     Rafal Kolanski and Gerwin Klein, NICTA
 
 Retrieve theorems from proof context.
@@ -166,7 +165,7 @@
 fun filter_simp ctxt t (_, thm) =
   let
     val (_, {mk_rews = {mk, ...}, ...}) =
-      MetaSimplifier.rep_ss (Simplifier.local_simpset_of ctxt);
+      Simplifier.rep_ss (Simplifier.local_simpset_of ctxt);
     val extract_simp =
       (map Thm.full_prop_of o mk, #1 o Logic.dest_equals o Logic.strip_imp_concl);
     val ss = is_matching_thm extract_simp ctxt false t thm
@@ -177,9 +176,9 @@
 
 (* filter_pattern *)
 
-fun get_names (_, thm) = let
-    val t = Thm.full_prop_of thm;
-  in (term_consts t) union (add_term_free_names (t, [])) end;
+fun get_names (_, thm) =
+  fold_aterms (fn Const (c, _) => insert (op =) c | Free (x, _) => insert (op =) x | _ => I)
+    (Thm.full_prop_of thm) [];
 
 fun add_pat_names (t, cs) =
       case strip_comb t of
@@ -287,7 +286,7 @@
 
 fun rem_thm_dups xs =
   xs ~~ (1 upto length xs)
-  |> sort (Term.fast_term_ord o pairself (Thm.prop_of o #2 o #1))
+  |> sort (TermOrd.fast_term_ord o pairself (Thm.prop_of o #2 o #1))
   |> rem_cdups
   |> sort (int_ord o pairself #2)
   |> map #1;
--- a/src/Pure/Isar/instance.ML	Mon Jan 19 20:37:08 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,28 +0,0 @@
-(*  Title:      Pure/Isar/instance.ML
-    ID:         $Id$
-    Author:     Florian Haftmann, TU Muenchen
-
-Wrapper for instantiation command.
-*)
-
-signature INSTANCE =
-sig
-  val instantiation_cmd: xstring list * sort * xstring -> theory -> local_theory
-end;
-
-structure Instance : INSTANCE =
-struct
-
-fun read_multi_arity thy (raw_tycos, raw_sorts, raw_sort) =
-  let
-    val all_arities = map (fn raw_tyco => Sign.read_arity thy
-      (raw_tyco, raw_sorts, raw_sort)) raw_tycos;
-    val tycos = map #1 all_arities;
-    val (_, sorts, sort) = hd all_arities;
-    val vs = Name.names Name.context Name.aT sorts;
-  in (tycos, vs, sort) end;
-
-fun instantiation_cmd raw_arities thy =
-  TheoryTarget.instantiation (read_multi_arity thy raw_arities) thy;
-
-end;
--- a/src/Pure/Isar/isar.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Isar/isar.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/isar.ML
-    ID:         $Id$
     Author:     Makarius
 
 The global Isabelle/Isar state and main read-eval-print loop.
@@ -7,15 +6,12 @@
 
 signature ISAR =
 sig
-  type id = string
-  val no_id: id
-  val create_command: Toplevel.transition -> id
-  val init_point: unit -> unit
+  val init: unit -> unit
+  val exn: unit -> (exn * string) option
   val state: unit -> Toplevel.state
   val context: unit -> Proof.context
   val goal: unit -> thm
   val print: unit -> unit
-  val exn: unit -> (exn * string) option
   val >> : Toplevel.transition -> bool
   val >>> : Toplevel.transition list -> unit
   val linear_undo: int -> unit
@@ -26,6 +22,10 @@
   val toplevel_loop: {init: bool, welcome: bool, sync: bool, secure: bool} -> unit
   val loop: unit -> unit
   val main: unit -> unit
+
+  type id = string
+  val no_id: id
+  val create_command: Toplevel.transition -> id
   val insert_command: id -> id -> unit
   val remove_command: id -> unit
 end;
@@ -34,6 +34,131 @@
 struct
 
 
+(** TTY model -- SINGLE-THREADED! **)
+
+(* the global state *)
+
+type history = (Toplevel.state * Toplevel.transition) list;
+  (*previous state, state transition -- regular commands only*)
+
+local
+  val global_history = ref ([]: history);
+  val global_state = ref Toplevel.toplevel;
+  val global_exn = ref (NONE: (exn * string) option);
+in
+
+fun edit_history count f = NAMED_CRITICAL "Isar" (fn () =>
+  let
+    fun edit 0 (st, hist) = (global_history := hist; global_state := st; global_exn := NONE)
+      | edit n (st, hist) = edit (n - 1) (f st hist);
+  in edit count (! global_state, ! global_history) end);
+
+fun state () = NAMED_CRITICAL "Isar" (fn () => ! global_state);
+fun set_state state = NAMED_CRITICAL "Isar" (fn () => global_state := state);
+
+fun exn () = NAMED_CRITICAL "Isar" (fn () => ! global_exn);
+fun set_exn exn =  NAMED_CRITICAL "Isar" (fn () => global_exn := exn);
+
+end;
+
+
+fun init () = edit_history 1 (K (K (Toplevel.toplevel, [])));
+
+fun context () = Toplevel.context_of (state ())
+  handle Toplevel.UNDEF => error "Unknown context";
+
+fun goal () = #2 (#2 (Proof.get_goal (Toplevel.proof_of (state ()))))
+  handle Toplevel.UNDEF => error "No goal present";
+
+fun print () = Toplevel.print_state false (state ());
+
+
+(* history navigation *)
+
+local
+
+fun find_and_undo _ [] = error "Undo history exhausted"
+  | find_and_undo which ((prev, tr) :: hist) =
+      ((case Toplevel.init_of tr of SOME name => ThyInfo.kill_thy name | NONE => ());
+        if which (Toplevel.name_of tr) then (prev, hist) else find_and_undo which hist);
+
+in
+
+fun linear_undo n = edit_history n (K (find_and_undo (K true)));
+
+fun undo n = edit_history n (fn st => fn hist =>
+  find_and_undo (if Toplevel.is_proof st then K true else OuterKeyword.is_theory) hist);
+
+fun kill () = edit_history 1 (fn st => fn hist =>
+  find_and_undo
+    (if Toplevel.is_proof st then OuterKeyword.is_theory else OuterKeyword.is_theory_begin) hist);
+
+fun kill_proof () = edit_history 1 (fn st => fn hist =>
+  if Toplevel.is_proof st then find_and_undo OuterKeyword.is_theory hist
+  else raise Toplevel.UNDEF);
+
+end;
+
+
+(* interactive state transformations *)
+
+fun op >> tr =
+  (case Toplevel.transition true tr (state ()) of
+    NONE => false
+  | SOME (_, SOME err) => (set_exn (SOME err); Toplevel.error_msg tr err; true)
+  | SOME (st', NONE) =>
+      let
+        val name = Toplevel.name_of tr;
+        val _ = if OuterKeyword.is_theory_begin name then init () else ();
+        val _ =
+          if OuterKeyword.is_regular name
+          then edit_history 1 (fn st => fn hist => (st', (st, tr) :: hist)) else ();
+      in true end);
+
+fun op >>> [] = ()
+  | op >>> (tr :: trs) = if op >> tr then op >>> trs else ();
+
+
+(* toplevel loop *)
+
+val crashes = ref ([]: exn list);
+
+local
+
+fun raw_loop secure src =
+  let
+    fun check_secure () =
+      (if secure then warning "Secure loop -- cannot exit to ML" else (); secure);
+  in
+    (case Source.get_single (Source.set_prompt Source.default_prompt src) of
+      NONE => if secure then quit () else ()
+    | SOME (tr, src') => if op >> tr orelse check_secure () then raw_loop secure src' else ())
+    handle exn =>
+      (Output.error_msg (Toplevel.exn_message exn)
+        handle crash =>
+          (CRITICAL (fn () => change crashes (cons crash));
+            warning "Recovering from Isar toplevel crash -- see also Isar.crashes");
+          raw_loop secure src)
+  end;
+
+in
+
+fun toplevel_loop {init = do_init, welcome, sync, secure} =
+ (Context.set_thread_data NONE;
+  if do_init then init () else ();  (* FIXME init editor model *)
+  if welcome then writeln (Session.welcome ()) else ();
+  uninterruptible (fn _ => fn () => raw_loop secure (OuterSyntax.isar sync)) ());
+
+end;
+
+fun loop () =
+  toplevel_loop {init = false, welcome = false, sync = false, secure = Secure.is_secure ()};
+
+fun main () =
+  toplevel_loop {init = true, welcome = true, sync = false, secure = Secure.is_secure ()};
+
+
+
 (** individual toplevel commands **)
 
 (* unique identification *)
@@ -41,15 +166,6 @@
 type id = string;
 val no_id : id = "";
 
-fun identify tr =
-  (case Toplevel.get_id tr of
-    SOME id => (id, tr)
-  | NONE =>
-      let val id =
-        if ! Toplevel.debug then "isabelle:" ^ Toplevel.name_of tr ^ serial_string ()
-        else "isabelle:" ^ serial_string ()
-      in (id, Toplevel.put_id id tr) end);
-
 
 (* command category *)
 
@@ -78,7 +194,7 @@
   Finished of Toplevel.state;
 
 fun status_markup Unprocessed = Markup.unprocessed
-  | status_markup Running = Markup.running
+  | status_markup Running = (Markup.runningN, [])
   | status_markup (Failed _) = Markup.failed
   | status_markup (Finished _) = Markup.finished;
 
@@ -168,7 +284,15 @@
 
 fun create_command raw_tr =
   let
-    val (id, tr) = identify raw_tr;
+    val (id, tr) =
+      (case Toplevel.get_id raw_tr of
+        SOME id => (id, raw_tr)
+      | NONE =>
+          let val id =
+            if ! Toplevel.debug then "isabelle:" ^ Toplevel.name_of raw_tr ^ serial_string ()
+            else "isabelle:" ^ serial_string ()
+          in (id, Toplevel.put_id id raw_tr) end);
+
     val cmd = make_command (category_of tr, tr, Unprocessed);
     val _ = change_commands (Graph.new_node (id, cmd));
   in id end;
@@ -190,153 +314,6 @@
 
 
 
-(** TTY model -- single-threaded **)
-
-(* global point *)
-
-local val global_point = ref no_id in
-
-fun change_point f = NAMED_CRITICAL "Isar" (fn () => change global_point f);
-fun point () = NAMED_CRITICAL "Isar" (fn () => ! global_point);
-
-end;
-
-
-fun set_point id = change_point (K id);
-fun init_point () = set_point no_id;
-
-fun point_state () = NAMED_CRITICAL "Isar" (fn () =>
-  let val id = point () in (id, the_state id) end);
-
-fun state () = #2 (point_state ());
-
-fun context () =
-  Toplevel.context_of (state ())
-    handle Toplevel.UNDEF => error "Unknown context";
-
-fun goal () =
-  #2 (#2 (Proof.get_goal (Toplevel.proof_of (state ()))))
-    handle Toplevel.UNDEF => error "No goal present";
-
-fun print () = Toplevel.print_state false (state ());
-
-
-(* global failure status *)
-
-local val global_exn = ref (NONE: (exn * string) option) in
-
-fun set_exn err = global_exn := err;
-fun exn () = ! global_exn;
-
-end;
-
-
-(* interactive state transformations *)
-
-fun op >> raw_tr =
-  let
-    val id = create_command raw_tr;
-    val {category, transition = tr, ...} = the_command id;
-    val (prev, prev_state) = point_state ();
-    val _ =
-      if is_regular category
-      then (dispose_commands (next_commands prev); change_commands (add_edge (prev, id)))
-      else ();
-  in
-    (case run true tr prev_state of
-      NONE => false
-    | SOME (status as Failed err) => (update_status status id; set_exn (SOME err); true)
-    | SOME status =>
-       (update_status status id; set_exn NONE;
-        if is_regular category then set_point id else ();
-        true))
-  end;
-
-fun op >>> [] = ()
-  | op >>> (tr :: trs) = if op >> tr then op >>> trs else ();
-
-
-(* implicit navigation wrt. proper commands *)
-
-local
-
-fun err_undo () = error "Undo history exhausted";
-
-fun find_category which id =
-  (case #category (the_command id) of
-    Empty => err_undo ()
-  | category => if which category then id else find_category which (prev_command id));
-
-fun find_begin_theory id =
-  if id = no_id then err_undo ()
-  else if is_some (Toplevel.init_of (#transition (the_command id))) then id
-  else find_begin_theory (prev_command id);
-
-fun undo_command id =
-  (case Toplevel.init_of (#transition (the_command id)) of
-    SOME name => prev_command id before ThyInfo.kill_thy name
-  | NONE => prev_command id);
-
-in
-
-fun linear_undo n = change_point (funpow n (fn id => undo_command (find_category is_proper id)));
-
-fun undo n = change_point (funpow n (fn id => undo_command
-  (find_category (if Toplevel.is_proof (the_state id) then is_proper else is_theory) id)));
-
-fun kill () = change_point (fn id => undo_command
-  (if Toplevel.is_proof (the_state id) then find_category is_theory id else find_begin_theory id));
-
-fun kill_proof () = change_point (fn id =>
-  if Toplevel.is_proof (the_state id) then undo_command (find_category is_theory id)
-  else raise Toplevel.UNDEF);
-
-end;
-
-
-(* toplevel loop *)
-
-val crashes = ref ([]: exn list);
-
-local
-
-fun raw_loop secure src =
-  let
-    fun check_secure () =
-      (if secure then warning "Secure loop -- cannot exit to ML" else (); secure);
-    val prev = point ();
-    val prev_name = Toplevel.name_of (#transition (the_command prev));
-    val prompt_markup =
-      prev <> no_id ? Markup.markup
-        (Markup.properties [(Markup.idN, prev), (Markup.nameN, prev_name)] Markup.prompt);
-  in
-    (case Source.get_single (Source.set_prompt (prompt_markup Source.default_prompt) src) of
-      NONE => if secure then quit () else ()
-    | SOME (tr, src') => if op >> tr orelse check_secure () then raw_loop secure src' else ())
-    handle exn => (Output.error_msg (Toplevel.exn_message exn) handle crash =>
-      (CRITICAL (fn () => change crashes (cons crash));
-        warning "Recovering after Isar toplevel crash -- see also Isar.crashes");
-      raw_loop secure src)
-  end;
-
-in
-
-fun toplevel_loop {init, welcome, sync, secure} =
- (Context.set_thread_data NONE;
-  if init then (init_point (); init_commands ()) else ();
-  if welcome then writeln (Session.welcome ()) else ();
-  uninterruptible (fn _ => fn () => raw_loop secure (OuterSyntax.isar sync)) ());
-
-end;
-
-fun loop () =
-  toplevel_loop {init = false, welcome = false, sync = false, secure = Secure.is_secure ()};
-
-fun main () =
-  toplevel_loop {init = true, welcome = true, sync = false, secure = Secure.is_secure ()};
-
-
-
 (** editor model **)
 
 (* run commands *)
@@ -377,26 +354,25 @@
 
 local
 
-structure P = OuterParse and K = OuterKeyword;
+structure P = OuterParse;
 val op >> = Scan.>>;
 
 in
 
 val _ =
-  OuterSyntax.improper_command "Isar.command" "define command (Isar editor model)" K.control
-    (P.props_text >> (fn (pos, str) =>
-      Toplevel.no_timing o Toplevel.imperative (fn () =>
-        ignore (create_command (OuterSyntax.prepare_command pos str)))));
+  OuterSyntax.internal_command "Isar.command"
+    (P.string -- P.string >> (fn (id, text) =>
+      Toplevel.imperative (fn () =>
+        ignore (create_command (OuterSyntax.prepare_command (Position.id id) text)))));
 
 val _ =
-  OuterSyntax.improper_command "Isar.insert" "insert command (Isar editor model)" K.control
+  OuterSyntax.internal_command "Isar.insert"
     (P.string -- P.string >> (fn (prev, id) =>
-      Toplevel.no_timing o Toplevel.imperative (fn () => insert_command prev id)));
+      Toplevel.imperative (fn () => insert_command prev id)));
 
 val _ =
-  OuterSyntax.improper_command "Isar.remove" "remove command (Isar editor model)" K.control
-    (P.string >> (fn id =>
-      Toplevel.no_timing o Toplevel.imperative (fn () => remove_command id)));
+  OuterSyntax.internal_command "Isar.remove"
+    (P.string >> (fn id => Toplevel.imperative (fn () => remove_command id)));
 
 end;
 
--- a/src/Pure/Isar/isar.scala	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Isar/isar.scala	Mon Jan 19 21:20:18 2009 +0100
@@ -1,22 +1,19 @@
 /*  Title:      Pure/Isar/isar.scala
     Author:     Makarius
 
-Isar toplevel editor model.
+Isar document model.
 */
 
 package isabelle
 
-import java.util.Properties
-
 
 class Isar(isabelle_system: IsabelleSystem, results: EventBus[IsabelleProcess.Result], args: String*)
   extends IsabelleProcess(isabelle_system, results, args: _*)
 {
-
   /* basic editor commands */
 
-  def create_command(props: Properties, text: String) =
-    output_sync("Isar.command " + IsabelleSyntax.encode_properties(props) + " " +
+  def create_command(id: String, text: String) =
+    output_sync("Isar.command " + IsabelleSyntax.encode_string(id) + " " +
       IsabelleSyntax.encode_string(text))
 
   def insert_command(prev: String, id: String) =
@@ -25,5 +22,4 @@
 
   def remove_command(id: String) =
     output_sync("Isar.remove " + IsabelleSyntax.encode_string(id))
-
 }
--- a/src/Pure/Isar/isar_cmd.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Isar/isar_cmd.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/isar_cmd.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Derived Isar commands.
@@ -219,7 +218,7 @@
 (* goals *)
 
 fun goal opt_chain goal stmt int =
-  opt_chain #> goal NONE (K Seq.single) stmt int;
+  opt_chain #> goal NONE (K I) stmt int;
 
 val have = goal I Proof.have;
 val hence = goal Proof.chain Proof.have;
@@ -229,12 +228,12 @@
 
 (* local endings *)
 
-fun local_qed m = Toplevel.proofs (Proof.local_qed (m, true));
-val local_terminal_proof = Toplevel.proofs o Proof.local_terminal_proof;
-val local_default_proof = Toplevel.proofs Proof.local_default_proof;
-val local_immediate_proof = Toplevel.proofs Proof.local_immediate_proof;
-val local_done_proof = Toplevel.proofs Proof.local_done_proof;
-val local_skip_proof = Toplevel.proofs' Proof.local_skip_proof;
+fun local_qed m = Toplevel.proof (Proof.local_qed (m, true));
+val local_terminal_proof = Toplevel.proof o Proof.local_terminal_proof;
+val local_default_proof = Toplevel.proof Proof.local_default_proof;
+val local_immediate_proof = Toplevel.proof Proof.local_immediate_proof;
+val local_done_proof = Toplevel.proof Proof.local_done_proof;
+val local_skip_proof = Toplevel.proof' Proof.local_skip_proof;
 
 val skip_local_qed = Toplevel.skip_proof (fn i => if i > 1 then i - 1 else raise Toplevel.UNDEF);
 
@@ -354,17 +353,17 @@
 val print_theorems = Toplevel.unknown_context o print_theorems_theory o print_theorems_proof;
 
 val print_locales = Toplevel.unknown_theory o
-  Toplevel.keep (NewLocale.print_locales o Toplevel.theory_of);
+  Toplevel.keep (Locale.print_locales o Toplevel.theory_of);
 
 fun print_locale (show_facts, name) = Toplevel.unknown_theory o
   Toplevel.keep (fn state =>
-    NewLocale.print_locale (Toplevel.theory_of state) show_facts name);
+    Locale.print_locale (Toplevel.theory_of state) show_facts name);
 
 fun print_registrations show_wits name = Toplevel.unknown_context o
   Toplevel.keep (Toplevel.node_case
-      (Context.cases (Locale.print_registrations show_wits name o ProofContext.init)
-        (Locale.print_registrations show_wits name))
-    (Locale.print_registrations show_wits name o Proof.context_of));
+      (Context.cases (Old_Locale.print_registrations show_wits name o ProofContext.init)
+        (Old_Locale.print_registrations show_wits name))
+    (Old_Locale.print_registrations show_wits name o Proof.context_of));
 
 val print_attributes = Toplevel.unknown_theory o
   Toplevel.keep (Attrib.print_attributes o Toplevel.theory_of);
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Isar/isar_document.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -0,0 +1,301 @@
+(*  Title:      Pure/Isar/isar_document.ML
+    Author:     Makarius
+
+Interactive Isar documents.
+*)
+
+signature ISAR_DOCUMENT =
+sig
+  type state_id = string
+  type command_id = string
+  type document_id = string
+  val define_command: command_id -> Toplevel.transition -> unit
+  val begin_document: document_id -> Path.T -> unit
+  val end_document: document_id -> unit
+  val edit_document: document_id -> document_id -> (command_id * command_id option) list -> unit
+end;
+
+structure IsarDocument: ISAR_DOCUMENT =
+struct
+
+(* unique identifiers *)
+
+type state_id = string;
+type command_id = string;
+type document_id = string;
+
+fun make_id () = "isabelle:" ^ serial_string ();
+
+fun err_dup kind id = error ("Duplicate " ^ kind ^ ": " ^ quote id);
+fun err_undef kind id = error ("Unknown " ^ kind ^ ": " ^ quote id);
+
+
+
+(** documents **)
+
+(* command entries *)
+
+datatype entry = Entry of {next: command_id option, state: state_id option};
+fun make_entry next state = Entry {next = next, state = state};
+
+fun the_entry entries (id: command_id) =
+  (case Symtab.lookup entries id of
+    NONE => err_undef "document entry" id
+  | SOME (Entry entry) => entry);
+
+fun put_entry (id: command_id, entry: entry) = Symtab.update (id, entry);
+
+fun put_entry_state (id: command_id) state entries =
+  let val {next, ...} = the_entry entries id
+  in put_entry (id, make_entry next state) entries end;
+
+fun reset_entry_state id = put_entry_state id NONE;
+fun set_entry_state (id, state_id) = put_entry_state id (SOME state_id);
+
+
+(* document *)
+
+datatype document = Document of
+ {dir: Path.T,                    (*master directory*)
+  name: string,                   (*theory name*)
+  start: command_id,              (*empty start command*)
+  entries: entry Symtab.table};   (*unique command entries indexed by command_id*)
+
+fun make_document dir name start entries =
+  Document {dir = dir, name = name, start = start, entries = entries};
+
+fun map_entries f (Document {dir, name, start, entries}) =
+  make_document dir name start (f entries);
+
+
+(* iterate entries *)
+
+fun fold_entries id0 f (Document {entries, ...}) =
+  let
+    fun apply NONE x = x
+      | apply (SOME id) x =
+          let val entry = the_entry entries id
+          in apply (#next entry) (f (id, entry) x) end;
+  in if Symtab.defined entries id0 then apply (SOME id0) else I end;
+
+fun first_entry P (Document {start, entries, ...}) =
+  let
+    fun first _ NONE = NONE
+      | first prev (SOME id) =
+          let val entry = the_entry entries id
+          in if P (id, entry) then SOME (prev, id, entry) else first (SOME id) (#next entry) end;
+  in first NONE (SOME start) end;
+
+
+(* modify entries *)
+
+fun insert_after (id: command_id) (id2: command_id) = map_entries (fn entries =>
+  let val {next, state} = the_entry entries id in
+    entries
+    |> put_entry (id, make_entry (SOME id2) state)
+    |> put_entry (id2, make_entry next NONE)
+  end);
+
+fun delete_after (id: command_id) = map_entries (fn entries =>
+  let val {next, state} = the_entry entries id in
+    (case next of
+      NONE => error ("No next entry to delete: " ^ quote id)
+    | SOME id2 =>
+        entries |>
+          (case #next (the_entry entries id2) of
+            NONE => put_entry (id, make_entry NONE state)
+          | SOME id3 => put_entry (id, make_entry (SOME id3) state) #> reset_entry_state id3))
+  end);
+
+
+
+(** global configuration **)
+
+local
+  val global_states = ref (Symtab.empty: Toplevel.state option future Symtab.table);
+  val global_commands = ref (Symtab.empty: Toplevel.transition Symtab.table);
+  val global_documents = ref (Symtab.empty: document Symtab.table);
+in
+
+fun change_states f = NAMED_CRITICAL "Isar" (fn () => change global_states f);
+fun get_states () = NAMED_CRITICAL "Isar" (fn () => ! global_states);
+
+fun change_commands f = NAMED_CRITICAL "Isar" (fn () => change global_commands f);
+fun get_commands () = NAMED_CRITICAL "Isar" (fn () => ! global_commands);
+
+fun change_documents f = NAMED_CRITICAL "Isar" (fn () => change global_documents f);
+fun get_documents () = NAMED_CRITICAL "Isar" (fn () => ! global_documents);
+
+fun init () = NAMED_CRITICAL "Isar" (fn () =>
+ (global_states := Symtab.empty;
+  global_commands := Symtab.empty;
+  global_documents := Symtab.empty));
+
+end;
+
+
+(* state *)
+
+val empty_state = Future.value (SOME Toplevel.toplevel);
+
+fun define_state (id: state_id) =
+  change_states (Symtab.update_new (id, empty_state))
+    handle Symtab.DUP dup => err_dup "state" dup;
+
+fun put_state (id: state_id) state = change_states (Symtab.update (id, state));
+
+fun the_state (id: state_id) =
+  (case Symtab.lookup (get_states ()) id of
+    NONE => err_undef "state" id
+  | SOME state => state);
+
+
+(* command *)
+
+fun define_command id tr =
+  change_commands (Symtab.update_new (id, Toplevel.put_id id tr))
+    handle Symtab.DUP dup => err_dup "command" dup;
+
+fun the_command (id: command_id) =
+  (case Symtab.lookup (get_commands ()) id of
+    NONE => err_undef "command" id
+  | SOME tr => tr);
+
+
+(* document *)
+
+fun define_document (id: document_id) document =
+  change_documents (Symtab.update_new (id, document))
+    handle Symtab.DUP dup => err_dup "document" dup;
+
+fun the_document (id: document_id) =
+  (case Symtab.lookup (get_documents ()) id of
+    NONE => err_undef "document" id
+  | SOME document => document);
+
+
+
+(** main operations **)
+
+(* begin/end document *)
+
+fun begin_document (id: document_id) path =
+  let
+    val (dir, name) = ThyLoad.split_thy_path path;
+    val _ = define_command id Toplevel.empty;
+    val _ = define_state id;
+    val entries = Symtab.make [(id, make_entry NONE (SOME id))];
+    val _ = define_document id (make_document dir name id entries);
+  in () end;
+
+fun end_document (id: document_id) =
+  let
+    val document as Document {name, ...} = the_document id;
+    val end_state =
+      the_state (the (#state (#3 (the
+        (first_entry (fn (_, {next, ...}) => is_none next) document)))));
+    val _ =
+      Future.fork_deps [end_state] (fn () =>
+        (case Future.join end_state of
+          SOME st =>
+           (Toplevel.run_command name (Toplevel.put_id id (Toplevel.commit_exit Position.none)) st;
+            ThyInfo.touch_child_thys name;
+            ThyInfo.register_thy name)
+        | NONE => error ("Failed to finish theory " ^ quote name)));
+  in () end;
+
+
+(* document editing *)
+
+local
+
+fun is_changed entries' (id, {next = _, state}) =
+  (case try (the_entry entries') id of
+    NONE => true
+  | SOME {next = _, state = state'} => state' <> state);
+
+fun new_state name (id, _) (state_id, updates) =
+  let
+    val state_id' = make_id ();
+    val _ = define_state state_id';
+    val tr = Toplevel.put_id state_id' (the_command id);
+    fun make_state' () =
+      let
+        val state = the_state state_id;
+        val state' =
+          Future.fork_deps [state] (fn () =>
+            (case Future.join state of
+              NONE => NONE
+            | SOME st => Toplevel.run_command name tr st));
+      in put_state state_id' state' end;
+  in (state_id', ((id, state_id'), make_state') :: updates) end;
+
+fun report_updates _ [] = ()
+  | report_updates (document_id: document_id) updates =
+      implode (map (fn (id, state_id) => Markup.markup (Markup.edit id state_id) "") updates)
+      |> Markup.markup (Markup.edits document_id)
+      |> Output.status;
+
+in
+
+fun edit_document (old_id: document_id) (new_id: document_id) edits =
+  let
+    val old_document as Document {name, entries = old_entries, ...} = the_document old_id;
+    val new_document as Document {entries = new_entries, ...} = old_document
+      |> fold (fn (id, SOME id2) => insert_after id id2 | (id, NONE) => delete_after id) edits;
+
+    fun cancel_old id = fold_entries id
+      (fn (_, {state = SOME state_id, ...}) => K (Future.cancel (the_state state_id)) | _ => K ())
+      old_document ();
+
+    val (updates, new_document') =
+      (case first_entry (is_changed old_entries) new_document of
+        NONE =>
+          (case first_entry (is_changed new_entries) old_document of
+            NONE => ([], new_document)
+          | SOME (_, id, _) => (cancel_old id; ([], new_document)))
+      | SOME (prev, id, _) =>
+          let
+            val _ = cancel_old id;
+            val prev_state_id = the (#state (the_entry new_entries (the prev)));
+            val (_, updates) = fold_entries id (new_state name) new_document (prev_state_id, []);
+            val new_document' = new_document |> map_entries (fold (set_entry_state o #1) updates);
+          in (rev updates, new_document') end);
+
+    val _ = define_document new_id new_document';
+    val _ = report_updates new_id (map #1 updates);
+    val _ = List.app (fn (_, run) => run ()) updates;
+  in () end;
+
+end;
+
+
+
+(** concrete syntax **)
+
+local structure P = OuterParse structure V = ValueParse in
+
+val _ =
+  OuterSyntax.internal_command "Isar.define_command"
+    (P.string -- P.string >> (fn (id, text) =>
+      Toplevel.imperative (fn () =>
+        define_command id (OuterSyntax.prepare_command (Position.id id) text))));
+
+val _ =
+  OuterSyntax.internal_command "Isar.begin_document"
+    (P.string -- P.string >> (fn (id, path) =>
+      Toplevel.imperative (fn () => begin_document id (Path.explode path))));
+
+val _ =
+  OuterSyntax.internal_command "Isar.end_document"
+    (P.string >> (fn id => Toplevel.imperative (fn () => end_document id)));
+
+val _ =
+  OuterSyntax.internal_command "Isar.edit_document"
+    (P.string -- P.string -- V.list (P.string -- (P.string >> SOME) || P.string >> rpair NONE)
+      >> (fn ((id, new_id), edits) => Toplevel.imperative (fn () => edit_document id new_id edits)));
+
+end;
+
+end;
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Isar/isar_document.scala	Mon Jan 19 21:20:18 2009 +0100
@@ -0,0 +1,61 @@
+/*  Title:      Pure/Isar/isar_document.scala
+    Author:     Makarius
+
+Interactive Isar documents.
+*/
+
+package isabelle
+
+
+object IsarDocument
+{
+  /* unique identifiers */
+
+  type State_ID = String
+  type Command_ID = String
+  type Document_ID = String
+
+
+  /* commands */
+
+  def define_command(proc: IsabelleProcess, id: Command_ID, text: String) {
+    proc.output_sync("Isar.define_command " + IsabelleSyntax.encode_string(id) + " " +
+      IsabelleSyntax.encode_string(text))
+  }
+
+
+  /* documents */
+
+  def begin_document(proc: IsabelleProcess, id: Document_ID, path: String) {
+    proc.output_sync("Isar.begin_document " + IsabelleSyntax.encode_string(id) + " " +
+       IsabelleSyntax.encode_string(path))
+  }
+
+  def end_document(proc: IsabelleProcess, id: Document_ID) {
+    proc.output_sync("Isar.end_document " + IsabelleSyntax.encode_string(id))
+  }
+
+  def edit_document(proc: IsabelleProcess, old_id: Document_ID, new_id: Document_ID,
+      edits: List[(Command_ID, Option[Command_ID])])
+  {
+    def append_edit(edit: (Command_ID, Option[Command_ID]), result: StringBuilder)
+    {
+      edit match {
+        case (id, None) => IsabelleSyntax.append_string(id, result)
+        case (id, Some(id2)) =>
+          IsabelleSyntax.append_string(id, result)
+          result.append(" ")
+          IsabelleSyntax.append_string(id2, result)
+      }
+    }
+
+    val buf = new StringBuilder(40)
+    buf.append("Isar.edit_document ")
+    IsabelleSyntax.append_string(old_id, buf)
+    buf.append(" ")
+    IsabelleSyntax.append_string(new_id, buf)
+    buf.append(" ")
+    IsabelleSyntax.append_list(append_edit, edits, buf)
+    proc.output_sync(buf.toString)
+  }
+}
--- a/src/Pure/Isar/isar_syn.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Isar/isar_syn.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/isar_syn.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Isar/Pure outer syntax.
@@ -394,9 +393,7 @@
     (P.name -- Scan.optional (P.$$$ "=" |-- P.!!! locale_val) (([], []), []) -- P.opt_begin
       >> (fn ((name, (expr, elems)), begin) =>
           (begin ? Toplevel.print) o Toplevel.begin_local_theory begin
-            (Expression.add_locale_cmd name name expr elems #>
-              (fn ((target, notes), ctxt) => TheoryTarget.begin target ctxt |>
-                fold (fn (kind, facts) => LocalTheory.notes kind facts #> snd) notes))));
+            (Expression.add_locale_cmd name "" expr elems #> snd)));
 
 val _ =
   OuterSyntax.command "sublocale"
@@ -430,24 +427,24 @@
 val locale_val =
   SpecParse.locale_expr --
     Scan.optional (P.$$$ "+" |-- P.!!! (Scan.repeat1 SpecParse.context_element)) [] ||
-  Scan.repeat1 SpecParse.context_element >> pair Locale.empty;
+  Scan.repeat1 SpecParse.context_element >> pair Old_Locale.empty;
 
 val _ =
   OuterSyntax.command "class_locale" "define named proof context based on classes" K.thy_decl
-    (P.name -- Scan.optional (P.$$$ "=" |-- P.!!! locale_val) (Locale.empty, []) -- P.opt_begin
+    (P.name -- Scan.optional (P.$$$ "=" |-- P.!!! locale_val) (Old_Locale.empty, []) -- P.opt_begin
       >> (fn ((name, (expr, elems)), begin) =>
           (begin ? Toplevel.print) o Toplevel.begin_local_theory begin
-            (Locale.add_locale_cmd name expr elems #-> TheoryTarget.begin)));
+            (Old_Locale.add_locale_cmd name expr elems #-> TheoryTarget.begin)));
 
 val _ =
   OuterSyntax.command "class_interpretation"
     "prove and register interpretation of locale expression in theory or locale" K.thy_goal
     (P.xname --| (P.$$$ "\\<subseteq>" || P.$$$ "<") -- P.!!! SpecParse.locale_expr
-      >> (Toplevel.print oo (Toplevel.theory_to_proof o Locale.interpretation_in_locale I)) ||
+      >> (Toplevel.print oo (Toplevel.theory_to_proof o Old_Locale.interpretation_in_locale I)) ||
       opt_prefix  -- SpecParse.locale_expr -- SpecParse.locale_insts
         >> (fn ((name, expr), insts) => Toplevel.print o
             Toplevel.theory_to_proof
-              (Locale.interpretation_cmd (Binding.base_name name) expr insts)));
+              (Old_Locale.interpretation_cmd (Binding.base_name name) expr insts)));
 
 val _ =
   OuterSyntax.command "class_interpret"
@@ -456,7 +453,7 @@
     (opt_prefix -- SpecParse.locale_expr -- SpecParse.locale_insts
       >> (fn ((name, expr), insts) => Toplevel.print o
           Toplevel.proof'
-            (fn int => Locale.interpret_cmd (Binding.base_name name) expr insts int)));
+            (fn int => Old_Locale.interpret_cmd (Binding.base_name name) expr insts int)));
 
 end;
 
@@ -473,17 +470,17 @@
    (P.name -- (P.$$$ "=" |-- class_val) -- P.opt_begin
     >> (fn ((name, (supclasses, elems)), begin) =>
         (begin ? Toplevel.print) o Toplevel.begin_local_theory begin
-          (Class.class_cmd name supclasses elems #-> TheoryTarget.begin)));
+          (Class.class_cmd name supclasses elems #> snd)));
 
 val _ =
   OuterSyntax.local_theory_to_proof "subclass" "prove a subclass relation" K.thy_goal
-    (P.xname >> Subclass.subclass_cmd);
+    (P.xname >> Class.subclass_cmd);
 
 val _ =
   OuterSyntax.command "instantiation" "instantiate and prove type arity" K.thy_decl
    (P.multi_arity --| P.begin
      >> (fn arities => Toplevel.print o
-         Toplevel.begin_local_theory true (Instance.instantiation_cmd arities)));
+         Toplevel.begin_local_theory true (TheoryTarget.instantiation_cmd arities)));
 
 val _ =
   OuterSyntax.command "instance" "prove type arity or subclass relation" K.thy_goal
@@ -757,9 +754,13 @@
 
 (* nested commands *)
 
+val props_text =
+  Scan.optional ValueParse.properties [] -- P.position P.string >> (fn (props, (str, pos)) =>
+    (Position.of_properties (Position.default_properties pos props), str));
+
 val _ =
   OuterSyntax.improper_command "Isabelle.command" "nested Isabelle command" K.control
-    (P.props_text :|-- (fn (pos, str) =>
+    (props_text :|-- (fn (pos, str) =>
       (case OuterSyntax.parse pos str of
         [tr] => Scan.succeed (K tr)
       | _ => Scan.fail_with (K "exactly one command expected"))
@@ -770,7 +771,7 @@
 
 val _ =
   OuterSyntax.improper_command "init_toplevel" "init toplevel point-of-interest" K.control
-    (Scan.succeed (Toplevel.no_timing o Toplevel.imperative Isar.init_point));
+    (Scan.succeed (Toplevel.no_timing o Toplevel.imperative Isar.init));
 
 val _ =
   OuterSyntax.improper_command "linear_undo" "undo commands" K.control
--- a/src/Pure/Isar/local_defs.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Isar/local_defs.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/local_defs.ML
-    ID:         $Id$
     Author:     Makarius
 
 Local definitions.
--- a/src/Pure/Isar/locale.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Isar/locale.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,6 +1,5 @@
 (*  Title:      Pure/Isar/locale.ML
-    ID:         $Id$
-    Author:     Clemens Ballarin, TU Muenchen; Markus Wenzel, LMU/TU Muenchen
+    Author:     Clemens Ballarin, TU Muenchen
 
 Locales -- Isar proof contexts as meta-level predicates, with local
 syntax and implicit structures.
@@ -28,1708 +27,438 @@
     pages 31-43, 2006.
 *)
 
-(* TODO:
-- beta-eta normalisation of interpretation parameters
-- dangling type frees in locales
-- test subsumption of interpretations when merging theories
-*)
-
 signature LOCALE =
 sig
-  datatype expr =
-    Locale of string |
-    Rename of expr * (string * mixfix option) option list |
-    Merge of expr list
-  val empty: expr
-
-  val intern: theory -> xstring -> string
-  val intern_expr: theory -> expr -> expr
-  val extern: theory -> string -> xstring
-  val init: string -> theory -> Proof.context
+  type locale
 
-  (* The specification of a locale *)
-  val parameters_of: theory -> string -> ((string * typ) * mixfix) list
-  val parameters_of_expr: theory -> expr -> ((string * typ) * mixfix) list
-  val local_asms_of: theory -> string -> (Attrib.binding * term list) list
-  val global_asms_of: theory -> string -> (Attrib.binding * term list) list
-
-  (* Theorems *)
-  val intros: theory -> string -> thm list * thm list
-  val dests: theory -> string -> thm list
-  (* Not part of the official interface.  DO NOT USE *)
-  val facts_of: theory -> string -> (Attrib.binding * (thm list * Attrib.src list) list) list list
-
-  (* Not part of the official interface.  DO NOT USE *)
-  val declarations_of: theory -> string -> declaration list * declaration list;
+  val register_locale: bstring ->
+    (string * sort) list * (Binding.T * typ option * mixfix) list ->
+    term option * term list ->
+    thm option * thm option -> thm list ->
+    (declaration * stamp) list * (declaration * stamp) list ->
+    ((string * (Attrib.binding * (thm list * Attrib.src list) list) list) * stamp) list ->
+    ((string * morphism) * stamp) list -> theory -> theory
 
-  (* Processing of locale statements *)
-  val read_context_statement: string option -> Element.context list ->
-    (string * string list) list list -> Proof.context ->
-    string option * Proof.context * Proof.context * (term * term list) list list
-  val read_context_statement_cmd: xstring option -> Element.context list ->
-    (string * string list) list list -> Proof.context ->
-    string option * Proof.context * Proof.context * (term * term list) list list
-  val cert_context_statement: string option -> Element.context_i list ->
-    (term * term list) list list -> Proof.context ->
-    string option * Proof.context * Proof.context * (term * term list) list list
-  val read_expr: expr -> Element.context list -> Proof.context ->
-    Element.context_i list * Proof.context
-  val cert_expr: expr -> Element.context_i list -> Proof.context ->
-    Element.context_i list * Proof.context
+  (* Locale name space *)
+  val intern: theory -> xstring -> string
+  val extern: theory -> string -> xstring
 
-  (* Diagnostic functions *)
-  val print_locales: theory -> unit
-  val print_locale: theory -> bool -> expr -> Element.context list -> unit
-  val print_registrations: bool -> string -> Proof.context -> unit
-
-  val add_locale: string -> bstring -> expr -> Element.context_i list -> theory
-    -> string * Proof.context
-  val add_locale_cmd: bstring -> expr -> Element.context list -> theory
-    -> string * Proof.context
-
-  (* Tactics *)
-  val intro_locales_tac: bool -> Proof.context -> thm list -> tactic
+  (* Specification *)
+  val defined: theory -> string -> bool
+  val params_of: theory -> string -> (Binding.T * typ option * mixfix) list
+  val intros_of: theory -> string -> thm option * thm option
+  val axioms_of: theory -> string -> thm list
+  val instance_of: theory -> string -> morphism -> term list
+  val specification_of: theory -> string -> term option * term list
+  val declarations_of: theory -> string -> declaration list * declaration list
+  val dependencies_of: theory -> string -> (string * morphism) list
 
   (* Storing results *)
-  val global_note_qualified: string ->
-    ((Binding.T * attribute list) * (thm list * attribute list) list) list ->
-    theory -> (string * thm list) list * theory
-  val local_note_qualified: string ->
-    ((Binding.T * attribute list) * (thm list * attribute list) list) list ->
-    Proof.context -> (string * thm list) list * Proof.context
   val add_thmss: string -> string -> (Attrib.binding * (thm list * Attrib.src list) list) list ->
     Proof.context -> Proof.context
   val add_type_syntax: string -> declaration -> Proof.context -> Proof.context
   val add_term_syntax: string -> declaration -> Proof.context -> Proof.context
   val add_declaration: string -> declaration -> Proof.context -> Proof.context
+  val add_dependency: string -> string * morphism -> theory -> theory
 
-  (* Interpretation *)
-  val get_interpret_morph: theory -> (Binding.T -> Binding.T) -> string * string ->
-    (Morphism.morphism * ((typ Vartab.table * typ list) * (term Vartab.table * term list))) ->
-    string -> term list -> Morphism.morphism
-  val interpretation: (Proof.context -> Proof.context) ->
-    (Binding.T -> Binding.T) -> expr ->
-    term option list * (Attrib.binding * term) list ->
-    theory ->
-    (Morphism.morphism * ((typ Vartab.table * typ list) * (term Vartab.table * term list))) * Proof.state
-  val interpretation_cmd: string -> expr -> string option list * (Attrib.binding * string) list ->
-    theory -> Proof.state
-  val interpretation_in_locale: (Proof.context -> Proof.context) ->
-    xstring * expr -> theory -> Proof.state
-  val interpret: (Proof.state -> Proof.state Seq.seq) ->
-    (Binding.T -> Binding.T) -> expr ->
-    term option list * (Attrib.binding * term) list ->
-    bool -> Proof.state ->
-    (Morphism.morphism * ((typ Vartab.table * typ list) * (term Vartab.table * term list))) * Proof.state
-  val interpret_cmd: string -> expr -> string option list * (Attrib.binding * string) list ->
-    bool -> Proof.state -> Proof.state
+  (* Activation *)
+  val activate_declarations: theory -> string * morphism ->
+    Proof.context -> Proof.context
+  val activate_global_facts: string * morphism -> theory -> theory
+  val activate_local_facts: string * morphism -> Proof.context -> Proof.context
+  val init: string -> theory -> Proof.context
+
+  (* Reasoning about locales *)
+  val witness_attrib: attribute
+  val intro_attrib: attribute
+  val unfold_attrib: attribute
+  val intro_locales_tac: bool -> Proof.context -> thm list -> tactic
+
+  (* Registrations *)
+  val add_registration: string * (morphism * morphism) ->
+    theory -> theory
+  val amend_registration: morphism -> string * morphism ->
+    theory -> theory
+  val get_registrations: theory -> (string * morphism) list
+
+  (* Diagnostic *)
+  val print_locales: theory -> unit
+  val print_locale: theory -> bool -> bstring -> unit
 end;
 
+
+(*** Theorem list extensible via attribute --- to control intro_locales_tac ***)
+
+(* FIXME: consider moving named_thms.ML up in the build hierarchy and using that *)
+functor ThmsFun() =
+struct
+
+structure Data = GenericDataFun
+(
+  type T = thm list;
+  val empty = [];
+  val extend = I;
+  fun merge _ = Thm.merge_thms;
+);
+
+val get = Data.get o Context.Proof;
+val add = Thm.declaration_attribute (Data.map o Thm.add_thm);
+
+end;
+
+
 structure Locale: LOCALE =
 struct
 
-(* legacy operations *)
-
-fun merge_lists _ xs [] = xs
-  | merge_lists _ [] ys = ys
-  | merge_lists eq xs ys = xs @ filter_out (member eq xs) ys;
-
-fun merge_alists eq xs = merge_lists (eq_fst eq) xs;
-
-
-(* auxiliary: noting name bindings with qualified base names *)
-
-fun global_note_qualified kind facts thy =
-  thy
-  |> Sign.qualified_names
-  |> PureThy.note_thmss kind facts
-  ||> Sign.restore_naming thy;
-
-fun local_note_qualified kind facts ctxt =
-  ctxt
-  |> ProofContext.qualified_names
-  |> ProofContext.note_thmss_i kind facts
-  ||> ProofContext.restore_naming ctxt;
-
-
-(** locale elements and expressions **)
-
 datatype ctxt = datatype Element.ctxt;
 
-datatype expr =
-  Locale of string |
-  Rename of expr * (string * mixfix option) option list |
-  Merge of expr list;
-
-val empty = Merge [];
-
-datatype 'a element =
-  Elem of 'a | Expr of expr;
-
-fun map_elem f (Elem e) = Elem (f e)
-  | map_elem _ (Expr e) = Expr e;
-
-type decl = declaration * stamp;
-
-type locale =
- {axiom: Element.witness list,
-    (* For locales that define predicates this is [A [A]], where A is the locale
-       specification.  Otherwise [].
-       Only required to generate the right witnesses for locales with predicates. *)
-  elems: (Element.context_i * stamp) list,
-    (* Static content, neither Fixes nor Constrains elements *)
-  params: ((string * typ) * mixfix) list,                        (*all term params*)
-  decls: decl list * decl list,                    (*type/term_syntax declarations*)
-  regs: ((string * string list) * Element.witness list) list,
-    (* Registrations: indentifiers and witnesses of locales interpreted in the locale. *)
-  intros: thm list * thm list,
-    (* Introduction rules: of delta predicate and locale predicate. *)
-  dests: thm list}
-    (* Destruction rules: projections from locale predicate to predicates of fragments. *)
-
-(* CB: an internal (Int) locale element was either imported or included,
-   an external (Ext) element appears directly in the locale text. *)
-
-datatype ('a, 'b) int_ext = Int of 'a | Ext of 'b;
-
 
 
-(** substitutions on Vars -- clone from element.ML **)
-
-(* instantiate types *)
-
-fun var_instT_type env =
-  if Vartab.is_empty env then I
-  else Term.map_type_tvar (fn (x, S) => the_default (TVar (x, S)) (Vartab.lookup env x));
-
-fun var_instT_term env =
-  if Vartab.is_empty env then I
-  else Term.map_types (var_instT_type env);
-
-fun var_inst_term (envT, env) =
-  if Vartab.is_empty env then var_instT_term envT
-  else
-    let
-      val instT = var_instT_type envT;
-      fun inst (Const (x, T)) = Const (x, instT T)
-        | inst (Free (x, T)) = Free(x, instT T)
-        | inst (Var (xi, T)) =
-            (case Vartab.lookup env xi of
-              NONE => Var (xi, instT T)
-            | SOME t => t)
-        | inst (b as Bound _) = b
-        | inst (Abs (x, T, t)) = Abs (x, instT T, inst t)
-        | inst (t $ u) = inst t $ inst u;
-    in Envir.beta_norm o inst end;
-
-
-(** management of registrations in theories and proof contexts **)
-
-type registration =
-  {prfx: (Binding.T -> Binding.T) * (string * string),
-      (* first component: interpretation name morphism;
-         second component: parameter prefix *)
-    exp: Morphism.morphism,
-      (* maps content to its originating context *)
-    imp: (typ Vartab.table * typ list) * (term Vartab.table * term list),
-      (* inverse of exp *)
-    wits: Element.witness list,
-      (* witnesses of the registration *)
-    eqns: thm Termtab.table,
-      (* theorems (equations) interpreting derived concepts and indexed by lhs *)
-    morph: unit
-      (* interpreting morphism *)
-  }
-
-structure Registrations :
-  sig
-    type T
-    val empty: T
-    val join: T * T -> T
-    val dest: theory -> T ->
-      (term list *
-        (((Binding.T -> Binding.T) * (string * string)) *
-         (Morphism.morphism * ((typ Vartab.table * typ list) * (term Vartab.table * term list))) *
-         Element.witness list *
-         thm Termtab.table)) list
-    val test: theory -> T * term list -> bool
-    val lookup: theory ->
-      T * (term list * ((typ Vartab.table * typ list) * (term Vartab.table * term list))) ->
-      (((Binding.T -> Binding.T) * (string * string)) * Element.witness list * thm Termtab.table) option
-    val insert: theory -> term list -> ((Binding.T -> Binding.T) * (string * string)) ->
-      (Morphism.morphism * ((typ Vartab.table * typ list) * (term Vartab.table * term list))) ->
-      T ->
-      T * (term list * (((Binding.T -> Binding.T) * (string * string)) * Element.witness list)) list
-    val add_witness: term list -> Element.witness -> T -> T
-    val add_equation: term list -> thm -> T -> T
-(*
-    val update_morph: term list -> Morphism.morphism -> T -> T
-    val get_morph: theory -> T ->
-      term list * ((typ Vartab.table * typ list) * (term Vartab.table * term list)) ->
-      Morphism.morphism
-*)
-  end =
-struct
-  (* A registration is indexed by parameter instantiation.
-     NB: index is exported whereas content is internalised. *)
-  type T = registration Termtab.table;
-
-  fun mk_reg prfx exp imp wits eqns morph =
-    {prfx = prfx, exp = exp, imp = imp, wits = wits, eqns = eqns, morph = morph};
-
-  fun map_reg f reg =
-    let
-      val {prfx, exp, imp, wits, eqns, morph} = reg;
-      val (prfx', exp', imp', wits', eqns', morph') = f (prfx, exp, imp, wits, eqns, morph);
-    in mk_reg prfx' exp' imp' wits' eqns' morph' end;
-
-  val empty = Termtab.empty;
+(*** Theory data ***)
 
-  (* term list represented as single term, for simultaneous matching *)
-  fun termify ts =
-    Term.list_comb (Const ("", map fastype_of ts ---> propT), ts);
-  fun untermify t =
-    let fun ut (Const _) ts = ts
-          | ut (s $ t) ts = ut s (t::ts)
-    in ut t [] end;
-
-  (* joining of registrations:
-     - prefix and morphisms of right theory;
-     - witnesses are equal, no attempt to subsumption testing;
-     - union of equalities, if conflicting (i.e. two eqns with equal lhs)
-       eqn of right theory takes precedence *)
-  fun join (r1, r2) = Termtab.join (fn _ => fn ({eqns = e1, ...}, {prfx = n, exp, imp, wits = w, eqns = e2, morph = m}) =>
-      mk_reg n exp imp w (Termtab.join (fn _ => fn (_, e) => e) (e1, e2)) m) (r1, r2);
-
-  fun dest_transfer thy regs =
-    Termtab.dest regs |> map (apsnd (map_reg (fn (n, e, i, ws, es, m) =>
-      (n, e, i, map (Element.transfer_witness thy) ws, Termtab.map (transfer thy) es, m))));
-
-  fun dest thy regs = dest_transfer thy regs |> map (apfst untermify) |>
-    map (apsnd (fn {prfx, exp, imp, wits, eqns, ...} => (prfx, (exp, imp), wits, eqns)));
-
-  (* registrations that subsume t *)
-  fun subsumers thy t regs =
-    filter (fn (t', _) => Pattern.matches thy (t', t)) (dest_transfer thy regs);
+datatype locale = Loc of {
+  (** static part **)
+  parameters: (string * sort) list * (Binding.T * typ option * mixfix) list,
+    (* type and term parameters *)
+  spec: term option * term list,
+    (* assumptions (as a single predicate expression) and defines *)
+  intros: thm option * thm option,
+  axioms: thm list,
+  (** dynamic part **)
+  decls: (declaration * stamp) list * (declaration * stamp) list,
+    (* type and term syntax declarations *)
+  notes: ((string * (Attrib.binding * (thm list * Attrib.src list) list) list) * stamp) list,
+    (* theorem declarations *)
+  dependencies: ((string * morphism) * stamp) list
+    (* locale dependencies (sublocale relation) *)
+};
 
-  (* test if registration that subsumes the query is present *)
-  fun test thy (regs, ts) =
-    not (null (subsumers thy (termify ts) regs));
-      
-  (* look up registration, pick one that subsumes the query *)
-  fun lookup thy (regs, (ts, ((impT, _), (imp, _)))) =
-    let
-      val t = termify ts;
-      val subs = subsumers thy t regs;
-    in
-      (case subs of
-        [] => NONE
-        | ((t', {prfx, exp = exp', imp = ((impT', domT'), (imp', dom')), wits, eqns, morph}) :: _) =>
-          let
-            val (tinst, inst) = Pattern.match thy (t', t) (Vartab.empty, Vartab.empty);
-            val tinst' = domT' |> map (fn (T as TFree (x, _)) =>
-                (x, T |> Morphism.typ exp' |> Envir.typ_subst_TVars tinst
-                      |> var_instT_type impT)) |> Symtab.make;
-            val inst' = dom' |> map (fn (t as Free (x, _)) =>
-                (x, t |> Morphism.term exp' |> Envir.subst_vars (tinst, inst)
-                      |> var_inst_term (impT, imp))) |> Symtab.make;
-            val inst'_morph = Element.inst_morphism thy (tinst', inst');
-          in SOME (prfx,
-            map (Element.morph_witness inst'_morph) wits,
-            Termtab.map (Morphism.thm inst'_morph) eqns)
-          end)
-    end;
-
-  (* add registration if not subsumed by ones already present,
-     additionally returns registrations that are strictly subsumed *)
-  fun insert thy ts prfx (exp, imp) regs =
-    let
-      val t = termify ts;
-      val subs = subsumers thy t regs ;
-    in (case subs of
-        [] => let
-                val sups =
-                  filter (fn (t', _) => Pattern.matches thy (t, t')) (dest_transfer thy regs);
-                val sups' = map (apfst untermify) sups |> map (fn (ts, {prfx, wits, ...}) => (ts, (prfx, wits)))
-              in (Termtab.update (t, mk_reg prfx exp imp [] Termtab.empty ()) regs, sups') end
-      | _ => (regs, []))
-    end;
-
-  fun gen_add f ts regs =
-    let
-      val t = termify ts;
-    in
-      Termtab.update (t, map_reg f (the (Termtab.lookup regs t))) regs
-    end;
-
-  (* add witness theorem to registration,
-     only if instantiation is exact, otherwise exception Option raised *)
-  fun add_witness ts wit regs =
-    gen_add (fn (x, e, i, wits, eqns, m) => (x, e, i, Element.close_witness wit :: wits, eqns, m))
-      ts regs;
-
-  (* add equation to registration, replaces previous equation with same lhs;
-     only if instantiation is exact, otherwise exception Option raised;
-     exception TERM raised if not a meta equality *)
-  fun add_equation ts thm regs =
-    gen_add (fn (x, e, i, thms, eqns, m) =>
-      (x, e, i, thms, Termtab.update (thm |> prop_of |> Logic.dest_equals |> fst, Thm.close_derivation thm) eqns, m))
-      ts regs;
-
-end;
-
-
-(** theory data : locales **)
+fun mk_locale ((parameters, spec, intros, axioms), ((decls, notes), dependencies)) =
+  Loc {parameters = parameters, intros = intros, axioms = axioms, spec = spec,
+    decls = decls, notes = notes, dependencies = dependencies};
+fun map_locale f (Loc {parameters, spec, intros, axioms, decls, notes, dependencies}) =
+  mk_locale (f ((parameters, spec, intros, axioms), ((decls, notes), dependencies)));
+fun merge_locale (Loc {parameters, spec, intros, axioms, decls = (decls1, decls2),
+  notes, dependencies}, Loc {decls = (decls1', decls2'), notes = notes',
+    dependencies = dependencies', ...}) = mk_locale
+      ((parameters, spec, intros, axioms),
+        (((merge (eq_snd op =) (decls1, decls1'), merge (eq_snd op =) (decls2, decls2')),
+          merge (eq_snd op =) (notes, notes')),
+            merge (eq_snd op =) (dependencies, dependencies')));
 
 structure LocalesData = TheoryDataFun
 (
-  type T = NameSpace.T * locale Symtab.table;
-    (* 1st entry: locale namespace,
-       2nd entry: locales of the theory *)
-
+  type T = locale NameSpace.table;
   val empty = NameSpace.empty_table;
   val copy = I;
   val extend = I;
-
-  fun join_locales _
-    ({axiom, elems, params, decls = (decls1, decls2), regs, intros, dests}: locale,
-      {elems = elems', decls = (decls1', decls2'), regs = regs', ...}: locale) =
-     {axiom = axiom,
-      elems = merge_lists (eq_snd (op =)) elems elems',
-      params = params,
-      decls =
-       (Library.merge (eq_snd (op =)) (decls1, decls1'),
-        Library.merge (eq_snd (op =)) (decls2, decls2')),
-      regs = merge_alists (op =) regs regs',
-      intros = intros,
-      dests = dests};
-  fun merge _ = NameSpace.join_tables join_locales;
+  fun merge _ = NameSpace.join_tables (K merge_locale);
 );
 
-
-
-(** context data : registrations **)
-
-structure RegistrationsData = GenericDataFun
-(
-  type T = Registrations.T Symtab.table;  (*registrations, indexed by locale name*)
-  val empty = Symtab.empty;
-  val extend = I;
-  fun merge _ = Symtab.join (K Registrations.join);
-);
-
-
-(** access locales **)
-
 val intern = NameSpace.intern o #1 o LocalesData.get;
 val extern = NameSpace.extern o #1 o LocalesData.get;
 
-fun get_locale thy name = Symtab.lookup (#2 (LocalesData.get thy)) name;
+fun get_locale thy = Symtab.lookup (#2 (LocalesData.get thy));
+
+fun defined thy = is_some o get_locale thy;
 
 fun the_locale thy name = case get_locale thy name
- of SOME loc => loc
+ of SOME (Loc loc) => loc
   | NONE => error ("Unknown locale " ^ quote name);
 
-fun register_locale bname loc thy =
-  thy |> LocalesData.map (NameSpace.bind (Sign.naming_of thy)
-    (Binding.name bname, loc) #> snd);
+fun register_locale bname parameters spec intros axioms decls notes dependencies thy =
+  thy |> LocalesData.map (NameSpace.bind (Sign.naming_of thy) (Binding.name bname,
+    mk_locale ((parameters, spec, intros, axioms), ((decls, notes), dependencies))) #> snd);
 
-fun change_locale name f thy =
-  let
-    val {axiom, elems, params, decls, regs, intros, dests} =
-        the_locale thy name;
-    val (axiom', elems', params', decls', regs', intros', dests') =
-      f (axiom, elems, params, decls, regs, intros, dests);
-  in
-    thy
-    |> (LocalesData.map o apsnd) (Symtab.update (name, {axiom = axiom',
-          elems = elems', params = params',
-          decls = decls', regs = regs', intros = intros', dests = dests'}))
-  end;
+fun change_locale name =
+  LocalesData.map o apsnd o Symtab.map_entry name o map_locale o apsnd;
 
 fun print_locales thy =
-  let val (space, locs) = LocalesData.get thy in
-    Pretty.strs ("locales:" :: map #1 (NameSpace.extern_table (space, locs)))
-    |> Pretty.writeln
-  end;
-
-
-(* access registrations *)
-
-(* retrieve registration from theory or context *)
-
-fun get_registrations ctxt name =
-  case Symtab.lookup (RegistrationsData.get ctxt) name of
-      NONE => []
-    | SOME reg => Registrations.dest (Context.theory_of ctxt) reg;
-
-fun get_global_registrations thy = get_registrations (Context.Theory thy);
-fun get_local_registrations ctxt = get_registrations (Context.Proof ctxt);
-
-
-fun get_registration ctxt imprt (name, ps) =
-  case Symtab.lookup (RegistrationsData.get ctxt) name of
-      NONE => NONE
-    | SOME reg => Registrations.lookup (Context.theory_of ctxt) (reg, (ps, imprt));
-
-fun get_global_registration thy = get_registration (Context.Theory thy);
-fun get_local_registration ctxt = get_registration (Context.Proof ctxt);
-
-
-fun test_registration ctxt (name, ps) =
-  case Symtab.lookup (RegistrationsData.get ctxt) name of
-      NONE => false
-    | SOME reg => Registrations.test (Context.theory_of ctxt) (reg, ps);
-
-fun test_global_registration thy = test_registration (Context.Theory thy);
-fun test_local_registration ctxt = test_registration (Context.Proof ctxt);
-
-
-(* add registration to theory or context, ignored if subsumed *)
-
-fun put_registration (name, ps) prfx morphs ctxt =
-  RegistrationsData.map (fn regs =>
-    let
-      val thy = Context.theory_of ctxt;
-      val reg = the_default Registrations.empty (Symtab.lookup regs name);
-      val (reg', sups) = Registrations.insert thy ps prfx morphs reg;
-      val _ = if not (null sups) then warning
-                ("Subsumed interpretation(s) of locale " ^
-                 quote (extern thy name) ^
-                 "\nwith the following prefix(es):" ^
-                  commas_quote (map (fn (_, ((_, (_, s)), _)) => s) sups))
-              else ();
-    in Symtab.update (name, reg') regs end) ctxt;
-
-fun put_global_registration id prfx morphs =
-  Context.theory_map (put_registration id prfx morphs);
-fun put_local_registration id prfx morphs =
-  Context.proof_map (put_registration id prfx morphs);
-
-fun put_registration_in_locale name id =
-  change_locale name (fn (axiom, elems, params, decls, regs, intros, dests) =>
-    (axiom, elems, params, decls, regs @ [(id, [])], intros, dests));
-
-
-(* add witness theorem to registration, ignored if registration not present *)
-
-fun add_witness (name, ps) thm =
-  RegistrationsData.map (Symtab.map_entry name (Registrations.add_witness ps thm));
-
-fun add_global_witness id thm = Context.theory_map (add_witness id thm);
-fun add_local_witness id thm = Context.proof_map (add_witness id thm);
-
-
-fun add_witness_in_locale name id thm =
-  change_locale name (fn (axiom, elems, params, decls, regs, intros, dests) =>
-    let
-      fun add (id', thms) =
-        if id = id' then (id', thm :: thms) else (id', thms);
-    in (axiom, elems, params, decls, map add regs, intros, dests) end);
-
-
-(* add equation to registration, ignored if registration not present *)
-
-fun add_equation (name, ps) thm =
-  RegistrationsData.map (Symtab.map_entry name (Registrations.add_equation ps thm));
-
-fun add_global_equation id thm = Context.theory_map (add_equation id thm);
-fun add_local_equation id thm = Context.proof_map (add_equation id thm);
-
-(*
-(* update morphism of registration, ignored if registration not present *)
-
-fun update_morph (name, ps) morph =
-  RegistrationsData.map (Symtab.map_entry name (Registrations.update_morph ps morph));
-
-fun update_global_morph id morph = Context.theory_map (update_morph id morph);
-fun update_local_morph id morph = Context.proof_map (update_morph id morph);
-*)
+  Pretty.strs ("locales:" :: map #1 (NameSpace.extern_table (LocalesData.get thy)))
+  |> Pretty.writeln;
 
 
-(* printing of registrations *)
+(*** Primitive operations ***)
+
+fun params_of thy = snd o #parameters o the_locale thy;
 
-fun print_registrations show_wits loc ctxt =
-  let
-    val thy = ProofContext.theory_of ctxt;
-    val prt_term = Pretty.quote o Syntax.pretty_term ctxt;
-    fun prt_term' t = if !show_types
-          then Pretty.block [prt_term t, Pretty.brk 1, Pretty.str "::",
-            Pretty.brk 1, (Pretty.quote o Syntax.pretty_typ ctxt) (type_of t)]
-          else prt_term t;
-    val prt_thm = prt_term o prop_of;
-    fun prt_inst ts =
-        Pretty.enclose "(" ")" (Pretty.breaks (map prt_term' ts));
-    fun prt_prfx ((false, prfx), param_prfx) = [Pretty.str prfx, Pretty.brk 1, Pretty.str "(optional)", Pretty.brk 1, Pretty.str param_prfx]
-      | prt_prfx ((true, prfx), param_prfx) = [Pretty.str prfx, Pretty.brk 1, Pretty.str param_prfx];
-    fun prt_eqns [] = Pretty.str "no equations."
-      | prt_eqns eqns = Pretty.block (Pretty.str "equations:" :: Pretty.brk 1 ::
-          Pretty.breaks (map prt_thm eqns));
-    fun prt_core ts eqns =
-          [prt_inst ts, Pretty.fbrk, prt_eqns (Termtab.dest eqns |> map snd)];
-    fun prt_witns [] = Pretty.str "no witnesses."
-      | prt_witns witns = Pretty.block (Pretty.str "witnesses:" :: Pretty.brk 1 ::
-          Pretty.breaks (map (Element.pretty_witness ctxt) witns))
-    fun prt_reg (ts, (_, _, witns, eqns)) =
-        if show_wits
-          then Pretty.block (prt_core ts eqns @ [Pretty.fbrk, prt_witns witns])
-          else Pretty.block (prt_core ts eqns)
+fun intros_of thy = #intros o the_locale thy;
+
+fun axioms_of thy = #axioms o the_locale thy;
 
-    val loc_int = intern thy loc;
-    val regs = RegistrationsData.get (Context.Proof ctxt);
-    val loc_regs = Symtab.lookup regs loc_int;
-  in
-    (case loc_regs of
-        NONE => Pretty.str ("no interpretations")
-      | SOME r => let
-            val r' = Registrations.dest thy r;
-            val r'' = Library.sort_wrt (fn (_, ((_, (_, prfx)), _, _, _)) => prfx) r';
-          in Pretty.big_list ("interpretations:") (map prt_reg r'') end)
-    |> Pretty.writeln
-  end;
+fun instance_of thy name morph = params_of thy name |>
+  map ((fn (b, T, _) => Free (Binding.base_name b, the T)) #> Morphism.term morph);
 
+fun specification_of thy = #spec o the_locale thy;
 
-(* diagnostics *)
+fun declarations_of thy name = the_locale thy name |>
+  #decls |> apfst (map fst) |> apsnd (map fst);
 
-fun err_in_locale ctxt msg ids =
-  let
-    val thy = ProofContext.theory_of ctxt;
-    fun prt_id (name, parms) =
-      [Pretty.block (Pretty.breaks (map Pretty.str (extern thy name :: parms)))];
-    val prt_ids = flat (separate [Pretty.str " +", Pretty.brk 1] (map prt_id ids));
-    val err_msg =
-      if forall (fn (s, _) => s = "") ids then msg
-      else msg ^ "\n" ^ Pretty.string_of (Pretty.block
-        (Pretty.str "The error(s) above occurred in locale:" :: Pretty.brk 1 :: prt_ids));
-  in error err_msg end;
-
-fun err_in_locale' ctxt msg ids' = err_in_locale ctxt msg (map fst ids');
+fun dependencies_of thy name = the_locale thy name |>
+  #dependencies |> map fst;
 
 
-fun pretty_ren NONE = Pretty.str "_"
-  | pretty_ren (SOME (x, NONE)) = Pretty.str x
-  | pretty_ren (SOME (x, SOME syn)) =
-      Pretty.block [Pretty.str x, Pretty.brk 1, Syntax.pretty_mixfix syn];
-
-fun pretty_expr thy (Locale name) = Pretty.str (extern thy name)
-  | pretty_expr thy (Rename (expr, xs)) =
-      Pretty.block [pretty_expr thy expr, Pretty.brk 1, Pretty.block (map pretty_ren xs |> Pretty.breaks)]
-  | pretty_expr thy (Merge es) =
-      Pretty.separate "+" (map (pretty_expr thy) es) |> Pretty.block;
-
-fun err_in_expr _ msg (Merge []) = error msg
-  | err_in_expr ctxt msg expr =
-    error (msg ^ "\n" ^ Pretty.string_of (Pretty.block
-      [Pretty.str "The error(s) above occured in locale expression:", Pretty.brk 1,
-       pretty_expr (ProofContext.theory_of ctxt) expr]));
+(*** Activate context elements of locale ***)
 
-
-(** structured contexts: rename + merge + implicit type instantiation **)
-
-(* parameter types *)
-
-fun frozen_tvars ctxt Ts =
-  #1 (Variable.importT_inst (map Logic.mk_type Ts) ctxt)
-  |> map (fn ((xi, S), T) => (xi, (S, T)));
-
-fun unify_frozen ctxt maxidx Ts Us =
-  let
-    fun paramify NONE i = (NONE, i)
-      | paramify (SOME T) i = apfst SOME (TypeInfer.paramify_dummies T i);
-
-    val (Ts', maxidx') = fold_map paramify Ts maxidx;
-    val (Us', maxidx'') = fold_map paramify Us maxidx';
-    val thy = ProofContext.theory_of ctxt;
+(** Identifiers: activated locales in theory or proof context **)
 
-    fun unify (SOME T, SOME U) env = (Sign.typ_unify thy (U, T) env
-          handle Type.TUNIFY => raise TYPE ("unify_frozen: failed to unify types", [U, T], []))
-      | unify _ env = env;
-    val (unifier, _) = fold unify (Ts' ~~ Us') (Vartab.empty, maxidx'');
-    val Vs = map (Option.map (Envir.norm_type unifier)) Us';
-    val unifier' = fold Vartab.update_new (frozen_tvars ctxt (map_filter I Vs)) unifier;
-  in map (Option.map (Envir.norm_type unifier')) Vs end;
-
-fun params_of elemss =
-  distinct (eq_fst (op = : string * string -> bool)) (maps (snd o fst) elemss);
-
-fun params_of' elemss =
-  distinct (eq_fst (op = : string * string -> bool)) (maps (snd o fst o fst) elemss);
-
-fun param_prefix locale_name params = (NameSpace.base locale_name ^ "_locale", space_implode "_" params);
-
-
-(* CB: param_types has the following type:
-  ('a * 'b option) list -> ('a * 'b) list *)
-fun param_types ps = map_filter (fn (_, NONE) => NONE | (x, SOME T) => SOME (x, T)) ps;
+type identifiers = (string * term list) list;
 
-
-fun merge_syntax ctxt ids ss = Symtab.merge (op = : mixfix * mixfix -> bool) ss
-  handle Symtab.DUP x => err_in_locale ctxt
-    ("Conflicting syntax for parameter: " ^ quote x) (map fst ids);
-
+val empty = ([] : identifiers);
 
-(* Distinction of assumed vs. derived identifiers.
-   The former may have axioms relating assumptions of the context to
-   assumptions of the specification fragment (for locales with
-   predicates).  The latter have witnesses relating assumptions of the
-   specification fragment to assumptions of other (assumed) specification
-   fragments. *)
-
-datatype 'a mode = Assumed of 'a | Derived of 'a;
-
-fun map_mode f (Assumed x) = Assumed (f x)
-  | map_mode f (Derived x) = Derived (f x);
-
-
-(* flatten expressions *)
+fun ident_eq thy ((n: string, ts), (m, ss)) = (m = n) andalso Pattern.matchess thy (ss, ts);
 
 local
 
-fun unify_parms ctxt fixed_parms raw_parmss =
-  let
-    val thy = ProofContext.theory_of ctxt;
-    val maxidx = length raw_parmss;
-    val idx_parmss = (0 upto maxidx - 1) ~~ raw_parmss;
-
-    fun varify i = Term.map_type_tfree (fn (a, S) => TVar ((a, i), S));
-    fun varify_parms (i, ps) = map (apsnd (varify i)) (param_types ps);
-    val parms = fixed_parms @ maps varify_parms idx_parmss;
+datatype 'a delayed = Ready of 'a | ToDo of ('a delayed * 'a delayed);
 
-    fun unify T U envir = Sign.typ_unify thy (U, T) envir
-      handle Type.TUNIFY =>
-        let
-          val T' = Envir.norm_type (fst envir) T;
-          val U' = Envir.norm_type (fst envir) U;
-          val prt = Syntax.string_of_typ ctxt;
-        in
-          raise TYPE ("unify_parms: failed to unify types " ^
-            prt U' ^ " and " ^ prt T', [U', T'], [])
-        end;
-    fun unify_list (T :: Us) = fold (unify T) Us
-      | unify_list [] = I;
-    val (unifier, _) = fold unify_list (map #2 (Symtab.dest (Symtab.make_list parms)))
-      (Vartab.empty, maxidx);
-
-    val parms' = map (apsnd (Envir.norm_type unifier)) (distinct (eq_fst (op =)) parms);
-    val unifier' = fold Vartab.update_new (frozen_tvars ctxt (map #2 parms')) unifier;
-
-    fun inst_parms (i, ps) =
-      List.foldr Term.add_typ_tfrees [] (map_filter snd ps)
-      |> map_filter (fn (a, S) =>
-          let val T = Envir.norm_type unifier' (TVar ((a, i), S))
-          in if T = TFree (a, S) then NONE else SOME (a, T) end)
-      |> Symtab.make;
-  in map inst_parms idx_parmss end;
+structure IdentifiersData = GenericDataFun
+(
+  type T = identifiers delayed;
+  val empty = Ready empty;
+  val extend = I;
+  fun merge _ = ToDo;
+);
 
 in
 
-fun unify_elemss _ _ [] = []
-  | unify_elemss _ [] [elems] = [elems]
-  | unify_elemss ctxt fixed_parms elemss =
-      let
-        val thy = ProofContext.theory_of ctxt;
-        val phis = unify_parms ctxt fixed_parms (map (snd o fst o fst) elemss)
-          |> map (Element.instT_morphism thy);
-        fun inst ((((name, ps), mode), elems), phi) =
-          (((name, map (apsnd (Option.map (Morphism.typ phi))) ps),
-              map_mode (map (Element.morph_witness phi)) mode),
-            map (Element.morph_ctxt phi) elems);
-      in map inst (elemss ~~ phis) end;
-
-
-fun renaming xs parms = zip_options parms xs
-  handle Library.UnequalLengths =>
-    error ("Too many arguments in renaming: " ^
-      commas (map (fn NONE => "_" | SOME x => quote (fst x)) xs));
-
-
-(* params_of_expr:
-   Compute parameters (with types and syntax) of locale expression.
-*)
-
-fun params_of_expr ctxt fixed_params expr (prev_parms, prev_types, prev_syn) =
-  let
-    val thy = ProofContext.theory_of ctxt;
-
-    fun merge_tenvs fixed tenv1 tenv2 =
-        let
-          val [env1, env2] = unify_parms ctxt fixed
-                [tenv1 |> Symtab.dest |> map (apsnd SOME),
-                 tenv2 |> Symtab.dest |> map (apsnd SOME)]
-        in
-          Symtab.merge (op =) (Symtab.map (Element.instT_type env1) tenv1,
-            Symtab.map (Element.instT_type env2) tenv2)
-        end;
-
-    fun merge_syn expr syn1 syn2 =
-        Symtab.merge (op = : mixfix * mixfix -> bool) (syn1, syn2)
-        handle Symtab.DUP x => err_in_expr ctxt
-          ("Conflicting syntax for parameter: " ^ quote x) expr;
-
-    fun params_of (expr as Locale name) =
-          let
-            val {params, ...} = the_locale thy name;
-          in (map (fst o fst) params, params |> map fst |> Symtab.make,
-               params |> map (apfst fst) |> Symtab.make) end
-      | params_of (expr as Rename (e, xs)) =
-          let
-            val (parms', types', syn') = params_of e;
-            val ren = renaming xs parms';
-            (* renaming may reduce number of parameters *)
-            val new_parms = map (Element.rename ren) parms' |> distinct (op =);
-            val ren_syn = syn' |> Symtab.dest |> map (Element.rename_var_name ren);
-            val new_syn = fold (Symtab.insert (op =)) ren_syn Symtab.empty
-                handle Symtab.DUP x =>
-                  err_in_expr ctxt ("Conflicting syntax for parameter: " ^ quote x) expr;
-            val syn_types = map (apsnd (fn mx =>
-                SOME (Type.freeze_type (#1 (TypeInfer.paramify_dummies (Syntax.mixfixT mx) 0)))))
-              (Symtab.dest new_syn);
-            val ren_types = types' |> Symtab.dest |> map (apfst (Element.rename ren));
-            val (env :: _) = unify_parms ctxt []
-                ((ren_types |> map (apsnd SOME)) :: map single syn_types);
-            val new_types = fold (Symtab.insert (op =))
-                (map (apsnd (Element.instT_type env)) ren_types) Symtab.empty;
-          in (new_parms, new_types, new_syn) end
-      | params_of (Merge es) =
-          fold (fn e => fn (parms, types, syn) =>
-                   let
-                     val (parms', types', syn') = params_of e
-                   in
-                     (merge_lists (op =) parms parms', merge_tenvs [] types types',
-                      merge_syn e syn syn')
-                   end)
-            es ([], Symtab.empty, Symtab.empty)
-
-      val (parms, types, syn) = params_of expr;
-    in
-      (merge_lists (op =) prev_parms parms, merge_tenvs fixed_params prev_types types,
-       merge_syn expr prev_syn syn)
-    end;
-
-fun make_params_ids params = [(("", params), ([], Assumed []))];
-fun make_raw_params_elemss (params, tenv, syn) =
-    [((("", map (fn p => (p, Symtab.lookup tenv p)) params), Assumed []),
-      Int [Fixes (map (fn p =>
-        (Binding.name p, Symtab.lookup tenv p, Symtab.lookup syn p |> the)) params)])];
-
-
-(* flatten_expr:
-   Extend list of identifiers by those new in locale expression expr.
-   Compute corresponding list of lists of locale elements (one entry per
-   identifier).
-
-   Identifiers represent locale fragments and are in an extended form:
-     ((name, ps), (ax_ps, axs))
-   (name, ps) is the locale name with all its parameters.
-   (ax_ps, axs) is the locale axioms with its parameters;
-     axs are always taken from the top level of the locale hierarchy,
-     hence axioms may contain additional parameters from later fragments:
-     ps subset of ax_ps.  axs is either singleton or empty.
-
-   Elements are enriched by identifier-like information:
-     (((name, ax_ps), axs), elems)
-   The parameters in ax_ps are the axiom parameters, but enriched by type
-   info: now each entry is a pair of string and typ option.  Axioms are
-   type-instantiated.
-
-*)
-
-fun flatten_expr ctxt ((prev_idents, prev_syntax), expr) =
-  let
-    val thy = ProofContext.theory_of ctxt;
-
-    fun rename_parms top ren ((name, ps), (parms, mode)) =
-        ((name, map (Element.rename ren) ps),
-         if top
-         then (map (Element.rename ren) parms,
-               map_mode (map (Element.morph_witness (Element.rename_morphism ren))) mode)
-         else (parms, mode));
+fun finish thy (ToDo (i1, i2)) = merge (ident_eq thy) (finish thy i1, finish thy i2)
+  | finish _ (Ready ids) = ids;
 
-    (* add (name, pTs) and its registrations, recursively; adjust hyps of witnesses *)
-
-    fun add_with_regs ((name, pTs), mode) (wits, ids, visited) =
-        if member (fn (a, (b, _)) => a = b) visited (name, map #1 pTs)
-        then (wits, ids, visited)
-        else
-          let
-            val {params, regs, ...} = the_locale thy name;
-            val pTs' = map #1 params;
-            val ren = map #1 pTs' ~~ map (fn (x, _) => (x, NONE)) pTs;
-              (* dummy syntax, since required by rename *)
-            val pTs'' = map (fn ((p, _), (_, T)) => (p, T)) (pTs ~~ pTs');
-            val [env] = unify_parms ctxt pTs [map (apsnd SOME) pTs''];
-              (* propagate parameter types, to keep them consistent *)
-            val regs' = map (fn ((name, ps), wits) =>
-                ((name, map (Element.rename ren) ps),
-                 map (Element.transfer_witness thy) wits)) regs;
-            val new_regs = regs';
-            val new_ids = map fst new_regs;
-            val new_idTs =
-              map (apsnd (map (fn p => (p, (the o AList.lookup (op =) pTs) p)))) new_ids;
-
-            val new_wits = new_regs |> map (#2 #> map
-              (Element.morph_witness
-                (Element.instT_morphism thy env $>
-                  Element.rename_morphism ren $>
-                  Element.satisfy_morphism wits)
-                #> Element.close_witness));
-            val new_ids' = map (fn (id, wits) =>
-                (id, ([], Derived wits))) (new_ids ~~ new_wits);
-            val new_idTs' = map (fn ((n, pTs), (_, ([], mode))) =>
-                ((n, pTs), mode)) (new_idTs ~~ new_ids');
-            val new_id = ((name, map #1 pTs), ([], mode));
-            val (wits', ids', visited') = fold add_with_regs new_idTs'
-              (wits @ flat new_wits, ids, visited @ [new_id]);
-          in
-            (wits', ids' @ [new_id], visited')
-          end;
-
-    (* distribute top-level axioms over assumed ids *)
-
-    fun axiomify all_ps ((name, parms), (_, Assumed _)) axioms =
-        let
-          val {elems, ...} = the_locale thy name;
-          val ts = maps
-            (fn (Assumes asms, _) => maps (map #1 o #2) asms
-              | _ => [])
-            elems;
-          val (axs1, axs2) = chop (length ts) axioms;
-        in (((name, parms), (all_ps, Assumed axs1)), axs2) end
-      | axiomify all_ps (id, (_, Derived ths)) axioms =
-          ((id, (all_ps, Derived ths)), axioms);
-
-    (* identifiers of an expression *)
+val _ = Context.>> (Context.map_theory (Theory.at_begin (fn thy =>
+  (case IdentifiersData.get (Context.Theory thy) of
+    Ready _ => NONE |
+    ids => SOME (Context.theory_map (IdentifiersData.put (Ready (finish thy ids))) thy))
+  )));
 
-    fun identify top (Locale name) =
-    (* CB: ids_ax is a list of tuples of the form ((name, ps), axs),
-       where name is a locale name, ps a list of parameter names and axs
-       a list of axioms relating to the identifier, axs is empty unless
-       identify at top level (top = true);
-       parms is accumulated list of parameters *)
-          let
-            val {axiom, params, ...} = the_locale thy name;
-            val ps = map (#1 o #1) params;
-            val (_, ids'', _) = add_with_regs ((name, map #1 params), Assumed []) ([], [], []);
-            val ids_ax = if top then fst (fold_map (axiomify ps) ids'' axiom) else ids'';
-            in (ids_ax, ps) end
-      | identify top (Rename (e, xs)) =
-          let
-            val (ids', parms') = identify top e;
-            val ren = renaming xs parms'
-              handle ERROR msg => err_in_locale' ctxt msg ids';
-
-            val ids'' = distinct (eq_fst (op =)) (map (rename_parms top ren) ids');
-            val parms'' = distinct (op =) (maps (#2 o #1) ids'');
-          in (ids'', parms'') end
-      | identify top (Merge es) =
-          fold (fn e => fn (ids, parms) =>
-                   let
-                     val (ids', parms') = identify top e
-                   in
-                     (merge_alists (op =) ids ids', merge_lists (op =) parms parms')
-                   end)
-            es ([], []);
+fun get_global_idents thy =
+  let val (Ready ids) = (IdentifiersData.get o Context.Theory) thy in ids end;
+val put_global_idents = Context.theory_map o IdentifiersData.put o Ready;
 
-    fun inst_wit all_params (t, th) = let
-         val {hyps, prop, ...} = Thm.rep_thm th;
-         val ps = map (apsnd SOME) (fold Term.add_frees (prop :: hyps) []);
-         val [env] = unify_parms ctxt all_params [ps];
-         val t' = Element.instT_term env t;
-         val th' = Element.instT_thm thy env th;
-       in (t', th') end;
-
-    fun eval all_params tenv syn ((name, params), (locale_params, mode)) =
-      let
-        val {params = ps_mx, elems = elems_stamped, ...} = the_locale thy name;
-        val elems = map fst elems_stamped;
-        val ps = map fst ps_mx;
-        fun lookup_syn x = (case Symtab.lookup syn x of SOME Structure => NONE | opt => opt);
-        val locale_params' = map (fn p => (p, Symtab.lookup tenv p |> the)) locale_params;
-        val mode' = map_mode (map (Element.map_witness (inst_wit all_params))) mode;
-        val ren = map fst ps ~~ map (fn p => (p, lookup_syn p)) params;
-        val [env] = unify_parms ctxt all_params [map (apfst (Element.rename ren) o apsnd SOME) ps];
-        val (lprfx, pprfx) = param_prefix name params;
-        val add_prefices = pprfx <> "" ? Binding.add_prefix false pprfx
-          #> Binding.add_prefix false lprfx;
-        val elem_morphism =
-          Element.rename_morphism ren $>
-          Morphism.binding_morphism add_prefices $>
-          Element.instT_morphism thy env;
-        val elems' = map (Element.morph_ctxt elem_morphism) elems;
-      in (((name, map (apsnd SOME) locale_params'), mode'), elems') end;
-
-    (* parameters, their types and syntax *)
-    val (all_params', tenv, syn) = params_of_expr ctxt [] expr ([], Symtab.empty, Symtab.empty);
-    val all_params = map (fn p => (p, Symtab.lookup tenv p |> the)) all_params';
-    (* compute identifiers and syntax, merge with previous ones *)
-    val (ids, _) = identify true expr;
-    val idents = subtract (eq_fst (op =)) prev_idents ids;
-    val syntax = merge_syntax ctxt ids (syn, prev_syntax);
-    (* type-instantiate elements *)
-    val final_elemss = map (eval all_params tenv syntax) idents;
-  in ((prev_idents @ idents, syntax), final_elemss) end;
+fun get_local_idents ctxt =
+  let val (Ready ids) = (IdentifiersData.get o Context.Proof) ctxt in ids end;
+val put_local_idents = Context.proof_map o IdentifiersData.put o Ready;
 
 end;
 
 
-(* activate elements *)
+(** Resolve locale dependencies in a depth-first fashion **)
 
 local
 
-fun axioms_export axs _ As =
-  (Element.satisfy_thm axs #> Drule.implies_intr_list (Library.drop (length axs, As)), fn t => t);
-
-
-(* NB: derived ids contain only facts at this stage *)
+val roundup_bound = 120;
 
-fun activate_elem _ _ (Fixes fixes) (ctxt, mode) =
-      ([], (ctxt |> ProofContext.add_fixes_i fixes |> snd, mode))
-  | activate_elem _ _ (Constrains _) (ctxt, mode) =
-      ([], (ctxt, mode))
-  | activate_elem ax_in_ctxt _ (Assumes asms) (ctxt, Assumed axs) =
-      let
-        val asms' = Attrib.map_specs (Attrib.attribute_i (ProofContext.theory_of ctxt)) asms;
-        val ts = maps (map #1 o #2) asms';
-        val (ps, qs) = chop (length ts) axs;
-        val (_, ctxt') =
-          ctxt |> fold Variable.auto_fixes ts
-          |> ProofContext.add_assms_i (axioms_export (if ax_in_ctxt then ps else [])) asms';
-      in ([], (ctxt', Assumed qs)) end
-  | activate_elem _ _ (Assumes asms) (ctxt, Derived ths) =
-      ([], (ctxt, Derived ths))
-  | activate_elem _ _ (Defines defs) (ctxt, Assumed axs) =
-      let
-        val defs' = Attrib.map_specs (Attrib.attribute_i (ProofContext.theory_of ctxt)) defs;
-        val asms = defs' |> map (fn ((name, atts), (t, ps)) =>
-            let val ((c, _), t') = LocalDefs.cert_def ctxt t
-            in (t', ((Binding.map_base (Thm.def_name_optional c) name, atts), [(t', ps)])) end);
-        val (_, ctxt') =
-          ctxt |> fold (Variable.auto_fixes o #1) asms
-          |> ProofContext.add_assms_i LocalDefs.def_export (map #2 asms);
-      in ([], (ctxt', Assumed axs)) end
-  | activate_elem _ _ (Defines defs) (ctxt, Derived ths) =
-      ([], (ctxt, Derived ths))
-  | activate_elem _ is_ext (Notes (kind, facts)) (ctxt, mode) =
-      let
-        val facts' = Attrib.map_facts (Attrib.attribute_i (ProofContext.theory_of ctxt)) facts;
-        val (res, ctxt') = ctxt |> local_note_qualified kind facts';
-      in (if is_ext then (map (#1 o #1) facts' ~~ map #2 res) else [], (ctxt', mode)) end;
-
-fun activate_elems ax_in_ctxt (((name, ps), mode), elems) ctxt =
-  let
-    val thy = ProofContext.theory_of ctxt;
-    val (res, (ctxt', _)) = fold_map (activate_elem ax_in_ctxt (name = ""))
-        elems (ProofContext.qualified_names ctxt, mode)
-      handle ERROR msg => err_in_locale ctxt msg [(name, map fst ps)];
-    val ctxt'' = if name = "" then ctxt'
-          else let
-              val ps' = map (fn (n, SOME T) => Free (n, T)) ps;
-            in if test_local_registration ctxt' (name, ps') then ctxt'
-              else let
-                  val ctxt'' = put_local_registration (name, ps') (I, (NameSpace.base name, ""))
-                    (Morphism.identity, ((Vartab.empty, []), (Vartab.empty, []) )) ctxt'
-                in case mode of
-                    Assumed axs =>
-                      fold (add_local_witness (name, ps') o
-                        Element.assume_witness thy o Element.witness_prop) axs ctxt''
-                  | Derived ths =>
-                     fold (add_local_witness (name, ps')) ths ctxt''
-                end
-            end
-  in (ProofContext.restore_naming ctxt ctxt'', res) end;
-
-fun activate_elemss ax_in_ctxt prep_facts =
-    fold_map (fn (((name, ps), mode), raw_elems) => fn ctxt =>
-      let
-        val elems = map (prep_facts ctxt) raw_elems;
-        val (ctxt', res) = apsnd flat
-            (activate_elems ax_in_ctxt (((name, ps), mode), elems) ctxt);
-        val elems' = elems |> map (Element.map_ctxt_attrib Args.closure);
-      in (((((name, ps), mode), elems'), res), ctxt') end);
+fun add thy depth (name, morph) (deps, marked) =
+  if depth > roundup_bound
+  then error "Roundup bound exceeded (sublocale relation probably not terminating)."
+  else
+    let
+      val {parameters = (_, params), dependencies, ...} = the_locale thy name;
+      val instance = instance_of thy name morph;
+    in
+      if member (ident_eq thy) marked (name, instance)
+      then (deps, marked)
+      else
+        let
+          val dependencies' =
+            map (fn ((name, morph'), _) => (name, morph' $>  morph)) dependencies;
+          val marked' = (name, instance) :: marked;
+          val (deps', marked'') = fold_rev (add thy (depth + 1)) dependencies' ([], marked');
+        in
+          ((name, morph) :: deps' @ deps, marked'')
+        end
+    end;
 
 in
 
-(* CB: activate_facts prep_facts elemss ctxt,
-   where elemss is a list of pairs consisting of identifiers and
-   context elements, extends ctxt by the context elements yielding
-   ctxt' and returns ((elemss', facts), ctxt').
-   Identifiers in the argument are of the form ((name, ps), axs) and
-   assumptions use the axioms in the identifiers to set up exporters
-   in ctxt'.  elemss' does not contain identifiers and is obtained
-   from elemss and the intermediate context with prep_facts.
-   If read_facts or cert_facts is used for prep_facts, these also remove
-   the internal/external markers from elemss. *)
-
-fun activate_facts ax_in_ctxt prep_facts args =
-  activate_elemss ax_in_ctxt prep_facts args
-  #>> (apsnd flat o split_list);
+fun roundup thy activate_dep (name, morph) (marked, input) =
+  let
+    (* Find all dependencies incuding new ones (which are dependencies enriching
+      existing registrations). *)
+    val (dependencies, marked') = add thy 0 (name, morph) ([], empty);
+    (* Filter out exisiting fragments. *)
+    val dependencies' = filter_out (fn (name, morph) =>
+      member (ident_eq thy) marked (name, instance_of thy name morph)) dependencies;
+  in
+    (merge (ident_eq thy) (marked, marked'), input |> fold_rev (activate_dep thy) dependencies')
+  end;
 
 end;
 
 
+(* Declarations, facts and entire locale content *)
 
-(** prepare locale elements **)
+fun activate_decls thy (name, morph) ctxt =
+  let
+    val {decls = (typ_decls, term_decls), ...} = the_locale thy name;
+  in
+    ctxt |> fold_rev (fn (decl, _) => Context.proof_map (decl morph)) typ_decls |>
+      fold_rev (fn (decl, _) => Context.proof_map (decl morph)) term_decls
+  end;
 
-(* expressions *)
+fun activate_notes activ_elem transfer thy (name, morph) input =
+  let
+    val {notes, ...} = the_locale thy name;
+    fun activate ((kind, facts), _) input =
+      let
+        val facts' = facts |> Element.facts_map (Element.morph_ctxt (transfer input $> morph))
+      in activ_elem (Notes (kind, facts')) input end;
+  in
+    fold_rev activate notes input
+  end;
 
-fun intern_expr thy (Locale xname) = Locale (intern thy xname)
-  | intern_expr thy (Merge exprs) = Merge (map (intern_expr thy) exprs)
-  | intern_expr thy (Rename (expr, xs)) = Rename (intern_expr thy expr, xs);
+fun activate_all name thy activ_elem transfer (marked, input) =
+  let
+    val {parameters = (_, params), spec = (asm, defs), ...} =
+      the_locale thy name;
+  in
+    input |>
+      (if not (null params) then activ_elem (Fixes params) else I) |>
+      (* FIXME type parameters *)
+      (if is_some asm then activ_elem (Assumes [(Attrib.empty_binding, [(the asm, [])])]) else I) |>
+      (if not (null defs)
+        then activ_elem (Defines (map (fn def => (Attrib.empty_binding, (def, []))) defs))
+        else I) |>
+      pair marked |> roundup thy (activate_notes activ_elem transfer) (name, Morphism.identity)
+  end;
 
 
-(* propositions and bindings *)
-
-(* flatten (ctxt, prep_expr) ((ids, syn), expr)
-   normalises expr (which is either a locale
-   expression or a single context element) wrt.
-   to the list ids of already accumulated identifiers.
-   It returns ((ids', syn'), elemss) where ids' is an extension of ids
-   with identifiers generated for expr, and elemss is the list of
-   context elements generated from expr.
-   syn and syn' are symtabs mapping parameter names to their syntax.  syn'
-   is an extension of syn.
-   For details, see flatten_expr.
-
-   Additionally, for a locale expression, the elems are grouped into a single
-   Int; individual context elements are marked Ext.  In this case, the
-   identifier-like information of the element is as follows:
-   - for Fixes: (("", ps), []) where the ps have type info NONE
-   - for other elements: (("", []), []).
-   The implementation of activate_facts relies on identifier names being
-   empty strings for external elements.
-*)
-
-fun flatten (ctxt, _) ((ids, syn), Elem (Fixes fixes)) = let
-        val ids' = ids @ [(("", map (Binding.base_name o #1) fixes), ([], Assumed []))]
-      in
-        ((ids',
-         merge_syntax ctxt ids'
-           (syn, Symtab.make (map (fn fx => (Binding.base_name (#1 fx), #3 fx)) fixes))
-           handle Symtab.DUP x => err_in_locale ctxt
-             ("Conflicting syntax for parameter: " ^ quote x)
-             (map #1 ids')),
-         [((("", map (rpair NONE o Binding.base_name o #1) fixes), Assumed []), Ext (Fixes fixes))])
-      end
-  | flatten _ ((ids, syn), Elem elem) =
-      ((ids @ [(("", []), ([], Assumed []))], syn), [((("", []), Assumed []), Ext elem)])
-  | flatten (ctxt, prep_expr) ((ids, syn), Expr expr) =
-      apsnd (map (apsnd Int)) (flatten_expr ctxt ((ids, syn), prep_expr expr));
-
-local
-
-local
-
-fun declare_int_elem (Fixes fixes) ctxt =
-      ([], ctxt |> ProofContext.add_fixes_i (map (fn (x, T, mx) =>
-        (x, Option.map (Term.map_type_tfree (TypeInfer.param 0)) T, mx)) fixes) |> snd)
-  | declare_int_elem _ ctxt = ([], ctxt);
-
-fun declare_ext_elem prep_vars (Fixes fixes) ctxt =
-      let val (vars, _) = prep_vars fixes ctxt
-      in ([], ctxt |> ProofContext.add_fixes_i vars |> snd) end
-  | declare_ext_elem prep_vars (Constrains csts) ctxt =
-      let val (_, ctxt') = prep_vars (map (fn (x, T) => (Binding.name x, SOME T, NoSyn)) csts) ctxt
-      in ([], ctxt') end
-  | declare_ext_elem _ (Assumes asms) ctxt = (map #2 asms, ctxt)
-  | declare_ext_elem _ (Defines defs) ctxt = (map (fn (_, (t, ps)) => [(t, ps)]) defs, ctxt)
-  | declare_ext_elem _ (Notes _) ctxt = ([], ctxt);
-
-fun declare_elems prep_vars (((name, ps), Assumed _), elems) ctxt = ((case elems
-     of Int es => fold_map declare_int_elem es ctxt
-      | Ext e => declare_ext_elem prep_vars e ctxt |>> single)
-          handle ERROR msg => err_in_locale ctxt msg [(name, map fst ps)])
-  | declare_elems _ ((_, Derived _), elems) ctxt = ([], ctxt);
-
-in
-
-fun declare_elemss prep_vars fixed_params raw_elemss ctxt =
-  let
-    (* CB: fix of type bug of goal in target with context elements.
-       Parameters new in context elements must receive types that are
-       distinct from types of parameters in target (fixed_params).  *)
-    val ctxt_with_fixed = 
-      fold Variable.declare_term (map Free fixed_params) ctxt;
-    val int_elemss =
-      raw_elemss
-      |> map_filter (fn (id, Int es) => SOME (id, es) | _ => NONE)
-      |> unify_elemss ctxt_with_fixed fixed_params;
-    val (raw_elemss', _) =
-      fold_map (curry (fn ((id, Int _), (_, es) :: elemss) => ((id, Int es), elemss) | x => x))
-        raw_elemss int_elemss;
-  in fold_map (declare_elems prep_vars) raw_elemss' ctxt end;
-
-end;
+(** Public activation functions **)
 
 local
 
-val norm_term = Envir.beta_norm oo Term.subst_atomic;
-
-fun abstract_thm thy eq =
-  Thm.assume (Thm.cterm_of thy eq) |> Drule.gen_all |> Drule.abs_def;
-
-fun bind_def ctxt (name, ps) eq (xs, env, ths) =
+fun init_global_elem (Notes (kind, facts)) thy =
   let
-    val ((y, T), b) = LocalDefs.abs_def eq;
-    val b' = norm_term env b;
-    val th = abstract_thm (ProofContext.theory_of ctxt) eq;
-    fun err msg = err_in_locale ctxt (msg ^ ": " ^ quote y) [(name, map fst ps)];
-  in
-    exists (fn (x, _) => x = y) xs andalso
-      err "Attempt to define previously specified variable";
-    exists (fn (Free (y', _), _) => y = y' | _ => false) env andalso
-      err "Attempt to redefine variable";
-    (Term.add_frees b' xs, (Free (y, T), b') :: env, th :: ths)
-  end;
-
-
-(* CB: for finish_elems (Int and Ext),
-   extracts specification, only of assumed elements *)
-
-fun eval_text _ _ _ (Fixes _) text = text
-  | eval_text _ _ _ (Constrains _) text = text
-  | eval_text _ (_, Assumed _) is_ext (Assumes asms)
-        (((exts, exts'), (ints, ints')), (xs, env, defs)) =
-      let
-        val ts = maps (map #1 o #2) asms;
-        val ts' = map (norm_term env) ts;
-        val spec' =
-          if is_ext then ((exts @ ts, exts' @ ts'), (ints, ints'))
-          else ((exts, exts'), (ints @ ts, ints' @ ts'));
-      in (spec', (fold Term.add_frees ts' xs, env, defs)) end
-  | eval_text _ (_, Derived _) _ (Assumes _) text = text
-  | eval_text ctxt (id, Assumed _) _ (Defines defs) (spec, binds) =
-      (spec, fold (bind_def ctxt id o #1 o #2) defs binds)
-  | eval_text _ (_, Derived _) _ (Defines _) text = text
-  | eval_text _ _ _ (Notes _) text = text;
-
-
-(* for finish_elems (Int),
-   remove redundant elements of derived identifiers,
-   turn assumptions and definitions into facts,
-   satisfy hypotheses of facts *)
-
-fun finish_derived _ _ (Assumed _) (Fixes fixes) = SOME (Fixes fixes)
-  | finish_derived _ _ (Assumed _) (Constrains csts) = SOME (Constrains csts)
-  | finish_derived _ _ (Assumed _) (Assumes asms) = SOME (Assumes asms)
-  | finish_derived _ _ (Assumed _) (Defines defs) = SOME (Defines defs)
+    val facts' = Attrib.map_facts (Attrib.attribute_i thy) facts
+  in Old_Locale.global_note_qualified kind facts' thy |> snd end
 
-  | finish_derived _ _ (Derived _) (Fixes _) = NONE
-  | finish_derived _ _ (Derived _) (Constrains _) = NONE
-  | finish_derived sign satisfy (Derived _) (Assumes asms) = asms
-      |> map (apsnd (map (fn (a, _) => ([Thm.assume (cterm_of sign a)], []))))
-      |> pair Thm.assumptionK |> Notes
-      |> Element.morph_ctxt satisfy |> SOME
-  | finish_derived sign satisfy (Derived _) (Defines defs) = defs
-      |> map (apsnd (fn (d, _) => [([Thm.assume (cterm_of sign d)], [])]))
-      |> pair Thm.definitionK |> Notes
-      |> Element.morph_ctxt satisfy |> SOME
-
-  | finish_derived _ satisfy _ (Notes facts) = Notes facts
-      |> Element.morph_ctxt satisfy |> SOME;
-
-(* CB: for finish_elems (Ext) *)
-
-fun closeup _ false elem = elem
-  | closeup ctxt true elem =
+fun init_local_elem (Fixes fixes) ctxt = ctxt |>
+      ProofContext.add_fixes_i fixes |> snd
+  | init_local_elem (Assumes assms) ctxt =
       let
-        fun close_frees t =
-          let
-            val rev_frees =
-              Term.fold_aterms (fn Free (x, T) =>
-                if Variable.is_fixed ctxt x then I else insert (op =) (x, T) | _ => I) t [];
-          in Term.list_all_free (rev rev_frees, t) end;
-
-        fun no_binds [] = []
-          | no_binds _ = error "Illegal term bindings in locale element";
+        val assms' = Attrib.map_specs (Attrib.attribute_i (ProofContext.theory_of ctxt)) assms
       in
-        (case elem of
-          Assumes asms => Assumes (asms |> map (fn (a, propps) =>
-            (a, map (fn (t, ps) => (close_frees t, no_binds ps)) propps)))
-        | Defines defs => Defines (defs |> map (fn (a, (t, ps)) =>
-            (a, (close_frees (#2 (LocalDefs.cert_def ctxt t)), no_binds ps))))
-        | e => e)
-      end;
-
-
-fun finish_ext_elem parms _ (Fixes fixes, _) = Fixes (map (fn (b, _, mx) =>
-      let val x = Binding.base_name b
-      in (b, AList.lookup (op =) parms x, mx) end) fixes)
-  | finish_ext_elem parms _ (Constrains _, _) = Constrains []
-  | finish_ext_elem _ close (Assumes asms, propp) =
-      close (Assumes (map #1 asms ~~ propp))
-  | finish_ext_elem _ close (Defines defs, propp) =
-      close (Defines (map #1 defs ~~ map (fn [(t, ps)] => (t, ps)) propp))
-  | finish_ext_elem _ _ (Notes facts, _) = Notes facts;
-
+        ctxt |> fold Variable.auto_fixes (maps (map fst o snd) assms') |>
+          ProofContext.add_assms_i Assumption.assume_export assms' |> snd
+     end
+  | init_local_elem (Defines defs) ctxt =
+      let
+        val defs' = Attrib.map_specs (Attrib.attribute_i (ProofContext.theory_of ctxt)) defs
+      in
+        ctxt |> fold Variable.auto_fixes (map (fst o snd) defs') |>
+          ProofContext.add_assms_i LocalDefs.def_export (map (fn (attn, t) => (attn, [t])) defs') |>
+          snd
+      end
+  | init_local_elem (Notes (kind, facts)) ctxt =
+      let
+        val facts' = Attrib.map_facts (Attrib.attribute_i (ProofContext.theory_of ctxt)) facts
+      in Old_Locale.local_note_qualified kind facts' ctxt |> snd end
 
-(* CB: finish_parms introduces type info from parms to identifiers *)
-(* CB: only needed for types that have been NONE so far???
-   If so, which are these??? *)
-
-fun finish_parms parms (((name, ps), mode), elems) =
-  (((name, map (fn (x, _) => (x, AList.lookup (op = : string * string -> bool) parms x)) ps), mode), elems);
-
-fun finish_elems ctxt parms _ ((text, wits), ((id, Int e), _)) =
-      let
-        val [(id' as (_, mode), es)] = unify_elemss ctxt parms [(id, e)];
-        val wits' = case mode of Assumed _ => wits | Derived ths => wits @ ths;
-        val text' = fold (eval_text ctxt id' false) es text;
-        val es' = map_filter
-          (finish_derived (ProofContext.theory_of ctxt) (Element.satisfy_morphism wits') mode) es;
-      in ((text', wits'), (id', map Int es')) end
-  | finish_elems ctxt parms do_close ((text, wits), ((id, Ext e), [propp])) =
-      let
-        val e' = finish_ext_elem parms (closeup ctxt do_close) (e, propp);
-        val text' = eval_text ctxt id true e' text;
-      in ((text', wits), (id, [Ext e'])) end
+fun cons_elem false (Notes notes) elems = elems
+  | cons_elem _ elem elems = elem :: elems
 
 in
 
-(* CB: only called by prep_elemss *)
+fun activate_declarations thy dep ctxt =
+  roundup thy activate_decls dep (get_local_idents ctxt, ctxt) |-> put_local_idents;
+
+fun activate_global_facts dep thy =
+  roundup thy (activate_notes init_global_elem Element.transfer_morphism)
+    dep (get_global_idents thy, thy) |-> put_global_idents;
+
+fun activate_local_facts dep ctxt =
+  roundup (ProofContext.theory_of ctxt)
+  (activate_notes init_local_elem (Element.transfer_morphism o ProofContext.theory_of)) dep
+    (get_local_idents ctxt, ctxt) |-> put_local_idents;
 
-fun finish_elemss ctxt parms do_close =
-  foldl_map (apsnd (finish_parms parms) o finish_elems ctxt parms do_close);
+fun init name thy =
+  activate_all name thy init_local_elem (Element.transfer_morphism o ProofContext.theory_of)
+    (empty, ProofContext.init thy) |-> put_local_idents;
+
+fun print_locale thy show_facts name =
+  let
+    val name' = intern thy name;
+    val ctxt = init name' thy
+  in
+    Pretty.big_list "locale elements:"
+      (activate_all name' thy (cons_elem show_facts) (K (Element.transfer_morphism thy))
+        (empty, []) |> snd |> rev |>
+        map (Element.pretty_ctxt ctxt) |> map Pretty.chunks) |> Pretty.writeln
+  end
 
 end;
 
 
-(* Remove duplicate Defines elements: temporary workaround to fix Afp/Category. *)
-
-fun defs_ord (defs1, defs2) =
-    list_ord (fn ((_, (d1, _)), (_, (d2, _))) =>
-      Term.fast_term_ord (d1, d2)) (defs1, defs2);
-structure Defstab =
-    TableFun(type key = (Attrib.binding * (term * term list)) list val ord = defs_ord);
-
-fun rem_dup_defs es ds =
-    fold_map (fn e as (Defines defs) => (fn ds =>
-                 if Defstab.defined ds defs
-                 then (Defines [], ds)
-                 else (e, Defstab.update (defs, ()) ds))
-               | e => (fn ds => (e, ds))) es ds;
-fun rem_dup_elemss (Int es) ds = apfst Int (rem_dup_defs es ds)
-  | rem_dup_elemss (Ext e) ds = (Ext e, ds);
-fun rem_dup_defines raw_elemss =
-    fold_map (fn (id as (_, (Assumed _)), es) => (fn ds =>
-                     apfst (pair id) (rem_dup_elemss es ds))
-               | (id as (_, (Derived _)), es) => (fn ds =>
-                     ((id, es), ds))) raw_elemss Defstab.empty |> #1;
-
-(* CB: type inference and consistency checks for locales.
-
-   Works by building a context (through declare_elemss), extracting the
-   required information and adjusting the context elements (finish_elemss).
-   Can also universally close free vars in assms and defs.  This is only
-   needed for Ext elements and controlled by parameter do_close.
-
-   Only elements of assumed identifiers are considered.
-*)
+(*** Registrations: interpretations in theories ***)
 
-fun prep_elemss prep_vars prepp do_close context fixed_params raw_elemss raw_concl =
-  let
-    (* CB: contexts computed in the course of this function are discarded.
-       They are used for type inference and consistency checks only. *)
-    (* CB: fixed_params are the parameters (with types) of the target locale,
-       empty list if there is no target. *)
-    (* CB: raw_elemss are list of pairs consisting of identifiers and
-       context elements, the latter marked as internal or external. *)
-    val raw_elemss = rem_dup_defines raw_elemss;
-    val (raw_proppss, raw_ctxt) = declare_elemss prep_vars fixed_params raw_elemss context;
-    (* CB: raw_ctxt is context with additional fixed variables derived from
-       the fixes elements in raw_elemss,
-       raw_proppss contains assumptions and definitions from the
-       external elements in raw_elemss. *)
-    fun prep_prop raw_propp (raw_ctxt, raw_concl)  =
-      let
-        (* CB: add type information from fixed_params to context (declare_term) *)
-        (* CB: process patterns (conclusion and external elements only) *)
-        val (ctxt, all_propp) =
-          prepp (fold Variable.declare_term (map Free fixed_params) raw_ctxt, raw_concl @ raw_propp);
-        (* CB: add type information from conclusion and external elements to context *)
-        val ctxt = fold Variable.declare_term (maps (map fst) all_propp) ctxt;
-        (* CB: resolve schematic variables (patterns) in conclusion and external elements. *)
-        val all_propp' = map2 (curry (op ~~))
-          (#1 (#2 (ProofContext.bind_propp_schematic_i (ctxt, all_propp)))) (map (map snd) all_propp);
-        val (concl, propp) = chop (length raw_concl) all_propp';
-      in (propp, (ctxt, concl)) end
+structure RegistrationsData = TheoryDataFun
+(
+  type T = ((string * (morphism * morphism)) * stamp) list;
+    (* FIXME mixins need to be stamped *)
+    (* registrations, in reverse order of declaration *)
+  val empty = [];
+  val extend = I;
+  val copy = I;
+  fun merge _ data : T = Library.merge (eq_snd op =) data;
+    (* FIXME consolidate with dependencies, consider one data slot only *)
+);
 
-    val (proppss, (ctxt, concl)) =
-      (fold_burrow o fold_burrow) prep_prop raw_proppss (raw_ctxt, raw_concl);
-
-    (* CB: obtain all parameters from identifier part of raw_elemss *)
-    val xs = map #1 (params_of' raw_elemss);
-    val typing = unify_frozen ctxt 0
-      (map (Variable.default_type raw_ctxt) xs)
-      (map (Variable.default_type ctxt) xs);
-    val parms = param_types (xs ~~ typing);
-    (* CB: parms are the parameters from raw_elemss, with correct typing. *)
+val get_registrations =
+  RegistrationsData.get #> map fst #> map (apsnd op $>);
 
-    (* CB: extract information from assumes and defines elements
-       (fixes, constrains and notes in raw_elemss don't have an effect on
-       text and elemss), compute final form of context elements. *)
-    val ((text, _), elemss) = finish_elemss ctxt parms do_close
-      ((((([], []), ([], [])), ([], [], [])), []), raw_elemss ~~ proppss);
-    (* CB: text has the following structure:
-           (((exts, exts'), (ints, ints')), (xs, env, defs))
-       where
-         exts: external assumptions (terms in external assumes elements)
-         exts': dito, normalised wrt. env
-         ints: internal assumptions (terms in internal assumes elements)
-         ints': dito, normalised wrt. env
-         xs: the free variables in exts' and ints' and rhss of definitions,
-           this includes parameters except defined parameters
-         env: list of term pairs encoding substitutions, where the first term
-           is a free variable; substitutions represent defines elements and
-           the rhs is normalised wrt. the previous env
-         defs: theorems representing the substitutions from defines elements
-           (thms are normalised wrt. env).
-       elemss is an updated version of raw_elemss:
-         - type info added to Fixes and modified in Constrains
-         - axiom and definition statement replaced by corresponding one
-           from proppss in Assumes and Defines
-         - Facts unchanged
-       *)
-  in ((parms, elemss, concl), text) end;
-
-in
-
-fun read_elemss x = prep_elemss ProofContext.read_vars ProofContext.read_propp_schematic x;
-fun cert_elemss x = prep_elemss ProofContext.cert_vars ProofContext.cert_propp_schematic x;
-
-end;
-
+fun add_registration (name, (base_morph, export)) thy =
+  roundup thy (fn _ => fn (name', morph') =>
+    (RegistrationsData.map o cons) ((name', (morph', export)), stamp ()))
+    (name, base_morph) (get_global_idents thy, thy) |> snd
+    (* FIXME |-> put_global_idents ?*);
 
-(* facts and attributes *)
-
-local
-
-fun check_name name =
-  if NameSpace.is_qualified name then error ("Illegal qualified name: " ^ quote name)
-  else name;
-
-fun prep_facts _ _ _ ctxt (Int elem) = elem
-      |> Element.morph_ctxt (Morphism.thm_morphism (Thm.transfer (ProofContext.theory_of ctxt)))
-  | prep_facts prep_name get intern ctxt (Ext elem) = elem |> Element.map_ctxt
-     {var = I, typ = I, term = I,
-      binding = Binding.map_base prep_name,
-      fact = get ctxt,
-      attrib = Args.assignable o intern (ProofContext.theory_of ctxt)};
-
-in
-
-fun read_facts x = prep_facts check_name ProofContext.get_fact Attrib.intern_src x;
-fun cert_facts x = prep_facts I (K I) (K I) x;
-
-end;
-
-
-(* Get the specification of a locale *)
-
-(*The global specification is made from the parameters and global
-  assumptions, the local specification from the parameters and the
-  local assumptions.*)
-
-local
-
-fun gen_asms_of get thy name =
+fun amend_registration morph (name, base_morph) thy =
   let
-    val ctxt = ProofContext.init thy;
-    val (_, raw_elemss) = flatten (ctxt, I) (([], Symtab.empty), Expr (Locale name));
-    val ((_, elemss, _), _) = read_elemss false ctxt [] raw_elemss [];
+    val regs = (RegistrationsData.get #> map fst) thy;
+    val base = instance_of thy name base_morph;
+    fun match (name', (morph', _)) =
+      name = name' andalso eq_list (op aconv) (base, instance_of thy name' morph');
+    val i = find_index match (rev regs);
+    val _ = if i = ~1 then error ("No registration of locale " ^
+        quote (extern thy name) ^ " and parameter instantiation " ^
+        space_implode " " (map (quote o Syntax.string_of_term_global thy) base) ^ " available.")
+      else ();
   in
-    elemss |> get
-      |> maps (fn (_, es) => map (fn Int e => e) es)
-      |> maps (fn Assumes asms => asms | _ => [])
-      |> map (apsnd (map fst))
-  end;
-
-in
-
-fun parameters_of thy = #params o the_locale thy;
-
-fun intros thy = #intros o the_locale thy;
-  (*returns introduction rule for delta predicate and locale predicate
-    as a pair of singleton lists*)
-
-fun dests thy = #dests o the_locale thy;
-
-fun facts_of thy = map_filter (fn (Element.Notes (_, facts), _) => SOME facts
-  | _ => NONE) o #elems o the_locale thy;
-
-fun parameters_of_expr thy expr =
-  let
-    val ctxt = ProofContext.init thy;
-    val pts = params_of_expr ctxt [] (intern_expr thy expr)
-        ([], Symtab.empty, Symtab.empty);
-    val raw_params_elemss = make_raw_params_elemss pts;
-    val ((_, syn), raw_elemss) = flatten (ctxt, intern_expr thy)
-        (([], Symtab.empty), Expr expr);
-    val ((parms, _, _), _) =
-        read_elemss false ctxt [] (raw_params_elemss @ raw_elemss) [];
-  in map (fn p as (n, _) => (p, Symtab.lookup syn n |> the)) parms end;
-
-fun local_asms_of thy name =
-  gen_asms_of (single o Library.last_elem) thy name;
-
-fun global_asms_of thy name =
-  gen_asms_of I thy name;
-
-end;
-
-
-(* full context statements: imports + elements + conclusion *)
-
-local
-
-fun prep_context_statement prep_expr prep_elemss prep_facts
-    do_close fixed_params imports elements raw_concl context =
-  let
-    val thy = ProofContext.theory_of context;
-
-    val (import_params, import_tenv, import_syn) =
-      params_of_expr context fixed_params (prep_expr thy imports)
-        ([], Symtab.empty, Symtab.empty);
-    val includes = map_filter (fn Expr e => SOME e | Elem _ => NONE) elements;
-    val (incl_params, incl_tenv, incl_syn) = fold (params_of_expr context fixed_params)
-      (map (prep_expr thy) includes) (import_params, import_tenv, import_syn);
-
-    val ((import_ids, _), raw_import_elemss) =
-      flatten (context, prep_expr thy) (([], Symtab.empty), Expr imports);
-    (* CB: normalise "includes" among elements *)
-    val ((ids, syn), raw_elemsss) = foldl_map (flatten (context, prep_expr thy))
-      ((import_ids, incl_syn), elements);
-
-    val raw_elemss = flat raw_elemsss;
-    (* CB: raw_import_elemss @ raw_elemss is the normalised list of
-       context elements obtained from import and elements. *)
-    (* Now additional elements for parameters are inserted. *)
-    val import_params_ids = make_params_ids import_params;
-    val incl_params_ids =
-        make_params_ids (incl_params \\ import_params);
-    val raw_import_params_elemss =
-        make_raw_params_elemss (import_params, incl_tenv, incl_syn);
-    val raw_incl_params_elemss =
-        make_raw_params_elemss (incl_params \\ import_params, incl_tenv, incl_syn);
-    val ((parms, all_elemss, concl), (spec, (_, _, defs))) = prep_elemss do_close
-      context fixed_params
-      (raw_import_params_elemss @ raw_import_elemss @ raw_incl_params_elemss @ raw_elemss) raw_concl;
-
-    (* replace extended ids (for axioms) by ids *)
-    val (import_ids', incl_ids) = chop (length import_ids) ids;
-    val all_ids = import_params_ids @ import_ids' @ incl_params_ids @ incl_ids;
-    val all_elemss' = map (fn (((_, ps), _), (((n, ps'), mode), elems)) =>
-        (((n, map (fn p => (p, (the o AList.lookup (op =) ps') p)) ps), mode), elems))
-      (all_ids ~~ all_elemss);
-    (* CB: all_elemss and parms contain the correct parameter types *)
-
-    val (ps, qs) = chop (length raw_import_params_elemss + length raw_import_elemss) all_elemss';
-    val ((import_elemss, _), import_ctxt) =
-      activate_facts false prep_facts ps context;
-
-    val ((elemss, _), ctxt) =
-      activate_facts false prep_facts qs (ProofContext.set_stmt true import_ctxt);
-  in
-    ((((import_ctxt, import_elemss), (ctxt, elemss, syn)),
-      (parms, spec, defs)), concl)
-  end;
-
-fun prep_statement prep_locale prep_ctxt raw_locale elems concl ctxt =
-  let
-    val thy = ProofContext.theory_of ctxt;
-    val locale = Option.map (prep_locale thy) raw_locale;
-    val (fixed_params, imports) =
-      (case locale of
-        NONE => ([], empty)
-      | SOME name =>
-          let val {params = ps, ...} = the_locale thy name
-          in (map fst ps, Locale name) end);
-    val ((((locale_ctxt, _), (elems_ctxt, _, _)), _), concl') =
-      prep_ctxt false fixed_params imports (map Elem elems) concl ctxt;
-  in (locale, locale_ctxt, elems_ctxt, concl') end;
-
-fun prep_expr prep imports body ctxt =
-  let
-    val (((_, import_elemss), (ctxt', elemss, _)), _) = prep imports body ctxt;
-    val all_elems = maps snd (import_elemss @ elemss);
-  in (all_elems, ctxt') end;
-
-in
-
-val read_ctxt = prep_context_statement intern_expr read_elemss read_facts;
-val cert_ctxt = prep_context_statement (K I) cert_elemss cert_facts;
-
-fun read_context imports body ctxt = #1 (read_ctxt true [] imports (map Elem body) [] ctxt);
-fun cert_context imports body ctxt = #1 (cert_ctxt true [] imports (map Elem body) [] ctxt);
-
-val read_expr = prep_expr read_context;
-val cert_expr = prep_expr cert_context;
-
-fun read_context_statement loc = prep_statement (K I) read_ctxt loc;
-fun read_context_statement_cmd loc = prep_statement intern read_ctxt loc;
-fun cert_context_statement loc = prep_statement (K I) cert_ctxt loc;
-
-end;
-
-
-(* init *)
-
-fun init loc =
-  ProofContext.init
-  #> #2 o cert_context_statement (SOME loc) [] [];
-
-
-(* print locale *)
-
-fun print_locale thy show_facts imports body =
-  let val (all_elems, ctxt) = read_expr imports body (ProofContext.init thy) in
-    Pretty.big_list "locale elements:" (all_elems
-      |> (if show_facts then I else filter (fn Notes _ => false | _ => true))
-      |> map (Element.pretty_ctxt ctxt) |> filter_out null
-      |> map Pretty.chunks)
-    |> Pretty.writeln
+    RegistrationsData.map (nth_map (length regs - 1 - i)
+      (fn ((name, (base, export)), stamp) => ((name, (base $> morph, export)), stamp))) thy
   end;
 
 
 
-(** store results **)
-
-(* join equations of an id with already accumulated ones *)
-
-fun join_eqns get_reg id eqns =
-  let
-    val eqns' = case get_reg id
-      of NONE => eqns
-        | SOME (_, _, eqns') => Termtab.join (fn _ => fn (_, e) => e) (eqns, eqns')
-            (* prefer equations from eqns' *)
-  in ((id, eqns'), eqns') end;
-
-
-(* collect witnesses and equations up to a particular target for a
-   registration; requires parameters and flattened list of identifiers
-   instead of recomputing it from the target *)
-
-fun collect_witnesses ctxt (imprt as ((impT, _), (imp, _))) parms ids ext_ts = let
-
-    val thy = ProofContext.theory_of ctxt;
-
-    val ts = map (var_inst_term (impT, imp)) ext_ts;
-    val (parms, parmTs) = split_list parms;
-    val parmvTs = map Logic.varifyT parmTs;
-    val vtinst = fold (Sign.typ_match thy) (parmvTs ~~ map Term.fastype_of ts) Vartab.empty;
-    val tinst = Vartab.dest vtinst |> map (fn ((x, 0), (_, T)) => (x, T))
-        |> Symtab.make;
-    val inst = Symtab.make (parms ~~ ts);
-
-    (* instantiate parameter names in ids *)
-    val ext_inst = Symtab.make (parms ~~ ext_ts);
-    fun ext_inst_names ps = map (the o Symtab.lookup ext_inst) ps;
-    val inst_ids = map (apfst (apsnd ext_inst_names)) ids;
-    val assumed_ids = map_filter (fn (id, (_, Assumed _)) => SOME id | _ => NONE) inst_ids;
-    val wits = maps (#2 o the o get_local_registration ctxt imprt) assumed_ids;
-    val eqns =
-      fold_map (join_eqns (get_local_registration ctxt imprt))
-        (map fst inst_ids) Termtab.empty |> snd |> Termtab.dest |> map snd;
-  in ((tinst, inst), wits, eqns) end;
-
-
-(* compute and apply morphism *)
-
-fun name_morph phi_name (lprfx, pprfx) b =
-  b
-  |> (if not (Binding.is_empty b) andalso pprfx <> ""
-        then Binding.add_prefix false pprfx else I)
-  |> (if not (Binding.is_empty b)
-        then Binding.add_prefix false lprfx else I)
-  |> phi_name;
+(*** Storing results ***)
 
-fun inst_morph thy phi_name param_prfx insts prems eqns export =
-  let
-    (* standardise export morphism *)
-    val exp_fact = Drule.zero_var_indexes_list o map Thm.strip_shyps o Morphism.fact export;
-    val exp_term = TermSubst.zero_var_indexes o Morphism.term export;
-      (* FIXME sync with exp_fact *)
-    val exp_typ = Logic.type_map exp_term;
-    val export' =
-      Morphism.morphism {binding = I, var = I, typ = exp_typ, term = exp_term, fact = exp_fact};
-  in
-    Morphism.binding_morphism (name_morph phi_name param_prfx) $>
-      Element.inst_morphism thy insts $>
-      Element.satisfy_morphism prems $>
-      Morphism.term_morphism (MetaSimplifier.rewrite_term thy eqns []) $>
-      Morphism.thm_morphism (MetaSimplifier.rewrite_rule eqns) $>
-      export'
-  end;
-
-fun activate_note thy phi_name param_prfx attrib insts prems eqns exp =
-  (Element.facts_map o Element.morph_ctxt)
-      (inst_morph thy phi_name param_prfx insts prems eqns exp)
-  #> Attrib.map_facts attrib;
-
-
-(* public interface to interpretation morphism *)
-
-fun get_interpret_morph thy phi_name param_prfx (exp, imp) target ext_ts =
-  let
-    val parms = the_locale thy target |> #params |> map fst;
-    val ids = flatten (ProofContext.init thy, intern_expr thy)
-      (([], Symtab.empty), Expr (Locale target)) |> fst |> fst;
-    val (insts, prems, eqns) = collect_witnesses (ProofContext.init thy) imp parms ids ext_ts;
-  in
-    inst_morph thy phi_name param_prfx insts prems eqns exp
-  end;
-
-(* store instantiations of args for all registered interpretations
-   of the theory *)
-
-fun note_thmss_registrations target (kind, args) thy =
-  let
-    val parms = the_locale thy target |> #params |> map fst;
-    val ids = flatten (ProofContext.init thy, intern_expr thy)
-      (([], Symtab.empty), Expr (Locale target)) |> fst |> fst;
-
-    val regs = get_global_registrations thy target;
-    (* add args to thy for all registrations *)
-
-    fun activate (ext_ts, ((phi_name, param_prfx), (exp, imp), _, _)) thy =
-      let
-        val (insts, prems, eqns) = collect_witnesses (ProofContext.init thy) imp parms ids ext_ts;
-        val args' = args
-          |> activate_note thy phi_name param_prfx
-               (Attrib.attribute_i thy) insts prems eqns exp;
-      in
-        thy
-        |> global_note_qualified kind args'
-        |> snd
-      end;
-  in fold activate regs thy end;
-
-
-(* locale results *)
+(* Theorems *)
 
 fun add_thmss loc kind args ctxt =
   let
-    val (([(_, [Notes args'])], _), ctxt') =
-      activate_facts true cert_facts
-        [((("", []), Assumed []), [Ext (Notes (kind, args))])] ctxt;
-    val ctxt'' = ctxt' |> ProofContext.theory
-      (change_locale loc
-        (fn (axiom, elems, params, decls, regs, intros, dests) =>
-          (axiom, elems @ [(Notes args', stamp ())],
-            params, decls, regs, intros, dests))
-      #> note_thmss_registrations loc args');
+    val ([Notes args'], ctxt') = Element.activate_i [Notes (kind, args)] ctxt;
+    val ctxt'' = ctxt' |> ProofContext.theory (
+      (change_locale loc o apfst o apsnd) (cons (args', stamp ()))
+        #>
+      (* Registrations *)
+      (fn thy => fold_rev (fn (name, morph) =>
+            let
+              val args'' = snd args' |> Element.facts_map (Element.morph_ctxt morph) |>
+                Attrib.map_facts (Attrib.attribute_i thy)
+            in Old_Locale.global_note_qualified kind args'' #> snd end)
+        (get_registrations thy |> filter (fn (name, _) => name = loc)) thy))
   in ctxt'' end;
 
 
-(* declarations *)
+(* Declarations *)
 
 local
 
 fun decl_attrib decl phi = Thm.declaration_attribute (K (decl phi));
 
 fun add_decls add loc decl =
-  ProofContext.theory (change_locale loc
-    (fn (axiom, elems, params, decls, regs, intros, dests) =>
-      (axiom, elems, params, add (decl, stamp ()) decls, regs, intros, dests))) #>
+  ProofContext.theory ((change_locale loc o apfst o apfst) (add (decl, stamp ())))
+    #>
   add_thmss loc Thm.internalK
     [((Binding.empty, [Attrib.internal (decl_attrib decl)]), [([Drule.dummy_thm], [])])];
 
@@ -1739,746 +468,41 @@
 val add_term_syntax = add_decls (apsnd o cons);
 val add_declaration = add_decls (K I);
 
-fun declarations_of thy loc =
-  the_locale thy loc |> #decls |> apfst (map fst) |> apsnd (map fst);
-
 end;
 
+(* Dependencies *)
+
+fun add_dependency loc dep = (change_locale loc o apsnd) (cons (dep, stamp ()));
 
 
-(** define locales **)
-
-(* predicate text *)
-(* CB: generate locale predicates and delta predicates *)
-
-local
-
-(* introN: name of theorems for introduction rules of locale and
-     delta predicates;
-   axiomsN: name of theorem set with destruct rules for locale predicates,
-     also name suffix of delta predicates. *)
-
-val introN = "intro";
-val axiomsN = "axioms";
-
-fun atomize_spec thy ts =
-  let
-    val t = Logic.mk_conjunction_balanced ts;
-    val body = ObjectLogic.atomize_term thy t;
-    val bodyT = Term.fastype_of body;
-  in
-    if bodyT = propT then (t, propT, Thm.reflexive (Thm.cterm_of thy t))
-    else (body, bodyT, ObjectLogic.atomize (Thm.cterm_of thy t))
-  end;
+(*** Reasoning about locales ***)
 
-fun aprop_tr' n c = (Syntax.constN ^ c, fn ctxt => fn args =>
-  if length args = n then
-    Syntax.const "_aprop" $
-      Term.list_comb (Syntax.free (Consts.extern (ProofContext.consts_of ctxt) c), args)
-  else raise Match);
-
-(* CB: define one predicate including its intro rule and axioms
-   - bname: predicate name
-   - parms: locale parameters
-   - defs: thms representing substitutions from defines elements
-   - ts: terms representing locale assumptions (not normalised wrt. defs)
-   - norm_ts: terms representing locale assumptions (normalised wrt. defs)
-   - thy: the theory
-*)
-
-fun def_pred bname parms defs ts norm_ts thy =
-  let
-    val name = Sign.full_bname thy bname;
-
-    val (body, bodyT, body_eq) = atomize_spec thy norm_ts;
-    val env = Term.add_term_free_names (body, []);
-    val xs = filter (member (op =) env o #1) parms;
-    val Ts = map #2 xs;
-    val extraTs = (Term.term_tfrees body \\ List.foldr Term.add_typ_tfrees [] Ts)
-      |> Library.sort_wrt #1 |> map TFree;
-    val predT = map Term.itselfT extraTs ---> Ts ---> bodyT;
-
-    val args = map Logic.mk_type extraTs @ map Free xs;
-    val head = Term.list_comb (Const (name, predT), args);
-    val statement = ObjectLogic.ensure_propT thy head;
+(** Storage for witnesses, intro and unfold rules **)
 
-    val ([pred_def], defs_thy) =
-      thy
-      |> bodyT = propT ? Sign.add_advanced_trfuns ([], [], [aprop_tr' (length args) name], [])
-      |> Sign.declare_const [] ((Binding.name bname, predT), NoSyn) |> snd
-      |> PureThy.add_defs false
-        [((Thm.def_name bname, Logic.mk_equals (head, body)), [Thm.kind_internal])];
-    val defs_ctxt = ProofContext.init defs_thy |> Variable.declare_term head;
-
-    val cert = Thm.cterm_of defs_thy;
-
-    val intro = Goal.prove_global defs_thy [] norm_ts statement (fn _ =>
-      MetaSimplifier.rewrite_goals_tac [pred_def] THEN
-      Tactic.compose_tac (false, body_eq RS Drule.equal_elim_rule1, 1) 1 THEN
-      Tactic.compose_tac (false,
-        Conjunction.intr_balanced (map (Thm.assume o cert) norm_ts), 0) 1);
-
-    val conjuncts =
-      (Drule.equal_elim_rule2 OF [body_eq,
-        MetaSimplifier.rewrite_rule [pred_def] (Thm.assume (cert statement))])
-      |> Conjunction.elim_balanced (length ts);
-    val axioms = ts ~~ conjuncts |> map (fn (t, ax) =>
-      Element.prove_witness defs_ctxt t
-       (MetaSimplifier.rewrite_goals_tac defs THEN
-        Tactic.compose_tac (false, ax, 0) 1));
-  in ((statement, intro, axioms), defs_thy) end;
-
-fun assumes_to_notes (Assumes asms) axms =
-      fold_map (fn (a, spec) => fn axs =>
-          let val (ps, qs) = chop (length spec) axs
-          in ((a, [(ps, [])]), qs) end) asms axms
-      |> apfst (curry Notes Thm.assumptionK)
-  | assumes_to_notes e axms = (e, axms);
-
-(* CB: the following two change only "new" elems, these have identifier ("", _). *)
-
-(* turn Assumes into Notes elements *)
-
-fun change_assumes_elemss axioms elemss =
-  let
-    val satisfy = Element.morph_ctxt (Element.satisfy_morphism axioms);
-    fun change (id as ("", _), es) =
-          fold_map assumes_to_notes (map satisfy es)
-          #-> (fn es' => pair (id, es'))
-      | change e = pair e;
-  in
-    fst (fold_map change elemss (map Element.conclude_witness axioms))
-  end;
-
-(* adjust hyps of Notes elements *)
-
-fun change_elemss_hyps axioms elemss =
-  let
-    val satisfy = Element.morph_ctxt (Element.satisfy_morphism axioms);
-    fun change (id as ("", _), es) = (id, map (fn e as Notes _ => satisfy e | e => e) es)
-      | change e = e;
-  in map change elemss end;
-
-in
-
-(* CB: main predicate definition function *)
+structure Witnesses = ThmsFun();
+structure Intros = ThmsFun();
+structure Unfolds = ThmsFun();
 
-fun define_preds pname (parms, ((exts, exts'), (ints, ints')), defs) elemss thy =
-  let
-    val ((elemss', more_ts), a_elem, a_intro, thy'') =
-      if null exts then ((elemss, []), [], [], thy)
-      else
-        let
-          val aname = if null ints then pname else pname ^ "_" ^ axiomsN;
-          val ((statement, intro, axioms), thy') =
-            thy
-            |> def_pred aname parms defs exts exts';
-          val elemss' = change_assumes_elemss axioms elemss;
-          val a_elem = [(("", []),
-            [Assumes [((Binding.name (pname ^ "_" ^ axiomsN), []), [(statement, [])])]])];
-          val (_, thy'') =
-            thy'
-            |> Sign.add_path aname
-            |> Sign.no_base_names
-            |> PureThy.note_thmss Thm.internalK [((Binding.name introN, []), [([intro], [])])]
-            ||> Sign.restore_naming thy';
-        in ((elemss', [statement]), a_elem, [intro], thy'') end;
-    val (predicate, stmt', elemss'', b_intro, thy'''') =
-      if null ints then (([], []), more_ts, elemss' @ a_elem, [], thy'')
-      else
-        let
-          val ((statement, intro, axioms), thy''') =
-            thy''
-            |> def_pred pname parms defs (ints @ more_ts) (ints' @ more_ts);
-          val cstatement = Thm.cterm_of thy''' statement;
-          val elemss'' = change_elemss_hyps axioms elemss';
-          val b_elem = [(("", []),
-               [Assumes [((Binding.name (pname ^ "_" ^ axiomsN), []), [(statement, [])])]])];
-          val (_, thy'''') =
-            thy'''
-            |> Sign.add_path pname
-            |> Sign.no_base_names
-            |> PureThy.note_thmss Thm.internalK
-                 [((Binding.name introN, []), [([intro], [])]),
-                  ((Binding.name axiomsN, []),
-                    [(map (Drule.standard o Element.conclude_witness) axioms, [])])]
-            ||> Sign.restore_naming thy''';
-        in (([cstatement], axioms), [statement], elemss'' @ b_elem, [intro], thy'''') end;
-  in (((elemss'', predicate, stmt'), (a_intro, b_intro)), thy'''') end;
-
-end;
-
-
-(* add_locale(_i) *)
-
-local
-
-(* turn Defines into Notes elements, accumulate definition terms *)
-
-fun defines_to_notes is_ext thy (Defines defs) defns =
-    let
-      val defs' = map (fn (_, (def, _)) => (Attrib.empty_binding, (def, []))) defs
-      val notes = map (fn (a, (def, _)) =>
-        (a, [([assume (cterm_of thy def)], [])])) defs
-    in
-      (if is_ext then SOME (Notes (Thm.definitionK, notes)) else NONE, defns @ [Defines defs'])
-    end
-  | defines_to_notes _ _ e defns = (SOME e, defns);
-
-fun change_defines_elemss thy elemss defns =
-  let
-    fun change (id as (n, _), es) defns =
-        let
-          val (es', defns') = fold_map (defines_to_notes (n="") thy) es defns
-        in ((id, map_filter I es'), defns') end
-  in fold_map change elemss defns end;
+val witness_attrib = Witnesses.add;
+val intro_attrib = Intros.add;
+val unfold_attrib = Unfolds.add;
 
-fun gen_add_locale prep_ctxt prep_expr
-    predicate_name bname raw_imports raw_body thy =
-    (* predicate_name: "" - locale with predicate named as locale
-        "name" - locale with predicate named "name" *)
-  let
-    val thy_ctxt = ProofContext.init thy;
-    val name = Sign.full_bname thy bname;
-    val _ = is_some (get_locale thy name) andalso
-      error ("Duplicate definition of locale " ^ quote name);
-
-    val (((import_ctxt, import_elemss), (body_ctxt, body_elemss, syn)),
-      text as (parms, ((_, exts'), _), defs)) =
-        prep_ctxt raw_imports raw_body thy_ctxt;
-    val elemss = import_elemss @ body_elemss |>
-      map_filter (fn ((id, Assumed axs), elems) => SOME (id, elems) | _ => NONE);
-
-    val extraTs = List.foldr Term.add_term_tfrees [] exts' \\
-      List.foldr Term.add_typ_tfrees [] (map snd parms);
-    val _ = if null extraTs then ()
-      else warning ("Additional type variable(s) in locale specification " ^ quote bname);
+(** Tactic **)
 
-    val predicate_name' = case predicate_name of "" => bname | _ => predicate_name;
-    val (elemss', defns) = change_defines_elemss thy elemss [];
-    val elemss'' = elemss' @ [(("", []), defns)];
-    val (((elemss''', predicate as (pred_statement, pred_axioms), stmt'), intros), thy') =
-      define_preds predicate_name' text elemss'' thy;
-    val regs = pred_axioms
-      |> fold_map (fn (id, elems) => fn wts => let
-             val ts = flat (map_filter (fn (Assumes asms) =>
-               SOME (maps (map #1 o #2) asms) | _ => NONE) elems);
-             val (wts1, wts2) = chop (length ts) wts;
-           in ((apsnd (map fst) id, wts1), wts2) end) elemss'''
-      |> fst
-      |> map_filter (fn (("", _), _) => NONE | e => SOME e);
-    fun axiomify axioms elemss =
-      (axioms, elemss) |> foldl_map (fn (axs, (id, elems)) => let
-                   val ts = flat (map_filter (fn (Assumes asms) =>
-                     SOME (maps (map #1 o #2) asms) | _ => NONE) elems);
-                   val (axs1, axs2) = chop (length ts) axs;
-                 in (axs2, ((id, Assumed axs1), elems)) end)
-      |> snd;
-    val ((_, facts), ctxt) = activate_facts true (K I)
-      (axiomify pred_axioms elemss''') (ProofContext.init thy');
-    val view_ctxt = Assumption.add_view thy_ctxt pred_statement ctxt;
-    val export = Thm.close_derivation o Goal.norm_result o
-      singleton (ProofContext.export view_ctxt thy_ctxt);
-    val facts' = facts |> map (fn (a, ths) => ((a, []), [(map export ths, [])]));
-    val elems' = maps #2 (filter (fn ((s, _), _) => s = "") elemss''');
-    val elems'' = map_filter (fn (Fixes _) => NONE | e => SOME e) elems';
-    val axs' = map (Element.assume_witness thy') stmt';
-    val loc_ctxt = thy'
-      |> Sign.add_path bname
-      |> Sign.no_base_names
-      |> PureThy.note_thmss Thm.assumptionK facts' |> snd
-      |> Sign.restore_naming thy'
-      |> register_locale bname {axiom = axs',
-        elems = map (fn e => (e, stamp ())) elems'',
-        params = params_of elemss''' |> map (fn (x, SOME T) => ((x, T), the (Symtab.lookup syn x))),
-        decls = ([], []),
-        regs = regs,
-        intros = intros,
-        dests = map Element.conclude_witness pred_axioms}
-      |> init name;
-  in (name, loc_ctxt) end;
-
-in
-
-val add_locale = gen_add_locale cert_context (K I);
-val add_locale_cmd = gen_add_locale read_context intern_expr "";
-
-end;
+fun intro_locales_tac eager ctxt facts st =
+  Method.intros_tac
+    (Witnesses.get ctxt @ Intros.get ctxt @ (if eager then Unfolds.get ctxt else [])) facts st;
 
 val _ = Context.>> (Context.map_theory
- (add_locale "" "var" empty [Fixes [(Binding.name (Name.internal "x"), NONE, NoSyn)]] #>
-  snd #> ProofContext.theory_of #>
-  add_locale "" "struct" empty [Fixes [(Binding.name (Name.internal "S"), NONE, Structure)]] #>
-  snd #> ProofContext.theory_of));
-
-
-
-
-(** Normalisation of locale statements ---
-    discharges goals implied by interpretations **)
-
-local
-
-fun locale_assm_intros thy =
-  Symtab.fold (fn (_, {intros = (a, _), ...}) => fn intros => (a @ intros))
-    (#2 (LocalesData.get thy)) [];
-fun locale_base_intros thy =
-  Symtab.fold (fn (_, {intros = (_, b), ...}) => fn intros => (b @ intros))
-    (#2 (LocalesData.get thy)) [];
-
-fun all_witnesses ctxt =
-  let
-    val thy = ProofContext.theory_of ctxt;
-    fun get registrations = Symtab.fold (fn (_, regs) => fn thms =>
-        (Registrations.dest thy regs |> map (fn (_, (_, (exp, _), wits, _)) =>
-          map (Element.conclude_witness #> Morphism.thm exp) wits) |> flat) @ thms)
-      registrations [];
-  in get (RegistrationsData.get (Context.Proof ctxt)) end;
-
-in
-
-fun intro_locales_tac eager ctxt facts st =
-  let
-    val wits = all_witnesses ctxt;
-    val thy = ProofContext.theory_of ctxt;
-    val intros = locale_base_intros thy @ (if eager then locale_assm_intros thy else []);
-  in
-    Method.intros_tac (wits @ intros) facts st
-  end;
+  (Method.add_methods
+    [("intro_locales",
+      Method.ctxt_args (fn ctxt => Method.METHOD (intro_locales_tac false ctxt ORELSE'
+        Old_Locale.intro_locales_tac false ctxt)),
+      "back-chain introduction rules of locales without unfolding predicates"),
+     ("unfold_locales",
+      Method.ctxt_args (fn ctxt => Method.METHOD (intro_locales_tac true ctxt ORELSE'
+        Old_Locale.intro_locales_tac true ctxt)),
+      "back-chain all introduction rules of locales")]));
 
 end;
 
-
-(** Interpretation commands **)
-
-local
-
-(* extract proof obligations (assms and defs) from elements *)
-
-fun extract_asms_elems ((id, Assumed _), elems) = (id, maps Element.prems_of elems)
-  | extract_asms_elems ((id, Derived _), _) = (id, []);
-
-
-(* activate instantiated facts in theory or context *)
-
-fun gen_activate_facts_elemss mk_ctxt note attrib put_reg add_wit add_eqn
-        phi_name all_elemss pss propss eq_attns (exp, imp) thmss thy_ctxt =
-  let
-    val ctxt = mk_ctxt thy_ctxt;
-    fun get_reg thy_ctxt = get_local_registration (mk_ctxt thy_ctxt);
-    fun test_reg thy_ctxt = test_local_registration (mk_ctxt thy_ctxt);
-
-    val (all_propss, eq_props) = chop (length all_elemss) propss;
-    val (all_thmss, eq_thms) = chop (length all_elemss) thmss;
-
-    (* Filter out fragments already registered. *)
-
-    val (new_elemss, xs) = split_list (filter_out (fn (((id, _), _), _) =>
-          test_reg thy_ctxt id) (all_elemss ~~ (pss ~~ (all_propss ~~ all_thmss))));
-    val (new_pss, ys) = split_list xs;
-    val (new_propss, new_thmss) = split_list ys;
-
-    val thy_ctxt' = thy_ctxt
-      (* add registrations *)
-      |> fold2 (fn ((id as (loc, _), _), _) => fn ps => put_reg id (phi_name, param_prefix loc ps) (exp, imp))
-           new_elemss new_pss
-      (* add witnesses of Assumed elements (only those generate proof obligations) *)
-      |> fold2 (fn (id, _) => fold (add_wit id)) new_propss new_thmss
-      (* add equations *)
-      |> fold2 (fn (id, _) => fold (add_eqn id)) eq_props
-          ((map o map) (Drule.abs_def o LocalDefs.meta_rewrite_rule ctxt o
-            Element.conclude_witness) eq_thms);
-
-    val prems = flat (map_filter
-          (fn ((id, Assumed _), _) => Option.map #2 (get_reg thy_ctxt' imp id)
-            | ((_, Derived _), _) => NONE) all_elemss);
-
-    val thy_ctxt'' = thy_ctxt'
-      (* add witnesses of Derived elements *)
-      |> fold (fn (id, thms) => fold
-           (add_wit id o Element.morph_witness (Element.satisfy_morphism prems)) thms)
-         (map_filter (fn ((_, Assumed _), _) => NONE
-            | ((id, Derived thms), _) => SOME (id, thms)) new_elemss)
-
-    fun activate_elem phi_name param_prfx insts prems eqns exp (Notes (kind, facts)) thy_ctxt =
-        let
-          val ctxt = mk_ctxt thy_ctxt;
-          val thy = ProofContext.theory_of ctxt;
-          val facts' = facts
-            |> activate_note thy phi_name param_prfx
-                 (attrib thy_ctxt) insts prems eqns exp;
-        in 
-          thy_ctxt
-          |> note kind facts'
-          |> snd
-        end
-      | activate_elem _ _ _ _ _ _ _ thy_ctxt = thy_ctxt;
-
-    fun activate_elems (((loc, ext_ts), _), _) ps thy_ctxt =
-      let
-        val ctxt = mk_ctxt thy_ctxt;
-        val thy = ProofContext.theory_of ctxt;
-        val {params, elems, ...} = the_locale thy loc;
-        val parms = map fst params;
-        val param_prfx = param_prefix loc ps;
-        val ids = flatten (ProofContext.init thy, intern_expr thy)
-          (([], Symtab.empty), Expr (Locale loc)) |> fst |> fst;
-        val (insts, prems, eqns) = collect_witnesses ctxt imp parms ids ext_ts;
-      in
-        thy_ctxt
-        |> fold (activate_elem phi_name param_prfx insts prems eqns exp o fst) elems
-      end;
-
-  in
-    thy_ctxt''
-    (* add equations as lemmas to context *)
-    |> (fold2 o fold2) (fn attn => fn thm => snd o yield_singleton (note Thm.lemmaK)
-         ((apsnd o map) (attrib thy_ctxt'') attn, [([Element.conclude_witness thm], [])]))
-            (unflat eq_thms eq_attns) eq_thms
-    (* add interpreted facts *)
-    |> fold2 activate_elems new_elemss new_pss
-  end;
-
-fun global_activate_facts_elemss x = gen_activate_facts_elemss
-  ProofContext.init
-  global_note_qualified
-  Attrib.attribute_i
-  put_global_registration
-  add_global_witness
-  add_global_equation
-  x;
-
-fun local_activate_facts_elemss x = gen_activate_facts_elemss
-  I
-  local_note_qualified
-  (Attrib.attribute_i o ProofContext.theory_of)
-  put_local_registration
-  add_local_witness
-  add_local_equation
-  x;
-
-fun prep_instantiations parse_term parse_prop ctxt parms (insts, eqns) =
-  let
-    (* parameters *)
-    val (parm_names, parm_types) = parms |> split_list
-      ||> map (TypeInfer.paramify_vars o Logic.varifyT);
-    val type_parms = fold Term.add_tvarsT parm_types [] |> map (Logic.mk_type o TVar);
-    val type_parm_names = fold Term.add_tfreesT (map snd parms) [] |> map fst;
-
-    (* parameter instantiations *)
-    val d = length parms - length insts;
-    val insts =
-      if d < 0 then error "More arguments than parameters in instantiation."
-      else insts @ replicate d NONE;
-    val (given_ps, given_insts) =
-      ((parm_names ~~ parm_types) ~~ insts) |> map_filter
-          (fn (_, NONE) => NONE
-            | ((n, T), SOME inst) => SOME ((n, T), inst))
-        |> split_list;
-    val (given_parm_names, given_parm_types) = given_ps |> split_list;
-
-    (* parse insts / eqns *)
-    val given_insts' = map (parse_term ctxt) given_insts;
-    val eqns' = map (parse_prop ctxt) eqns;
-
-    (* type inference and contexts *)
-    val arg = type_parms @ map2 TypeInfer.constrain given_parm_types given_insts' @ eqns';
-    val res = Syntax.check_terms ctxt arg;
-    val ctxt' = ctxt |> fold Variable.auto_fixes res;
-
-    (* instantiation *)
-    val (type_parms'', res') = chop (length type_parms) res;
-    val (given_insts'', eqns'') = chop (length given_insts) res';
-    val instT = Symtab.make (type_parm_names ~~ map Logic.dest_type type_parms'');
-    val inst = Symtab.make (given_parm_names ~~ given_insts'');
-
-    (* export from eigencontext *)
-    val export = Variable.export_morphism ctxt' ctxt;
-
-    (* import, its inverse *)
-    val domT = fold Term.add_tfrees res [] |> map TFree;
-    val importT = domT |> map (fn x => (Morphism.typ export x, x))
-      |> map_filter (fn (TFree _, _) => NONE  (* fixed point of export *)
-               | (TVar y, x) => SOME (fst y, x)
-               | _ => error "internal: illegal export in interpretation")
-      |> Vartab.make;
-    val dom = fold Term.add_frees res [] |> map Free;
-    val imprt = dom |> map (fn x => (Morphism.term export x, x))
-      |> map_filter (fn (Free _, _) => NONE  (* fixed point of export *)
-               | (Var y, x) => SOME (fst y, x)
-               | _ => error "internal: illegal export in interpretation")
-      |> Vartab.make;
-  in (((instT, inst), eqns''), (export, ((importT, domT), (imprt, dom)))) end;
-
-val read_instantiations = prep_instantiations Syntax.parse_term Syntax.parse_prop;
-val check_instantiations = prep_instantiations (K I) (K I);
-
-fun gen_prep_registration mk_ctxt test_reg activate
-    prep_attr prep_expr prep_insts
-    thy_ctxt phi_name raw_expr raw_insts =
-  let
-    val ctxt = mk_ctxt thy_ctxt;
-    val thy = ProofContext.theory_of ctxt;
-    val ctxt' = ProofContext.init thy;
-    fun prep_attn attn = (apsnd o map)
-      (Attrib.crude_closure ctxt o Args.assignable o prep_attr thy) attn;
-
-    val expr = prep_expr thy raw_expr;
-
-    val pts = params_of_expr ctxt' [] expr ([], Symtab.empty, Symtab.empty);
-    val params_ids = make_params_ids (#1 pts);
-    val raw_params_elemss = make_raw_params_elemss pts;
-    val ((ids, _), raw_elemss) = flatten (ctxt', I) (([], Symtab.empty), Expr expr);
-    val ((parms, all_elemss, _), (_, (_, defs, _))) =
-      read_elemss false ctxt' [] (raw_params_elemss @ raw_elemss) [];
-
-    (** compute instantiation **)
-
-    (* consistency check: equations need to be stored in a particular locale,
-       therefore if equations are present locale expression must be a name *)
-
-    val _ = case (expr, snd raw_insts) of
-        (Locale _, _) => () | (_, []) => ()
-      | (_, _) => error "Interpretations with `where' only permitted if locale expression is a name.";
-
-    (* read or certify instantiation *)
-    val (raw_insts', raw_eqns) = raw_insts;
-    val (raw_eq_attns, raw_eqns') = split_list raw_eqns;
-    val (((instT, inst1), eqns), morphs) = prep_insts ctxt parms (raw_insts', raw_eqns');
-    val eq_attns = map prep_attn raw_eq_attns;
-
-    (* defined params without given instantiation *)
-    val not_given = filter_out (Symtab.defined inst1 o fst) parms;
-    fun add_def (p, pT) inst =
-      let
-        val (t, T) = case find_first (fn (Free (a, _), _) => a = p) defs of
-               NONE => error ("Instance missing for parameter " ^ quote p)
-             | SOME (Free (_, T), t) => (t, T);
-        val d = Element.inst_term (instT, inst) t;
-      in Symtab.update_new (p, d) inst end;
-    val inst2 = fold add_def not_given inst1;
-    val inst_morphism = Element.inst_morphism thy (instT, inst2);
-    (* Note: insts contain no vars. *)
-
-    (** compute proof obligations **)
-
-    (* restore "small" ids *)
-    val ids' = map (fn ((n, ps), (_, mode)) =>
-          ((n, map (fn p => Free (p, (the o AList.lookup (op =) parms) p)) ps), mode))
-        ids;
-    val (_, all_elemss') = chop (length raw_params_elemss) all_elemss
-    (* instantiate ids and elements *)
-    val inst_elemss = (ids' ~~ all_elemss') |> map (fn (((n, ps), _), ((_, mode), elems)) =>
-      ((n, map (Morphism.term (inst_morphism $> fst morphs)) ps),
-        map (fn Int e => Element.morph_ctxt inst_morphism e) elems)
-      |> apfst (fn id => (id, map_mode (map (Element.morph_witness inst_morphism)) mode)));
-
-    (* equations *)
-    val eqn_elems = if null eqns then []
-      else [(Library.last_elem inst_elemss |> fst |> fst, eqns)];
-
-    val propss = map extract_asms_elems inst_elemss @ eqn_elems;
-
-  in
-    (propss, activate phi_name inst_elemss (map (snd o fst) ids) propss eq_attns morphs, morphs)
-  end;
-
-fun gen_prep_global_registration mk_ctxt = gen_prep_registration ProofContext.init
-  test_global_registration
-  global_activate_facts_elemss mk_ctxt;
-
-fun gen_prep_local_registration mk_ctxt = gen_prep_registration I
-  test_local_registration
-  local_activate_facts_elemss mk_ctxt;
-
-val prep_global_registration = gen_prep_global_registration
-  (K I) (K I) check_instantiations;
-val prep_global_registration_cmd = gen_prep_global_registration
-  Attrib.intern_src intern_expr read_instantiations;
-
-val prep_local_registration = gen_prep_local_registration
-  (K I) (K I) check_instantiations;
-val prep_local_registration_cmd = gen_prep_local_registration
-  Attrib.intern_src intern_expr read_instantiations;
-
-fun prep_registration_in_locale target expr thy =
-  (* target already in internal form *)
-  let
-    val ctxt = ProofContext.init thy;
-    val ((raw_target_ids, target_syn), _) = flatten (ctxt, I)
-        (([], Symtab.empty), Expr (Locale target));
-    val fixed = the_locale thy target |> #params |> map #1;
-    val ((all_ids, syn), raw_elemss) = flatten (ctxt, intern_expr thy)
-        ((raw_target_ids, target_syn), Expr expr);
-    val (target_ids, ids) = chop (length raw_target_ids) all_ids;
-    val ((parms, elemss, _), _) = read_elemss false ctxt fixed raw_elemss [];
-
-    (** compute proof obligations **)
-
-    (* restore "small" ids, with mode *)
-    val ids' = map (apsnd snd) ids;
-    (* remove Int markers *)
-    val elemss' = map (fn (_, es) =>
-        map (fn Int e => e) es) elemss
-    (* extract assumptions and defs *)
-    val ids_elemss = ids' ~~ elemss';
-    val propss = map extract_asms_elems ids_elemss;
-
-    (** activation function:
-        - add registrations to the target locale
-        - add induced registrations for all global registrations of
-          the target, unless already present
-        - add facts of induced registrations to theory **)
-
-    fun activate thmss thy =
-      let
-        val satisfy = Element.satisfy_thm (flat thmss);
-        val ids_elemss_thmss = ids_elemss ~~ thmss;
-        val regs = get_global_registrations thy target;
-
-        fun activate_id (((id, Assumed _), _), thms) thy =
-            thy |> put_registration_in_locale target id
-                |> fold (add_witness_in_locale target id) thms
-          | activate_id _ thy = thy;
-
-        fun activate_reg (ext_ts, ((phi_name, param_prfx), (exp, imp), _, _)) thy =
-          let
-            val (insts, wits, _) = collect_witnesses (ProofContext.init thy) imp fixed target_ids ext_ts;
-            val inst_parms = map (the o AList.lookup (op =) (map #1 fixed ~~ ext_ts));
-            val disch = Element.satisfy_thm wits;
-            val new_elemss = filter (fn (((name, ps), _), _) =>
-                not (test_global_registration thy (name, inst_parms ps))) (ids_elemss);
-            fun activate_assumed_id (((_, Derived _), _), _) thy = thy
-              | activate_assumed_id ((((name, ps), Assumed _), _), thms) thy = let
-                val ps' = inst_parms ps;
-              in
-                if test_global_registration thy (name, ps')
-                then thy
-                else thy
-                  |> put_global_registration (name, ps') (phi_name, param_prefix name ps) (exp, imp)
-                  |> fold (fn witn => fn thy => add_global_witness (name, ps')
-                     (Element.morph_witness (Element.inst_morphism thy insts) witn) thy) thms
-              end;
-
-            fun activate_derived_id ((_, Assumed _), _) thy = thy
-              | activate_derived_id (((name, ps), Derived ths), _) thy = let
-                val ps' = inst_parms ps;
-              in
-                if test_global_registration thy (name, ps')
-                then thy
-                else thy
-                  |> put_global_registration (name, ps') (phi_name, param_prefix name ps) (exp, imp)
-                  |> fold (fn witn => fn thy => add_global_witness (name, ps')
-                       (witn |> Element.map_witness (fn (t, th) =>  (* FIXME *)
-                       (Element.inst_term insts t,
-                        disch (Element.inst_thm thy insts (satisfy th))))) thy) ths
-              end;
-
-            fun activate_elem (loc, ps) (Notes (kind, facts)) thy =
-                let
-                  val att_morphism =
-                    Morphism.binding_morphism (name_morph phi_name param_prfx) $>
-                    Morphism.thm_morphism satisfy $>
-                    Element.inst_morphism thy insts $>
-                    Morphism.thm_morphism disch;
-                  val facts' = facts
-                    |> Attrib.map_facts (Attrib.attribute_i thy o Args.morph_values att_morphism)
-                    |> (map o apsnd o map o apfst o map) (disch o Element.inst_thm thy insts o satisfy)
-                    |> (map o apfst o apfst) (name_morph phi_name param_prfx);
-                in
-                  thy
-                  |> global_note_qualified kind facts'
-                  |> snd
-                end
-              | activate_elem _ _ thy = thy;
-
-            fun activate_elems ((id, _), elems) thy = fold (activate_elem id) elems thy;
-
-          in thy |> fold activate_assumed_id ids_elemss_thmss
-                 |> fold activate_derived_id ids_elemss
-                 |> fold activate_elems new_elemss end;
-      in
-        thy |> fold activate_id ids_elemss_thmss
-            |> fold activate_reg regs
-      end;
-
-  in (propss, activate) end;
-
-fun prep_propp propss = propss |> map (fn (_, props) =>
-  map (rpair [] o Element.mark_witness) props);
-
-fun prep_result propps thmss =
-  ListPair.map (fn ((_, props), thms) => map2 Element.make_witness props thms) (propps, thmss);
-
-fun gen_interpretation prep_registration after_qed prfx raw_expr raw_insts thy =
-  let
-    val (propss, activate, morphs) = prep_registration thy prfx raw_expr raw_insts;
-    fun after_qed' results =
-      ProofContext.theory (activate (prep_result propss results))
-      #> after_qed;
-  in
-    thy
-    |> ProofContext.init
-    |> Proof.theorem_i NONE after_qed' (prep_propp propss)
-    |> Element.refine_witness
-    |> Seq.hd
-    |> pair morphs
-  end;
-
-fun gen_interpret prep_registration after_qed name_morph expr insts int state =
-  let
-    val _ = Proof.assert_forward_or_chain state;
-    val ctxt = Proof.context_of state;
-    val (propss, activate, morphs) = prep_registration ctxt name_morph expr insts;
-    fun after_qed' results =
-      Proof.map_context (K (ctxt |> activate (prep_result propss results)))
-      #> Proof.put_facts NONE
-      #> after_qed;
-  in
-    state
-    |> Proof.local_goal (ProofDisplay.print_results int) (K I) ProofContext.bind_propp_i
-      "interpret" NONE after_qed' (map (pair (Binding.empty, [])) (prep_propp propss))
-    |> Element.refine_witness |> Seq.hd
-    |> pair morphs
-  end;
-
-fun standard_name_morph interp_prfx b =
-  if Binding.is_empty b then b
-  else Binding.map_prefix (fn ((lprfx, _) :: pprfx) =>
-    fold (Binding.add_prefix false o fst) pprfx
-    #> interp_prfx <> "" ? Binding.add_prefix true interp_prfx
-    #> Binding.add_prefix false lprfx
-  ) b;
-
-in
-
-val interpretation = gen_interpretation prep_global_registration;
-fun interpretation_cmd interp_prfx = snd ooo gen_interpretation prep_global_registration_cmd
-  I (standard_name_morph interp_prfx);
-
-fun interpretation_in_locale after_qed (raw_target, expr) thy =
-  let
-    val target = intern thy raw_target;
-    val (propss, activate) = prep_registration_in_locale target expr thy;
-    val raw_propp = prep_propp propss;
-
-    val (_, _, goal_ctxt, propp) = thy
-      |> ProofContext.init
-      |> cert_context_statement (SOME target) [] raw_propp;
-
-    fun after_qed' results =
-      ProofContext.theory (activate (prep_result propss results))
-      #> after_qed;
-  in
-    goal_ctxt
-    |> Proof.theorem_i NONE after_qed' propp
-    |> Element.refine_witness |> Seq.hd
-  end;
-
-val interpret = gen_interpret prep_local_registration;
-fun interpret_cmd interp_prfx = snd oooo gen_interpret prep_local_registration_cmd
-  Seq.single (standard_name_morph interp_prfx);
-
-end;
-
-end;
--- a/src/Pure/Isar/method.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Isar/method.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/method.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Isar proof methods.
@@ -223,15 +222,9 @@
   if cond (Logic.strip_assums_concl prop)
   then Tactic.rtac rule i else no_tac);
 
-fun legacy_tac st =
-  (legacy_feature
-      ("Implicit use of prems in assumption proof" ^ Position.str_of (Position.thread_data ()));
-    all_tac st);
-
 fun assm_tac ctxt =
   assume_tac APPEND'
   Goal.assume_rule_tac ctxt APPEND'
-  (CHANGED o solve_tac (Assumption.prems_of ctxt) THEN' K legacy_tac) APPEND'
   cond_rtac (can Logic.dest_equals) Drule.reflexive_thm APPEND'
   cond_rtac (can Logic.dest_term) Drule.termI;
 
--- a/src/Pure/Isar/new_locale.ML	Mon Jan 19 20:37:08 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,508 +0,0 @@
-(*  Title:      Pure/Isar/new_locale.ML
-    Author:     Clemens Ballarin, TU Muenchen
-
-New locale development --- experimental.
-*)
-
-signature NEW_LOCALE =
-sig
-  type locale
-
-  val test_locale: theory -> string -> bool
-  val register_locale: bstring ->
-    (string * sort) list * (Binding.T * typ option * mixfix) list ->
-    term option * term list ->
-    (declaration * stamp) list * (declaration * stamp) list ->
-    ((string * (Attrib.binding * (thm list * Attrib.src list) list) list) * stamp) list ->
-    ((string * Morphism.morphism) * stamp) list -> theory -> theory
-
-  (* Locale name space *)
-  val intern: theory -> xstring -> string
-  val extern: theory -> string -> xstring
-
-  (* Specification *)
-  val params_of: theory -> string -> (Binding.T * typ option * mixfix) list
-  val instance_of: theory -> string -> Morphism.morphism -> term list
-  val specification_of: theory -> string -> term option * term list
-  val declarations_of: theory -> string -> declaration list * declaration list
-
-  (* Storing results *)
-  val add_thmss: string -> string -> (Attrib.binding * (thm list * Attrib.src list) list) list ->
-    Proof.context -> Proof.context
-  val add_type_syntax: string -> declaration -> Proof.context -> Proof.context
-  val add_term_syntax: string -> declaration -> Proof.context -> Proof.context
-  val add_declaration: string -> declaration -> Proof.context -> Proof.context
-  val add_dependency: string -> (string * Morphism.morphism) -> theory -> theory
-
-  (* Activation *)
-  val activate_declarations: theory -> string * Morphism.morphism ->
-    Proof.context -> Proof.context
-  val activate_global_facts: string * Morphism.morphism -> theory -> theory
-  val activate_local_facts: string * Morphism.morphism -> Proof.context -> Proof.context
-  val init: string -> theory -> Proof.context
-
-  (* Reasoning about locales *)
-  val witness_attrib: attribute
-  val intro_attrib: attribute
-  val unfold_attrib: attribute
-  val intro_locales_tac: bool -> Proof.context -> thm list -> tactic
-
-  (* Registrations *)
-  val add_global_registration: (string * (Morphism.morphism * Morphism.morphism)) ->
-    theory -> theory
-  val amend_global_registration: Morphism.morphism -> (string * Morphism.morphism) ->
-    theory -> theory
-  val get_global_registrations: theory -> (string * Morphism.morphism) list
-
-  (* Diagnostic *)
-  val print_locales: theory -> unit
-  val print_locale: theory -> bool -> bstring -> unit
-end;
-
-
-(*** Theorem list extensible via attribute --- to control intro_locales_tac ***)
-
-(* FIXME: consider moving named_thms.ML up in the build hierarchy and using that *)
-functor ThmsFun() =
-struct
-
-structure Data = GenericDataFun
-(
-  type T = thm list;
-  val empty = [];
-  val extend = I;
-  fun merge _ = Thm.merge_thms;
-);
-
-val get = Data.get o Context.Proof;
-val add = Thm.declaration_attribute (Data.map o Thm.add_thm);
-
-end;
-
-
-structure NewLocale: NEW_LOCALE =
-struct
-
-datatype ctxt = datatype Element.ctxt;
-
-
-(*** Basics ***)
-
-datatype locale = Loc of {
-  (* extensible lists are in reverse order: decls, notes, dependencies *)
-  parameters: (string * sort) list * (Binding.T * typ option * mixfix) list,
-    (* type and term parameters *)
-  spec: term option * term list,
-    (* assumptions (as a single predicate expression) and defines *)
-  decls: (declaration * stamp) list * (declaration * stamp) list,
-    (* type and term syntax declarations *)
-  notes: ((string * (Attrib.binding * (thm list * Attrib.src list) list) list) * stamp) list,
-    (* theorem declarations *)
-  dependencies: ((string * Morphism.morphism) * stamp) list
-    (* locale dependencies (sublocale relation) *)
-}
-
-
-(*** Theory data ***)
-
-structure LocalesData = TheoryDataFun
-(
-  type T = NameSpace.T * locale Symtab.table;
-    (* locale namespace and locales of the theory *)
-
-  val empty = NameSpace.empty_table;
-  val copy = I;
-  val extend = I;
-
-  fun join_locales _
-    (Loc {parameters, spec, decls = (decls1, decls2), notes, dependencies},
-      Loc {decls = (decls1', decls2'), notes = notes',
-        dependencies = dependencies', ...}) =
-        let fun s_merge x = merge (eq_snd (op =)) x in
-          Loc {parameters = parameters,
-            spec = spec,
-            decls = (s_merge (decls1, decls1'), s_merge (decls2, decls2')),
-            notes = s_merge (notes, notes'),
-            dependencies = s_merge (dependencies, dependencies')
-          }          
-        end;
-  fun merge _ = NameSpace.join_tables join_locales;
-);
-
-val intern = NameSpace.intern o #1 o LocalesData.get;
-val extern = NameSpace.extern o #1 o LocalesData.get;
-
-fun get_locale thy name = Symtab.lookup (#2 (LocalesData.get thy)) name;
-
-fun the_locale thy name = case get_locale thy name
- of SOME loc => loc
-  | NONE => error ("Unknown locale " ^ quote name);
-
-fun test_locale thy name = case get_locale thy name
- of SOME _ => true | NONE => false;
-
-fun register_locale bname parameters spec decls notes dependencies thy =
-  thy |> LocalesData.map (NameSpace.bind (Sign.naming_of thy) (Binding.name bname,
-    Loc {parameters = parameters, spec = spec, decls = decls, notes = notes,
-      dependencies = dependencies}) #> snd);
-
-fun change_locale name f thy =
-  let
-    val Loc {parameters, spec, decls, notes, dependencies} =
-        the_locale thy name;
-    val (parameters', spec', decls', notes', dependencies') =
-      f (parameters, spec, decls, notes, dependencies);
-  in
-    thy
-    |> (LocalesData.map o apsnd) (Symtab.update (name, Loc {parameters = parameters',
-      spec = spec', decls = decls', notes = notes', dependencies = dependencies'}))
-  end;
-
-fun print_locales thy =
-  let val (space, locs) = LocalesData.get thy in
-    Pretty.strs ("locales:" :: map #1 (NameSpace.extern_table (space, locs)))
-    |> Pretty.writeln
-  end;
-
-
-(*** Primitive operations ***)
-
-fun params_of thy name =
-  let
-    val Loc {parameters = (_, params), ...} = the_locale thy name
-  in params end;
-
-fun instance_of thy name morph =
-  params_of thy name |>
-    map ((fn (b, T, _) => Free (Binding.base_name b, the T)) #> Morphism.term morph);
-
-fun specification_of thy name =
-  let
-    val Loc {spec, ...} = the_locale thy name
-  in spec end;
-
-fun declarations_of thy name =
-  let
-    val Loc {decls, ...} = the_locale thy name
-  in
-    decls |> apfst (map fst) |> apsnd (map fst)
-  end;
-
-
-(*** Activate context elements of locale ***)
-
-(** Identifiers: activated locales in theory or proof context **)
-
-type identifiers = (string * term list) list;
-
-val empty = ([] : identifiers);
-
-fun ident_eq thy ((n, ts), (m, ss)) = (m = n) andalso Pattern.matchess thy (ss, ts);
-
-local
-
-datatype 'a delayed = Ready of 'a | ToDo of ('a delayed * 'a delayed);
-
-structure IdentifiersData = GenericDataFun
-(
-  type T = identifiers delayed;
-  val empty = Ready empty;
-  val extend = I;
-  fun merge _ = ToDo;
-);
-
-in
-
-fun finish thy (ToDo (i1, i2)) = merge (ident_eq thy) (finish thy i1, finish thy i2)
-  | finish _ (Ready ids) = ids;
-
-val _ = Context.>> (Context.map_theory (Theory.at_begin (fn thy =>
-  (case IdentifiersData.get (Context.Theory thy) of
-    Ready _ => NONE |
-    ids => SOME (Context.theory_map (IdentifiersData.put (Ready (finish thy ids))) thy))
-  )));
-  
-fun get_global_idents thy =
-  let val (Ready ids) = (IdentifiersData.get o Context.Theory) thy in ids end;
-val put_global_idents = Context.theory_map o IdentifiersData.put o Ready;
-
-fun get_local_idents ctxt =
-  let val (Ready ids) = (IdentifiersData.get o Context.Proof) ctxt in ids end;
-val put_local_idents = Context.proof_map o IdentifiersData.put o Ready;
-
-end;
-
-
-(** Resolve locale dependencies in a depth-first fashion **)
-
-local
-
-val roundup_bound = 120;
-
-fun add thy depth (name, morph) (deps, marked) =
-  if depth > roundup_bound
-  then error "Roundup bound exceeded (sublocale relation probably not terminating)."
-  else
-    let
-      val Loc {parameters = (_, params), dependencies, ...} = the_locale thy name;
-      val instance = instance_of thy name morph;
-    in
-      if member (ident_eq thy) marked (name, instance)
-      then (deps, marked)
-      else
-        let
-          val dependencies' =
-            map (fn ((name, morph'), _) => (name, morph' $>  morph)) dependencies;
-          val marked' = (name, instance) :: marked;
-          val (deps', marked'') = fold_rev (add thy (depth + 1)) dependencies' ([], marked');
-        in
-          ((name, morph) :: deps' @ deps, marked'')
-        end
-    end;
-
-in
-
-fun roundup thy activate_dep (name, morph) (marked, input) =
-  let
-    (* Find all dependencies incuding new ones (which are dependencies enriching
-      existing registrations). *)
-    val (dependencies, marked') = add thy 0 (name, morph) ([], empty);
-    (* Filter out exisiting fragments. *)
-    val dependencies' = filter_out (fn (name, morph) =>
-      member (ident_eq thy) marked (name, instance_of thy name morph)) dependencies;
-  in
-    (merge (ident_eq thy) (marked, marked'), input |> fold_rev (activate_dep thy) dependencies')
-  end;
-
-end;
-
-
-(* Declarations, facts and entire locale content *)
-
-fun activate_decls thy (name, morph) ctxt =
-  let
-    val Loc {decls = (typ_decls, term_decls), ...} = the_locale thy name;
-  in
-    ctxt |> fold_rev (fn (decl, _) => Context.proof_map (decl morph)) typ_decls |>
-      fold_rev (fn (decl, _) => Context.proof_map (decl morph)) term_decls
-  end;
-
-fun activate_notes activ_elem transfer thy (name, morph) input =
-  let
-    val Loc {notes, ...} = the_locale thy name;
-    fun activate ((kind, facts), _) input =
-      let
-        val facts' = facts |> Element.facts_map (Element.morph_ctxt (transfer input $> morph))
-      in activ_elem (Notes (kind, facts')) input end;
-  in
-    fold_rev activate notes input
-  end;
-
-fun activate_all name thy activ_elem transfer (marked, input) =
-  let
-    val Loc {parameters = (_, params), spec = (asm, defs), ...} =
-      the_locale thy name;
-  in
-    input |>
-      (if not (null params) then activ_elem (Fixes params) else I) |>
-      (* FIXME type parameters *)
-      (if is_some asm then activ_elem (Assumes [(Attrib.empty_binding, [(the asm, [])])]) else I) |>
-      (if not (null defs)
-        then activ_elem (Defines (map (fn def => (Attrib.empty_binding, (def, []))) defs))
-        else I) |>
-      pair marked |> roundup thy (activate_notes activ_elem transfer) (name, Morphism.identity)
-  end;
-
-
-(** Public activation functions **)
-
-local
-
-fun init_global_elem (Notes (kind, facts)) thy =
-      let
-        val facts' = Attrib.map_facts (Attrib.attribute_i thy) facts
-      in Locale.global_note_qualified kind facts' thy |> snd end
-
-fun init_local_elem (Fixes fixes) ctxt = ctxt |>
-      ProofContext.add_fixes_i fixes |> snd
-  | init_local_elem (Assumes assms) ctxt =
-      let
-        val assms' = Attrib.map_specs (Attrib.attribute_i (ProofContext.theory_of ctxt)) assms
-      in
-        ctxt |> fold Variable.auto_fixes (maps (map fst o snd) assms') |>
-          ProofContext.add_assms_i Assumption.assume_export assms' |> snd
-     end 
-  | init_local_elem (Defines defs) ctxt =
-      let
-        val defs' = Attrib.map_specs (Attrib.attribute_i (ProofContext.theory_of ctxt)) defs
-      in
-        ctxt |> fold Variable.auto_fixes (map (fst o snd) defs') |>
-          ProofContext.add_assms_i LocalDefs.def_export (map (fn (attn, t) => (attn, [t])) defs') |>
-          snd
-      end
-  | init_local_elem (Notes (kind, facts)) ctxt =
-      let
-        val facts' = Attrib.map_facts (Attrib.attribute_i (ProofContext.theory_of ctxt)) facts
-      in Locale.local_note_qualified kind facts' ctxt |> snd end
-
-fun cons_elem false (Notes notes) elems = elems
-  | cons_elem _ elem elems = elem :: elems
-
-in
-
-fun activate_declarations thy dep ctxt =
-  roundup thy activate_decls dep (get_local_idents ctxt, ctxt) |> uncurry put_local_idents;
-
-fun activate_global_facts dep thy =
-  roundup thy (activate_notes init_global_elem Element.transfer_morphism)
-    dep (get_global_idents thy, thy) |>
-  uncurry put_global_idents;
-
-fun activate_local_facts dep ctxt =
-  roundup (ProofContext.theory_of ctxt)
-  (activate_notes init_local_elem (Element.transfer_morphism o ProofContext.theory_of)) dep
-    (get_local_idents ctxt, ctxt) |>
-  uncurry put_local_idents;
-
-fun init name thy =
-  activate_all name thy init_local_elem (Element.transfer_morphism o ProofContext.theory_of)
-    (empty, ProofContext.init thy) |>
-  uncurry put_local_idents;
-
-fun print_locale thy show_facts name =
-  let
-    val name' = intern thy name;
-    val ctxt = init name' thy
-  in
-    Pretty.big_list "locale elements:"
-      (activate_all name' thy (cons_elem show_facts) (K (Element.transfer_morphism thy))
-        (empty, []) |> snd |> rev |> 
-        map (Element.pretty_ctxt ctxt) |> map Pretty.chunks) |> Pretty.writeln
-  end
-
-end;
-
-
-(*** Registrations: interpretations in theories ***)
-
-(* FIXME only global variant needed *)
-structure RegistrationsData = GenericDataFun
-(
-  type T = ((string * (Morphism.morphism * Morphism.morphism)) * stamp) list;
-(* FIXME mixins need to be stamped *)
-    (* registrations, in reverse order of declaration *)
-  val empty = [];
-  val extend = I;
-  fun merge _ = Library.merge (eq_snd (op =));
-    (* FIXME consolidate with dependencies, consider one data slot only *)
-);
-
-val get_global_registrations =
-  Context.Theory #> RegistrationsData.get #> map fst #> map (apsnd op $>);
-
-fun add_global reg =
-  (Context.theory_map o RegistrationsData.map) (cons (reg, stamp ()));
-
-fun add_global_registration (name, (base_morph, export)) thy =
-  roundup thy (fn _ => fn (name', morph') => fn thy => add_global (name', (morph', export)) thy)
-    (name, base_morph) (get_global_idents thy, thy) |>
-    snd (* FIXME ?? uncurry put_global_idents *);
-
-fun amend_global_registration morph (name, base_morph) thy =
-  let
-    val regs = (Context.Theory #> RegistrationsData.get #> map fst) thy;
-    val base = instance_of thy name base_morph;
-    fun match (name', (morph', _)) =
-      name = name' andalso eq_list (op aconv) (base, instance_of thy name' morph');
-    val i = find_index match (rev regs);
-    val _ = if i = ~1 then error ("No interpretation of locale " ^
-        quote (extern thy name) ^ " and parameter instantiation " ^
-        space_implode " " (map (quote o Syntax.string_of_term_global thy) base) ^ " available.")
-      else ();
-  in
-    (Context.theory_map o RegistrationsData.map) (nth_map (length regs - 1 - i)
-      (fn ((name, (base, export)), stamp) => ((name, (base $> morph, export)), stamp))) thy
-  end;
-
-
-(*** Storing results ***)
-
-(* Theorems *)
-
-fun add_thmss loc kind args ctxt =
-  let
-    val ([Notes args'], ctxt') = Element.activate_i [Notes (kind, args)] ctxt;
-    val ctxt'' = ctxt' |> ProofContext.theory (
-      change_locale loc
-        (fn (parameters, spec, decls, notes, dependencies) =>
-          (parameters, spec, decls, (args', stamp ()) :: notes, dependencies)) #>
-      (* Registrations *)
-      (fn thy => fold_rev (fn (name, morph) =>
-            let
-              val args'' = snd args' |> Element.facts_map (Element.morph_ctxt morph) |>
-                Attrib.map_facts (Attrib.attribute_i thy)
-            in Locale.global_note_qualified kind args'' #> snd end)
-        (get_global_registrations thy |> filter (fn (name, _) => name = loc)) thy))
-  in ctxt'' end;
-
-
-(* Declarations *)
-
-local
-
-fun decl_attrib decl phi = Thm.declaration_attribute (K (decl phi));
-
-fun add_decls add loc decl =
-  ProofContext.theory (change_locale loc
-    (fn (parameters, spec, decls, notes, dependencies) =>
-      (parameters, spec, add (decl, stamp ()) decls, notes, dependencies))) #>
-  add_thmss loc Thm.internalK
-    [((Binding.empty, [Attrib.internal (decl_attrib decl)]), [([Drule.dummy_thm], [])])];
-
-in
-
-val add_type_syntax = add_decls (apfst o cons);
-val add_term_syntax = add_decls (apsnd o cons);
-val add_declaration = add_decls (K I);
-
-end;
-
-(* Dependencies *)
-
-fun add_dependency loc dep =
-  change_locale loc
-    (fn (parameters, spec, decls, notes, dependencies) =>
-      (parameters, spec, decls, notes, (dep, stamp ()) :: dependencies));
-
-
-(*** Reasoning about locales ***)
-
-(** Storage for witnesses, intro and unfold rules **)
-
-structure Witnesses = ThmsFun();
-structure Intros = ThmsFun();
-structure Unfolds = ThmsFun();
-
-val witness_attrib = Witnesses.add;
-val intro_attrib = Intros.add;
-val unfold_attrib = Unfolds.add;
-
-(** Tactic **)
-
-fun intro_locales_tac eager ctxt facts st =
-  Method.intros_tac
-    (Witnesses.get ctxt @ Intros.get ctxt @ (if eager then Unfolds.get ctxt else [])) facts st;
-
-val _ = Context.>> (Context.map_theory
-  (Method.add_methods
-    [("intro_locales",
-      Method.ctxt_args (fn ctxt => Method.METHOD (intro_locales_tac false ctxt ORELSE'
-        Locale.intro_locales_tac false ctxt)),
-      "back-chain introduction rules of locales without unfolding predicates"),
-     ("unfold_locales",
-      Method.ctxt_args (fn ctxt => Method.METHOD (intro_locales_tac true ctxt ORELSE'
-        Locale.intro_locales_tac true ctxt)),
-      "back-chain all introduction rules of locales")]));
-
-end;
-
--- a/src/Pure/Isar/obtain.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Isar/obtain.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/obtain.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 The 'obtain' and 'guess' language elements -- generalized existence at
@@ -150,13 +149,13 @@
 
     fun after_qed _ =
       Proof.local_qed (NONE, false)
-      #> Seq.map (`Proof.the_fact #-> (fn rule =>
+      #> `Proof.the_fact #-> (fn rule =>
         Proof.fix_i vars
-        #> Proof.assm_i (obtain_export fix_ctxt rule (map (cert o Free) parms)) asms));
+        #> Proof.assm_i (obtain_export fix_ctxt rule (map (cert o Free) parms)) asms);
   in
     state
     |> Proof.enter_forward
-    |> Proof.have_i NONE (K Seq.single) [((Binding.empty, []), [(obtain_prop, [])])] int
+    |> Proof.have_i NONE (K I) [((Binding.empty, []), [(obtain_prop, [])])] int
     |> Proof.proof (SOME Method.succeed_text) |> Seq.hd
     |> Proof.fix_i [(Binding.name thesisN, NONE, NoSyn)]
     |> Proof.assume_i
@@ -306,7 +305,7 @@
     val before_qed = SOME (Method.primitive_text (Goal.conclude #> MetaSimplifier.norm_hhf #>
         (fn th => Goal.protect (Conjunction.intr (Drule.mk_term (Thm.cprop_of th)) th))));
     fun after_qed [[_, res]] =
-      Proof.end_block #> guess_context (check_result ctxt thesis res) #> Seq.single;
+      Proof.end_block #> guess_context (check_result ctxt thesis res);
   in
     state
     |> Proof.enter_forward
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Isar/old_locale.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -0,0 +1,2485 @@
+(*  Title:      Pure/Isar/locale.ML
+    Author:     Clemens Ballarin, TU Muenchen
+    Author:     Markus Wenzel, LMU/TU Muenchen
+
+Locales -- Isar proof contexts as meta-level predicates, with local
+syntax and implicit structures.
+
+Draws basic ideas from Florian Kammueller's original version of
+locales, but uses the richer infrastructure of Isar instead of the raw
+meta-logic.  Furthermore, structured import of contexts (with merge
+and rename operations) are provided, as well as type-inference of the
+signature parts, and predicate definitions of the specification text.
+
+Interpretation enables the reuse of theorems of locales in other
+contexts, namely those defined by theories, structured proofs and
+locales themselves.
+
+See also:
+
+[1] Clemens Ballarin. Locales and Locale Expressions in Isabelle/Isar.
+    In Stefano Berardi et al., Types for Proofs and Programs: International
+    Workshop, TYPES 2003, Torino, Italy, LNCS 3085, pages 34-50, 2004.
+[2] Clemens Ballarin. Interpretation of Locales in Isabelle: Managing
+    Dependencies between Locales. Technical Report TUM-I0607, Technische
+    Universitaet Muenchen, 2006.
+[3] Clemens Ballarin. Interpretation of Locales in Isabelle: Theories and
+    Proof Contexts. In J.M. Borwein and W.M. Farmer, MKM 2006, LNAI 4108,
+    pages 31-43, 2006.
+*)
+
+(* TODO:
+- beta-eta normalisation of interpretation parameters
+- dangling type frees in locales
+- test subsumption of interpretations when merging theories
+*)
+
+signature OLD_LOCALE =
+sig
+  datatype expr =
+    Locale of string |
+    Rename of expr * (string * mixfix option) option list |
+    Merge of expr list
+  val empty: expr
+
+  val intern: theory -> xstring -> string
+  val intern_expr: theory -> expr -> expr
+  val extern: theory -> string -> xstring
+  val init: string -> theory -> Proof.context
+
+  (* The specification of a locale *)
+  val parameters_of: theory -> string -> ((string * typ) * mixfix) list
+  val parameters_of_expr: theory -> expr -> ((string * typ) * mixfix) list
+  val local_asms_of: theory -> string -> (Attrib.binding * term list) list
+  val global_asms_of: theory -> string -> (Attrib.binding * term list) list
+
+  (* Theorems *)
+  val intros: theory -> string -> thm list * thm list
+  val dests: theory -> string -> thm list
+  (* Not part of the official interface.  DO NOT USE *)
+  val facts_of: theory -> string -> (Attrib.binding * (thm list * Attrib.src list) list) list list
+
+  (* Not part of the official interface.  DO NOT USE *)
+  val declarations_of: theory -> string -> declaration list * declaration list;
+
+  (* Processing of locale statements *)
+  val read_context_statement: string option -> Element.context list ->
+    (string * string list) list list -> Proof.context ->
+    string option * Proof.context * Proof.context * (term * term list) list list
+  val read_context_statement_cmd: xstring option -> Element.context list ->
+    (string * string list) list list -> Proof.context ->
+    string option * Proof.context * Proof.context * (term * term list) list list
+  val cert_context_statement: string option -> Element.context_i list ->
+    (term * term list) list list -> Proof.context ->
+    string option * Proof.context * Proof.context * (term * term list) list list
+  val read_expr: expr -> Element.context list -> Proof.context ->
+    Element.context_i list * Proof.context
+  val cert_expr: expr -> Element.context_i list -> Proof.context ->
+    Element.context_i list * Proof.context
+
+  (* Diagnostic functions *)
+  val print_locales: theory -> unit
+  val print_locale: theory -> bool -> expr -> Element.context list -> unit
+  val print_registrations: bool -> string -> Proof.context -> unit
+
+  val add_locale: string -> bstring -> expr -> Element.context_i list -> theory
+    -> string * Proof.context
+  val add_locale_cmd: bstring -> expr -> Element.context list -> theory
+    -> string * Proof.context
+
+  (* Tactics *)
+  val intro_locales_tac: bool -> Proof.context -> thm list -> tactic
+
+  (* Storing results *)
+  val global_note_qualified: string ->
+    ((Binding.T * attribute list) * (thm list * attribute list) list) list ->
+    theory -> (string * thm list) list * theory
+  val local_note_qualified: string ->
+    ((Binding.T * attribute list) * (thm list * attribute list) list) list ->
+    Proof.context -> (string * thm list) list * Proof.context
+  val add_thmss: string -> string -> (Attrib.binding * (thm list * Attrib.src list) list) list ->
+    Proof.context -> Proof.context
+  val add_type_syntax: string -> declaration -> Proof.context -> Proof.context
+  val add_term_syntax: string -> declaration -> Proof.context -> Proof.context
+  val add_declaration: string -> declaration -> Proof.context -> Proof.context
+
+  (* Interpretation *)
+  val get_interpret_morph: theory -> (Binding.T -> Binding.T) -> string * string ->
+    (Morphism.morphism * ((typ Vartab.table * typ list) * (term Vartab.table * term list))) ->
+    string -> term list -> Morphism.morphism
+  val interpretation: (Proof.context -> Proof.context) ->
+    (Binding.T -> Binding.T) -> expr ->
+    term option list * (Attrib.binding * term) list ->
+    theory ->
+    (Morphism.morphism * ((typ Vartab.table * typ list) * (term Vartab.table * term list))) * Proof.state
+  val interpretation_cmd: string -> expr -> string option list * (Attrib.binding * string) list ->
+    theory -> Proof.state
+  val interpretation_in_locale: (Proof.context -> Proof.context) ->
+    xstring * expr -> theory -> Proof.state
+  val interpret: (Proof.state -> Proof.state) ->
+    (Binding.T -> Binding.T) -> expr ->
+    term option list * (Attrib.binding * term) list ->
+    bool -> Proof.state ->
+    (Morphism.morphism * ((typ Vartab.table * typ list) * (term Vartab.table * term list))) * Proof.state
+  val interpret_cmd: string -> expr -> string option list * (Attrib.binding * string) list ->
+    bool -> Proof.state -> Proof.state
+end;
+
+structure Old_Locale: OLD_LOCALE =
+struct
+
+(* legacy operations *)
+
+fun merge_lists _ xs [] = xs
+  | merge_lists _ [] ys = ys
+  | merge_lists eq xs ys = xs @ filter_out (member eq xs) ys;
+
+fun merge_alists eq xs = merge_lists (eq_fst eq) xs;
+
+
+(* auxiliary: noting name bindings with qualified base names *)
+
+fun global_note_qualified kind facts thy =
+  thy
+  |> Sign.qualified_names
+  |> PureThy.note_thmss kind facts
+  ||> Sign.restore_naming thy;
+
+fun local_note_qualified kind facts ctxt =
+  ctxt
+  |> ProofContext.qualified_names
+  |> ProofContext.note_thmss_i kind facts
+  ||> ProofContext.restore_naming ctxt;
+
+
+(** locale elements and expressions **)
+
+datatype ctxt = datatype Element.ctxt;
+
+datatype expr =
+  Locale of string |
+  Rename of expr * (string * mixfix option) option list |
+  Merge of expr list;
+
+val empty = Merge [];
+
+datatype 'a element =
+  Elem of 'a | Expr of expr;
+
+fun map_elem f (Elem e) = Elem (f e)
+  | map_elem _ (Expr e) = Expr e;
+
+type decl = declaration * stamp;
+
+type locale =
+ {axiom: Element.witness list,
+    (* For locales that define predicates this is [A [A]], where A is the locale
+       specification.  Otherwise [].
+       Only required to generate the right witnesses for locales with predicates. *)
+  elems: (Element.context_i * stamp) list,
+    (* Static content, neither Fixes nor Constrains elements *)
+  params: ((string * typ) * mixfix) list,                        (*all term params*)
+  decls: decl list * decl list,                    (*type/term_syntax declarations*)
+  regs: ((string * string list) * Element.witness list) list,
+    (* Registrations: indentifiers and witnesses of locales interpreted in the locale. *)
+  intros: thm list * thm list,
+    (* Introduction rules: of delta predicate and locale predicate. *)
+  dests: thm list}
+    (* Destruction rules: projections from locale predicate to predicates of fragments. *)
+
+(* CB: an internal (Int) locale element was either imported or included,
+   an external (Ext) element appears directly in the locale text. *)
+
+datatype ('a, 'b) int_ext = Int of 'a | Ext of 'b;
+
+
+
+(** substitutions on Vars -- clone from element.ML **)
+
+(* instantiate types *)
+
+fun var_instT_type env =
+  if Vartab.is_empty env then I
+  else Term.map_type_tvar (fn (x, S) => the_default (TVar (x, S)) (Vartab.lookup env x));
+
+fun var_instT_term env =
+  if Vartab.is_empty env then I
+  else Term.map_types (var_instT_type env);
+
+fun var_inst_term (envT, env) =
+  if Vartab.is_empty env then var_instT_term envT
+  else
+    let
+      val instT = var_instT_type envT;
+      fun inst (Const (x, T)) = Const (x, instT T)
+        | inst (Free (x, T)) = Free(x, instT T)
+        | inst (Var (xi, T)) =
+            (case Vartab.lookup env xi of
+              NONE => Var (xi, instT T)
+            | SOME t => t)
+        | inst (b as Bound _) = b
+        | inst (Abs (x, T, t)) = Abs (x, instT T, inst t)
+        | inst (t $ u) = inst t $ inst u;
+    in Envir.beta_norm o inst end;
+
+
+(** management of registrations in theories and proof contexts **)
+
+type registration =
+  {prfx: (Binding.T -> Binding.T) * (string * string),
+      (* first component: interpretation name morphism;
+         second component: parameter prefix *)
+    exp: Morphism.morphism,
+      (* maps content to its originating context *)
+    imp: (typ Vartab.table * typ list) * (term Vartab.table * term list),
+      (* inverse of exp *)
+    wits: Element.witness list,
+      (* witnesses of the registration *)
+    eqns: thm Termtab.table,
+      (* theorems (equations) interpreting derived concepts and indexed by lhs *)
+    morph: unit
+      (* interpreting morphism *)
+  }
+
+structure Registrations :
+  sig
+    type T
+    val empty: T
+    val join: T * T -> T
+    val dest: theory -> T ->
+      (term list *
+        (((Binding.T -> Binding.T) * (string * string)) *
+         (Morphism.morphism * ((typ Vartab.table * typ list) * (term Vartab.table * term list))) *
+         Element.witness list *
+         thm Termtab.table)) list
+    val test: theory -> T * term list -> bool
+    val lookup: theory ->
+      T * (term list * ((typ Vartab.table * typ list) * (term Vartab.table * term list))) ->
+      (((Binding.T -> Binding.T) * (string * string)) * Element.witness list * thm Termtab.table) option
+    val insert: theory -> term list -> ((Binding.T -> Binding.T) * (string * string)) ->
+      (Morphism.morphism * ((typ Vartab.table * typ list) * (term Vartab.table * term list))) ->
+      T ->
+      T * (term list * (((Binding.T -> Binding.T) * (string * string)) * Element.witness list)) list
+    val add_witness: term list -> Element.witness -> T -> T
+    val add_equation: term list -> thm -> T -> T
+(*
+    val update_morph: term list -> Morphism.morphism -> T -> T
+    val get_morph: theory -> T ->
+      term list * ((typ Vartab.table * typ list) * (term Vartab.table * term list)) ->
+      Morphism.morphism
+*)
+  end =
+struct
+  (* A registration is indexed by parameter instantiation.
+     NB: index is exported whereas content is internalised. *)
+  type T = registration Termtab.table;
+
+  fun mk_reg prfx exp imp wits eqns morph =
+    {prfx = prfx, exp = exp, imp = imp, wits = wits, eqns = eqns, morph = morph};
+
+  fun map_reg f reg =
+    let
+      val {prfx, exp, imp, wits, eqns, morph} = reg;
+      val (prfx', exp', imp', wits', eqns', morph') = f (prfx, exp, imp, wits, eqns, morph);
+    in mk_reg prfx' exp' imp' wits' eqns' morph' end;
+
+  val empty = Termtab.empty;
+
+  (* term list represented as single term, for simultaneous matching *)
+  fun termify ts =
+    Term.list_comb (Const ("", map fastype_of ts ---> propT), ts);
+  fun untermify t =
+    let fun ut (Const _) ts = ts
+          | ut (s $ t) ts = ut s (t::ts)
+    in ut t [] end;
+
+  (* joining of registrations:
+     - prefix and morphisms of right theory;
+     - witnesses are equal, no attempt to subsumption testing;
+     - union of equalities, if conflicting (i.e. two eqns with equal lhs)
+       eqn of right theory takes precedence *)
+  fun join (r1, r2) = Termtab.join (fn _ => fn ({eqns = e1, ...}, {prfx = n, exp, imp, wits = w, eqns = e2, morph = m}) =>
+      mk_reg n exp imp w (Termtab.join (fn _ => fn (_, e) => e) (e1, e2)) m) (r1, r2);
+
+  fun dest_transfer thy regs =
+    Termtab.dest regs |> map (apsnd (map_reg (fn (n, e, i, ws, es, m) =>
+      (n, e, i, map (Element.transfer_witness thy) ws, Termtab.map (transfer thy) es, m))));
+
+  fun dest thy regs = dest_transfer thy regs |> map (apfst untermify) |>
+    map (apsnd (fn {prfx, exp, imp, wits, eqns, ...} => (prfx, (exp, imp), wits, eqns)));
+
+  (* registrations that subsume t *)
+  fun subsumers thy t regs =
+    filter (fn (t', _) => Pattern.matches thy (t', t)) (dest_transfer thy regs);
+
+  (* test if registration that subsumes the query is present *)
+  fun test thy (regs, ts) =
+    not (null (subsumers thy (termify ts) regs));
+      
+  (* look up registration, pick one that subsumes the query *)
+  fun lookup thy (regs, (ts, ((impT, _), (imp, _)))) =
+    let
+      val t = termify ts;
+      val subs = subsumers thy t regs;
+    in
+      (case subs of
+        [] => NONE
+        | ((t', {prfx, exp = exp', imp = ((impT', domT'), (imp', dom')), wits, eqns, morph}) :: _) =>
+          let
+            val (tinst, inst) = Pattern.match thy (t', t) (Vartab.empty, Vartab.empty);
+            val tinst' = domT' |> map (fn (T as TFree (x, _)) =>
+                (x, T |> Morphism.typ exp' |> Envir.typ_subst_TVars tinst
+                      |> var_instT_type impT)) |> Symtab.make;
+            val inst' = dom' |> map (fn (t as Free (x, _)) =>
+                (x, t |> Morphism.term exp' |> Envir.subst_vars (tinst, inst)
+                      |> var_inst_term (impT, imp))) |> Symtab.make;
+            val inst'_morph = Element.inst_morphism thy (tinst', inst');
+          in SOME (prfx,
+            map (Element.morph_witness inst'_morph) wits,
+            Termtab.map (Morphism.thm inst'_morph) eqns)
+          end)
+    end;
+
+  (* add registration if not subsumed by ones already present,
+     additionally returns registrations that are strictly subsumed *)
+  fun insert thy ts prfx (exp, imp) regs =
+    let
+      val t = termify ts;
+      val subs = subsumers thy t regs ;
+    in (case subs of
+        [] => let
+                val sups =
+                  filter (fn (t', _) => Pattern.matches thy (t, t')) (dest_transfer thy regs);
+                val sups' = map (apfst untermify) sups |> map (fn (ts, {prfx, wits, ...}) => (ts, (prfx, wits)))
+              in (Termtab.update (t, mk_reg prfx exp imp [] Termtab.empty ()) regs, sups') end
+      | _ => (regs, []))
+    end;
+
+  fun gen_add f ts regs =
+    let
+      val t = termify ts;
+    in
+      Termtab.update (t, map_reg f (the (Termtab.lookup regs t))) regs
+    end;
+
+  (* add witness theorem to registration,
+     only if instantiation is exact, otherwise exception Option raised *)
+  fun add_witness ts wit regs =
+    gen_add (fn (x, e, i, wits, eqns, m) => (x, e, i, Element.close_witness wit :: wits, eqns, m))
+      ts regs;
+
+  (* add equation to registration, replaces previous equation with same lhs;
+     only if instantiation is exact, otherwise exception Option raised;
+     exception TERM raised if not a meta equality *)
+  fun add_equation ts thm regs =
+    gen_add (fn (x, e, i, thms, eqns, m) =>
+      (x, e, i, thms, Termtab.update (thm |> prop_of |> Logic.dest_equals |> fst, Thm.close_derivation thm) eqns, m))
+      ts regs;
+
+end;
+
+
+(** theory data : locales **)
+
+structure LocalesData = TheoryDataFun
+(
+  type T = NameSpace.T * locale Symtab.table;
+    (* 1st entry: locale namespace,
+       2nd entry: locales of the theory *)
+
+  val empty = NameSpace.empty_table;
+  val copy = I;
+  val extend = I;
+
+  fun join_locales _
+    ({axiom, elems, params, decls = (decls1, decls2), regs, intros, dests}: locale,
+      {elems = elems', decls = (decls1', decls2'), regs = regs', ...}: locale) =
+     {axiom = axiom,
+      elems = merge_lists (eq_snd (op =)) elems elems',
+      params = params,
+      decls =
+       (Library.merge (eq_snd (op =)) (decls1, decls1'),
+        Library.merge (eq_snd (op =)) (decls2, decls2')),
+      regs = merge_alists (op =) regs regs',
+      intros = intros,
+      dests = dests};
+  fun merge _ = NameSpace.join_tables join_locales;
+);
+
+
+
+(** context data : registrations **)
+
+structure RegistrationsData = GenericDataFun
+(
+  type T = Registrations.T Symtab.table;  (*registrations, indexed by locale name*)
+  val empty = Symtab.empty;
+  val extend = I;
+  fun merge _ = Symtab.join (K Registrations.join);
+);
+
+
+(** access locales **)
+
+val intern = NameSpace.intern o #1 o LocalesData.get;
+val extern = NameSpace.extern o #1 o LocalesData.get;
+
+fun get_locale thy name = Symtab.lookup (#2 (LocalesData.get thy)) name;
+
+fun the_locale thy name = case get_locale thy name
+ of SOME loc => loc
+  | NONE => error ("Unknown locale " ^ quote name);
+
+fun register_locale bname loc thy =
+  thy |> LocalesData.map (NameSpace.bind (Sign.naming_of thy)
+    (Binding.name bname, loc) #> snd);
+
+fun change_locale name f thy =
+  let
+    val {axiom, elems, params, decls, regs, intros, dests} =
+        the_locale thy name;
+    val (axiom', elems', params', decls', regs', intros', dests') =
+      f (axiom, elems, params, decls, regs, intros, dests);
+  in
+    thy
+    |> (LocalesData.map o apsnd) (Symtab.update (name, {axiom = axiom',
+          elems = elems', params = params',
+          decls = decls', regs = regs', intros = intros', dests = dests'}))
+  end;
+
+fun print_locales thy =
+  let val (space, locs) = LocalesData.get thy in
+    Pretty.strs ("locales:" :: map #1 (NameSpace.extern_table (space, locs)))
+    |> Pretty.writeln
+  end;
+
+
+(* access registrations *)
+
+(* retrieve registration from theory or context *)
+
+fun get_registrations ctxt name =
+  case Symtab.lookup (RegistrationsData.get ctxt) name of
+      NONE => []
+    | SOME reg => Registrations.dest (Context.theory_of ctxt) reg;
+
+fun get_global_registrations thy = get_registrations (Context.Theory thy);
+fun get_local_registrations ctxt = get_registrations (Context.Proof ctxt);
+
+
+fun get_registration ctxt imprt (name, ps) =
+  case Symtab.lookup (RegistrationsData.get ctxt) name of
+      NONE => NONE
+    | SOME reg => Registrations.lookup (Context.theory_of ctxt) (reg, (ps, imprt));
+
+fun get_global_registration thy = get_registration (Context.Theory thy);
+fun get_local_registration ctxt = get_registration (Context.Proof ctxt);
+
+
+fun test_registration ctxt (name, ps) =
+  case Symtab.lookup (RegistrationsData.get ctxt) name of
+      NONE => false
+    | SOME reg => Registrations.test (Context.theory_of ctxt) (reg, ps);
+
+fun test_global_registration thy = test_registration (Context.Theory thy);
+fun test_local_registration ctxt = test_registration (Context.Proof ctxt);
+
+
+(* add registration to theory or context, ignored if subsumed *)
+
+fun put_registration (name, ps) prfx morphs ctxt =
+  RegistrationsData.map (fn regs =>
+    let
+      val thy = Context.theory_of ctxt;
+      val reg = the_default Registrations.empty (Symtab.lookup regs name);
+      val (reg', sups) = Registrations.insert thy ps prfx morphs reg;
+      val _ = if not (null sups) then warning
+                ("Subsumed interpretation(s) of locale " ^
+                 quote (extern thy name) ^
+                 "\nwith the following prefix(es):" ^
+                  commas_quote (map (fn (_, ((_, (_, s)), _)) => s) sups))
+              else ();
+    in Symtab.update (name, reg') regs end) ctxt;
+
+fun put_global_registration id prfx morphs =
+  Context.theory_map (put_registration id prfx morphs);
+fun put_local_registration id prfx morphs =
+  Context.proof_map (put_registration id prfx morphs);
+
+fun put_registration_in_locale name id =
+  change_locale name (fn (axiom, elems, params, decls, regs, intros, dests) =>
+    (axiom, elems, params, decls, regs @ [(id, [])], intros, dests));
+
+
+(* add witness theorem to registration, ignored if registration not present *)
+
+fun add_witness (name, ps) thm =
+  RegistrationsData.map (Symtab.map_entry name (Registrations.add_witness ps thm));
+
+fun add_global_witness id thm = Context.theory_map (add_witness id thm);
+fun add_local_witness id thm = Context.proof_map (add_witness id thm);
+
+
+fun add_witness_in_locale name id thm =
+  change_locale name (fn (axiom, elems, params, decls, regs, intros, dests) =>
+    let
+      fun add (id', thms) =
+        if id = id' then (id', thm :: thms) else (id', thms);
+    in (axiom, elems, params, decls, map add regs, intros, dests) end);
+
+
+(* add equation to registration, ignored if registration not present *)
+
+fun add_equation (name, ps) thm =
+  RegistrationsData.map (Symtab.map_entry name (Registrations.add_equation ps thm));
+
+fun add_global_equation id thm = Context.theory_map (add_equation id thm);
+fun add_local_equation id thm = Context.proof_map (add_equation id thm);
+
+(*
+(* update morphism of registration, ignored if registration not present *)
+
+fun update_morph (name, ps) morph =
+  RegistrationsData.map (Symtab.map_entry name (Registrations.update_morph ps morph));
+
+fun update_global_morph id morph = Context.theory_map (update_morph id morph);
+fun update_local_morph id morph = Context.proof_map (update_morph id morph);
+*)
+
+
+(* printing of registrations *)
+
+fun print_registrations show_wits loc ctxt =
+  let
+    val thy = ProofContext.theory_of ctxt;
+    val prt_term = Pretty.quote o Syntax.pretty_term ctxt;
+    fun prt_term' t = if !show_types
+          then Pretty.block [prt_term t, Pretty.brk 1, Pretty.str "::",
+            Pretty.brk 1, (Pretty.quote o Syntax.pretty_typ ctxt) (type_of t)]
+          else prt_term t;
+    val prt_thm = prt_term o prop_of;
+    fun prt_inst ts =
+        Pretty.enclose "(" ")" (Pretty.breaks (map prt_term' ts));
+    fun prt_prfx ((false, prfx), param_prfx) = [Pretty.str prfx, Pretty.brk 1, Pretty.str "(optional)", Pretty.brk 1, Pretty.str param_prfx]
+      | prt_prfx ((true, prfx), param_prfx) = [Pretty.str prfx, Pretty.brk 1, Pretty.str param_prfx];
+    fun prt_eqns [] = Pretty.str "no equations."
+      | prt_eqns eqns = Pretty.block (Pretty.str "equations:" :: Pretty.brk 1 ::
+          Pretty.breaks (map prt_thm eqns));
+    fun prt_core ts eqns =
+          [prt_inst ts, Pretty.fbrk, prt_eqns (Termtab.dest eqns |> map snd)];
+    fun prt_witns [] = Pretty.str "no witnesses."
+      | prt_witns witns = Pretty.block (Pretty.str "witnesses:" :: Pretty.brk 1 ::
+          Pretty.breaks (map (Element.pretty_witness ctxt) witns))
+    fun prt_reg (ts, (_, _, witns, eqns)) =
+        if show_wits
+          then Pretty.block (prt_core ts eqns @ [Pretty.fbrk, prt_witns witns])
+          else Pretty.block (prt_core ts eqns)
+
+    val loc_int = intern thy loc;
+    val regs = RegistrationsData.get (Context.Proof ctxt);
+    val loc_regs = Symtab.lookup regs loc_int;
+  in
+    (case loc_regs of
+        NONE => Pretty.str ("no interpretations")
+      | SOME r => let
+            val r' = Registrations.dest thy r;
+            val r'' = Library.sort_wrt (fn (_, ((_, (_, prfx)), _, _, _)) => prfx) r';
+          in Pretty.big_list ("interpretations:") (map prt_reg r'') end)
+    |> Pretty.writeln
+  end;
+
+
+(* diagnostics *)
+
+fun err_in_locale ctxt msg ids =
+  let
+    val thy = ProofContext.theory_of ctxt;
+    fun prt_id (name, parms) =
+      [Pretty.block (Pretty.breaks (map Pretty.str (extern thy name :: parms)))];
+    val prt_ids = flat (separate [Pretty.str " +", Pretty.brk 1] (map prt_id ids));
+    val err_msg =
+      if forall (fn (s, _) => s = "") ids then msg
+      else msg ^ "\n" ^ Pretty.string_of (Pretty.block
+        (Pretty.str "The error(s) above occurred in locale:" :: Pretty.brk 1 :: prt_ids));
+  in error err_msg end;
+
+fun err_in_locale' ctxt msg ids' = err_in_locale ctxt msg (map fst ids');
+
+
+fun pretty_ren NONE = Pretty.str "_"
+  | pretty_ren (SOME (x, NONE)) = Pretty.str x
+  | pretty_ren (SOME (x, SOME syn)) =
+      Pretty.block [Pretty.str x, Pretty.brk 1, Syntax.pretty_mixfix syn];
+
+fun pretty_expr thy (Locale name) = Pretty.str (extern thy name)
+  | pretty_expr thy (Rename (expr, xs)) =
+      Pretty.block [pretty_expr thy expr, Pretty.brk 1, Pretty.block (map pretty_ren xs |> Pretty.breaks)]
+  | pretty_expr thy (Merge es) =
+      Pretty.separate "+" (map (pretty_expr thy) es) |> Pretty.block;
+
+fun err_in_expr _ msg (Merge []) = error msg
+  | err_in_expr ctxt msg expr =
+    error (msg ^ "\n" ^ Pretty.string_of (Pretty.block
+      [Pretty.str "The error(s) above occured in locale expression:", Pretty.brk 1,
+       pretty_expr (ProofContext.theory_of ctxt) expr]));
+
+
+(** structured contexts: rename + merge + implicit type instantiation **)
+
+(* parameter types *)
+
+fun frozen_tvars ctxt Ts =
+  #1 (Variable.importT_inst (map Logic.mk_type Ts) ctxt)
+  |> map (fn ((xi, S), T) => (xi, (S, T)));
+
+fun unify_frozen ctxt maxidx Ts Us =
+  let
+    fun paramify NONE i = (NONE, i)
+      | paramify (SOME T) i = apfst SOME (TypeInfer.paramify_dummies T i);
+
+    val (Ts', maxidx') = fold_map paramify Ts maxidx;
+    val (Us', maxidx'') = fold_map paramify Us maxidx';
+    val thy = ProofContext.theory_of ctxt;
+
+    fun unify (SOME T, SOME U) env = (Sign.typ_unify thy (U, T) env
+          handle Type.TUNIFY => raise TYPE ("unify_frozen: failed to unify types", [U, T], []))
+      | unify _ env = env;
+    val (unifier, _) = fold unify (Ts' ~~ Us') (Vartab.empty, maxidx'');
+    val Vs = map (Option.map (Envir.norm_type unifier)) Us';
+    val unifier' = fold Vartab.update_new (frozen_tvars ctxt (map_filter I Vs)) unifier;
+  in map (Option.map (Envir.norm_type unifier')) Vs end;
+
+fun params_of elemss =
+  distinct (eq_fst (op = : string * string -> bool)) (maps (snd o fst) elemss);
+
+fun params_of' elemss =
+  distinct (eq_fst (op = : string * string -> bool)) (maps (snd o fst o fst) elemss);
+
+fun param_prefix locale_name params = (NameSpace.base locale_name ^ "_locale", space_implode "_" params);
+
+
+(* CB: param_types has the following type:
+  ('a * 'b option) list -> ('a * 'b) list *)
+fun param_types ps = map_filter (fn (_, NONE) => NONE | (x, SOME T) => SOME (x, T)) ps;
+
+
+fun merge_syntax ctxt ids ss = Symtab.merge (op = : mixfix * mixfix -> bool) ss
+  handle Symtab.DUP x => err_in_locale ctxt
+    ("Conflicting syntax for parameter: " ^ quote x) (map fst ids);
+
+
+(* Distinction of assumed vs. derived identifiers.
+   The former may have axioms relating assumptions of the context to
+   assumptions of the specification fragment (for locales with
+   predicates).  The latter have witnesses relating assumptions of the
+   specification fragment to assumptions of other (assumed) specification
+   fragments. *)
+
+datatype 'a mode = Assumed of 'a | Derived of 'a;
+
+fun map_mode f (Assumed x) = Assumed (f x)
+  | map_mode f (Derived x) = Derived (f x);
+
+
+(* flatten expressions *)
+
+local
+
+fun unify_parms ctxt fixed_parms raw_parmss =
+  let
+    val thy = ProofContext.theory_of ctxt;
+    val maxidx = length raw_parmss;
+    val idx_parmss = (0 upto maxidx - 1) ~~ raw_parmss;
+
+    fun varify i = Term.map_type_tfree (fn (a, S) => TVar ((a, i), S));
+    fun varify_parms (i, ps) = map (apsnd (varify i)) (param_types ps);
+    val parms = fixed_parms @ maps varify_parms idx_parmss;
+
+    fun unify T U envir = Sign.typ_unify thy (U, T) envir
+      handle Type.TUNIFY =>
+        let
+          val T' = Envir.norm_type (fst envir) T;
+          val U' = Envir.norm_type (fst envir) U;
+          val prt = Syntax.string_of_typ ctxt;
+        in
+          raise TYPE ("unify_parms: failed to unify types " ^
+            prt U' ^ " and " ^ prt T', [U', T'], [])
+        end;
+    fun unify_list (T :: Us) = fold (unify T) Us
+      | unify_list [] = I;
+    val (unifier, _) = fold unify_list (map #2 (Symtab.dest (Symtab.make_list parms)))
+      (Vartab.empty, maxidx);
+
+    val parms' = map (apsnd (Envir.norm_type unifier)) (distinct (eq_fst (op =)) parms);
+    val unifier' = fold Vartab.update_new (frozen_tvars ctxt (map #2 parms')) unifier;
+
+    fun inst_parms (i, ps) =
+      List.foldr OldTerm.add_typ_tfrees [] (map_filter snd ps)
+      |> map_filter (fn (a, S) =>
+          let val T = Envir.norm_type unifier' (TVar ((a, i), S))
+          in if T = TFree (a, S) then NONE else SOME (a, T) end)
+      |> Symtab.make;
+  in map inst_parms idx_parmss end;
+
+in
+
+fun unify_elemss _ _ [] = []
+  | unify_elemss _ [] [elems] = [elems]
+  | unify_elemss ctxt fixed_parms elemss =
+      let
+        val thy = ProofContext.theory_of ctxt;
+        val phis = unify_parms ctxt fixed_parms (map (snd o fst o fst) elemss)
+          |> map (Element.instT_morphism thy);
+        fun inst ((((name, ps), mode), elems), phi) =
+          (((name, map (apsnd (Option.map (Morphism.typ phi))) ps),
+              map_mode (map (Element.morph_witness phi)) mode),
+            map (Element.morph_ctxt phi) elems);
+      in map inst (elemss ~~ phis) end;
+
+
+fun renaming xs parms = zip_options parms xs
+  handle Library.UnequalLengths =>
+    error ("Too many arguments in renaming: " ^
+      commas (map (fn NONE => "_" | SOME x => quote (fst x)) xs));
+
+
+(* params_of_expr:
+   Compute parameters (with types and syntax) of locale expression.
+*)
+
+fun params_of_expr ctxt fixed_params expr (prev_parms, prev_types, prev_syn) =
+  let
+    val thy = ProofContext.theory_of ctxt;
+
+    fun merge_tenvs fixed tenv1 tenv2 =
+        let
+          val [env1, env2] = unify_parms ctxt fixed
+                [tenv1 |> Symtab.dest |> map (apsnd SOME),
+                 tenv2 |> Symtab.dest |> map (apsnd SOME)]
+        in
+          Symtab.merge (op =) (Symtab.map (Element.instT_type env1) tenv1,
+            Symtab.map (Element.instT_type env2) tenv2)
+        end;
+
+    fun merge_syn expr syn1 syn2 =
+        Symtab.merge (op = : mixfix * mixfix -> bool) (syn1, syn2)
+        handle Symtab.DUP x => err_in_expr ctxt
+          ("Conflicting syntax for parameter: " ^ quote x) expr;
+
+    fun params_of (expr as Locale name) =
+          let
+            val {params, ...} = the_locale thy name;
+          in (map (fst o fst) params, params |> map fst |> Symtab.make,
+               params |> map (apfst fst) |> Symtab.make) end
+      | params_of (expr as Rename (e, xs)) =
+          let
+            val (parms', types', syn') = params_of e;
+            val ren = renaming xs parms';
+            (* renaming may reduce number of parameters *)
+            val new_parms = map (Element.rename ren) parms' |> distinct (op =);
+            val ren_syn = syn' |> Symtab.dest |> map (Element.rename_var_name ren);
+            val new_syn = fold (Symtab.insert (op =)) ren_syn Symtab.empty
+                handle Symtab.DUP x =>
+                  err_in_expr ctxt ("Conflicting syntax for parameter: " ^ quote x) expr;
+            val syn_types = map (apsnd (fn mx =>
+                SOME (Type.freeze_type (#1 (TypeInfer.paramify_dummies (Syntax.mixfixT mx) 0)))))
+              (Symtab.dest new_syn);
+            val ren_types = types' |> Symtab.dest |> map (apfst (Element.rename ren));
+            val (env :: _) = unify_parms ctxt []
+                ((ren_types |> map (apsnd SOME)) :: map single syn_types);
+            val new_types = fold (Symtab.insert (op =))
+                (map (apsnd (Element.instT_type env)) ren_types) Symtab.empty;
+          in (new_parms, new_types, new_syn) end
+      | params_of (Merge es) =
+          fold (fn e => fn (parms, types, syn) =>
+                   let
+                     val (parms', types', syn') = params_of e
+                   in
+                     (merge_lists (op =) parms parms', merge_tenvs [] types types',
+                      merge_syn e syn syn')
+                   end)
+            es ([], Symtab.empty, Symtab.empty)
+
+      val (parms, types, syn) = params_of expr;
+    in
+      (merge_lists (op =) prev_parms parms, merge_tenvs fixed_params prev_types types,
+       merge_syn expr prev_syn syn)
+    end;
+
+fun make_params_ids params = [(("", params), ([], Assumed []))];
+fun make_raw_params_elemss (params, tenv, syn) =
+    [((("", map (fn p => (p, Symtab.lookup tenv p)) params), Assumed []),
+      Int [Fixes (map (fn p =>
+        (Binding.name p, Symtab.lookup tenv p, Symtab.lookup syn p |> the)) params)])];
+
+
+(* flatten_expr:
+   Extend list of identifiers by those new in locale expression expr.
+   Compute corresponding list of lists of locale elements (one entry per
+   identifier).
+
+   Identifiers represent locale fragments and are in an extended form:
+     ((name, ps), (ax_ps, axs))
+   (name, ps) is the locale name with all its parameters.
+   (ax_ps, axs) is the locale axioms with its parameters;
+     axs are always taken from the top level of the locale hierarchy,
+     hence axioms may contain additional parameters from later fragments:
+     ps subset of ax_ps.  axs is either singleton or empty.
+
+   Elements are enriched by identifier-like information:
+     (((name, ax_ps), axs), elems)
+   The parameters in ax_ps are the axiom parameters, but enriched by type
+   info: now each entry is a pair of string and typ option.  Axioms are
+   type-instantiated.
+
+*)
+
+fun flatten_expr ctxt ((prev_idents, prev_syntax), expr) =
+  let
+    val thy = ProofContext.theory_of ctxt;
+
+    fun rename_parms top ren ((name, ps), (parms, mode)) =
+        ((name, map (Element.rename ren) ps),
+         if top
+         then (map (Element.rename ren) parms,
+               map_mode (map (Element.morph_witness (Element.rename_morphism ren))) mode)
+         else (parms, mode));
+
+    (* add (name, pTs) and its registrations, recursively; adjust hyps of witnesses *)
+
+    fun add_with_regs ((name, pTs), mode) (wits, ids, visited) =
+        if member (fn (a, (b, _)) => a = b) visited (name, map #1 pTs)
+        then (wits, ids, visited)
+        else
+          let
+            val {params, regs, ...} = the_locale thy name;
+            val pTs' = map #1 params;
+            val ren = map #1 pTs' ~~ map (fn (x, _) => (x, NONE)) pTs;
+              (* dummy syntax, since required by rename *)
+            val pTs'' = map (fn ((p, _), (_, T)) => (p, T)) (pTs ~~ pTs');
+            val [env] = unify_parms ctxt pTs [map (apsnd SOME) pTs''];
+              (* propagate parameter types, to keep them consistent *)
+            val regs' = map (fn ((name, ps), wits) =>
+                ((name, map (Element.rename ren) ps),
+                 map (Element.transfer_witness thy) wits)) regs;
+            val new_regs = regs';
+            val new_ids = map fst new_regs;
+            val new_idTs =
+              map (apsnd (map (fn p => (p, (the o AList.lookup (op =) pTs) p)))) new_ids;
+
+            val new_wits = new_regs |> map (#2 #> map
+              (Element.morph_witness
+                (Element.instT_morphism thy env $>
+                  Element.rename_morphism ren $>
+                  Element.satisfy_morphism wits)
+                #> Element.close_witness));
+            val new_ids' = map (fn (id, wits) =>
+                (id, ([], Derived wits))) (new_ids ~~ new_wits);
+            val new_idTs' = map (fn ((n, pTs), (_, ([], mode))) =>
+                ((n, pTs), mode)) (new_idTs ~~ new_ids');
+            val new_id = ((name, map #1 pTs), ([], mode));
+            val (wits', ids', visited') = fold add_with_regs new_idTs'
+              (wits @ flat new_wits, ids, visited @ [new_id]);
+          in
+            (wits', ids' @ [new_id], visited')
+          end;
+
+    (* distribute top-level axioms over assumed ids *)
+
+    fun axiomify all_ps ((name, parms), (_, Assumed _)) axioms =
+        let
+          val {elems, ...} = the_locale thy name;
+          val ts = maps
+            (fn (Assumes asms, _) => maps (map #1 o #2) asms
+              | _ => [])
+            elems;
+          val (axs1, axs2) = chop (length ts) axioms;
+        in (((name, parms), (all_ps, Assumed axs1)), axs2) end
+      | axiomify all_ps (id, (_, Derived ths)) axioms =
+          ((id, (all_ps, Derived ths)), axioms);
+
+    (* identifiers of an expression *)
+
+    fun identify top (Locale name) =
+    (* CB: ids_ax is a list of tuples of the form ((name, ps), axs),
+       where name is a locale name, ps a list of parameter names and axs
+       a list of axioms relating to the identifier, axs is empty unless
+       identify at top level (top = true);
+       parms is accumulated list of parameters *)
+          let
+            val {axiom, params, ...} = the_locale thy name;
+            val ps = map (#1 o #1) params;
+            val (_, ids'', _) = add_with_regs ((name, map #1 params), Assumed []) ([], [], []);
+            val ids_ax = if top then fst (fold_map (axiomify ps) ids'' axiom) else ids'';
+            in (ids_ax, ps) end
+      | identify top (Rename (e, xs)) =
+          let
+            val (ids', parms') = identify top e;
+            val ren = renaming xs parms'
+              handle ERROR msg => err_in_locale' ctxt msg ids';
+
+            val ids'' = distinct (eq_fst (op =)) (map (rename_parms top ren) ids');
+            val parms'' = distinct (op =) (maps (#2 o #1) ids'');
+          in (ids'', parms'') end
+      | identify top (Merge es) =
+          fold (fn e => fn (ids, parms) =>
+                   let
+                     val (ids', parms') = identify top e
+                   in
+                     (merge_alists (op =) ids ids', merge_lists (op =) parms parms')
+                   end)
+            es ([], []);
+
+    fun inst_wit all_params (t, th) = let
+         val {hyps, prop, ...} = Thm.rep_thm th;
+         val ps = map (apsnd SOME) (fold Term.add_frees (prop :: hyps) []);
+         val [env] = unify_parms ctxt all_params [ps];
+         val t' = Element.instT_term env t;
+         val th' = Element.instT_thm thy env th;
+       in (t', th') end;
+
+    fun eval all_params tenv syn ((name, params), (locale_params, mode)) =
+      let
+        val {params = ps_mx, elems = elems_stamped, ...} = the_locale thy name;
+        val elems = map fst elems_stamped;
+        val ps = map fst ps_mx;
+        fun lookup_syn x = (case Symtab.lookup syn x of SOME Structure => NONE | opt => opt);
+        val locale_params' = map (fn p => (p, Symtab.lookup tenv p |> the)) locale_params;
+        val mode' = map_mode (map (Element.map_witness (inst_wit all_params))) mode;
+        val ren = map fst ps ~~ map (fn p => (p, lookup_syn p)) params;
+        val [env] = unify_parms ctxt all_params [map (apfst (Element.rename ren) o apsnd SOME) ps];
+        val (lprfx, pprfx) = param_prefix name params;
+        val add_prefices = pprfx <> "" ? Binding.add_prefix false pprfx
+          #> Binding.add_prefix false lprfx;
+        val elem_morphism =
+          Element.rename_morphism ren $>
+          Morphism.binding_morphism add_prefices $>
+          Element.instT_morphism thy env;
+        val elems' = map (Element.morph_ctxt elem_morphism) elems;
+      in (((name, map (apsnd SOME) locale_params'), mode'), elems') end;
+
+    (* parameters, their types and syntax *)
+    val (all_params', tenv, syn) = params_of_expr ctxt [] expr ([], Symtab.empty, Symtab.empty);
+    val all_params = map (fn p => (p, Symtab.lookup tenv p |> the)) all_params';
+    (* compute identifiers and syntax, merge with previous ones *)
+    val (ids, _) = identify true expr;
+    val idents = subtract (eq_fst (op =)) prev_idents ids;
+    val syntax = merge_syntax ctxt ids (syn, prev_syntax);
+    (* type-instantiate elements *)
+    val final_elemss = map (eval all_params tenv syntax) idents;
+  in ((prev_idents @ idents, syntax), final_elemss) end;
+
+end;
+
+
+(* activate elements *)
+
+local
+
+fun axioms_export axs _ As =
+  (Element.satisfy_thm axs #> Drule.implies_intr_list (Library.drop (length axs, As)), fn t => t);
+
+
+(* NB: derived ids contain only facts at this stage *)
+
+fun activate_elem _ _ (Fixes fixes) (ctxt, mode) =
+      ([], (ctxt |> ProofContext.add_fixes_i fixes |> snd, mode))
+  | activate_elem _ _ (Constrains _) (ctxt, mode) =
+      ([], (ctxt, mode))
+  | activate_elem ax_in_ctxt _ (Assumes asms) (ctxt, Assumed axs) =
+      let
+        val asms' = Attrib.map_specs (Attrib.attribute_i (ProofContext.theory_of ctxt)) asms;
+        val ts = maps (map #1 o #2) asms';
+        val (ps, qs) = chop (length ts) axs;
+        val (_, ctxt') =
+          ctxt |> fold Variable.auto_fixes ts
+          |> ProofContext.add_assms_i (axioms_export (if ax_in_ctxt then ps else [])) asms';
+      in ([], (ctxt', Assumed qs)) end
+  | activate_elem _ _ (Assumes asms) (ctxt, Derived ths) =
+      ([], (ctxt, Derived ths))
+  | activate_elem _ _ (Defines defs) (ctxt, Assumed axs) =
+      let
+        val defs' = Attrib.map_specs (Attrib.attribute_i (ProofContext.theory_of ctxt)) defs;
+        val asms = defs' |> map (fn ((name, atts), (t, ps)) =>
+            let val ((c, _), t') = LocalDefs.cert_def ctxt t
+            in (t', ((Binding.map_base (Thm.def_name_optional c) name, atts), [(t', ps)])) end);
+        val (_, ctxt') =
+          ctxt |> fold (Variable.auto_fixes o #1) asms
+          |> ProofContext.add_assms_i LocalDefs.def_export (map #2 asms);
+      in ([], (ctxt', Assumed axs)) end
+  | activate_elem _ _ (Defines defs) (ctxt, Derived ths) =
+      ([], (ctxt, Derived ths))
+  | activate_elem _ is_ext (Notes (kind, facts)) (ctxt, mode) =
+      let
+        val facts' = Attrib.map_facts (Attrib.attribute_i (ProofContext.theory_of ctxt)) facts;
+        val (res, ctxt') = ctxt |> local_note_qualified kind facts';
+      in (if is_ext then (map (#1 o #1) facts' ~~ map #2 res) else [], (ctxt', mode)) end;
+
+fun activate_elems ax_in_ctxt (((name, ps), mode), elems) ctxt =
+  let
+    val thy = ProofContext.theory_of ctxt;
+    val (res, (ctxt', _)) = fold_map (activate_elem ax_in_ctxt (name = ""))
+        elems (ProofContext.qualified_names ctxt, mode)
+      handle ERROR msg => err_in_locale ctxt msg [(name, map fst ps)];
+    val ctxt'' = if name = "" then ctxt'
+          else let
+              val ps' = map (fn (n, SOME T) => Free (n, T)) ps;
+            in if test_local_registration ctxt' (name, ps') then ctxt'
+              else let
+                  val ctxt'' = put_local_registration (name, ps') (I, (NameSpace.base name, ""))
+                    (Morphism.identity, ((Vartab.empty, []), (Vartab.empty, []) )) ctxt'
+                in case mode of
+                    Assumed axs =>
+                      fold (add_local_witness (name, ps') o
+                        Element.assume_witness thy o Element.witness_prop) axs ctxt''
+                  | Derived ths =>
+                     fold (add_local_witness (name, ps')) ths ctxt''
+                end
+            end
+  in (ProofContext.restore_naming ctxt ctxt'', res) end;
+
+fun activate_elemss ax_in_ctxt prep_facts =
+    fold_map (fn (((name, ps), mode), raw_elems) => fn ctxt =>
+      let
+        val elems = map (prep_facts ctxt) raw_elems;
+        val (ctxt', res) = apsnd flat
+            (activate_elems ax_in_ctxt (((name, ps), mode), elems) ctxt);
+        val elems' = elems |> map (Element.map_ctxt_attrib Args.closure);
+      in (((((name, ps), mode), elems'), res), ctxt') end);
+
+in
+
+(* CB: activate_facts prep_facts elemss ctxt,
+   where elemss is a list of pairs consisting of identifiers and
+   context elements, extends ctxt by the context elements yielding
+   ctxt' and returns ((elemss', facts), ctxt').
+   Identifiers in the argument are of the form ((name, ps), axs) and
+   assumptions use the axioms in the identifiers to set up exporters
+   in ctxt'.  elemss' does not contain identifiers and is obtained
+   from elemss and the intermediate context with prep_facts.
+   If read_facts or cert_facts is used for prep_facts, these also remove
+   the internal/external markers from elemss. *)
+
+fun activate_facts ax_in_ctxt prep_facts args =
+  activate_elemss ax_in_ctxt prep_facts args
+  #>> (apsnd flat o split_list);
+
+end;
+
+
+
+(** prepare locale elements **)
+
+(* expressions *)
+
+fun intern_expr thy (Locale xname) = Locale (intern thy xname)
+  | intern_expr thy (Merge exprs) = Merge (map (intern_expr thy) exprs)
+  | intern_expr thy (Rename (expr, xs)) = Rename (intern_expr thy expr, xs);
+
+
+(* propositions and bindings *)
+
+(* flatten (ctxt, prep_expr) ((ids, syn), expr)
+   normalises expr (which is either a locale
+   expression or a single context element) wrt.
+   to the list ids of already accumulated identifiers.
+   It returns ((ids', syn'), elemss) where ids' is an extension of ids
+   with identifiers generated for expr, and elemss is the list of
+   context elements generated from expr.
+   syn and syn' are symtabs mapping parameter names to their syntax.  syn'
+   is an extension of syn.
+   For details, see flatten_expr.
+
+   Additionally, for a locale expression, the elems are grouped into a single
+   Int; individual context elements are marked Ext.  In this case, the
+   identifier-like information of the element is as follows:
+   - for Fixes: (("", ps), []) where the ps have type info NONE
+   - for other elements: (("", []), []).
+   The implementation of activate_facts relies on identifier names being
+   empty strings for external elements.
+*)
+
+fun flatten (ctxt, _) ((ids, syn), Elem (Fixes fixes)) = let
+        val ids' = ids @ [(("", map (Binding.base_name o #1) fixes), ([], Assumed []))]
+      in
+        ((ids',
+         merge_syntax ctxt ids'
+           (syn, Symtab.make (map (fn fx => (Binding.base_name (#1 fx), #3 fx)) fixes))
+           handle Symtab.DUP x => err_in_locale ctxt
+             ("Conflicting syntax for parameter: " ^ quote x)
+             (map #1 ids')),
+         [((("", map (rpair NONE o Binding.base_name o #1) fixes), Assumed []), Ext (Fixes fixes))])
+      end
+  | flatten _ ((ids, syn), Elem elem) =
+      ((ids @ [(("", []), ([], Assumed []))], syn), [((("", []), Assumed []), Ext elem)])
+  | flatten (ctxt, prep_expr) ((ids, syn), Expr expr) =
+      apsnd (map (apsnd Int)) (flatten_expr ctxt ((ids, syn), prep_expr expr));
+
+local
+
+local
+
+fun declare_int_elem (Fixes fixes) ctxt =
+      ([], ctxt |> ProofContext.add_fixes_i (map (fn (x, T, mx) =>
+        (x, Option.map (Term.map_type_tfree (TypeInfer.param 0)) T, mx)) fixes) |> snd)
+  | declare_int_elem _ ctxt = ([], ctxt);
+
+fun declare_ext_elem prep_vars (Fixes fixes) ctxt =
+      let val (vars, _) = prep_vars fixes ctxt
+      in ([], ctxt |> ProofContext.add_fixes_i vars |> snd) end
+  | declare_ext_elem prep_vars (Constrains csts) ctxt =
+      let val (_, ctxt') = prep_vars (map (fn (x, T) => (Binding.name x, SOME T, NoSyn)) csts) ctxt
+      in ([], ctxt') end
+  | declare_ext_elem _ (Assumes asms) ctxt = (map #2 asms, ctxt)
+  | declare_ext_elem _ (Defines defs) ctxt = (map (fn (_, (t, ps)) => [(t, ps)]) defs, ctxt)
+  | declare_ext_elem _ (Notes _) ctxt = ([], ctxt);
+
+fun declare_elems prep_vars (((name, ps), Assumed _), elems) ctxt = ((case elems
+     of Int es => fold_map declare_int_elem es ctxt
+      | Ext e => declare_ext_elem prep_vars e ctxt |>> single)
+          handle ERROR msg => err_in_locale ctxt msg [(name, map fst ps)])
+  | declare_elems _ ((_, Derived _), elems) ctxt = ([], ctxt);
+
+in
+
+fun declare_elemss prep_vars fixed_params raw_elemss ctxt =
+  let
+    (* CB: fix of type bug of goal in target with context elements.
+       Parameters new in context elements must receive types that are
+       distinct from types of parameters in target (fixed_params).  *)
+    val ctxt_with_fixed = 
+      fold Variable.declare_term (map Free fixed_params) ctxt;
+    val int_elemss =
+      raw_elemss
+      |> map_filter (fn (id, Int es) => SOME (id, es) | _ => NONE)
+      |> unify_elemss ctxt_with_fixed fixed_params;
+    val (raw_elemss', _) =
+      fold_map (curry (fn ((id, Int _), (_, es) :: elemss) => ((id, Int es), elemss) | x => x))
+        raw_elemss int_elemss;
+  in fold_map (declare_elems prep_vars) raw_elemss' ctxt end;
+
+end;
+
+local
+
+val norm_term = Envir.beta_norm oo Term.subst_atomic;
+
+fun abstract_thm thy eq =
+  Thm.assume (Thm.cterm_of thy eq) |> Drule.gen_all |> Drule.abs_def;
+
+fun bind_def ctxt (name, ps) eq (xs, env, ths) =
+  let
+    val ((y, T), b) = LocalDefs.abs_def eq;
+    val b' = norm_term env b;
+    val th = abstract_thm (ProofContext.theory_of ctxt) eq;
+    fun err msg = err_in_locale ctxt (msg ^ ": " ^ quote y) [(name, map fst ps)];
+  in
+    exists (fn (x, _) => x = y) xs andalso
+      err "Attempt to define previously specified variable";
+    exists (fn (Free (y', _), _) => y = y' | _ => false) env andalso
+      err "Attempt to redefine variable";
+    (Term.add_frees b' xs, (Free (y, T), b') :: env, th :: ths)
+  end;
+
+
+(* CB: for finish_elems (Int and Ext),
+   extracts specification, only of assumed elements *)
+
+fun eval_text _ _ _ (Fixes _) text = text
+  | eval_text _ _ _ (Constrains _) text = text
+  | eval_text _ (_, Assumed _) is_ext (Assumes asms)
+        (((exts, exts'), (ints, ints')), (xs, env, defs)) =
+      let
+        val ts = maps (map #1 o #2) asms;
+        val ts' = map (norm_term env) ts;
+        val spec' =
+          if is_ext then ((exts @ ts, exts' @ ts'), (ints, ints'))
+          else ((exts, exts'), (ints @ ts, ints' @ ts'));
+      in (spec', (fold Term.add_frees ts' xs, env, defs)) end
+  | eval_text _ (_, Derived _) _ (Assumes _) text = text
+  | eval_text ctxt (id, Assumed _) _ (Defines defs) (spec, binds) =
+      (spec, fold (bind_def ctxt id o #1 o #2) defs binds)
+  | eval_text _ (_, Derived _) _ (Defines _) text = text
+  | eval_text _ _ _ (Notes _) text = text;
+
+
+(* for finish_elems (Int),
+   remove redundant elements of derived identifiers,
+   turn assumptions and definitions into facts,
+   satisfy hypotheses of facts *)
+
+fun finish_derived _ _ (Assumed _) (Fixes fixes) = SOME (Fixes fixes)
+  | finish_derived _ _ (Assumed _) (Constrains csts) = SOME (Constrains csts)
+  | finish_derived _ _ (Assumed _) (Assumes asms) = SOME (Assumes asms)
+  | finish_derived _ _ (Assumed _) (Defines defs) = SOME (Defines defs)
+
+  | finish_derived _ _ (Derived _) (Fixes _) = NONE
+  | finish_derived _ _ (Derived _) (Constrains _) = NONE
+  | finish_derived sign satisfy (Derived _) (Assumes asms) = asms
+      |> map (apsnd (map (fn (a, _) => ([Thm.assume (cterm_of sign a)], []))))
+      |> pair Thm.assumptionK |> Notes
+      |> Element.morph_ctxt satisfy |> SOME
+  | finish_derived sign satisfy (Derived _) (Defines defs) = defs
+      |> map (apsnd (fn (d, _) => [([Thm.assume (cterm_of sign d)], [])]))
+      |> pair Thm.definitionK |> Notes
+      |> Element.morph_ctxt satisfy |> SOME
+
+  | finish_derived _ satisfy _ (Notes facts) = Notes facts
+      |> Element.morph_ctxt satisfy |> SOME;
+
+(* CB: for finish_elems (Ext) *)
+
+fun closeup _ false elem = elem
+  | closeup ctxt true elem =
+      let
+        fun close_frees t =
+          let
+            val rev_frees =
+              Term.fold_aterms (fn Free (x, T) =>
+                if Variable.is_fixed ctxt x then I else insert (op =) (x, T) | _ => I) t [];
+          in Term.list_all_free (rev rev_frees, t) end;
+
+        fun no_binds [] = []
+          | no_binds _ = error "Illegal term bindings in locale element";
+      in
+        (case elem of
+          Assumes asms => Assumes (asms |> map (fn (a, propps) =>
+            (a, map (fn (t, ps) => (close_frees t, no_binds ps)) propps)))
+        | Defines defs => Defines (defs |> map (fn (a, (t, ps)) =>
+            (a, (close_frees (#2 (LocalDefs.cert_def ctxt t)), no_binds ps))))
+        | e => e)
+      end;
+
+
+fun finish_ext_elem parms _ (Fixes fixes, _) = Fixes (map (fn (b, _, mx) =>
+      let val x = Binding.base_name b
+      in (b, AList.lookup (op =) parms x, mx) end) fixes)
+  | finish_ext_elem parms _ (Constrains _, _) = Constrains []
+  | finish_ext_elem _ close (Assumes asms, propp) =
+      close (Assumes (map #1 asms ~~ propp))
+  | finish_ext_elem _ close (Defines defs, propp) =
+      close (Defines (map #1 defs ~~ map (fn [(t, ps)] => (t, ps)) propp))
+  | finish_ext_elem _ _ (Notes facts, _) = Notes facts;
+
+
+(* CB: finish_parms introduces type info from parms to identifiers *)
+(* CB: only needed for types that have been NONE so far???
+   If so, which are these??? *)
+
+fun finish_parms parms (((name, ps), mode), elems) =
+  (((name, map (fn (x, _) => (x, AList.lookup (op = : string * string -> bool) parms x)) ps), mode), elems);
+
+fun finish_elems ctxt parms _ ((text, wits), ((id, Int e), _)) =
+      let
+        val [(id' as (_, mode), es)] = unify_elemss ctxt parms [(id, e)];
+        val wits' = case mode of Assumed _ => wits | Derived ths => wits @ ths;
+        val text' = fold (eval_text ctxt id' false) es text;
+        val es' = map_filter
+          (finish_derived (ProofContext.theory_of ctxt) (Element.satisfy_morphism wits') mode) es;
+      in ((text', wits'), (id', map Int es')) end
+  | finish_elems ctxt parms do_close ((text, wits), ((id, Ext e), [propp])) =
+      let
+        val e' = finish_ext_elem parms (closeup ctxt do_close) (e, propp);
+        val text' = eval_text ctxt id true e' text;
+      in ((text', wits), (id, [Ext e'])) end
+
+in
+
+(* CB: only called by prep_elemss *)
+
+fun finish_elemss ctxt parms do_close =
+  foldl_map (apsnd (finish_parms parms) o finish_elems ctxt parms do_close);
+
+end;
+
+
+(* Remove duplicate Defines elements: temporary workaround to fix Afp/Category. *)
+
+fun defs_ord (defs1, defs2) =
+    list_ord (fn ((_, (d1, _)), (_, (d2, _))) =>
+      TermOrd.fast_term_ord (d1, d2)) (defs1, defs2);
+structure Defstab =
+    TableFun(type key = (Attrib.binding * (term * term list)) list val ord = defs_ord);
+
+fun rem_dup_defs es ds =
+    fold_map (fn e as (Defines defs) => (fn ds =>
+                 if Defstab.defined ds defs
+                 then (Defines [], ds)
+                 else (e, Defstab.update (defs, ()) ds))
+               | e => (fn ds => (e, ds))) es ds;
+fun rem_dup_elemss (Int es) ds = apfst Int (rem_dup_defs es ds)
+  | rem_dup_elemss (Ext e) ds = (Ext e, ds);
+fun rem_dup_defines raw_elemss =
+    fold_map (fn (id as (_, (Assumed _)), es) => (fn ds =>
+                     apfst (pair id) (rem_dup_elemss es ds))
+               | (id as (_, (Derived _)), es) => (fn ds =>
+                     ((id, es), ds))) raw_elemss Defstab.empty |> #1;
+
+(* CB: type inference and consistency checks for locales.
+
+   Works by building a context (through declare_elemss), extracting the
+   required information and adjusting the context elements (finish_elemss).
+   Can also universally close free vars in assms and defs.  This is only
+   needed for Ext elements and controlled by parameter do_close.
+
+   Only elements of assumed identifiers are considered.
+*)
+
+fun prep_elemss prep_vars prepp do_close context fixed_params raw_elemss raw_concl =
+  let
+    (* CB: contexts computed in the course of this function are discarded.
+       They are used for type inference and consistency checks only. *)
+    (* CB: fixed_params are the parameters (with types) of the target locale,
+       empty list if there is no target. *)
+    (* CB: raw_elemss are list of pairs consisting of identifiers and
+       context elements, the latter marked as internal or external. *)
+    val raw_elemss = rem_dup_defines raw_elemss;
+    val (raw_proppss, raw_ctxt) = declare_elemss prep_vars fixed_params raw_elemss context;
+    (* CB: raw_ctxt is context with additional fixed variables derived from
+       the fixes elements in raw_elemss,
+       raw_proppss contains assumptions and definitions from the
+       external elements in raw_elemss. *)
+    fun prep_prop raw_propp (raw_ctxt, raw_concl)  =
+      let
+        (* CB: add type information from fixed_params to context (declare_term) *)
+        (* CB: process patterns (conclusion and external elements only) *)
+        val (ctxt, all_propp) =
+          prepp (fold Variable.declare_term (map Free fixed_params) raw_ctxt, raw_concl @ raw_propp);
+        (* CB: add type information from conclusion and external elements to context *)
+        val ctxt = fold Variable.declare_term (maps (map fst) all_propp) ctxt;
+        (* CB: resolve schematic variables (patterns) in conclusion and external elements. *)
+        val all_propp' = map2 (curry (op ~~))
+          (#1 (#2 (ProofContext.bind_propp_schematic_i (ctxt, all_propp)))) (map (map snd) all_propp);
+        val (concl, propp) = chop (length raw_concl) all_propp';
+      in (propp, (ctxt, concl)) end
+
+    val (proppss, (ctxt, concl)) =
+      (fold_burrow o fold_burrow) prep_prop raw_proppss (raw_ctxt, raw_concl);
+
+    (* CB: obtain all parameters from identifier part of raw_elemss *)
+    val xs = map #1 (params_of' raw_elemss);
+    val typing = unify_frozen ctxt 0
+      (map (Variable.default_type raw_ctxt) xs)
+      (map (Variable.default_type ctxt) xs);
+    val parms = param_types (xs ~~ typing);
+    (* CB: parms are the parameters from raw_elemss, with correct typing. *)
+
+    (* CB: extract information from assumes and defines elements
+       (fixes, constrains and notes in raw_elemss don't have an effect on
+       text and elemss), compute final form of context elements. *)
+    val ((text, _), elemss) = finish_elemss ctxt parms do_close
+      ((((([], []), ([], [])), ([], [], [])), []), raw_elemss ~~ proppss);
+    (* CB: text has the following structure:
+           (((exts, exts'), (ints, ints')), (xs, env, defs))
+       where
+         exts: external assumptions (terms in external assumes elements)
+         exts': dito, normalised wrt. env
+         ints: internal assumptions (terms in internal assumes elements)
+         ints': dito, normalised wrt. env
+         xs: the free variables in exts' and ints' and rhss of definitions,
+           this includes parameters except defined parameters
+         env: list of term pairs encoding substitutions, where the first term
+           is a free variable; substitutions represent defines elements and
+           the rhs is normalised wrt. the previous env
+         defs: theorems representing the substitutions from defines elements
+           (thms are normalised wrt. env).
+       elemss is an updated version of raw_elemss:
+         - type info added to Fixes and modified in Constrains
+         - axiom and definition statement replaced by corresponding one
+           from proppss in Assumes and Defines
+         - Facts unchanged
+       *)
+  in ((parms, elemss, concl), text) end;
+
+in
+
+fun read_elemss x = prep_elemss ProofContext.read_vars ProofContext.read_propp_schematic x;
+fun cert_elemss x = prep_elemss ProofContext.cert_vars ProofContext.cert_propp_schematic x;
+
+end;
+
+
+(* facts and attributes *)
+
+local
+
+fun check_name name =
+  if NameSpace.is_qualified name then error ("Illegal qualified name: " ^ quote name)
+  else name;
+
+fun prep_facts _ _ _ ctxt (Int elem) = elem
+      |> Element.morph_ctxt (Morphism.thm_morphism (Thm.transfer (ProofContext.theory_of ctxt)))
+  | prep_facts prep_name get intern ctxt (Ext elem) = elem |> Element.map_ctxt
+     {var = I, typ = I, term = I,
+      binding = Binding.map_base prep_name,
+      fact = get ctxt,
+      attrib = Args.assignable o intern (ProofContext.theory_of ctxt)};
+
+in
+
+fun read_facts x = prep_facts check_name ProofContext.get_fact Attrib.intern_src x;
+fun cert_facts x = prep_facts I (K I) (K I) x;
+
+end;
+
+
+(* Get the specification of a locale *)
+
+(*The global specification is made from the parameters and global
+  assumptions, the local specification from the parameters and the
+  local assumptions.*)
+
+local
+
+fun gen_asms_of get thy name =
+  let
+    val ctxt = ProofContext.init thy;
+    val (_, raw_elemss) = flatten (ctxt, I) (([], Symtab.empty), Expr (Locale name));
+    val ((_, elemss, _), _) = read_elemss false ctxt [] raw_elemss [];
+  in
+    elemss |> get
+      |> maps (fn (_, es) => map (fn Int e => e) es)
+      |> maps (fn Assumes asms => asms | _ => [])
+      |> map (apsnd (map fst))
+  end;
+
+in
+
+fun parameters_of thy = #params o the_locale thy;
+
+fun intros thy = #intros o the_locale thy;
+  (*returns introduction rule for delta predicate and locale predicate
+    as a pair of singleton lists*)
+
+fun dests thy = #dests o the_locale thy;
+
+fun facts_of thy = map_filter (fn (Element.Notes (_, facts), _) => SOME facts
+  | _ => NONE) o #elems o the_locale thy;
+
+fun parameters_of_expr thy expr =
+  let
+    val ctxt = ProofContext.init thy;
+    val pts = params_of_expr ctxt [] (intern_expr thy expr)
+        ([], Symtab.empty, Symtab.empty);
+    val raw_params_elemss = make_raw_params_elemss pts;
+    val ((_, syn), raw_elemss) = flatten (ctxt, intern_expr thy)
+        (([], Symtab.empty), Expr expr);
+    val ((parms, _, _), _) =
+        read_elemss false ctxt [] (raw_params_elemss @ raw_elemss) [];
+  in map (fn p as (n, _) => (p, Symtab.lookup syn n |> the)) parms end;
+
+fun local_asms_of thy name =
+  gen_asms_of (single o Library.last_elem) thy name;
+
+fun global_asms_of thy name =
+  gen_asms_of I thy name;
+
+end;
+
+
+(* full context statements: imports + elements + conclusion *)
+
+local
+
+fun prep_context_statement prep_expr prep_elemss prep_facts
+    do_close fixed_params imports elements raw_concl context =
+  let
+    val thy = ProofContext.theory_of context;
+
+    val (import_params, import_tenv, import_syn) =
+      params_of_expr context fixed_params (prep_expr thy imports)
+        ([], Symtab.empty, Symtab.empty);
+    val includes = map_filter (fn Expr e => SOME e | Elem _ => NONE) elements;
+    val (incl_params, incl_tenv, incl_syn) = fold (params_of_expr context fixed_params)
+      (map (prep_expr thy) includes) (import_params, import_tenv, import_syn);
+
+    val ((import_ids, _), raw_import_elemss) =
+      flatten (context, prep_expr thy) (([], Symtab.empty), Expr imports);
+    (* CB: normalise "includes" among elements *)
+    val ((ids, syn), raw_elemsss) = foldl_map (flatten (context, prep_expr thy))
+      ((import_ids, incl_syn), elements);
+
+    val raw_elemss = flat raw_elemsss;
+    (* CB: raw_import_elemss @ raw_elemss is the normalised list of
+       context elements obtained from import and elements. *)
+    (* Now additional elements for parameters are inserted. *)
+    val import_params_ids = make_params_ids import_params;
+    val incl_params_ids =
+        make_params_ids (incl_params \\ import_params);
+    val raw_import_params_elemss =
+        make_raw_params_elemss (import_params, incl_tenv, incl_syn);
+    val raw_incl_params_elemss =
+        make_raw_params_elemss (incl_params \\ import_params, incl_tenv, incl_syn);
+    val ((parms, all_elemss, concl), (spec, (_, _, defs))) = prep_elemss do_close
+      context fixed_params
+      (raw_import_params_elemss @ raw_import_elemss @ raw_incl_params_elemss @ raw_elemss) raw_concl;
+
+    (* replace extended ids (for axioms) by ids *)
+    val (import_ids', incl_ids) = chop (length import_ids) ids;
+    val all_ids = import_params_ids @ import_ids' @ incl_params_ids @ incl_ids;
+    val all_elemss' = map (fn (((_, ps), _), (((n, ps'), mode), elems)) =>
+        (((n, map (fn p => (p, (the o AList.lookup (op =) ps') p)) ps), mode), elems))
+      (all_ids ~~ all_elemss);
+    (* CB: all_elemss and parms contain the correct parameter types *)
+
+    val (ps, qs) = chop (length raw_import_params_elemss + length raw_import_elemss) all_elemss';
+    val ((import_elemss, _), import_ctxt) =
+      activate_facts false prep_facts ps context;
+
+    val ((elemss, _), ctxt) =
+      activate_facts false prep_facts qs (ProofContext.set_stmt true import_ctxt);
+  in
+    ((((import_ctxt, import_elemss), (ctxt, elemss, syn)),
+      (parms, spec, defs)), concl)
+  end;
+
+fun prep_statement prep_locale prep_ctxt raw_locale elems concl ctxt =
+  let
+    val thy = ProofContext.theory_of ctxt;
+    val locale = Option.map (prep_locale thy) raw_locale;
+    val (fixed_params, imports) =
+      (case locale of
+        NONE => ([], empty)
+      | SOME name =>
+          let val {params = ps, ...} = the_locale thy name
+          in (map fst ps, Locale name) end);
+    val ((((locale_ctxt, _), (elems_ctxt, _, _)), _), concl') =
+      prep_ctxt false fixed_params imports (map Elem elems) concl ctxt;
+  in (locale, locale_ctxt, elems_ctxt, concl') end;
+
+fun prep_expr prep imports body ctxt =
+  let
+    val (((_, import_elemss), (ctxt', elemss, _)), _) = prep imports body ctxt;
+    val all_elems = maps snd (import_elemss @ elemss);
+  in (all_elems, ctxt') end;
+
+in
+
+val read_ctxt = prep_context_statement intern_expr read_elemss read_facts;
+val cert_ctxt = prep_context_statement (K I) cert_elemss cert_facts;
+
+fun read_context imports body ctxt = #1 (read_ctxt true [] imports (map Elem body) [] ctxt);
+fun cert_context imports body ctxt = #1 (cert_ctxt true [] imports (map Elem body) [] ctxt);
+
+val read_expr = prep_expr read_context;
+val cert_expr = prep_expr cert_context;
+
+fun read_context_statement loc = prep_statement (K I) read_ctxt loc;
+fun read_context_statement_cmd loc = prep_statement intern read_ctxt loc;
+fun cert_context_statement loc = prep_statement (K I) cert_ctxt loc;
+
+end;
+
+
+(* init *)
+
+fun init loc =
+  ProofContext.init
+  #> #2 o cert_context_statement (SOME loc) [] [];
+
+
+(* print locale *)
+
+fun print_locale thy show_facts imports body =
+  let val (all_elems, ctxt) = read_expr imports body (ProofContext.init thy) in
+    Pretty.big_list "locale elements:" (all_elems
+      |> (if show_facts then I else filter (fn Notes _ => false | _ => true))
+      |> map (Element.pretty_ctxt ctxt) |> filter_out null
+      |> map Pretty.chunks)
+    |> Pretty.writeln
+  end;
+
+
+
+(** store results **)
+
+(* join equations of an id with already accumulated ones *)
+
+fun join_eqns get_reg id eqns =
+  let
+    val eqns' = case get_reg id
+      of NONE => eqns
+        | SOME (_, _, eqns') => Termtab.join (fn _ => fn (_, e) => e) (eqns, eqns')
+            (* prefer equations from eqns' *)
+  in ((id, eqns'), eqns') end;
+
+
+(* collect witnesses and equations up to a particular target for a
+   registration; requires parameters and flattened list of identifiers
+   instead of recomputing it from the target *)
+
+fun collect_witnesses ctxt (imprt as ((impT, _), (imp, _))) parms ids ext_ts = let
+
+    val thy = ProofContext.theory_of ctxt;
+
+    val ts = map (var_inst_term (impT, imp)) ext_ts;
+    val (parms, parmTs) = split_list parms;
+    val parmvTs = map Logic.varifyT parmTs;
+    val vtinst = fold (Sign.typ_match thy) (parmvTs ~~ map Term.fastype_of ts) Vartab.empty;
+    val tinst = Vartab.dest vtinst |> map (fn ((x, 0), (_, T)) => (x, T))
+        |> Symtab.make;
+    val inst = Symtab.make (parms ~~ ts);
+
+    (* instantiate parameter names in ids *)
+    val ext_inst = Symtab.make (parms ~~ ext_ts);
+    fun ext_inst_names ps = map (the o Symtab.lookup ext_inst) ps;
+    val inst_ids = map (apfst (apsnd ext_inst_names)) ids;
+    val assumed_ids = map_filter (fn (id, (_, Assumed _)) => SOME id | _ => NONE) inst_ids;
+    val wits = maps (#2 o the o get_local_registration ctxt imprt) assumed_ids;
+    val eqns =
+      fold_map (join_eqns (get_local_registration ctxt imprt))
+        (map fst inst_ids) Termtab.empty |> snd |> Termtab.dest |> map snd;
+  in ((tinst, inst), wits, eqns) end;
+
+
+(* compute and apply morphism *)
+
+fun name_morph phi_name (lprfx, pprfx) b =
+  b
+  |> (if not (Binding.is_empty b) andalso pprfx <> ""
+        then Binding.add_prefix false pprfx else I)
+  |> (if not (Binding.is_empty b)
+        then Binding.add_prefix false lprfx else I)
+  |> phi_name;
+
+fun inst_morph thy phi_name param_prfx insts prems eqns export =
+  let
+    (* standardise export morphism *)
+    val exp_fact = Drule.zero_var_indexes_list o map Thm.strip_shyps o Morphism.fact export;
+    val exp_term = TermSubst.zero_var_indexes o Morphism.term export;
+      (* FIXME sync with exp_fact *)
+    val exp_typ = Logic.type_map exp_term;
+    val export' =
+      Morphism.morphism {binding = I, var = I, typ = exp_typ, term = exp_term, fact = exp_fact};
+  in
+    Morphism.binding_morphism (name_morph phi_name param_prfx) $>
+      Element.inst_morphism thy insts $>
+      Element.satisfy_morphism prems $>
+      Morphism.term_morphism (MetaSimplifier.rewrite_term thy eqns []) $>
+      Morphism.thm_morphism (MetaSimplifier.rewrite_rule eqns) $>
+      export'
+  end;
+
+fun activate_note thy phi_name param_prfx attrib insts prems eqns exp =
+  (Element.facts_map o Element.morph_ctxt)
+      (inst_morph thy phi_name param_prfx insts prems eqns exp)
+  #> Attrib.map_facts attrib;
+
+
+(* public interface to interpretation morphism *)
+
+fun get_interpret_morph thy phi_name param_prfx (exp, imp) target ext_ts =
+  let
+    val parms = the_locale thy target |> #params |> map fst;
+    val ids = flatten (ProofContext.init thy, intern_expr thy)
+      (([], Symtab.empty), Expr (Locale target)) |> fst |> fst;
+    val (insts, prems, eqns) = collect_witnesses (ProofContext.init thy) imp parms ids ext_ts;
+  in
+    inst_morph thy phi_name param_prfx insts prems eqns exp
+  end;
+
+(* store instantiations of args for all registered interpretations
+   of the theory *)
+
+fun note_thmss_registrations target (kind, args) thy =
+  let
+    val parms = the_locale thy target |> #params |> map fst;
+    val ids = flatten (ProofContext.init thy, intern_expr thy)
+      (([], Symtab.empty), Expr (Locale target)) |> fst |> fst;
+
+    val regs = get_global_registrations thy target;
+    (* add args to thy for all registrations *)
+
+    fun activate (ext_ts, ((phi_name, param_prfx), (exp, imp), _, _)) thy =
+      let
+        val (insts, prems, eqns) = collect_witnesses (ProofContext.init thy) imp parms ids ext_ts;
+        val args' = args
+          |> activate_note thy phi_name param_prfx
+               (Attrib.attribute_i thy) insts prems eqns exp;
+      in
+        thy
+        |> global_note_qualified kind args'
+        |> snd
+      end;
+  in fold activate regs thy end;
+
+
+(* locale results *)
+
+fun add_thmss loc kind args ctxt =
+  let
+    val (([(_, [Notes args'])], _), ctxt') =
+      activate_facts true cert_facts
+        [((("", []), Assumed []), [Ext (Notes (kind, args))])] ctxt;
+    val ctxt'' = ctxt' |> ProofContext.theory
+      (change_locale loc
+        (fn (axiom, elems, params, decls, regs, intros, dests) =>
+          (axiom, elems @ [(Notes args', stamp ())],
+            params, decls, regs, intros, dests))
+      #> note_thmss_registrations loc args');
+  in ctxt'' end;
+
+
+(* declarations *)
+
+local
+
+fun decl_attrib decl phi = Thm.declaration_attribute (K (decl phi));
+
+fun add_decls add loc decl =
+  ProofContext.theory (change_locale loc
+    (fn (axiom, elems, params, decls, regs, intros, dests) =>
+      (axiom, elems, params, add (decl, stamp ()) decls, regs, intros, dests))) #>
+  add_thmss loc Thm.internalK
+    [((Binding.empty, [Attrib.internal (decl_attrib decl)]), [([Drule.dummy_thm], [])])];
+
+in
+
+val add_type_syntax = add_decls (apfst o cons);
+val add_term_syntax = add_decls (apsnd o cons);
+val add_declaration = add_decls (K I);
+
+fun declarations_of thy loc =
+  the_locale thy loc |> #decls |> apfst (map fst) |> apsnd (map fst);
+
+end;
+
+
+
+(** define locales **)
+
+(* predicate text *)
+(* CB: generate locale predicates and delta predicates *)
+
+local
+
+(* introN: name of theorems for introduction rules of locale and
+     delta predicates;
+   axiomsN: name of theorem set with destruct rules for locale predicates,
+     also name suffix of delta predicates. *)
+
+val introN = "intro";
+val axiomsN = "axioms";
+
+fun atomize_spec thy ts =
+  let
+    val t = Logic.mk_conjunction_balanced ts;
+    val body = ObjectLogic.atomize_term thy t;
+    val bodyT = Term.fastype_of body;
+  in
+    if bodyT = propT then (t, propT, Thm.reflexive (Thm.cterm_of thy t))
+    else (body, bodyT, ObjectLogic.atomize (Thm.cterm_of thy t))
+  end;
+
+fun aprop_tr' n c = (Syntax.constN ^ c, fn ctxt => fn args =>
+  if length args = n then
+    Syntax.const "_aprop" $
+      Term.list_comb (Syntax.free (Consts.extern (ProofContext.consts_of ctxt) c), args)
+  else raise Match);
+
+(* CB: define one predicate including its intro rule and axioms
+   - bname: predicate name
+   - parms: locale parameters
+   - defs: thms representing substitutions from defines elements
+   - ts: terms representing locale assumptions (not normalised wrt. defs)
+   - norm_ts: terms representing locale assumptions (normalised wrt. defs)
+   - thy: the theory
+*)
+
+fun def_pred bname parms defs ts norm_ts thy =
+  let
+    val name = Sign.full_bname thy bname;
+
+    val (body, bodyT, body_eq) = atomize_spec thy norm_ts;
+    val env = Term.add_free_names body [];
+    val xs = filter (member (op =) env o #1) parms;
+    val Ts = map #2 xs;
+    val extraTs =
+      (Term.add_tfrees body [] \\ fold Term.add_tfreesT Ts [])
+      |> Library.sort_wrt #1 |> map TFree;
+    val predT = map Term.itselfT extraTs ---> Ts ---> bodyT;
+
+    val args = map Logic.mk_type extraTs @ map Free xs;
+    val head = Term.list_comb (Const (name, predT), args);
+    val statement = ObjectLogic.ensure_propT thy head;
+
+    val ([pred_def], defs_thy) =
+      thy
+      |> bodyT = propT ? Sign.add_advanced_trfuns ([], [], [aprop_tr' (length args) name], [])
+      |> Sign.declare_const [] ((Binding.name bname, predT), NoSyn) |> snd
+      |> PureThy.add_defs false
+        [((Thm.def_name bname, Logic.mk_equals (head, body)), [Thm.kind_internal])];
+    val defs_ctxt = ProofContext.init defs_thy |> Variable.declare_term head;
+
+    val cert = Thm.cterm_of defs_thy;
+
+    val intro = Goal.prove_global defs_thy [] norm_ts statement (fn _ =>
+      MetaSimplifier.rewrite_goals_tac [pred_def] THEN
+      Tactic.compose_tac (false, body_eq RS Drule.equal_elim_rule1, 1) 1 THEN
+      Tactic.compose_tac (false,
+        Conjunction.intr_balanced (map (Thm.assume o cert) norm_ts), 0) 1);
+
+    val conjuncts =
+      (Drule.equal_elim_rule2 OF [body_eq,
+        MetaSimplifier.rewrite_rule [pred_def] (Thm.assume (cert statement))])
+      |> Conjunction.elim_balanced (length ts);
+    val axioms = ts ~~ conjuncts |> map (fn (t, ax) =>
+      Element.prove_witness defs_ctxt t
+       (MetaSimplifier.rewrite_goals_tac defs THEN
+        Tactic.compose_tac (false, ax, 0) 1));
+  in ((statement, intro, axioms), defs_thy) end;
+
+fun assumes_to_notes (Assumes asms) axms =
+      fold_map (fn (a, spec) => fn axs =>
+          let val (ps, qs) = chop (length spec) axs
+          in ((a, [(ps, [])]), qs) end) asms axms
+      |> apfst (curry Notes Thm.assumptionK)
+  | assumes_to_notes e axms = (e, axms);
+
+(* CB: the following two change only "new" elems, these have identifier ("", _). *)
+
+(* turn Assumes into Notes elements *)
+
+fun change_assumes_elemss axioms elemss =
+  let
+    val satisfy = Element.morph_ctxt (Element.satisfy_morphism axioms);
+    fun change (id as ("", _), es) =
+          fold_map assumes_to_notes (map satisfy es)
+          #-> (fn es' => pair (id, es'))
+      | change e = pair e;
+  in
+    fst (fold_map change elemss (map Element.conclude_witness axioms))
+  end;
+
+(* adjust hyps of Notes elements *)
+
+fun change_elemss_hyps axioms elemss =
+  let
+    val satisfy = Element.morph_ctxt (Element.satisfy_morphism axioms);
+    fun change (id as ("", _), es) = (id, map (fn e as Notes _ => satisfy e | e => e) es)
+      | change e = e;
+  in map change elemss end;
+
+in
+
+(* CB: main predicate definition function *)
+
+fun define_preds pname (parms, ((exts, exts'), (ints, ints')), defs) elemss thy =
+  let
+    val ((elemss', more_ts), a_elem, a_intro, thy'') =
+      if null exts then ((elemss, []), [], [], thy)
+      else
+        let
+          val aname = if null ints then pname else pname ^ "_" ^ axiomsN;
+          val ((statement, intro, axioms), thy') =
+            thy
+            |> def_pred aname parms defs exts exts';
+          val elemss' = change_assumes_elemss axioms elemss;
+          val a_elem = [(("", []),
+            [Assumes [((Binding.name (pname ^ "_" ^ axiomsN), []), [(statement, [])])]])];
+          val (_, thy'') =
+            thy'
+            |> Sign.add_path aname
+            |> Sign.no_base_names
+            |> PureThy.note_thmss Thm.internalK [((Binding.name introN, []), [([intro], [])])]
+            ||> Sign.restore_naming thy';
+        in ((elemss', [statement]), a_elem, [intro], thy'') end;
+    val (predicate, stmt', elemss'', b_intro, thy'''') =
+      if null ints then (([], []), more_ts, elemss' @ a_elem, [], thy'')
+      else
+        let
+          val ((statement, intro, axioms), thy''') =
+            thy''
+            |> def_pred pname parms defs (ints @ more_ts) (ints' @ more_ts);
+          val cstatement = Thm.cterm_of thy''' statement;
+          val elemss'' = change_elemss_hyps axioms elemss';
+          val b_elem = [(("", []),
+               [Assumes [((Binding.name (pname ^ "_" ^ axiomsN), []), [(statement, [])])]])];
+          val (_, thy'''') =
+            thy'''
+            |> Sign.add_path pname
+            |> Sign.no_base_names
+            |> PureThy.note_thmss Thm.internalK
+                 [((Binding.name introN, []), [([intro], [])]),
+                  ((Binding.name axiomsN, []),
+                    [(map (Drule.standard o Element.conclude_witness) axioms, [])])]
+            ||> Sign.restore_naming thy''';
+        in (([cstatement], axioms), [statement], elemss'' @ b_elem, [intro], thy'''') end;
+  in (((elemss'', predicate, stmt'), (a_intro, b_intro)), thy'''') end;
+
+end;
+
+
+(* add_locale(_i) *)
+
+local
+
+(* turn Defines into Notes elements, accumulate definition terms *)
+
+fun defines_to_notes is_ext thy (Defines defs) defns =
+    let
+      val defs' = map (fn (_, (def, _)) => (Attrib.empty_binding, (def, []))) defs
+      val notes = map (fn (a, (def, _)) =>
+        (a, [([assume (cterm_of thy def)], [])])) defs
+    in
+      (if is_ext then SOME (Notes (Thm.definitionK, notes)) else NONE, defns @ [Defines defs'])
+    end
+  | defines_to_notes _ _ e defns = (SOME e, defns);
+
+fun change_defines_elemss thy elemss defns =
+  let
+    fun change (id as (n, _), es) defns =
+        let
+          val (es', defns') = fold_map (defines_to_notes (n="") thy) es defns
+        in ((id, map_filter I es'), defns') end
+  in fold_map change elemss defns end;
+
+fun gen_add_locale prep_ctxt prep_expr
+    predicate_name bname raw_imports raw_body thy =
+    (* predicate_name: "" - locale with predicate named as locale
+        "name" - locale with predicate named "name" *)
+  let
+    val thy_ctxt = ProofContext.init thy;
+    val name = Sign.full_bname thy bname;
+    val _ = is_some (get_locale thy name) andalso
+      error ("Duplicate definition of locale " ^ quote name);
+
+    val (((import_ctxt, import_elemss), (body_ctxt, body_elemss, syn)),
+      text as (parms, ((_, exts'), _), defs)) =
+        prep_ctxt raw_imports raw_body thy_ctxt;
+    val elemss = import_elemss @ body_elemss |>
+      map_filter (fn ((id, Assumed axs), elems) => SOME (id, elems) | _ => NONE);
+
+    val extraTs = List.foldr OldTerm.add_term_tfrees [] exts' \\
+      List.foldr OldTerm.add_typ_tfrees [] (map snd parms);
+    val _ = if null extraTs then ()
+      else warning ("Additional type variable(s) in locale specification " ^ quote bname);
+
+    val predicate_name' = case predicate_name of "" => bname | _ => predicate_name;
+    val (elemss', defns) = change_defines_elemss thy elemss [];
+    val elemss'' = elemss' @ [(("", []), defns)];
+    val (((elemss''', predicate as (pred_statement, pred_axioms), stmt'), intros), thy') =
+      define_preds predicate_name' text elemss'' thy;
+    val regs = pred_axioms
+      |> fold_map (fn (id, elems) => fn wts => let
+             val ts = flat (map_filter (fn (Assumes asms) =>
+               SOME (maps (map #1 o #2) asms) | _ => NONE) elems);
+             val (wts1, wts2) = chop (length ts) wts;
+           in ((apsnd (map fst) id, wts1), wts2) end) elemss'''
+      |> fst
+      |> map_filter (fn (("", _), _) => NONE | e => SOME e);
+    fun axiomify axioms elemss =
+      (axioms, elemss) |> foldl_map (fn (axs, (id, elems)) => let
+                   val ts = flat (map_filter (fn (Assumes asms) =>
+                     SOME (maps (map #1 o #2) asms) | _ => NONE) elems);
+                   val (axs1, axs2) = chop (length ts) axs;
+                 in (axs2, ((id, Assumed axs1), elems)) end)
+      |> snd;
+    val ((_, facts), ctxt) = activate_facts true (K I)
+      (axiomify pred_axioms elemss''') (ProofContext.init thy');
+    val view_ctxt = Assumption.add_view thy_ctxt pred_statement ctxt;
+    val export = Thm.close_derivation o Goal.norm_result o
+      singleton (ProofContext.export view_ctxt thy_ctxt);
+    val facts' = facts |> map (fn (a, ths) => ((a, []), [(map export ths, [])]));
+    val elems' = maps #2 (filter (fn ((s, _), _) => s = "") elemss''');
+    val elems'' = map_filter (fn (Fixes _) => NONE | e => SOME e) elems';
+    val axs' = map (Element.assume_witness thy') stmt';
+    val loc_ctxt = thy'
+      |> Sign.add_path bname
+      |> Sign.no_base_names
+      |> PureThy.note_thmss Thm.assumptionK facts' |> snd
+      |> Sign.restore_naming thy'
+      |> register_locale bname {axiom = axs',
+        elems = map (fn e => (e, stamp ())) elems'',
+        params = params_of elemss''' |> map (fn (x, SOME T) => ((x, T), the (Symtab.lookup syn x))),
+        decls = ([], []),
+        regs = regs,
+        intros = intros,
+        dests = map Element.conclude_witness pred_axioms}
+      |> init name;
+  in (name, loc_ctxt) end;
+
+in
+
+val add_locale = gen_add_locale cert_context (K I);
+val add_locale_cmd = gen_add_locale read_context intern_expr "";
+
+end;
+
+val _ = Context.>> (Context.map_theory
+ (add_locale "" "var" empty [Fixes [(Binding.name (Name.internal "x"), NONE, NoSyn)]] #>
+  snd #> ProofContext.theory_of #>
+  add_locale "" "struct" empty [Fixes [(Binding.name (Name.internal "S"), NONE, Structure)]] #>
+  snd #> ProofContext.theory_of));
+
+
+
+
+(** Normalisation of locale statements ---
+    discharges goals implied by interpretations **)
+
+local
+
+fun locale_assm_intros thy =
+  Symtab.fold (fn (_, {intros = (a, _), ...}) => fn intros => (a @ intros))
+    (#2 (LocalesData.get thy)) [];
+fun locale_base_intros thy =
+  Symtab.fold (fn (_, {intros = (_, b), ...}) => fn intros => (b @ intros))
+    (#2 (LocalesData.get thy)) [];
+
+fun all_witnesses ctxt =
+  let
+    val thy = ProofContext.theory_of ctxt;
+    fun get registrations = Symtab.fold (fn (_, regs) => fn thms =>
+        (Registrations.dest thy regs |> map (fn (_, (_, (exp, _), wits, _)) =>
+          map (Element.conclude_witness #> Morphism.thm exp) wits) |> flat) @ thms)
+      registrations [];
+  in get (RegistrationsData.get (Context.Proof ctxt)) end;
+
+in
+
+fun intro_locales_tac eager ctxt facts st =
+  let
+    val wits = all_witnesses ctxt;
+    val thy = ProofContext.theory_of ctxt;
+    val intros = locale_base_intros thy @ (if eager then locale_assm_intros thy else []);
+  in
+    Method.intros_tac (wits @ intros) facts st
+  end;
+
+end;
+
+
+(** Interpretation commands **)
+
+local
+
+(* extract proof obligations (assms and defs) from elements *)
+
+fun extract_asms_elems ((id, Assumed _), elems) = (id, maps Element.prems_of elems)
+  | extract_asms_elems ((id, Derived _), _) = (id, []);
+
+
+(* activate instantiated facts in theory or context *)
+
+fun gen_activate_facts_elemss mk_ctxt note attrib put_reg add_wit add_eqn
+        phi_name all_elemss pss propss eq_attns (exp, imp) thmss thy_ctxt =
+  let
+    val ctxt = mk_ctxt thy_ctxt;
+    fun get_reg thy_ctxt = get_local_registration (mk_ctxt thy_ctxt);
+    fun test_reg thy_ctxt = test_local_registration (mk_ctxt thy_ctxt);
+
+    val (all_propss, eq_props) = chop (length all_elemss) propss;
+    val (all_thmss, eq_thms) = chop (length all_elemss) thmss;
+
+    (* Filter out fragments already registered. *)
+
+    val (new_elemss, xs) = split_list (filter_out (fn (((id, _), _), _) =>
+          test_reg thy_ctxt id) (all_elemss ~~ (pss ~~ (all_propss ~~ all_thmss))));
+    val (new_pss, ys) = split_list xs;
+    val (new_propss, new_thmss) = split_list ys;
+
+    val thy_ctxt' = thy_ctxt
+      (* add registrations *)
+      |> fold2 (fn ((id as (loc, _), _), _) => fn ps => put_reg id (phi_name, param_prefix loc ps) (exp, imp))
+           new_elemss new_pss
+      (* add witnesses of Assumed elements (only those generate proof obligations) *)
+      |> fold2 (fn (id, _) => fold (add_wit id)) new_propss new_thmss
+      (* add equations *)
+      |> fold2 (fn (id, _) => fold (add_eqn id)) eq_props
+          ((map o map) (Drule.abs_def o LocalDefs.meta_rewrite_rule ctxt o
+            Element.conclude_witness) eq_thms);
+
+    val prems = flat (map_filter
+          (fn ((id, Assumed _), _) => Option.map #2 (get_reg thy_ctxt' imp id)
+            | ((_, Derived _), _) => NONE) all_elemss);
+
+    val thy_ctxt'' = thy_ctxt'
+      (* add witnesses of Derived elements *)
+      |> fold (fn (id, thms) => fold
+           (add_wit id o Element.morph_witness (Element.satisfy_morphism prems)) thms)
+         (map_filter (fn ((_, Assumed _), _) => NONE
+            | ((id, Derived thms), _) => SOME (id, thms)) new_elemss)
+
+    fun activate_elem phi_name param_prfx insts prems eqns exp (Notes (kind, facts)) thy_ctxt =
+        let
+          val ctxt = mk_ctxt thy_ctxt;
+          val thy = ProofContext.theory_of ctxt;
+          val facts' = facts
+            |> activate_note thy phi_name param_prfx
+                 (attrib thy_ctxt) insts prems eqns exp;
+        in 
+          thy_ctxt
+          |> note kind facts'
+          |> snd
+        end
+      | activate_elem _ _ _ _ _ _ _ thy_ctxt = thy_ctxt;
+
+    fun activate_elems (((loc, ext_ts), _), _) ps thy_ctxt =
+      let
+        val ctxt = mk_ctxt thy_ctxt;
+        val thy = ProofContext.theory_of ctxt;
+        val {params, elems, ...} = the_locale thy loc;
+        val parms = map fst params;
+        val param_prfx = param_prefix loc ps;
+        val ids = flatten (ProofContext.init thy, intern_expr thy)
+          (([], Symtab.empty), Expr (Locale loc)) |> fst |> fst;
+        val (insts, prems, eqns) = collect_witnesses ctxt imp parms ids ext_ts;
+      in
+        thy_ctxt
+        |> fold (activate_elem phi_name param_prfx insts prems eqns exp o fst) elems
+      end;
+
+  in
+    thy_ctxt''
+    (* add equations as lemmas to context *)
+    |> (fold2 o fold2) (fn attn => fn thm => snd o yield_singleton (note Thm.lemmaK)
+         ((apsnd o map) (attrib thy_ctxt'') attn, [([Element.conclude_witness thm], [])]))
+            (unflat eq_thms eq_attns) eq_thms
+    (* add interpreted facts *)
+    |> fold2 activate_elems new_elemss new_pss
+  end;
+
+fun global_activate_facts_elemss x = gen_activate_facts_elemss
+  ProofContext.init
+  global_note_qualified
+  Attrib.attribute_i
+  put_global_registration
+  add_global_witness
+  add_global_equation
+  x;
+
+fun local_activate_facts_elemss x = gen_activate_facts_elemss
+  I
+  local_note_qualified
+  (Attrib.attribute_i o ProofContext.theory_of)
+  put_local_registration
+  add_local_witness
+  add_local_equation
+  x;
+
+fun prep_instantiations parse_term parse_prop ctxt parms (insts, eqns) =
+  let
+    (* parameters *)
+    val (parm_names, parm_types) = parms |> split_list
+      ||> map (TypeInfer.paramify_vars o Logic.varifyT);
+    val type_parms = fold Term.add_tvarsT parm_types [] |> map (Logic.mk_type o TVar);
+    val type_parm_names = fold Term.add_tfreesT (map snd parms) [] |> map fst;
+
+    (* parameter instantiations *)
+    val d = length parms - length insts;
+    val insts =
+      if d < 0 then error "More arguments than parameters in instantiation."
+      else insts @ replicate d NONE;
+    val (given_ps, given_insts) =
+      ((parm_names ~~ parm_types) ~~ insts) |> map_filter
+          (fn (_, NONE) => NONE
+            | ((n, T), SOME inst) => SOME ((n, T), inst))
+        |> split_list;
+    val (given_parm_names, given_parm_types) = given_ps |> split_list;
+
+    (* parse insts / eqns *)
+    val given_insts' = map (parse_term ctxt) given_insts;
+    val eqns' = map (parse_prop ctxt) eqns;
+
+    (* type inference and contexts *)
+    val arg = type_parms @ map2 TypeInfer.constrain given_parm_types given_insts' @ eqns';
+    val res = Syntax.check_terms ctxt arg;
+    val ctxt' = ctxt |> fold Variable.auto_fixes res;
+
+    (* instantiation *)
+    val (type_parms'', res') = chop (length type_parms) res;
+    val (given_insts'', eqns'') = chop (length given_insts) res';
+    val instT = Symtab.make (type_parm_names ~~ map Logic.dest_type type_parms'');
+    val inst = Symtab.make (given_parm_names ~~ given_insts'');
+
+    (* export from eigencontext *)
+    val export = Variable.export_morphism ctxt' ctxt;
+
+    (* import, its inverse *)
+    val domT = fold Term.add_tfrees res [] |> map TFree;
+    val importT = domT |> map (fn x => (Morphism.typ export x, x))
+      |> map_filter (fn (TFree _, _) => NONE  (* fixed point of export *)
+               | (TVar y, x) => SOME (fst y, x)
+               | _ => error "internal: illegal export in interpretation")
+      |> Vartab.make;
+    val dom = fold Term.add_frees res [] |> map Free;
+    val imprt = dom |> map (fn x => (Morphism.term export x, x))
+      |> map_filter (fn (Free _, _) => NONE  (* fixed point of export *)
+               | (Var y, x) => SOME (fst y, x)
+               | _ => error "internal: illegal export in interpretation")
+      |> Vartab.make;
+  in (((instT, inst), eqns''), (export, ((importT, domT), (imprt, dom)))) end;
+
+val read_instantiations = prep_instantiations Syntax.parse_term Syntax.parse_prop;
+val check_instantiations = prep_instantiations (K I) (K I);
+
+fun gen_prep_registration mk_ctxt test_reg activate
+    prep_attr prep_expr prep_insts
+    thy_ctxt phi_name raw_expr raw_insts =
+  let
+    val ctxt = mk_ctxt thy_ctxt;
+    val thy = ProofContext.theory_of ctxt;
+    val ctxt' = ProofContext.init thy;
+    fun prep_attn attn = (apsnd o map)
+      (Attrib.crude_closure ctxt o Args.assignable o prep_attr thy) attn;
+
+    val expr = prep_expr thy raw_expr;
+
+    val pts = params_of_expr ctxt' [] expr ([], Symtab.empty, Symtab.empty);
+    val params_ids = make_params_ids (#1 pts);
+    val raw_params_elemss = make_raw_params_elemss pts;
+    val ((ids, _), raw_elemss) = flatten (ctxt', I) (([], Symtab.empty), Expr expr);
+    val ((parms, all_elemss, _), (_, (_, defs, _))) =
+      read_elemss false ctxt' [] (raw_params_elemss @ raw_elemss) [];
+
+    (** compute instantiation **)
+
+    (* consistency check: equations need to be stored in a particular locale,
+       therefore if equations are present locale expression must be a name *)
+
+    val _ = case (expr, snd raw_insts) of
+        (Locale _, _) => () | (_, []) => ()
+      | (_, _) => error "Interpretations with `where' only permitted if locale expression is a name.";
+
+    (* read or certify instantiation *)
+    val (raw_insts', raw_eqns) = raw_insts;
+    val (raw_eq_attns, raw_eqns') = split_list raw_eqns;
+    val (((instT, inst1), eqns), morphs) = prep_insts ctxt parms (raw_insts', raw_eqns');
+    val eq_attns = map prep_attn raw_eq_attns;
+
+    (* defined params without given instantiation *)
+    val not_given = filter_out (Symtab.defined inst1 o fst) parms;
+    fun add_def (p, pT) inst =
+      let
+        val (t, T) = case find_first (fn (Free (a, _), _) => a = p) defs of
+               NONE => error ("Instance missing for parameter " ^ quote p)
+             | SOME (Free (_, T), t) => (t, T);
+        val d = Element.inst_term (instT, inst) t;
+      in Symtab.update_new (p, d) inst end;
+    val inst2 = fold add_def not_given inst1;
+    val inst_morphism = Element.inst_morphism thy (instT, inst2);
+    (* Note: insts contain no vars. *)
+
+    (** compute proof obligations **)
+
+    (* restore "small" ids *)
+    val ids' = map (fn ((n, ps), (_, mode)) =>
+          ((n, map (fn p => Free (p, (the o AList.lookup (op =) parms) p)) ps), mode))
+        ids;
+    val (_, all_elemss') = chop (length raw_params_elemss) all_elemss
+    (* instantiate ids and elements *)
+    val inst_elemss = (ids' ~~ all_elemss') |> map (fn (((n, ps), _), ((_, mode), elems)) =>
+      ((n, map (Morphism.term (inst_morphism $> fst morphs)) ps),
+        map (fn Int e => Element.morph_ctxt inst_morphism e) elems)
+      |> apfst (fn id => (id, map_mode (map (Element.morph_witness inst_morphism)) mode)));
+
+    (* equations *)
+    val eqn_elems = if null eqns then []
+      else [(Library.last_elem inst_elemss |> fst |> fst, eqns)];
+
+    val propss = map extract_asms_elems inst_elemss @ eqn_elems;
+
+  in
+    (propss, activate phi_name inst_elemss (map (snd o fst) ids) propss eq_attns morphs, morphs)
+  end;
+
+fun gen_prep_global_registration mk_ctxt = gen_prep_registration ProofContext.init
+  test_global_registration
+  global_activate_facts_elemss mk_ctxt;
+
+fun gen_prep_local_registration mk_ctxt = gen_prep_registration I
+  test_local_registration
+  local_activate_facts_elemss mk_ctxt;
+
+val prep_global_registration = gen_prep_global_registration
+  (K I) (K I) check_instantiations;
+val prep_global_registration_cmd = gen_prep_global_registration
+  Attrib.intern_src intern_expr read_instantiations;
+
+val prep_local_registration = gen_prep_local_registration
+  (K I) (K I) check_instantiations;
+val prep_local_registration_cmd = gen_prep_local_registration
+  Attrib.intern_src intern_expr read_instantiations;
+
+fun prep_registration_in_locale target expr thy =
+  (* target already in internal form *)
+  let
+    val ctxt = ProofContext.init thy;
+    val ((raw_target_ids, target_syn), _) = flatten (ctxt, I)
+        (([], Symtab.empty), Expr (Locale target));
+    val fixed = the_locale thy target |> #params |> map #1;
+    val ((all_ids, syn), raw_elemss) = flatten (ctxt, intern_expr thy)
+        ((raw_target_ids, target_syn), Expr expr);
+    val (target_ids, ids) = chop (length raw_target_ids) all_ids;
+    val ((parms, elemss, _), _) = read_elemss false ctxt fixed raw_elemss [];
+
+    (** compute proof obligations **)
+
+    (* restore "small" ids, with mode *)
+    val ids' = map (apsnd snd) ids;
+    (* remove Int markers *)
+    val elemss' = map (fn (_, es) =>
+        map (fn Int e => e) es) elemss
+    (* extract assumptions and defs *)
+    val ids_elemss = ids' ~~ elemss';
+    val propss = map extract_asms_elems ids_elemss;
+
+    (** activation function:
+        - add registrations to the target locale
+        - add induced registrations for all global registrations of
+          the target, unless already present
+        - add facts of induced registrations to theory **)
+
+    fun activate thmss thy =
+      let
+        val satisfy = Element.satisfy_thm (flat thmss);
+        val ids_elemss_thmss = ids_elemss ~~ thmss;
+        val regs = get_global_registrations thy target;
+
+        fun activate_id (((id, Assumed _), _), thms) thy =
+            thy |> put_registration_in_locale target id
+                |> fold (add_witness_in_locale target id) thms
+          | activate_id _ thy = thy;
+
+        fun activate_reg (ext_ts, ((phi_name, param_prfx), (exp, imp), _, _)) thy =
+          let
+            val (insts, wits, _) = collect_witnesses (ProofContext.init thy) imp fixed target_ids ext_ts;
+            val inst_parms = map (the o AList.lookup (op =) (map #1 fixed ~~ ext_ts));
+            val disch = Element.satisfy_thm wits;
+            val new_elemss = filter (fn (((name, ps), _), _) =>
+                not (test_global_registration thy (name, inst_parms ps))) (ids_elemss);
+            fun activate_assumed_id (((_, Derived _), _), _) thy = thy
+              | activate_assumed_id ((((name, ps), Assumed _), _), thms) thy = let
+                val ps' = inst_parms ps;
+              in
+                if test_global_registration thy (name, ps')
+                then thy
+                else thy
+                  |> put_global_registration (name, ps') (phi_name, param_prefix name ps) (exp, imp)
+                  |> fold (fn witn => fn thy => add_global_witness (name, ps')
+                     (Element.morph_witness (Element.inst_morphism thy insts) witn) thy) thms
+              end;
+
+            fun activate_derived_id ((_, Assumed _), _) thy = thy
+              | activate_derived_id (((name, ps), Derived ths), _) thy = let
+                val ps' = inst_parms ps;
+              in
+                if test_global_registration thy (name, ps')
+                then thy
+                else thy
+                  |> put_global_registration (name, ps') (phi_name, param_prefix name ps) (exp, imp)
+                  |> fold (fn witn => fn thy => add_global_witness (name, ps')
+                       (witn |> Element.map_witness (fn (t, th) =>  (* FIXME *)
+                       (Element.inst_term insts t,
+                        disch (Element.inst_thm thy insts (satisfy th))))) thy) ths
+              end;
+
+            fun activate_elem (loc, ps) (Notes (kind, facts)) thy =
+                let
+                  val att_morphism =
+                    Morphism.binding_morphism (name_morph phi_name param_prfx) $>
+                    Morphism.thm_morphism satisfy $>
+                    Element.inst_morphism thy insts $>
+                    Morphism.thm_morphism disch;
+                  val facts' = facts
+                    |> Attrib.map_facts (Attrib.attribute_i thy o Args.morph_values att_morphism)
+                    |> (map o apsnd o map o apfst o map) (disch o Element.inst_thm thy insts o satisfy)
+                    |> (map o apfst o apfst) (name_morph phi_name param_prfx);
+                in
+                  thy
+                  |> global_note_qualified kind facts'
+                  |> snd
+                end
+              | activate_elem _ _ thy = thy;
+
+            fun activate_elems ((id, _), elems) thy = fold (activate_elem id) elems thy;
+
+          in thy |> fold activate_assumed_id ids_elemss_thmss
+                 |> fold activate_derived_id ids_elemss
+                 |> fold activate_elems new_elemss end;
+      in
+        thy |> fold activate_id ids_elemss_thmss
+            |> fold activate_reg regs
+      end;
+
+  in (propss, activate) end;
+
+fun prep_propp propss = propss |> map (fn (_, props) =>
+  map (rpair [] o Element.mark_witness) props);
+
+fun prep_result propps thmss =
+  ListPair.map (fn ((_, props), thms) => map2 Element.make_witness props thms) (propps, thmss);
+
+fun gen_interpretation prep_registration after_qed prfx raw_expr raw_insts thy =
+  let
+    val (propss, activate, morphs) = prep_registration thy prfx raw_expr raw_insts;
+    fun after_qed' results =
+      ProofContext.theory (activate (prep_result propss results))
+      #> after_qed;
+  in
+    thy
+    |> ProofContext.init
+    |> Proof.theorem_i NONE after_qed' (prep_propp propss)
+    |> Element.refine_witness
+    |> Seq.hd
+    |> pair morphs
+  end;
+
+fun gen_interpret prep_registration after_qed name_morph expr insts int state =
+  let
+    val _ = Proof.assert_forward_or_chain state;
+    val ctxt = Proof.context_of state;
+    val (propss, activate, morphs) = prep_registration ctxt name_morph expr insts;
+    fun after_qed' results =
+      Proof.map_context (K (ctxt |> activate (prep_result propss results)))
+      #> Proof.put_facts NONE
+      #> after_qed;
+  in
+    state
+    |> Proof.local_goal (ProofDisplay.print_results int) (K I) ProofContext.bind_propp_i
+      "interpret" NONE after_qed' (map (pair (Binding.empty, [])) (prep_propp propss))
+    |> Element.refine_witness |> Seq.hd
+    |> pair morphs
+  end;
+
+fun standard_name_morph interp_prfx b =
+  if Binding.is_empty b then b
+  else Binding.map_prefix (fn ((lprfx, _) :: pprfx) =>
+    fold (Binding.add_prefix false o fst) pprfx
+    #> interp_prfx <> "" ? Binding.add_prefix true interp_prfx
+    #> Binding.add_prefix false lprfx
+  ) b;
+
+in
+
+val interpretation = gen_interpretation prep_global_registration;
+fun interpretation_cmd interp_prfx = snd ooo gen_interpretation prep_global_registration_cmd
+  I (standard_name_morph interp_prfx);
+
+fun interpretation_in_locale after_qed (raw_target, expr) thy =
+  let
+    val target = intern thy raw_target;
+    val (propss, activate) = prep_registration_in_locale target expr thy;
+    val raw_propp = prep_propp propss;
+
+    val (_, _, goal_ctxt, propp) = thy
+      |> ProofContext.init
+      |> cert_context_statement (SOME target) [] raw_propp;
+
+    fun after_qed' results =
+      ProofContext.theory (activate (prep_result propss results))
+      #> after_qed;
+  in
+    goal_ctxt
+    |> Proof.theorem_i NONE after_qed' propp
+    |> Element.refine_witness |> Seq.hd
+  end;
+
+val interpret = gen_interpret prep_local_registration;
+fun interpret_cmd interp_prfx = snd oooo gen_interpret prep_local_registration_cmd
+  I (standard_name_morph interp_prfx);
+
+end;
+
+end;
--- a/src/Pure/Isar/outer_keyword.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Isar/outer_keyword.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/outer_keyword.ML
-    ID:         $Id$
     Author:     Makarius
 
 Isar command keyword classification and global keyword tables.
@@ -47,9 +46,12 @@
   val report: unit -> unit
   val keyword: string -> unit
   val command: string -> T -> unit
+  val is_diag: string -> bool
+  val is_control: string -> bool
+  val is_regular: string -> bool
+  val is_theory_begin: string -> bool
   val is_theory: string -> bool
   val is_proof: string -> bool
-  val is_diag: string -> bool
   val is_theory_goal: string -> bool
   val is_proof_goal: string -> bool
   val is_qed: string -> bool
@@ -174,6 +176,12 @@
     NONE => false
   | SOME k => member (op = o pairself kind_of) ks k);
 
+val is_diag = command_category [diag];
+val is_control = command_category [control];
+fun is_regular name = not (is_diag name orelse is_control name);
+
+val is_theory_begin = command_category [thy_begin];
+
 val is_theory = command_category
   [thy_begin, thy_switch, thy_end, thy_heading, thy_decl, thy_script, thy_goal];
 
@@ -181,8 +189,6 @@
   [qed, qed_block, qed_global, prf_heading, prf_goal, prf_block, prf_open, prf_close,
     prf_chain, prf_decl, prf_asm, prf_asm_goal, prf_script];
 
-val is_diag = command_category [diag];
-
 val is_theory_goal = command_category [thy_goal];
 val is_proof_goal = command_category [prf_goal, prf_asm_goal];
 val is_qed = command_category [qed, qed_block];
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Isar/outer_keyword.scala	Mon Jan 19 21:20:18 2009 +0100
@@ -0,0 +1,46 @@
+/*  Title:      Pure/Isar/outer_keyword.scala
+    Author:     Makarius
+
+Isar command keyword classification.
+*/
+
+package isabelle
+
+object OuterKeyword {
+
+  val MINOR = "minor"
+  val CONTROL = "control"
+  val DIAG = "diag"
+  val THY_BEGIN = "theory-begin"
+  val THY_SWITCH = "theory-switch"
+  val THY_END = "theory-end"
+  val THY_HEADING = "theory-heading"
+  val THY_DECL = "theory-decl"
+  val THY_SCRIPT = "theory-script"
+  val THY_GOAL = "theory-goal"
+  val QED = "qed"
+  val QED_BLOCK = "qed-block"
+  val QED_GLOBAL = "qed-global"
+  val PRF_HEADING = "proof-heading"
+  val PRF_GOAL = "proof-goal"
+  val PRF_BLOCK = "proof-block"
+  val PRF_OPEN = "proof-open"
+  val PRF_CLOSE = "proof-close"
+  val PRF_CHAIN = "proof-chain"
+  val PRF_DECL = "proof-decl"
+  val PRF_ASM = "proof-asm"
+  val PRF_ASM_GOAL = "proof-asm-goal"
+  val PRF_SCRIPT = "proof-script"
+
+  val minor = Set(MINOR)
+  val control = Set(CONTROL)
+  val diag = Set(DIAG)
+  val heading = Set(THY_HEADING, PRF_HEADING)
+  val theory1 = Set(THY_BEGIN, THY_SWITCH, THY_END)
+  val theory2 = Set(THY_DECL, THY_GOAL)
+  val proof1 =
+    Set(QED, QED_BLOCK, QED_GLOBAL, PRF_GOAL, PRF_BLOCK, PRF_OPEN, PRF_CLOSE, PRF_CHAIN, PRF_DECL)
+  val proof2 = Set(PRF_ASM, PRF_ASM_GOAL)
+  val improper = Set(THY_SCRIPT, PRF_SCRIPT)
+}
+
--- a/src/Pure/Isar/outer_parse.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Isar/outer_parse.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/outer_parse.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Generic parsers for Isabelle/Isar outer syntax.
@@ -8,48 +7,49 @@
 signature OUTER_PARSE =
 sig
   type token = OuterLex.token
+  type 'a parser = token list -> 'a * token list
   val group: string -> (token list -> 'a) -> token list -> 'a
   val !!! : (token list -> 'a) -> token list -> 'a
   val !!!! : (token list -> 'a) -> token list -> 'a
   val triple1: ('a * 'b) * 'c -> 'a * 'b * 'c
   val triple2: 'a * ('b * 'c) -> 'a * 'b * 'c
   val triple_swap: ('a * 'b) * 'c -> ('a * 'c) * 'b
-  val not_eof: token list -> token * token list
+  val not_eof: token parser
   val position: (token list -> 'a * 'b) -> token list -> ('a * Position.T) * 'b
-  val command: token list -> string * token list
-  val keyword: token list -> string * token list
-  val short_ident: token list -> string * token list
-  val long_ident: token list -> string * token list
-  val sym_ident: token list -> string * token list
-  val minus: token list -> string * token list
-  val term_var: token list -> string * token list
-  val type_ident: token list -> string * token list
-  val type_var: token list -> string * token list
-  val number: token list -> string * token list
-  val string: token list -> string * token list
-  val alt_string: token list -> string * token list
-  val verbatim: token list -> string * token list
-  val sync: token list -> string * token list
-  val eof: token list -> string * token list
-  val keyword_with: (string -> bool) -> token list -> string * token list
-  val keyword_ident_or_symbolic: token list -> string * token list
-  val $$$ : string -> token list -> string * token list
-  val reserved: string -> token list -> string * token list
-  val semicolon: token list -> string * token list
-  val underscore: token list -> string * token list
-  val maybe: (token list -> 'a * token list) -> token list -> 'a option * token list
-  val tag_name: token list -> string * token list
-  val tags: token list -> string list * token list
-  val opt_unit: token list -> unit * token list
-  val opt_keyword: string -> token list -> bool * token list
-  val begin: token list -> string * token list
-  val opt_begin: token list -> bool * token list
-  val nat: token list -> int * token list
-  val int: token list -> int * token list
-  val enum: string -> (token list -> 'a * token list) -> token list -> 'a list * token list
-  val enum1: string -> (token list -> 'a * token list) -> token list -> 'a list * token list
-  val and_list: (token list -> 'a * token list) -> token list -> 'a list * token list
-  val and_list1: (token list -> 'a * token list) -> token list -> 'a list * token list
+  val command: string parser
+  val keyword: string parser
+  val short_ident: string parser
+  val long_ident: string parser
+  val sym_ident: string parser
+  val minus: string parser
+  val term_var: string parser
+  val type_ident: string parser
+  val type_var: string parser
+  val number: string parser
+  val string: string parser
+  val alt_string: string parser
+  val verbatim: string parser
+  val sync: string parser
+  val eof: string parser
+  val keyword_with: (string -> bool) -> string parser
+  val keyword_ident_or_symbolic: string parser
+  val $$$ : string -> string parser
+  val reserved: string -> string parser
+  val semicolon: string parser
+  val underscore: string parser
+  val maybe: 'a parser -> 'a option parser
+  val tag_name: string parser
+  val tags: string list parser
+  val opt_unit: unit parser
+  val opt_keyword: string -> bool parser
+  val begin: string parser
+  val opt_begin: bool parser
+  val nat: int parser
+  val int: int parser
+  val enum: string -> 'a parser -> 'a list parser
+  val enum1: string -> 'a parser -> 'a list parser
+  val and_list: 'a parser -> 'a list parser
+  val and_list1: 'a parser -> 'a list parser
   val enum': string -> ('a * token list -> 'b * ('a * token list)) ->
     'a * token list -> 'b list * ('a * token list)
   val enum1': string -> ('a * token list -> 'b * ('a * token list)) ->
@@ -58,46 +58,44 @@
     'a * token list -> 'b list * ('a * token list)
   val and_list1': ('a * token list -> 'b * ('a * token list)) ->
     'a * token list -> 'b list * ('a * token list)
-  val list: (token list -> 'a * token list) -> token list -> 'a list * token list
-  val list1: (token list -> 'a * token list) -> token list -> 'a list * token list
-  val name: token list -> bstring * token list
-  val binding: token list -> Binding.T * token list
-  val xname: token list -> xstring * token list
-  val text: token list -> string * token list
-  val path: token list -> Path.T * token list
-  val parname: token list -> string * token list
-  val parbinding: token list -> Binding.T * token list
-  val sort: token list -> string * token list
-  val arity: token list -> (string * string list * string) * token list
-  val multi_arity: token list -> (string list * string list * string) * token list
-  val type_args: token list -> string list * token list
-  val typ_group: token list -> string * token list
-  val typ: token list -> string * token list
-  val mixfix: token list -> mixfix * token list
-  val mixfix': token list -> mixfix * token list
-  val opt_infix: token list -> mixfix * token list
-  val opt_infix': token list -> mixfix * token list
-  val opt_mixfix: token list -> mixfix * token list
-  val opt_mixfix': token list -> mixfix * token list
-  val where_: token list -> string * token list
-  val const: token list -> (string * string * mixfix) * token list
-  val params: token list -> (Binding.T * string option) list * token list
-  val simple_fixes: token list -> (Binding.T * string option) list * token list
-  val fixes: token list -> (Binding.T * string option * mixfix) list * token list
-  val for_fixes: token list -> (Binding.T * string option * mixfix) list * token list
-  val for_simple_fixes: token list -> (Binding.T * string option) list * token list
-  val ML_source: token list -> (SymbolPos.text * Position.T) * token list
-  val doc_source: token list -> (SymbolPos.text * Position.T) * token list
-  val properties: token list -> Properties.T * token list
-  val props_text: token list -> (Position.T * string) * token list
-  val term_group: token list -> string * token list
-  val prop_group: token list -> string * token list
-  val term: token list -> string * token list
-  val prop: token list -> string * token list
-  val propp: token list -> (string * string list) * token list
-  val termp: token list -> (string * string list) * token list
-  val target: token list -> xstring * token list
-  val opt_target: token list -> xstring option * token list
+  val list: 'a parser -> 'a list parser
+  val list1: 'a parser -> 'a list parser
+  val name: bstring parser
+  val binding: Binding.T parser
+  val xname: xstring parser
+  val text: string parser
+  val path: Path.T parser
+  val parname: string parser
+  val parbinding: Binding.T parser
+  val sort: string parser
+  val arity: (string * string list * string) parser
+  val multi_arity: (string list * string list * string) parser
+  val type_args: string list parser
+  val typ_group: string parser
+  val typ: string parser
+  val mixfix: mixfix parser
+  val mixfix': mixfix parser
+  val opt_infix: mixfix parser
+  val opt_infix': mixfix parser
+  val opt_mixfix: mixfix parser
+  val opt_mixfix': mixfix parser
+  val where_: string parser
+  val const: (string * string * mixfix) parser
+  val params: (Binding.T * string option) list parser
+  val simple_fixes: (Binding.T * string option) list parser
+  val fixes: (Binding.T * string option * mixfix) list parser
+  val for_fixes: (Binding.T * string option * mixfix) list parser
+  val for_simple_fixes: (Binding.T * string option) list parser
+  val ML_source: (SymbolPos.text * Position.T) parser
+  val doc_source: (SymbolPos.text * Position.T) parser
+  val term_group: string parser
+  val prop_group: string parser
+  val term: string parser
+  val prop: string parser
+  val propp: (string * string list) parser
+  val termp: (string * string list) parser
+  val target: xstring parser
+  val opt_target: xstring option parser
 end;
 
 structure OuterParse: OUTER_PARSE =
@@ -106,6 +104,8 @@
 structure T = OuterLex;
 type token = T.token;
 
+type 'a parser = token list -> 'a * token list;
+
 
 (** error handling **)
 
@@ -310,12 +310,6 @@
 val ML_source = source_position (group "ML source" text);
 val doc_source = source_position (group "document source" text);
 
-val properties = $$$ "(" |-- !!! (list1 (string -- ($$$ "=" |-- string)) --| $$$ ")");
-
-val props_text =
-  Scan.optional properties [] -- position string >> (fn (props, (str, pos)) =>
-    (Position.of_properties (Position.default_properties pos props), str));
-
 
 (* terms *)
 
--- a/src/Pure/Isar/outer_syntax.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Isar/outer_syntax.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/outer_syntax.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 The global Isabelle/Isar outer syntax.
@@ -10,17 +9,22 @@
 
 signature OUTER_SYNTAX =
 sig
-  type parser_fn = OuterLex.token list ->
-    (Toplevel.transition -> Toplevel.transition) * OuterLex.token list
-  val command: string -> string -> OuterKeyword.T -> parser_fn -> unit
-  val markup_command: ThyOutput.markup -> string -> string -> OuterKeyword.T -> parser_fn -> unit
-  val improper_command: string -> string -> OuterKeyword.T -> parser_fn -> unit
+  type 'a parser = 'a OuterParse.parser
+  val command: string -> string -> OuterKeyword.T ->
+    (Toplevel.transition -> Toplevel.transition) parser -> unit
+  val markup_command: ThyOutput.markup -> string -> string -> OuterKeyword.T ->
+    (Toplevel.transition -> Toplevel.transition) parser -> unit
+  val improper_command: string -> string -> OuterKeyword.T ->
+    (Toplevel.transition -> Toplevel.transition) parser -> unit
+  val internal_command: string -> (Toplevel.transition -> Toplevel.transition) parser -> unit
+  val local_theory': string -> string -> OuterKeyword.T ->
+    (bool -> local_theory -> local_theory) parser -> unit
   val local_theory: string -> string -> OuterKeyword.T ->
-    (OuterParse.token list -> (local_theory -> local_theory) * OuterLex.token list) -> unit
+    (local_theory -> local_theory) parser -> unit
   val local_theory_to_proof': string -> string -> OuterKeyword.T ->
-    (OuterParse.token list -> (bool -> local_theory -> Proof.state) * OuterLex.token list) -> unit
+    (bool -> local_theory -> Proof.state) parser -> unit
   val local_theory_to_proof: string -> string -> OuterKeyword.T ->
-    (OuterParse.token list -> (local_theory -> Proof.state) * OuterLex.token list) -> unit
+    (local_theory -> Proof.state) parser -> unit
   val print_outer_syntax: unit -> unit
   val scan: Position.T -> string -> OuterLex.token list
   val parse: Position.T -> string -> Toplevel.transition list
@@ -28,7 +32,7 @@
   type isar
   val isar: bool -> isar
   val prepare_command: Position.T -> string -> Toplevel.transition
-  val load_thy: string -> Position.T -> string list -> bool -> unit
+  val load_thy: string -> Position.T -> string list -> bool -> unit -> unit
 end;
 
 structure OuterSyntax: OUTER_SYNTAX =
@@ -36,22 +40,22 @@
 
 structure T = OuterLex;
 structure P = OuterParse;
+type 'a parser = 'a P.parser;
+
 
 
 (** outer syntax **)
 
-(* parsers *)
+(* command parsers *)
 
-type parser_fn = T.token list -> (Toplevel.transition -> Toplevel.transition) * T.token list;
-
-datatype parser = Parser of
+datatype command = Command of
  {comment: string,
   markup: ThyOutput.markup option,
   int_only: bool,
-  parse: parser_fn};
+  parse: (Toplevel.transition -> Toplevel.transition) parser};
 
-fun make_parser comment markup int_only parse =
-  Parser {comment = comment, markup = markup, int_only = int_only, parse = parse};
+fun make_command comment markup int_only parse =
+  Command {comment = comment, markup = markup, int_only = int_only, parse = parse};
 
 
 (* parse command *)
@@ -63,7 +67,7 @@
 
 fun body cmd (name, _) =
   (case cmd name of
-    SOME (Parser {int_only, parse, ...}) =>
+    SOME (Command {int_only, parse, ...}) =>
       P.!!! (Scan.prompt (name ^ "# ") (P.tags |-- parse >> pair int_only))
   | NONE => sys_error ("no parser for outer syntax command " ^ quote name));
 
@@ -85,47 +89,50 @@
 
 local
 
-val global_parsers = ref (Symtab.empty: parser Symtab.table);
+val global_commands = ref (Symtab.empty: command Symtab.table);
 val global_markups = ref ([]: (string * ThyOutput.markup) list);
 
-fun change_parsers f = CRITICAL (fn () =>
- (change global_parsers f;
+fun change_commands f = CRITICAL (fn () =>
+ (change global_commands f;
   global_markups :=
-    Symtab.fold (fn (name, Parser {markup = SOME m, ...}) => cons (name, m) | _ => I)
-      (! global_parsers) []));
+    Symtab.fold (fn (name, Command {markup = SOME m, ...}) => cons (name, m) | _ => I)
+      (! global_commands) []));
 
 in
 
 (* access current syntax *)
 
-fun get_parsers () = CRITICAL (fn () => ! global_parsers);
+fun get_commands () = CRITICAL (fn () => ! global_commands);
 fun get_markups () = CRITICAL (fn () => ! global_markups);
 
-fun get_parser () = Symtab.lookup (get_parsers ());
-fun get_syntax () = CRITICAL (fn () => (OuterKeyword.get_lexicons (), get_parser ()));
+fun get_command () = Symtab.lookup (get_commands ());
+fun get_syntax () = CRITICAL (fn () => (OuterKeyword.get_lexicons (), get_command ()));
 
 fun is_markup kind name = AList.lookup (op =) (get_markups ()) name = SOME kind;
 
 
 (* augment syntax *)
 
-fun add_parser name kind parser = CRITICAL (fn () =>
+fun add_command name kind cmd = CRITICAL (fn () =>
  (OuterKeyword.command name kind;
-  if not (Symtab.defined (get_parsers ()) name) then ()
+  if not (Symtab.defined (get_commands ()) name) then ()
   else warning ("Redefining outer syntax command " ^ quote name);
-  change_parsers (Symtab.update (name, parser))));
+  change_commands (Symtab.update (name, cmd))));
 
 fun command name comment kind parse =
-  add_parser name kind (make_parser comment NONE false parse);
+  add_command name kind (make_command comment NONE false parse);
 
 fun markup_command markup name comment kind parse =
-  add_parser name kind (make_parser comment (SOME markup) false parse);
+  add_command name kind (make_command comment (SOME markup) false parse);
 
 fun improper_command name comment kind parse =
-  add_parser name kind (make_parser comment NONE true parse);
+  add_command name kind (make_command comment NONE true parse);
 
 end;
 
+fun internal_command name parse =
+  command name "(internal)" OuterKeyword.control (parse >> (fn tr => Toplevel.no_timing o tr));
+
 
 (* local_theory commands *)
 
@@ -133,22 +140,23 @@
   command name comment kind (P.opt_target -- parse
     >> (fn (loc, f) => (if do_print then Toplevel.print else I) o trans loc f));
 
-val local_theory           = local_theory_command false Toplevel.local_theory;
+val local_theory' = local_theory_command false Toplevel.local_theory';
+val local_theory = local_theory_command false Toplevel.local_theory;
 val local_theory_to_proof' = local_theory_command true Toplevel.local_theory_to_proof';
-val local_theory_to_proof  = local_theory_command true Toplevel.local_theory_to_proof;
+val local_theory_to_proof = local_theory_command true Toplevel.local_theory_to_proof;
 
 
 (* inspect syntax *)
 
-fun dest_parsers () =
-  get_parsers () |> Symtab.dest |> sort_wrt #1
-  |> map (fn (name, Parser {comment, int_only, ...}) => (name, comment, int_only));
+fun dest_commands () =
+  get_commands () |> Symtab.dest |> sort_wrt #1
+  |> map (fn (name, Command {comment, int_only, ...}) => (name, comment, int_only));
 
 fun print_outer_syntax () =
   let
     fun pretty_cmd (name, comment, _) =
       Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
-    val (int_cmds, cmds) = List.partition #3 (dest_parsers ());
+    val (int_cmds, cmds) = List.partition #3 (dest_commands ());
   in
     [Pretty.strs ("syntax keywords:" :: map quote (OuterKeyword.dest_keywords ())),
       Pretty.big_list "commands:" (map pretty_cmd cmds),
@@ -194,7 +202,7 @@
   Source.of_string str
   |> Symbol.source {do_recover = false}
   |> T.source {do_recover = SOME false} OuterKeyword.get_lexicons pos
-  |> toplevel_source false NONE get_parser
+  |> toplevel_source false NONE get_command
   |> Source.exhaust;
 
 
@@ -225,39 +233,39 @@
   Source.tty
   |> Symbol.source {do_recover = true}
   |> T.source {do_recover = SOME true} OuterKeyword.get_lexicons Position.none
-  |> toplevel_source term (SOME true) get_parser;
+  |> toplevel_source term (SOME true) get_command;
 
 
 (* prepare toplevel commands -- fail-safe *)
 
 val not_singleton = "Exactly one command expected";
 
-fun prepare_span parser span =
+fun prepare_span commands span =
   let
-    val range_pos = Position.encode_range (ThyEdit.span_range span);
-    val toks = ThyEdit.span_content span;
-    val _ = List.app ThyEdit.report_token toks;
+    val range_pos = Position.encode_range (ThySyntax.span_range span);
+    val toks = ThySyntax.span_content span;
+    val _ = List.app ThySyntax.report_token toks;
   in
-    (case Source.exhaust (toplevel_source false NONE (K parser) (Source.of_list toks)) of
+    (case Source.exhaust (toplevel_source false NONE (K commands) (Source.of_list toks)) of
       [tr] => (tr, true)
     | [] => (Toplevel.ignored range_pos, false)
     | _ => (Toplevel.malformed range_pos not_singleton, true))
     handle ERROR msg => (Toplevel.malformed range_pos msg, true)
   end;
 
-fun prepare_unit parser (cmd, proof, proper_proof) =
+fun prepare_unit commands (cmd, proof, proper_proof) =
   let
-    val (tr, proper_cmd) = prepare_span parser cmd;
-    val proof_trs = map (prepare_span parser) proof |> filter #2 |> map #1;
+    val (tr, proper_cmd) = prepare_span commands cmd;
+    val proof_trs = map (prepare_span commands) proof |> filter #2 |> map #1;
   in
     if proper_cmd andalso proper_proof then [(tr, proof_trs)]
     else map (rpair []) (if proper_cmd then tr :: proof_trs else proof_trs)
   end;
 
 fun prepare_command pos str =
-  let val (lexs, parser) = get_syntax () in
-    (case ThyEdit.parse_spans lexs pos str of
-      [span] => #1 (prepare_span parser span)
+  let val (lexs, commands) = get_syntax () in
+    (case ThySyntax.parse_spans lexs pos str of
+      [span] => #1 (prepare_span commands span)
     | _ => Toplevel.malformed pos not_singleton)
   end;
 
@@ -266,30 +274,27 @@
 
 fun load_thy name pos text time =
   let
-    val (lexs, parser) = get_syntax ();
+    val (lexs, commands) = get_syntax ();
 
     val _ = Present.init_theory name;
 
-    val toks = Source.exhausted (ThyEdit.token_source lexs pos (Source.of_list text));
-    val spans = Source.exhaust (ThyEdit.span_source toks);
-    val _ = List.app ThyEdit.report_span spans;
-    val units = Source.exhaust (ThyEdit.unit_source (Source.of_list spans))
-      |> maps (prepare_unit parser);
+    val toks = Source.exhausted (ThySyntax.token_source lexs pos (Source.of_list text));
+    val spans = Source.exhaust (ThySyntax.span_source toks);
+    val _ = List.app ThySyntax.report_span spans;
+    val units = Source.exhaust (ThySyntax.unit_source (Source.of_list spans))
+      |> maps (prepare_unit commands);
 
     val _ = Present.theory_source name
-      (fn () => HTML.html_mode (implode o map ThyEdit.present_span) spans);
+      (fn () => HTML.html_mode (implode o map ThySyntax.present_span) spans);
 
     val _ = if time then writeln ("\n**** Starting theory " ^ quote name ^ " ****") else ();
-    val _ = cond_timeit time "" (fn () =>
-      let
-        val (results, commit_exit) = Toplevel.excursion units;
-        val _ =
-          ThyOutput.present_thy (#1 lexs) OuterKeyword.command_tags is_markup results toks
-          |> Buffer.content
-          |> Present.theory_output name
-        val _ = commit_exit ();
-      in () end);
+    val results = cond_timeit time "" (fn () => Toplevel.excursion units);
     val _ = if time then writeln ("**** Finished theory " ^ quote name ^ " ****\n") else ();
-  in () end;
+
+    fun after_load () =
+      ThyOutput.present_thy (#1 lexs) OuterKeyword.command_tags is_markup (Lazy.force results) toks
+      |> Buffer.content
+      |> Present.theory_output name;
+  in after_load end;
 
 end;
--- a/src/Pure/Isar/proof.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Isar/proof.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -30,7 +30,6 @@
   val enter_forward: state -> state
   val goal_message: (unit -> Pretty.T) -> state -> state
   val get_goal: state -> context * (thm list * thm)
-  val schematic_goal: state -> bool
   val show_main_goal: bool ref
   val verbose: bool ref
   val pretty_state: int -> state -> Pretty.T list
@@ -87,34 +86,40 @@
   val local_goal: (context -> ((string * string) * (string * thm list) list) -> unit) ->
     (theory -> 'a -> attribute) ->
     (context * 'b list -> context * (term list list * (context -> context))) ->
-    string -> Method.text option -> (thm list list -> state -> state Seq.seq) ->
+    string -> Method.text option -> (thm list list -> state -> state) ->
     ((Binding.T * 'a list) * 'b) list -> state -> state
-  val local_qed: Method.text option * bool -> state -> state Seq.seq
+  val local_qed: Method.text option * bool -> state -> state
   val theorem: Method.text option -> (thm list list -> context -> context) ->
     (string * string list) list list -> context -> state
   val theorem_i: Method.text option -> (thm list list -> context -> context) ->
     (term * term list) list list -> context -> state
   val global_qed: Method.text option * bool -> state -> context
-  val local_terminal_proof: Method.text * Method.text option -> state -> state Seq.seq
-  val local_default_proof: state -> state Seq.seq
-  val local_immediate_proof: state -> state Seq.seq
-  val local_done_proof: state -> state Seq.seq
-  val local_skip_proof: bool -> state -> state Seq.seq
+  val local_terminal_proof: Method.text * Method.text option -> state -> state
+  val local_default_proof: state -> state
+  val local_immediate_proof: state -> state
+  val local_skip_proof: bool -> state -> state
+  val local_done_proof: state -> state
   val global_terminal_proof: Method.text * Method.text option -> state -> context
   val global_default_proof: state -> context
   val global_immediate_proof: state -> context
+  val global_skip_proof: bool -> state -> context
   val global_done_proof: state -> context
-  val global_skip_proof: bool -> state -> context
-  val have: Method.text option -> (thm list list -> state -> state Seq.seq) ->
+  val have: Method.text option -> (thm list list -> state -> state) ->
     (Attrib.binding * (string * string list) list) list -> bool -> state -> state
-  val have_i: Method.text option -> (thm list list -> state -> state Seq.seq) ->
+  val have_i: Method.text option -> (thm list list -> state -> state) ->
     ((Binding.T * attribute list) * (term * term list) list) list -> bool -> state -> state
-  val show: Method.text option -> (thm list list -> state -> state Seq.seq) ->
+  val show: Method.text option -> (thm list list -> state -> state) ->
     (Attrib.binding * (string * string list) list) list -> bool -> state -> state
-  val show_i: Method.text option -> (thm list list -> state -> state Seq.seq) ->
+  val show_i: Method.text option -> (thm list list -> state -> state) ->
     ((Binding.T * attribute list) * (term * term list) list) list -> bool -> state -> state
+  val schematic_goal: state -> bool
   val is_relevant: state -> bool
-  val future_proof: (state -> ('a * context) future) -> state -> 'a future * context
+  val local_future_proof: (state -> ('a * state) Future.future) ->
+    state -> 'a Future.future * state
+  val global_future_proof: (state -> ('a * Proof.context) Future.future) ->
+    state -> 'a Future.future * context
+  val local_future_terminal_proof: Method.text * Method.text option -> bool -> state -> state
+  val global_future_terminal_proof: Method.text * Method.text option -> bool -> state -> context
 end;
 
 structure Proof: PROOF =
@@ -140,14 +145,14 @@
     goal: goal option}
 and goal =
   Goal of
-   {statement: (string * Position.T) * term list list * cterm,
+   {statement: (string * Position.T) * term list list * term,
       (*goal kind and statement (starting with vars), initial proposition*)
     messages: (unit -> Pretty.T) list,    (*persistent messages (hints etc.)*)
     using: thm list,                      (*goal facts*)
     goal: thm,                            (*subgoals ==> statement*)
     before_qed: Method.text option,
     after_qed:
-      (thm list list -> state -> state Seq.seq) *
+      (thm list list -> state -> state) *
       (thm list list -> context -> context)};
 
 fun make_goal (statement, messages, using, goal, before_qed, after_qed) =
@@ -311,12 +316,6 @@
   let val (ctxt, (_, {using, goal, ...})) = find_goal state
   in (ctxt, (using, goal)) end;
 
-fun schematic_goal state =
-  let val (_, (_, {statement = (_, _, prop), ...})) = find_goal state in
-    Term.exists_subterm Term.is_Var (Thm.term_of prop) orelse
-    Term.exists_type (Term.exists_subtype Term.is_TVar) (Thm.term_of prop)
-  end;
-
 
 
 (** pretty_state **)
@@ -347,7 +346,7 @@
       | strs' => enclose " (" ")" (commas strs'));
 
     fun prt_goal (SOME (ctxt, (i,
-            {statement = ((_, pos), _, _), messages, using, goal, before_qed, after_qed}))) =
+      {statement = ((_, pos), _, _), messages, using, goal, before_qed = _, after_qed = _}))) =
           pretty_facts "using " ctxt
             (if mode <> Backward orelse null using then NONE else SOME using) @
           [Pretty.str ("goal" ^
@@ -750,6 +749,13 @@
   |> `before_qed |-> (refine o the_default Method.succeed_text)
   |> Seq.maps (refine (Method.finish_text txt (Position.thread_data ())));
 
+fun check_result msg sq =
+  (case Seq.pull sq of
+    NONE => error msg
+  | SOME (s, _) => s);
+
+fun check_finish sq = check_result "Failed to finish proof" sq;
+
 
 (* unstructured refinement *)
 
@@ -806,9 +812,10 @@
 
     val propss' = map (Logic.mk_term o Var) vars :: propss;
     val goal_propss = filter_out null propss';
-    val goal = cert (Logic.mk_conjunction_balanced (map Logic.mk_conjunction_balanced goal_propss))
+    val goal =
+      cert (Logic.mk_conjunction_balanced (map Logic.mk_conjunction_balanced goal_propss))
       |> Thm.weaken_sorts (Variable.sorts_of (context_of goal_state));
-    val statement = ((kind, pos), propss', goal);
+    val statement = ((kind, pos), propss', Thm.term_of goal);
     val after_qed' = after_qed |>> (fn after_local =>
       fn results => map_context after_ctxt #> after_local results);
   in
@@ -840,7 +847,7 @@
       |> burrow (ProofContext.export goal_ctxt outer_ctxt);
 
     val (after_local, after_global) = after_qed;
-    fun after_local' x y = Position.setmp_thread_data_seq pos (fn () => after_local x y) ();
+    fun after_local' x y = Position.setmp_thread_data pos (fn () => after_local x y) ();
     fun after_global' x y = Position.setmp_thread_data pos (fn () => after_global x y) ();
   in
     outer_state
@@ -869,61 +876,47 @@
     |> tap (Variable.warn_extra_tfrees (context_of state) o context_of)
   end;
 
-fun local_qed txt =
-  end_proof false txt #>
-  Seq.maps (generic_qed ProofContext.auto_bind_facts #->
+fun local_qeds txt =
+  end_proof false txt
+  #> Seq.map (generic_qed ProofContext.auto_bind_facts #->
     (fn ((after_qed, _), results) => after_qed results));
 
+fun local_qed txt = local_qeds txt #> check_finish;
+
 
 (* global goals *)
 
 fun global_goal prepp before_qed after_qed propp ctxt =
   init ctxt
-  |> generic_goal (prepp #> ProofContext.auto_fixes) ""
-    before_qed (K Seq.single, after_qed) propp;
+  |> generic_goal (prepp #> ProofContext.auto_fixes) "" before_qed (K I, after_qed) propp;
 
 val theorem = global_goal ProofContext.bind_propp_schematic;
 val theorem_i = global_goal ProofContext.bind_propp_schematic_i;
 
-fun check_result msg sq =
-  (case Seq.pull sq of
-    NONE => error msg
-  | SOME (s', sq') => Seq.cons s' sq');
-
 fun global_qeds txt =
   end_proof true txt
   #> Seq.map (generic_qed (K I) #> (fn (((_, after_qed), results), state) =>
-    after_qed results (context_of state)))
-  |> Seq.DETERM;   (*backtracking may destroy theory!*)
+    after_qed results (context_of state)));
 
-fun global_qed txt =
-  global_qeds txt
-  #> check_result "Failed to finish proof"
-  #> Seq.hd;
+fun global_qed txt = global_qeds txt #> check_finish;
 
 
 (* concluding steps *)
 
-fun local_terminal_proof (text, opt_text) =
-  proof (SOME text) #> Seq.maps (local_qed (opt_text, true));
+fun terminal_proof qed initial terminal =
+  proof (SOME initial) #> Seq.maps (qed terminal) #> check_finish;
 
+fun local_terminal_proof (text, opt_text) = terminal_proof local_qeds text (opt_text, true);
 val local_default_proof = local_terminal_proof (Method.default_text, NONE);
 val local_immediate_proof = local_terminal_proof (Method.this_text, NONE);
-val local_done_proof = proof (SOME Method.done_text) #> Seq.maps (local_qed (NONE, false));
 fun local_skip_proof int = local_terminal_proof (Method.sorry_text int, NONE);
+val local_done_proof = terminal_proof local_qeds Method.done_text (NONE, false);
 
-fun global_term_proof immed (text, opt_text) =
-  proof (SOME text)
-  #> check_result "Terminal proof method failed"
-  #> Seq.maps (global_qeds (opt_text, immed))
-  #> check_result "Failed to finish proof (after successful terminal method)"
-  #> Seq.hd;
-
-val global_terminal_proof = global_term_proof true;
+fun global_terminal_proof (text, opt_text) = terminal_proof global_qeds text (opt_text, true);
 val global_default_proof = global_terminal_proof (Method.default_text, NONE);
 val global_immediate_proof = global_terminal_proof (Method.this_text, NONE);
-val global_done_proof = global_term_proof false (Method.done_text, NONE);
 fun global_skip_proof int = global_terminal_proof (Method.sorry_text int, NONE);
+val global_done_proof = terminal_proof global_qeds Method.done_text (NONE, false);
 
 
 (* common goal statements *)
@@ -938,7 +931,7 @@
     val testing = ref false;
     val rule = ref (NONE: thm option);
     fun fail_msg ctxt =
-      "Local statement will fail to solve any pending goal" ::
+      "Local statement will fail to refine any pending goal" ::
       (case ! rule of NONE => [] | SOME th => [ProofDisplay.string_of_rule ctxt "Failed" th])
       |> cat_lines;
 
@@ -949,13 +942,14 @@
       else if int then priority (ProofDisplay.string_of_rule ctxt "Successful" th)
       else ();
     val test_proof =
-      (Seq.pull oo local_skip_proof) true
+      try (local_skip_proof true)
       |> setmp_noncritical testing true
       |> Exn.capture;
 
     fun after_qed' results =
       refine_goals print_rule (context_of state) (flat results)
-      #> Seq.maps (after_qed results);
+      #> check_result "Failed to refine any pending goal"
+      #> after_qed results;
   in
     state
     |> local_goal print_results prep_att prepp "show" before_qed after_qed' stmt
@@ -977,46 +971,89 @@
 end;
 
 
-(* fork global proofs *)
+
+(** future proofs **)
+
+(* relevant proof states *)
 
-fun is_initial state =
-  (case try find_goal state of
-    NONE => false
-  | SOME (_, (_, {statement = (_, _, prop), goal, ...})) => Thm.eq_thm_prop (Goal.init prop, goal));
+fun is_schematic t =
+  Term.exists_subterm Term.is_Var t orelse
+  Term.exists_type (Term.exists_subtype Term.is_TVar) t;
+
+fun schematic_goal state =
+  let val (_, (_, {statement = (_, _, prop), ...})) = find_goal state
+  in is_schematic prop end;
 
 fun is_relevant state =
-  can (assert_bottom false) state orelse
-  can assert_forward state orelse
-  not (is_initial state) orelse
-  schematic_goal state;
+  (case try find_goal state of
+    NONE => true
+  | SOME (_, (_, {statement = (_, _, prop), goal, ...})) =>
+      is_schematic prop orelse not (Logic.protect prop aconv Thm.concl_of goal));
+
+
+(* full proofs *)
 
-fun future_proof proof state =
+local
+
+fun future_proof done get_context fork_proof state =
   let
+    val _ = assert_backward state;
+    val (goal_ctxt, (_, goal)) = find_goal state;
+    val {statement as (kind, propss, prop), messages, using, goal, before_qed, after_qed} = goal;
+
     val _ = is_relevant state andalso error "Cannot fork relevant proof";
 
-    val (goal_ctxt, (_, goal)) = find_goal state;
-    val {statement as (kind, propss, cprop), messages, using, goal, before_qed, after_qed} = goal;
-    val prop = Thm.term_of cprop;
-    val cprop_protected = #2 (Thm.dest_implies (Thm.cprop_of goal));
+    val prop' = Logic.protect prop;
+    val statement' = (kind, [[], [prop']], prop');
+    val goal' = Thm.adjust_maxidx_thm (Thm.maxidx_of goal)
+      (Drule.comp_no_flatten (goal, Thm.nprems_of goal) 1 Drule.protectI);
 
-    val statement' = (kind, [[], [Thm.term_of cprop_protected]], cprop_protected);
-    val goal' = Thm.adjust_maxidx_thm ~1 (Drule.protectI RS Goal.init cprop_protected);
-    fun after_qed' [[th]] = ProofContext.put_thms false (AutoBind.thisN, SOME [Goal.conclude th]);
-    val this_name = ProofContext.full_bname (ProofContext.reset_naming goal_ctxt) AutoBind.thisN;
+    fun after_local' [[th]] = put_thms false (AutoBind.thisN, SOME [th]);
+    fun after_global' [[th]] = ProofContext.put_thms false (AutoBind.thisN, SOME [th]);
+    val after_qed' = (after_local', after_global');
+    val this_name = ProofContext.full_bname goal_ctxt AutoBind.thisN;
 
     val result_ctxt =
       state
-      |> map_contexts (Variable.auto_fixes prop)
-      |> map_goal I (K (statement', messages, using, goal', before_qed, (#1 after_qed, after_qed')))
-      |> proof;
+      |> map_contexts (Variable.declare_term prop)
+      |> map_goal I (K (statement', messages, using, goal', before_qed, after_qed'))
+      |> fork_proof;
 
-    val thm = result_ctxt |> Future.map (fn (_, ctxt) =>
-      ProofContext.get_fact_single ctxt (Facts.named this_name));
-    val finished_goal = Goal.protect (Goal.future_result goal_ctxt thm prop);
+    val future_thm = result_ctxt |> Future.map (fn (_, x) =>
+      ProofContext.get_fact_single (get_context x) (Facts.named this_name));
+    val finished_goal = exception_trace (fn () => Goal.future_result goal_ctxt future_thm prop');
     val state' =
       state
       |> map_goal I (K (statement, messages, using, finished_goal, NONE, after_qed))
-      |> global_done_proof;
+      |> done;
   in (Future.map #1 result_ctxt, state') end;
 
+in
+
+fun local_future_proof x = future_proof local_done_proof context_of x;
+fun global_future_proof x = future_proof global_done_proof I x;
+
 end;
+
+
+(* terminal proofs *)
+
+local
+
+fun future_terminal_proof proof1 proof2 meths int state =
+  if int orelse is_relevant state orelse not (Goal.future_enabled ())
+  then proof1 meths state
+  else snd (state |> proof2 (fn state' => Future.fork_pri ~1 (fn () => ((), proof1 meths state'))));
+
+in
+
+fun local_future_terminal_proof x =
+  future_terminal_proof local_terminal_proof local_future_proof x;
+
+fun global_future_terminal_proof x =
+  future_terminal_proof global_terminal_proof global_future_proof x;
+
+end;
+
+end;
+
--- a/src/Pure/Isar/proof_context.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Isar/proof_context.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -408,9 +408,8 @@
    ("free", free_or_skolem),
    ("bound", plain_markup Markup.bound),
    ("var", var_or_skolem),
-   ("num", plain_markup Markup.num),
-   ("xnum", plain_markup Markup.xnum),
-   ("xstr", plain_markup Markup.xstr)];
+   ("numeral", plain_markup Markup.numeral),
+   ("inner_string", plain_markup Markup.inner_string)];
 
 in val _ = Context.>> (Context.map_theory (Sign.add_tokentrfuns token_trans)) end;
 
--- a/src/Pure/Isar/rule_cases.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Isar/rule_cases.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/rule_cases.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Annotations and local contexts of rules.
@@ -319,7 +318,7 @@
 local
 
 fun equal_cterms ts us =
-  is_equal (list_ord (Term.fast_term_ord o pairself Thm.term_of) (ts, us));
+  is_equal (list_ord (TermOrd.fast_term_ord o pairself Thm.term_of) (ts, us));
 
 fun prep_rule n th =
   let
--- a/src/Pure/Isar/session.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Isar/session.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/session.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Session management -- maintain state of logic images.
@@ -11,7 +10,7 @@
   val name: unit -> string
   val welcome: unit -> string
   val use_dir: string -> bool -> string list -> bool -> bool -> string -> bool -> string list ->
-    string -> string -> bool * string -> string -> int -> bool -> int -> int -> unit
+    string -> string -> bool * string -> string -> int -> bool -> int -> int -> bool -> unit
   val finish: unit -> unit
 end;
 
@@ -87,7 +86,7 @@
   | dumping (cp, path) = SOME (cp, Path.explode path);
 
 fun use_dir root build modes reset info doc doc_graph doc_versions
-    parent name dump rpath level verbose max_threads trace_threads =
+    parent name dump rpath level verbose max_threads trace_threads parallel_proofs =
   ((fn () =>
      (init reset parent name;
       Present.init build info doc doc_graph doc_versions (path ()) name
@@ -96,6 +95,7 @@
       finish ()))
     |> setmp_noncritical Proofterm.proofs level
     |> setmp_noncritical print_mode (modes @ print_mode_value ())
+    |> setmp_noncritical Goal.parallel_proofs parallel_proofs
     |> setmp_noncritical Multithreading.trace trace_threads
     |> setmp_noncritical Multithreading.max_threads
       (if Multithreading.available then max_threads
--- a/src/Pure/Isar/skip_proof.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Isar/skip_proof.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -34,7 +34,7 @@
   ALLGOALS (Tactic.rtac (make_thm thy (Var (("A", 0), propT)))) st;
 
 fun prove ctxt xs asms prop tac =
-  (if Future.enabled () then Goal.prove_future else Goal.prove) ctxt xs asms prop
+  (if Goal.future_enabled () then Goal.prove_future else Goal.prove) ctxt xs asms prop
     (fn args => fn st =>
       if ! quick_and_dirty
       then setmp quick_and_dirty true (cheat_tac (ProofContext.theory_of ctxt)) st
--- a/src/Pure/Isar/spec_parse.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Isar/spec_parse.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/spec_parse.ML
-    ID:         $Id$
     Author:     Makarius
 
 Parsers for complex specifications.
@@ -7,35 +6,33 @@
 
 signature SPEC_PARSE =
 sig
-  type token
-  val attrib: OuterLex.token list -> Attrib.src * token list
-  val attribs: token list -> Attrib.src list * token list
-  val opt_attribs: token list -> Attrib.src list * token list
-  val thm_name: string -> token list -> Attrib.binding * token list
-  val opt_thm_name: string -> token list -> Attrib.binding * token list
-  val spec: token list -> (Attrib.binding * string list) * token list
-  val named_spec: token list -> (Attrib.binding * string list) * token list
-  val spec_name: token list -> ((Binding.T * string) * Attrib.src list) * token list
-  val spec_opt_name: token list -> ((Binding.T * string) * Attrib.src list) * token list
-  val xthm: token list -> (Facts.ref * Attrib.src list) * token list
-  val xthms1: token list -> (Facts.ref * Attrib.src list) list * token list
-  val name_facts: token list ->
-    (Attrib.binding * (Facts.ref * Attrib.src list) list) list * token list
-  val locale_mixfix: token list -> mixfix * token list
-  val locale_fixes: token list -> (Binding.T * string option * mixfix) list * token list
-  val locale_insts: token list -> (string option list * (Attrib.binding * string) list) * token list
-  val class_expr: token list -> string list * token list
-  val locale_expr: token list -> Locale.expr * token list
-  val locale_expression: OuterParse.token list -> Expression.expression * OuterParse.token list
-  val locale_keyword: token list -> string * token list
-  val context_element: token list -> Element.context * token list
-  val statement: token list -> (Attrib.binding * (string * string list) list) list * token list
-  val general_statement: token list ->
-    (Element.context list * Element.statement) * OuterLex.token list
-  val statement_keyword: token list -> string * token list
-  val specification: token list ->
-    (Binding.T *
-      ((Attrib.binding * string list) list * (Binding.T * string option) list)) list * token list
+  type token = OuterParse.token
+  type 'a parser = 'a OuterParse.parser
+  val attrib: Attrib.src parser
+  val attribs: Attrib.src list parser
+  val opt_attribs: Attrib.src list parser
+  val thm_name: string -> Attrib.binding parser
+  val opt_thm_name: string -> Attrib.binding parser
+  val spec: (Attrib.binding * string list) parser
+  val named_spec: (Attrib.binding * string list) parser
+  val spec_name: ((Binding.T * string) * Attrib.src list) parser
+  val spec_opt_name: ((Binding.T * string) * Attrib.src list) parser
+  val xthm: (Facts.ref * Attrib.src list) parser
+  val xthms1: (Facts.ref * Attrib.src list) list parser
+  val name_facts: (Attrib.binding * (Facts.ref * Attrib.src list) list) list parser
+  val locale_mixfix: mixfix parser
+  val locale_fixes: (Binding.T * string option * mixfix) list parser
+  val locale_insts: (string option list * (Attrib.binding * string) list) parser
+  val class_expr: string list parser
+  val locale_expr: Old_Locale.expr parser
+  val locale_expression: Expression.expression parser
+  val locale_keyword: string parser
+  val context_element: Element.context parser
+  val statement: (Attrib.binding * (string * string list) list) list parser
+  val general_statement: (Element.context list * Element.statement) parser
+  val statement_keyword: string parser
+  val specification:
+    (Binding.T * ((Attrib.binding * string list) list * (Binding.T * string option) list)) list parser
 end;
 
 structure SpecParse: SPEC_PARSE =
@@ -43,6 +40,7 @@
 
 structure P = OuterParse;
 type token = P.token;
+type 'a parser = 'a P.parser;
 
 
 (* theorem specifications *)
@@ -119,9 +117,9 @@
 
 val locale_expr =
   let
-    fun expr2 x = (P.xname >> Locale.Locale || P.$$$ "(" |-- P.!!! (expr0 --| P.$$$ ")")) x
-    and expr1 x = (expr2 -- Scan.repeat1 (P.maybe rename) >> Locale.Rename || expr2) x
-    and expr0 x = (plus1_unless locale_keyword expr1 >> (fn [e] => e | es => Locale.Merge es)) x;
+    fun expr2 x = (P.xname >> Old_Locale.Locale || P.$$$ "(" |-- P.!!! (expr0 --| P.$$$ ")")) x
+    and expr1 x = (expr2 -- Scan.repeat1 (P.maybe rename) >> Old_Locale.Rename || expr2) x
+    and expr0 x = (plus1_unless locale_keyword expr1 >> (fn [e] => e | es => Old_Locale.Merge es)) x;
   in expr0 end;
 
 val locale_expression =
--- a/src/Pure/Isar/subclass.ML	Mon Jan 19 20:37:08 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,67 +0,0 @@
-(*  Title:      Pure/Isar/subclass.ML
-    Author:     Florian Haftmann, TU Muenchen
-
-User interface for proving subclass relationship between type classes.
-*)
-
-signature SUBCLASS =
-sig
-  val subclass: class -> local_theory -> Proof.state
-  val subclass_cmd: xstring -> local_theory -> Proof.state
-  val prove_subclass: tactic -> class -> local_theory -> local_theory
-end;
-
-structure Subclass : SUBCLASS =
-struct
-
-local
-
-fun gen_subclass prep_class do_proof raw_sup lthy =
-  let
-    val thy = ProofContext.theory_of lthy;
-    val sup = prep_class thy raw_sup;
-    val sub = case TheoryTarget.peek lthy
-     of {is_class = false, ...} => error "Not a class context"
-      | {target, ...} => target;
-    val _ = if Sign.subsort thy ([sup], [sub])
-      then error ("Class " ^ Syntax.string_of_sort lthy [sup]
-        ^ " is subclass of class " ^ Syntax.string_of_sort lthy [sub])
-      else ();
-    val sub_params = map fst (Class.these_params thy [sub]);
-    val sup_params = map fst (Class.these_params thy [sup]);
-    val err_params = subtract (op =) sub_params sup_params;
-    val _ = if null err_params then [] else
-      error ("Class " ^ Syntax.string_of_sort lthy [sub] ^ " lacks parameter(s) " ^
-        commas_quote err_params ^ " of " ^ Syntax.string_of_sort lthy [sup]);
-    val sublocale_prop =
-      Locale.global_asms_of thy sup
-      |> maps snd
-      |> try the_single
-      |> Option.map (ObjectLogic.ensure_propT thy);
-    fun after_qed some_thm =
-      LocalTheory.theory (Class.prove_subclass (sub, sup) some_thm)
-      #> (TheoryTarget.init (SOME sub) o ProofContext.theory_of);
-  in
-    do_proof after_qed sublocale_prop lthy
-  end;
-
-fun user_proof after_qed NONE =
-      Proof.theorem_i NONE (K (after_qed NONE)) [[]]
-  | user_proof after_qed (SOME prop) =
-      Proof.theorem_i NONE (after_qed o try the_single o the_single) [[(prop, [])]];
-
-fun tactic_proof tac after_qed NONE lthy =
-      after_qed NONE lthy
-  | tactic_proof tac after_qed (SOME prop) lthy =
-      after_qed (SOME (Goal.prove (LocalTheory.target_of lthy) [] [] prop
-        (K tac))) lthy;
-
-in
-
-val subclass = gen_subclass (K I) user_proof;
-val subclass_cmd = gen_subclass Sign.read_class user_proof;
-fun prove_subclass tac = gen_subclass (K I) (tactic_proof tac);
-
-end; (*local*)
-
-end;
--- a/src/Pure/Isar/theory_target.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Isar/theory_target.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -13,6 +13,7 @@
   val begin: string -> Proof.context -> local_theory
   val context: xstring -> theory -> local_theory
   val instantiation: string list * (string * sort) list * sort -> theory -> local_theory
+  val instantiation_cmd: xstring list * sort * xstring -> theory -> local_theory
   val overloading: (string * (string * typ) * bool) list -> theory -> local_theory
   val overloading_cmd: (string * string * bool) list -> theory -> local_theory
 end;
@@ -23,19 +24,19 @@
 (* new locales *)
 
 fun locale_extern new_locale x = 
-  if new_locale then NewLocale.extern x else Locale.extern x;
+  if new_locale then Locale.extern x else Old_Locale.extern x;
 fun locale_add_type_syntax new_locale x =
-  if new_locale then NewLocale.add_type_syntax x else Locale.add_type_syntax x;
+  if new_locale then Locale.add_type_syntax x else Old_Locale.add_type_syntax x;
 fun locale_add_term_syntax new_locale x =
-  if new_locale then NewLocale.add_term_syntax x else Locale.add_term_syntax x;
+  if new_locale then Locale.add_term_syntax x else Old_Locale.add_term_syntax x;
 fun locale_add_declaration new_locale x =
-  if new_locale then NewLocale.add_declaration x else Locale.add_declaration x;
+  if new_locale then Locale.add_declaration x else Old_Locale.add_declaration x;
 fun locale_add_thmss new_locale x =
-  if new_locale then NewLocale.add_thmss x else Locale.add_thmss x;
+  if new_locale then Locale.add_thmss x else Old_Locale.add_thmss x;
 fun locale_init new_locale x =
-  if new_locale then NewLocale.init x else Locale.init x;
+  if new_locale then Locale.init x else Old_Locale.init x;
 fun locale_intern new_locale x =
-  if new_locale then NewLocale.intern x else Locale.intern x;
+  if new_locale then Locale.intern x else Old_Locale.intern x;
 
 (* context data *)
 
@@ -82,7 +83,7 @@
   Pretty.block [Pretty.str "theory", Pretty.brk 1,
       Pretty.str (Context.theory_name (ProofContext.theory_of ctxt))] ::
     (if not (null overloading) then [Overloading.pretty ctxt]
-     else if not (null (#1 instantiation)) then [Class.pretty_instantiation ctxt]
+     else if not (null (#1 instantiation)) then [Class_Target.pretty_instantiation ctxt]
      else pretty_thy ctxt target is_locale is_class);
 
 
@@ -108,7 +109,7 @@
 
 fun class_target (Target {target, ...}) f =
   LocalTheory.raw_theory f #>
-  LocalTheory.target (Class.refresh_syntax target);
+  LocalTheory.target (Class_Target.refresh_syntax target);
 
 
 (* notes *)
@@ -207,7 +208,7 @@
     val (prefix', _) = Binding.dest b';
     val class_global = Binding.base_name b = Binding.base_name b'
       andalso not (null prefix')
-      andalso (fst o snd o split_last) prefix' = Class.class_prefix target;
+      andalso (fst o snd o split_last) prefix' = Class_Target.class_prefix target;
   in
     not (is_class andalso (similar_body orelse class_global)) ?
       (Context.mapping_result
@@ -231,11 +232,11 @@
     val (mx1, mx2, mx3) = fork_mixfix ta mx;
     fun syntax_error c = error ("Illegal mixfix syntax for overloaded constant " ^ quote c);
     val declare_const =
-      (case Class.instantiation_param lthy c of
+      (case Class_Target.instantiation_param lthy c of
         SOME c' =>
           if mx3 <> NoSyn then syntax_error c'
           else LocalTheory.theory_result (AxClass.declare_overloaded (c', U))
-            ##> Class.confirm_declaration c
+            ##> Class_Target.confirm_declaration c
         | NONE =>
             (case Overloading.operation lthy c of
               SOME (c', _) =>
@@ -248,7 +249,7 @@
   in
     lthy'
     |> is_locale ? term_syntax ta (locale_const ta Syntax.mode_default tags ((b, mx2), t))
-    |> is_class ? class_target ta (Class.declare target tags ((c, mx1), t))
+    |> is_class ? class_target ta (Class_Target.declare target tags ((c, mx1), t))
     |> LocalDefs.add_def ((b, NoSyn), t)
   end;
 
@@ -275,7 +276,7 @@
         #-> (fn (lhs, _) =>
           let val lhs' = Term.list_comb (Logic.unvarify lhs, xs) in
             term_syntax ta (locale_const ta prmode tags ((b, mx2), lhs')) #>
-            is_class ? class_target ta (Class.abbrev target prmode tags ((c, mx1), t'))
+            is_class ? class_target ta (Class_Target.abbrev target prmode tags ((c, mx1), t'))
           end)
       else
         LocalTheory.theory
@@ -311,7 +312,7 @@
         SOME (_, checked) =>
           (fn name => fn (Const (c, _), rhs) => Overloading.define checked name (c, rhs))
       | NONE =>
-          if is_none (Class.instantiation_param lthy c)
+          if is_none (Class_Target.instantiation_param lthy c)
           then (fn name => fn eq => Thm.add_def false false (name, Logic.mk_equals eq))
           else (fn name => fn (Const (c, _), rhs) => AxClass.define_overloaded name (c, rhs)));
     val (global_def, lthy3) = lthy2
@@ -333,15 +334,15 @@
 
 fun init_target _ NONE = global_target
   | init_target thy (SOME target) =
-      make_target target (NewLocale.test_locale thy (NewLocale.intern thy target))
-      true (Class.is_class thy target) ([], [], []) [];
+      make_target target (Locale.defined thy (Locale.intern thy target))
+      true (Class_Target.is_class thy target) ([], [], []) [];
 
 fun init_ctxt (Target {target, new_locale, is_locale, is_class, instantiation, overloading}) =
-  if not (null (#1 instantiation)) then Class.init_instantiation instantiation
+  if not (null (#1 instantiation)) then Class_Target.init_instantiation instantiation
   else if not (null overloading) then Overloading.init overloading
   else if not is_locale then ProofContext.init
   else if not is_class then locale_init new_locale target
-  else Class.init target;
+  else Class_Target.init target;
 
 fun init_lthy (ta as Target {target, instantiation, overloading, ...}) =
   Data.put ta #>
@@ -355,11 +356,20 @@
     declaration = declaration ta,
     reinit = fn lthy => init_lthy_ctxt ta (ProofContext.theory_of lthy),
     exit = LocalTheory.target_of o
-      (if not (null (#1 instantiation)) then Class.conclude_instantiation
+      (if not (null (#1 instantiation)) then Class_Target.conclude_instantiation
        else if not (null overloading) then Overloading.conclude
        else I)}
 and init_lthy_ctxt ta = init_lthy ta o init_ctxt ta;
 
+fun read_multi_arity thy (raw_tycos, raw_sorts, raw_sort) =
+  let
+    val all_arities = map (fn raw_tyco => Sign.read_arity thy
+      (raw_tyco, raw_sorts, raw_sort)) raw_tycos;
+    val tycos = map #1 all_arities;
+    val (_, sorts, sort) = hd all_arities;
+    val vs = Name.names Name.context Name.aT sorts;
+  in (tycos, vs, sort) end;
+
 fun gen_overloading prep_const raw_ops thy =
   let
     val ctxt = ProofContext.init thy;
@@ -374,9 +384,11 @@
 
 fun context "-" thy = init NONE thy
   | context target thy = init (SOME (locale_intern
-      (NewLocale.test_locale thy (NewLocale.intern thy target)) thy target)) thy;
+      (Locale.defined thy (Locale.intern thy target)) thy target)) thy;
 
 fun instantiation arities = init_lthy_ctxt (make_target "" false false false arities []);
+fun instantiation_cmd raw_arities thy =
+  instantiation (read_multi_arity thy raw_arities) thy;
 
 val overloading = gen_overloading (fn ctxt => Syntax.check_term ctxt o Const);
 val overloading_cmd = gen_overloading Syntax.read_term;
--- a/src/Pure/Isar/toplevel.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Isar/toplevel.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/toplevel.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Isabelle/Isar toplevel transactions.
@@ -66,6 +65,8 @@
   val theory': (bool -> theory -> theory) -> transition -> transition
   val begin_local_theory: bool -> (theory -> local_theory) -> transition -> transition
   val end_local_theory: transition -> transition
+  val local_theory': xstring option -> (bool -> local_theory -> local_theory) ->
+    transition -> transition
   val local_theory: xstring option -> (local_theory -> local_theory) -> transition -> transition
   val present_local_theory: xstring option -> (node -> unit) -> transition -> transition
   val local_theory_to_proof': xstring option -> (bool -> local_theory -> Proof.state) ->
@@ -95,7 +96,8 @@
   val transition: bool -> transition -> state -> (state * (exn * string) option) option
   val commit_exit: Position.T -> transition
   val command: transition -> state -> state
-  val excursion: (transition * transition list) list -> (transition * state) list * (unit -> unit)
+  val run_command: string -> transition -> state -> state option
+  val excursion: (transition * transition list) list -> (transition * state) list lazy
 end;
 
 structure Toplevel: TOPLEVEL =
@@ -512,14 +514,15 @@
   (fn Theory (gthy, _) =>
         let
           val finish = loc_finish loc gthy;
-          val lthy' = f (loc_begin loc gthy);
+          val lthy' = f int (loc_begin loc gthy);
         in Theory (finish lthy', SOME lthy') end
     | _ => raise UNDEF) #> tap g);
 
 in
 
-fun local_theory loc f = local_theory_presentation loc f (K I);
-fun present_local_theory loc g = local_theory_presentation loc I g;
+fun local_theory' loc f = local_theory_presentation loc f (K I);
+fun local_theory loc f = local_theory' loc (K f);
+fun present_local_theory loc g = local_theory_presentation loc (K I) g;
 
 end;
 
@@ -691,6 +694,22 @@
   | SOME (_, SOME exn_info) => raise EXCURSION_FAIL exn_info
   | NONE => raise EXCURSION_FAIL (TERMINATE, at_command tr));
 
+fun command_result tr st =
+  let val st' = command tr st
+  in (st', st') end;
+
+fun run_command thy_name tr st =
+  (case
+      (case init_of tr of
+        SOME name => Exn.capture (fn () => ThyLoad.check_name thy_name name) ()
+      | NONE => Exn.Result ()) of
+    Exn.Result () =>
+      (case transition true tr st of
+        SOME (st', NONE) => (status tr Markup.finished; SOME st')
+      | SOME (_, SOME exn_info) => (error_msg tr exn_info; status tr Markup.failed; NONE)
+      | NONE => (error_msg tr (TERMINATE, at_command tr); status tr Markup.failed; NONE))
+  | Exn.Exn exn => (error_msg tr (exn, at_command tr); status tr Markup.failed; NONE));
+
 
 (* excursion of units, consisting of commands with proof *)
 
@@ -700,10 +719,6 @@
   fun init _ = NONE;
 );
 
-fun command_result tr st =
-  let val st' = command tr st
-  in (st', st') end;
-
 fun proof_result immediate (tr, proof_trs) st =
   let val st' = command tr st in
     if immediate orelse null proof_trs orelse
@@ -716,9 +731,9 @@
         val (body_trs, end_tr) = split_last proof_trs;
         val finish = Context.Theory o ProofContext.theory_of;
 
-        val future_proof = Proof.future_proof
+        val future_proof = Proof.global_future_proof
           (fn prf =>
-            Future.fork_pri 1 (fn () =>
+            Future.fork_pri ~1 (fn () =>
               let val (states, State (result_node, _)) =
                 (case st' of State (SOME (Proof (_, (_, orig_gthy)), exit), prev)
                   => State (SOME (Proof (ProofNode.init prf, (finish, orig_gthy)), exit), prev))
@@ -743,13 +758,13 @@
   let
     val end_pos = if null input then error "No input" else pos_of (fst (List.last input));
 
-    val immediate = not (Future.enabled ());
-    val (future_results, end_state) = fold_map (proof_result immediate) input toplevel;
+    val immediate = not (Goal.future_enabled ());
+    val (results, end_state) = fold_map (proof_result immediate) input toplevel;
     val _ =
       (case end_state of
-        State (NONE, SOME (Theory (Context.Theory _, _), _)) => ()
+        State (NONE, SOME (Theory (Context.Theory _, _), _)) =>
+          command (commit_exit end_pos) end_state
       | _ => error "Unfinished development at end of input");
-    val results = maps Lazy.force future_results;
-  in (results, fn () => ignore (command (commit_exit end_pos) end_state)) end;
+  in Lazy.lazy (fn () => maps Lazy.force results) end;
 
 end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Isar/value_parse.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -0,0 +1,46 @@
+(*  Title:      Pure/Isar/value_parse.ML
+    Author:     Makarius
+
+Outer syntax parsers for basic ML values.
+*)
+
+signature VALUE_PARSE =
+sig
+  type 'a parser = 'a OuterParse.parser
+  val comma: 'a parser -> 'a parser
+  val equal: 'a parser -> 'a parser
+  val parens: 'a parser -> 'a parser
+  val pair: 'a parser -> 'b parser -> ('a * 'b) parser
+  val triple: 'a parser -> 'b parser -> 'c parser -> ('a * 'b * 'c) parser
+  val list: 'a parser -> 'a list parser
+  val properties: Properties.T parser
+end;
+
+structure ValueParse: VALUE_PARSE =
+struct
+
+structure P = OuterParse;
+type 'a parser = 'a P.parser;
+
+
+(* syntax utilities *)
+
+fun comma p = P.$$$ "," |-- P.!!! p;
+fun equal p = P.$$$ "=" |-- P.!!! p;
+fun parens p = P.$$$ "(" |-- P.!!! (p --| P.$$$ ")");
+
+
+(* tuples *)
+
+val unit = parens (Scan.succeed ());
+fun pair p1 p2 = parens (p1 -- comma p2);
+fun triple p1 p2 p3 = parens (p1 -- comma p2 -- comma p3) >> P.triple1;
+
+
+(* lists *)
+
+fun list p = parens (P.enum "," p);
+val properties = list (P.string -- equal P.string);
+
+end;
+
--- a/src/Pure/ML-Systems/alice.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/ML-Systems/alice.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML-Systems/alice.ML
-    ID:         $Id$
 
 Compatibility file for Alice 1.4.
 
--- a/src/Pure/ML-Systems/exn.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/ML-Systems/exn.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML-Systems/exn.ML
-    ID:         $Id$
     Author:     Makarius
 
 Extra support for exceptions.
--- a/src/Pure/ML-Systems/install_pp_polyml.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/ML-Systems/install_pp_polyml.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML-Systems/install_pp_polyml.ML
-    ID:         $Id$
 
 Extra toplevel pretty-printing for Poly/ML.
 *)
--- a/src/Pure/ML-Systems/ml_name_space.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/ML-Systems/ml_name_space.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML-Systems/ml_name_space.ML
-    ID:         $Id$
     Author:     Makarius
 
 ML name space -- dummy version of Poly/ML 5.2 facility.
--- a/src/Pure/ML-Systems/mosml.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/ML-Systems/mosml.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,14 +1,13 @@
 (*  Title:      Pure/ML-Systems/mosml.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1999  University of Cambridge
+    Author:     Makarius
 
 Compatibility file for Moscow ML 2.01
 
 NOTE: saving images does not work; may run it interactively as follows:
 
 $ cd Isabelle/src/Pure
-$ isabelle RAW_ML_SYSTEM
+$ isabelle-process RAW_ML_SYSTEM
 > val ml_system = "mosml";
 > use "ML-Systems/mosml.ML";
 > use "ROOT.ML";
@@ -27,6 +26,7 @@
 fun forget_structure _ = ();
 
 load "Obj";
+load "Word";
 load "Bool";
 load "Int";
 load "Real";
@@ -35,6 +35,7 @@
 load "Process";
 load "FileSys";
 load "IO";
+load "CharVector";
 
 exception Interrupt;
 
@@ -54,6 +55,14 @@
 structure IntInf =
 struct
   fun divMod (x, y) = (x div y, x mod y);
+
+  local
+    fun log 1 a = a
+      | log n a = log (n div 2) (a + 1);
+  in
+    fun log2 n = if n > 0 then log n 0 else raise Domain;
+  end;
+
   open Int;
 end;
 
--- a/src/Pure/ML-Systems/multithreading.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/ML-Systems/multithreading.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML-Systems/multithreading.ML
-    ID:         $Id$
     Author:     Makarius
 
 Dummy implementation of multithreading setup.
--- a/src/Pure/ML-Systems/multithreading_polyml.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/ML-Systems/multithreading_polyml.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML-Systems/multithreading_polyml.ML
-    ID:         $Id$
     Author:     Makarius
 
 Multithreading in Poly/ML 5.2 or later (cf. polyml/basis/Thread.sml).
@@ -77,12 +76,12 @@
 
 fun with_attributes new_atts f x =
   let
-    val orig_atts = Thread.getAttributes ();
+    val orig_atts = safe_interrupts (Thread.getAttributes ());
     fun restore () = Thread.setAttributes orig_atts;
     val result =
       (Thread.setAttributes (safe_interrupts new_atts);
         Exn.Result (f orig_atts x) before restore ())
-      handle e => Exn.Exn e before restore ()
+      handle e => Exn.Exn e before restore ();
   in Exn.release result end;
 
 fun interruptible f = with_attributes regular_interrupts (fn _ => f);
--- a/src/Pure/ML-Systems/overloading_smlnj.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/ML-Systems/overloading_smlnj.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML-Systems/overloading_smlnj.ML
-    ID:         $Id$
     Author:     Makarius
 
 Overloading in SML/NJ (cf. smlnj/base/system/smlnj/init/pervasive.sml).
--- a/src/Pure/ML-Systems/polyml-4.1.3.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/ML-Systems/polyml-4.1.3.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML-Systems/polyml-4.1.3.ML
-    ID:         $Id$
 
 Compatibility wrapper for Poly/ML 4.1.3.
 *)
--- a/src/Pure/ML-Systems/polyml-4.1.4.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/ML-Systems/polyml-4.1.4.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML-Systems/polyml-4.1.4.ML
-    ID:         $Id$
 
 Compatibility wrapper for Poly/ML 4.1.4.
 *)
--- a/src/Pure/ML-Systems/polyml-4.2.0.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/ML-Systems/polyml-4.2.0.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML-Systems/polyml-4.2.0.ML
-    ID:         $Id$
 
 Compatibility wrapper for Poly/ML 4.2.0.
 *)
--- a/src/Pure/ML-Systems/polyml-5.0.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/ML-Systems/polyml-5.0.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML-Systems/polyml-5.0.ML
-    ID:         $Id$
 
 Compatibility wrapper for Poly/ML 5.0.
 *)
--- a/src/Pure/ML-Systems/polyml-5.1.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/ML-Systems/polyml-5.1.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML-Systems/polyml-5.1.ML
-    ID:         $Id$
 
 Compatibility wrapper for Poly/ML 5.1.
 *)
--- a/src/Pure/ML-Systems/polyml.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/ML-Systems/polyml.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML-Systems/polyml.ML
-    ID:         $Id$
 
 Compatibility wrapper for Poly/ML 5.2 or later.
 *)
--- a/src/Pure/ML-Systems/polyml_common.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/ML-Systems/polyml_common.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML-Systems/polyml_common.ML
-    ID:         $Id$
 
 Compatibility file for Poly/ML -- common part for 4.x and 5.x.
 *)
--- a/src/Pure/ML-Systems/polyml_old_basis.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/ML-Systems/polyml_old_basis.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML-Systems/polyml_old_basis.ML
-    ID:         $Id$
 
 Fixes for the old SML basis library (before Poly/ML 4.2.0).
 *)
--- a/src/Pure/ML-Systems/polyml_old_compiler4.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/ML-Systems/polyml_old_compiler4.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML-Systems/polyml_old_compiler4.ML
-    ID:         $Id$
 
 Runtime compilation -- for old PolyML.compiler (version 4.x).
 *)
--- a/src/Pure/ML-Systems/polyml_old_compiler5.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/ML-Systems/polyml_old_compiler5.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML-Systems/polyml_old_compiler5.ML
-    ID:         $Id$
 
 Runtime compilation -- for old PolyML.compilerEx (version 5.0, 5.1).
 *)
--- a/src/Pure/ML-Systems/proper_int.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/ML-Systems/proper_int.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML-Systems/proper_int.ML
-    ID:         $Id$
     Author:     Makarius
 
 SML basis with type int representing proper integers, not machine
--- a/src/Pure/ML-Systems/smlnj.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/ML-Systems/smlnj.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML-Systems/smlnj.ML
-    ID:         $Id$
 
 Compatibility file for Standard ML of New Jersey 110 or later.
 *)
--- a/src/Pure/ML-Systems/system_shell.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/ML-Systems/system_shell.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML-Systems/system_shell.ML
-    ID:         $Id$
     Author:     Makarius
 
 Generic system shell processes (no provisions to propagate interrupts;
--- a/src/Pure/ML-Systems/thread_dummy.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/ML-Systems/thread_dummy.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML-Systems/thread_dummy.ML
-    ID:         $Id$
     Author:     Makarius
 
 Default (mostly dummy) implementation of thread structures
--- a/src/Pure/ML-Systems/time_limit.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/ML-Systems/time_limit.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML-Systems/time_limit.ML
-    ID:         $Id$
     Author:     Makarius
 
 Dummy implementation of NJ's TimeLimit structure.
--- a/src/Pure/ML-Systems/universal.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/ML-Systems/universal.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML-Systems/universal.ML
-    ID:         $Id$
     Author:     Makarius
 
 Universal values via tagged union.  Emulates structure Universal
--- a/src/Pure/Proof/extraction.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Proof/extraction.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Proof/extraction.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Extraction of programs from proofs.
@@ -549,7 +548,7 @@
           let
             val prf = force_proof body;
             val (vs', tye) = find_inst prop Ts ts vs;
-            val tye' = (map fst (term_tvars prop) ~~ Ts') @ tye;
+            val tye' = (map fst (OldTerm.term_tvars prop) ~~ Ts') @ tye;
             val T = etype_of thy' vs' [] prop;
             val defs' = if T = nullT then defs
               else fst (extr d defs vs ts Ts hs prf0)
@@ -570,7 +569,7 @@
                     val corr_prf' = List.foldr forall_intr_prf
                       (proof_combt
                          (PThm (serial (),
-                          ((corr_name name vs', corr_prop, SOME (map TVar (term_tvars corr_prop))),
+                          ((corr_name name vs', corr_prop, SOME (map TVar (OldTerm.term_tvars corr_prop))),
                             Lazy.value (make_proof_body corr_prf))), vfs_of corr_prop))
                       (map (get_var_type corr_prop) (vfs_of prop))
                   in
@@ -588,7 +587,7 @@
       | corr d defs vs ts Ts hs (prf0 as PAxm (s, prop, SOME Ts')) _ _ =
           let
             val (vs', tye) = find_inst prop Ts ts vs;
-            val tye' = (map fst (term_tvars prop) ~~ Ts') @ tye
+            val tye' = (map fst (OldTerm.term_tvars prop) ~~ Ts') @ tye
           in
             if etype_of thy' vs' [] prop = nullT andalso
               realizes_null vs' prop aconv prop then (defs, prf0)
@@ -639,7 +638,7 @@
           let
             val prf = force_proof body;
             val (vs', tye) = find_inst prop Ts ts vs;
-            val tye' = (map fst (term_tvars prop) ~~ Ts') @ tye
+            val tye' = (map fst (OldTerm.term_tvars prop) ~~ Ts') @ tye
           in
             case Symtab.lookup realizers s of
               NONE => (case find vs' (find' s defs) of
@@ -676,12 +675,12 @@
                          (chtype [propT, T] combination_axm %> f %> f %> c %> t' %%
                            (chtype [T --> propT] reflexive_axm %> f) %%
                            PAxm (cname ^ "_def", eqn,
-                             SOME (map TVar (term_tvars eqn))))) %% corr_prf;
+                             SOME (map TVar (OldTerm.term_tvars eqn))))) %% corr_prf;
                     val corr_prop = Reconstruct.prop_of corr_prf';
                     val corr_prf'' = List.foldr forall_intr_prf
                       (proof_combt
                         (PThm (serial (),
-                         ((corr_name s vs', corr_prop, SOME (map TVar (term_tvars corr_prop))),
+                         ((corr_name s vs', corr_prop, SOME (map TVar (OldTerm.term_tvars corr_prop))),
                            Lazy.value (make_proof_body corr_prf'))), vfs_of corr_prop))
                       (map (get_var_type corr_prop) (vfs_of prop));
                   in
@@ -699,7 +698,7 @@
       | extr d defs vs ts Ts hs (prf0 as PAxm (s, prop, SOME Ts')) =
           let
             val (vs', tye) = find_inst prop Ts ts vs;
-            val tye' = (map fst (term_tvars prop) ~~ Ts') @ tye
+            val tye' = (map fst (OldTerm.term_tvars prop) ~~ Ts') @ tye
           in
             case find vs' (Symtab.lookup_list realizers s) of
               SOME (t, _) => (defs, subst_TVars tye' t)
@@ -739,7 +738,7 @@
            in
              thy'
              |> PureThy.store_thm (corr_name s vs,
-                  Thm.varifyT (funpow (length (term_vars corr_prop))
+                  Thm.varifyT (funpow (length (OldTerm.term_vars corr_prop))
                     (Thm.forall_elim_var 0) (forall_intr_frees
                       (ProofChecker.thm_of_proof thy'
                        (fst (Proofterm.freeze_thaw_prf prf))))))
--- a/src/Pure/Proof/proof_rewrite_rules.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Proof/proof_rewrite_rules.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Proof/proof_rewrite_rules.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Simplification functions for proof terms involving meta level rules.
@@ -196,7 +195,8 @@
   let
     fun rew_term Ts t =
       let
-        val frees = map Free (Name.invent_list (add_term_names (t, [])) "xa" (length Ts) ~~ Ts);
+        val frees =
+          map Free (Name.invent_list (OldTerm.add_term_names (t, [])) "xa" (length Ts) ~~ Ts);
         val t' = r (subst_bounds (frees, t));
         fun strip [] t = t
           | strip (_ :: xs) (Abs (_, _, t)) = strip xs t;
@@ -228,7 +228,7 @@
           if member (op =) defs s then
             let
               val vs = vars_of prop;
-              val tvars = term_tvars prop;
+              val tvars = OldTerm.term_tvars prop;
               val (_, rhs) = Logic.dest_equals prop;
               val rhs' = Term.betapplys (subst_TVars (map fst tvars ~~ Ts)
                 (fold_rev (fn x => fn b => Abs ("", dummyT, abstract_over (x, b))) vs rhs),
@@ -249,7 +249,9 @@
         val cnames = map (fst o dest_Const o fst) defs';
         val thms = fold_proof_atoms true
           (fn PThm (_, ((name, prop, _), _)) =>
-              if member (op =) defnames name orelse null (term_consts prop inter cnames) then I
+              if member (op =) defnames name orelse
+                not (exists_Const (member (op =) cnames o #1) prop)
+              then I
               else cons (name, SOME prop)
             | _ => I) [prf] [];
       in Reconstruct.expand_proof thy thms end;
--- a/src/Pure/Proof/proofchecker.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Proof/proofchecker.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Proof/proofchecker.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Simple proof checker based only on the core inference rules
@@ -31,13 +30,12 @@
 
 fun thm_of_proof thy prf =
   let
-    val prf_names = Proofterm.fold_proof_terms
-      (fold_aterms (fn Free (x, _) => Name.declare x | _ => I)) (K I) prf Name.context;
+    val prf_names = Proofterm.fold_proof_terms Term.declare_term_frees (K I) prf Name.context;
     val lookup = lookup_thm thy;
 
     fun thm_of_atom thm Ts =
       let
-        val tvars = term_tvars (Thm.full_prop_of thm);
+        val tvars = OldTerm.term_tvars (Thm.full_prop_of thm);
         val (fmap, thm') = Thm.varifyT' [] thm;
         val ctye = map (pairself (Thm.ctyp_of thy))
           (map TVar tvars @ map (fn ((_, S), ixn) => TVar (ixn, S)) fmap ~~ Ts)
--- a/src/Pure/Proof/reconstruct.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Proof/reconstruct.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Proof/reconstruct.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Reconstruction of partial proof terms.
@@ -139,8 +138,8 @@
 
     fun mk_cnstrts_atom env vTs prop opTs prf =
           let
-            val tvars = term_tvars prop;
-            val tfrees = term_tfrees prop;
+            val tvars = OldTerm.term_tvars prop;
+            val tfrees = OldTerm.term_tfrees prop;
             val (env', Ts) = (case opTs of
                 NONE => foldl_map mk_tvar (env, map snd tvars @ map snd tfrees)
               | SOME Ts => (env, Ts));
@@ -232,9 +231,6 @@
 
 (**** update list of free variables of constraints ****)
 
-val add_term_ixns = fold_aterms (fn Var (xi, _) => insert (op =) xi | _ => I);
-val add_typ_ixns = fold_atyps (fn TVar (ai, _) => insert (op =) ai | _ => I);
-
 fun upd_constrs env cs =
   let
     val Envir.Envir {asol, iTs, ...} = env;
@@ -242,8 +238,8 @@
       |> Vartab.fold (cons o #1) asol
       |> Vartab.fold (cons o #1) iTs;
     val vran = []
-      |> Vartab.fold (add_term_ixns o #2 o #2) asol
-      |> Vartab.fold (add_typ_ixns o #2 o #2) iTs;
+      |> Vartab.fold (Term.add_var_names o #2 o #2) asol
+      |> Vartab.fold (Term.add_tvar_namesT o #2 o #2) iTs;
     fun check_cs [] = []
       | check_cs ((u, p, vs)::ps) =
           let val vs' = subtract (op =) dom vs;
@@ -294,7 +290,7 @@
     val (t, prf, cs, env, _) = make_constraints_cprf thy
       (Envir.empty (maxidx_proof cprf ~1)) cprf';
     val cs' = map (fn p => (true, p, op union
-        (pairself (map (fst o dest_Var) o term_vars) p)))
+        (pairself (map (fst o dest_Var) o OldTerm.term_vars) p)))
       (map (pairself (Envir.norm_term env)) ((t, prop')::cs));
     val _ = message ("Solving remaining constraints (" ^ string_of_int (length cs') ^ ") ...");
     val env' = solve thy cs' env
@@ -303,7 +299,7 @@
   end;
 
 fun prop_of_atom prop Ts = subst_atomic_types
-  (map TVar (term_tvars prop) @ map TFree (term_tfrees prop) ~~ Ts)
+  (map TVar (OldTerm.term_tvars prop) @ map TFree (OldTerm.term_tfrees prop) ~~ Ts)
   (forall_intr_vfs prop);
 
 val head_norm = Envir.head_norm (Envir.empty 0);
@@ -370,9 +366,9 @@
                   end
               | SOME (maxidx', prf) => (maxidx' + maxidx + 1,
                   inc (maxidx + 1) prf, prfs));
-            val tfrees = term_tfrees prop;
+            val tfrees = OldTerm.term_tfrees prop;
             val tye = map (fn ((s, j), _) => (s, maxidx + 1 + j))
-              (term_tvars prop) @ map (rpair ~1 o fst) tfrees ~~ Ts;
+              (OldTerm.term_tvars prop) @ map (rpair ~1 o fst) tfrees ~~ Ts;
             val varify = map_type_tfree (fn p as (a, S) =>
               if member (op =) tfrees p then TVar ((a, ~1), S) else TFree p)
           in
--- a/src/Pure/ProofGeneral/pgip_parser.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/ProofGeneral/pgip_parser.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -92,18 +92,18 @@
 
 fun parse_span span =
   let
-    val kind = ThyEdit.span_kind span;
-    val toks = ThyEdit.span_content span;
-    val text = implode (map (PrintMode.setmp [] ThyEdit.present_token) toks);
+    val kind = ThySyntax.span_kind span;
+    val toks = ThySyntax.span_content span;
+    val text = implode (map (PrintMode.setmp [] ThySyntax.present_token) toks);
   in
     (case kind of
-      ThyEdit.Command name => parse_command name text
-    | ThyEdit.Ignored => [D.Whitespace {text = text}]
-    | ThyEdit.Malformed => [D.Unparseable {text = text}])
+      ThySyntax.Command name => parse_command name text
+    | ThySyntax.Ignored => [D.Whitespace {text = text}]
+    | ThySyntax.Malformed => [D.Unparseable {text = text}])
   end;
 
 
 fun pgip_parser pos str =
-  maps parse_span (ThyEdit.parse_spans (OuterKeyword.get_lexicons ()) pos str);
+  maps parse_span (ThySyntax.parse_spans (OuterKeyword.get_lexicons ()) pos str);
 
 end;
--- a/src/Pure/ProofGeneral/preferences.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/ProofGeneral/preferences.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ProofGeneral/preferences.ML
-    ID:         $Id$
     Author:     David Aspinall and Markus Wenzel
 
 User preferences for Isabelle which are maintained by the interface.
@@ -164,7 +163,10 @@
   proof_pref,
   nat_pref Multithreading.max_threads
     "max-threads"
-    "Maximum number of threads"];
+    "Maximum number of threads",
+  bool_pref Goal.parallel_proofs
+    "parallel-proofs"
+    "Check proofs in parallel"];
 
 val pure_preferences =
  [("Display", display_preferences),
--- a/src/Pure/ProofGeneral/proof_general_emacs.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/ProofGeneral/proof_general_emacs.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ProofGeneral/proof_general_emacs.ML
-    ID:         $Id$
     Author:     David Aspinall and Markus Wenzel
 
 Isabelle/Isar configuration for Emacs Proof General.
@@ -21,81 +20,77 @@
 (* print modes *)
 
 val proof_generalN = "ProofGeneralEmacs";  (*token markup (colouring vars, etc.)*)
-val pgasciiN = "PGASCII";                  (*plain 7-bit ASCII communication*)
 val thm_depsN = "thm_deps";                (*meta-information about theorem deps*)
 val test_markupN = "test_markup";          (*XML markup for everything*)
 
-val _ = Markup.add_mode test_markupN (fn (name, props) =>
-  if name = Markup.promptN then ("", "") else XML.output_markup (name, props));
-
-fun special oct =
-  if print_mode_active pgasciiN then Symbol.SOH ^ chr (ord (oct_char oct) - 167)
-  else oct_char oct;
+fun special ch = Symbol.SOH ^ ch;
 
 
-(* text output: print modes for xsymbol *)
+(* render markup *)
 
 local
 
-fun xsym_output "\\" = "\\\\"
-  | xsym_output s = if Symbol.is_raw s then Symbol.decode_raw s else s;
-
-fun xsymbols_output s =
-  if print_mode_active Symbol.xsymbolsN andalso exists_string (fn c => c = "\\") s then
-    let val syms = Symbol.explode s
-    in (implode (map xsym_output syms), Symbol.length syms) end
-  else Output.default_output s;
+fun render_trees ts = fold render_tree ts
+and render_tree (XML.Text s) = Buffer.add s
+  | render_tree (XML.Elem (name, props, ts)) =
+      let
+        val (bg1, en1) =
+          if name <> Markup.promptN andalso print_mode_active test_markupN
+          then XML.output_markup (name, props)
+          else Markup.no_output;
+        val (bg2, en2) =
+          if (case ts of [XML.Text _] => false | _ => true) then Markup.no_output
+          else if name = Markup.stateN then (special "O" ^ "\n", "\n" ^ special "P")
+          else if name = Markup.sendbackN then (special "W", special "X")
+          else if name = Markup.hiliteN then (special "0", special "1")
+          else if name = Markup.tclassN then (special "B", special "A")
+          else if name = Markup.tfreeN then (special "C", special "A")
+          else if name = Markup.tvarN then (special "D", special "A")
+          else if name = Markup.freeN then (special "E", special "A")
+          else if name = Markup.boundN then (special "F", special "A")
+          else if name = Markup.varN then (special "G", special "A")
+          else if name = Markup.skolemN then (special "H", special "A")
+          else Markup.no_output;
+      in
+        Buffer.add bg1 #>
+        Buffer.add bg2 #>
+        render_trees ts #>
+        Buffer.add en2 #>
+        Buffer.add en1
+      end;
 
 in
 
-fun setup_xsymbols_output () =
-  Output.add_mode Symbol.xsymbolsN xsymbols_output Symbol.encode_raw;
+fun render text =
+  Buffer.content (render_trees (YXML.parse_body text) Buffer.empty);
 
 end;
 
 
-(* common markup *)
+(* messages *)
+
+fun message bg en prfx text =
+  (case render text of
+    "" => ()
+  | s => Output.writeln_default (enclose bg en (prefix_lines prfx s)));
 
-val _ = Markup.add_mode proof_generalN (fn (name, props) =>
-  let
-    val (bg1, en1) =
-      if name = Markup.stateN then (special "366" ^ "\n", "\n" ^ special "367")
-      else if name = Markup.sendbackN then (special "376", special "377")
-      else if name = Markup.hiliteN then (special "327", special "330")
-      else if name = Markup.tclassN then (special "351", special "350")
-      else if name = Markup.tfreeN then (special "352", special "350")
-      else if name = Markup.tvarN then (special "353", special "350")
-      else if name = Markup.freeN then (special "354", special "350")
-      else if name = Markup.boundN then (special "355", special "350")
-      else if name = Markup.varN then (special "356", special "350")
-      else if name = Markup.skolemN then (special "357", special "350")
-      else ("", "");
-    val (bg2, en2) =
-      if name <> Markup.promptN andalso print_mode_active test_markupN
-      then XML.output_markup (name, props)
-      else ("", "");
-  in (bg1 ^ bg2, en2 ^ en1) end);
+fun setup_messages () =
+ (Output.writeln_fn := message "" "" "";
+  Output.status_fn := (fn s => ! Output.priority_fn s);
+  Output.priority_fn := message (special "I") (special "J") "";
+  Output.tracing_fn := message (special "I" ^ special "V") (special "J") "";
+  Output.debug_fn := message (special "K") (special "L") "+++ ";
+  Output.warning_fn := message (special "K") (special "L") "### ";
+  Output.error_fn := message (special "M") (special "N") "*** ";
+  Output.prompt_fn := (fn s => Output.std_output (render s ^ special "S")));
+
+fun panic s =
+  (message (special "M") (special "N") "!!! " ("## SYSTEM EXIT ##\n" ^ s); exit 1);
 
 
-(* messages and notification *)
-
-fun decorate bg en prfx =
-  Output.writeln_default o enclose bg en o prefix_lines prfx;
+(* notification *)
 
-fun setup_messages () =
- (Output.writeln_fn := Output.writeln_default;
-  Output.status_fn := (fn "" => () | s => ! Output.priority_fn s);
-  Output.priority_fn := (fn s => decorate (special "360") (special "361") "" s);
-  Output.tracing_fn := (fn s => decorate (special "360" ^ special "375") (special "361") "" s);
-  Output.debug_fn := (fn s => decorate (special "362") (special "363") "+++ " s);
-  Output.warning_fn := (fn s => decorate (special "362") (special "363") "### " s);
-  Output.error_fn := (fn s => decorate (special "364") (special "365") "*** " s);
-  Output.prompt_fn := (fn s => Output.std_output (s ^ special "372")));
-
-fun panic s =
-  (decorate (special "364") (special "365") "!!! " ("## SYSTEM EXIT ##\n" ^ s); exit 1);
-
-fun emacs_notify s = decorate (special "360") (special "361") "" s;
+val emacs_notify = message (special "I") (special "J") "";
 
 fun tell_clear_goals () =
   emacs_notify "Proof General, please clear the goals buffer.";
@@ -149,7 +144,7 @@
             (warning (cat_lines ["Failed to register theory: " ^ quote name, msg]);
               tell_file_retracted (ThyLoad.thy_path name))
       else ();
-    val _ = Isar.init_point ();
+    val _ = Isar.init ();
   in () end;
 
 
@@ -161,7 +156,7 @@
  (sync_thy_loader ();
   tell_clear_goals ();
   tell_clear_response ();
-  Isar.init_point ();
+  Isar.init ();
   welcome ());
 
 
@@ -237,7 +232,9 @@
   | init true =
       (! initialized orelse
         (Output.no_warnings init_outer_syntax ();
-          setup_xsymbols_output ();
+          Output.add_mode Symbol.xsymbolsN Symbol.output Output.default_escape;
+          Output.add_mode proof_generalN Output.default_output Output.default_escape;
+          Markup.add_mode proof_generalN YXML.output_markup;
           setup_messages ();
           ProofGeneralPgip.init_pgip_channel (! Output.priority_fn);
           setup_thy_loader ();
--- a/src/Pure/ProofGeneral/proof_general_pgip.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/ProofGeneral/proof_general_pgip.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -242,7 +242,7 @@
     (sync_thy_loader ();
      tell_clear_goals ();
      tell_clear_response ();
-     Isar.init_point ();
+     Isar.init ();
      welcome ());
 
 
--- a/src/Pure/ROOT.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/ROOT.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -25,7 +25,9 @@
 (*fundamental structures*)
 use "name.ML";
 use "term.ML";
+use "term_ord.ML";
 use "term_subst.ML";
+use "old_term.ML";
 use "logic.ML";
 use "General/pretty.ML";
 use "context.ML";
--- a/src/Pure/Syntax/ast.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Syntax/ast.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Syntax/ast.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Abstract syntax trees, translation rules, matching and normalization of asts.
--- a/src/Pure/Syntax/lexicon.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Syntax/lexicon.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Syntax/lexicon.ML
-    ID:         $Id$
     Author:     Tobias Nipkow and Markus Wenzel, TU Muenchen
 
 Lexer for the inner Isabelle syntax (terms and types).
@@ -169,10 +168,10 @@
   | VarSy       => Markup.var
   | TFreeSy     => Markup.tfree
   | TVarSy      => Markup.tvar
-  | NumSy       => Markup.num
-  | FloatSy     => Markup.float
-  | XNumSy      => Markup.xnum
-  | StrSy       => Markup.xstr
+  | NumSy       => Markup.numeral
+  | FloatSy     => Markup.numeral
+  | XNumSy      => Markup.numeral
+  | StrSy       => Markup.inner_string
   | Space       => Markup.none
   | Comment     => Markup.inner_comment
   | EOF         => Markup.none;
--- a/src/Pure/Syntax/mixfix.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Syntax/mixfix.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Syntax/mixfix.ML
-    ID:         $Id$
     Author:     Tobias Nipkow and Markus Wenzel, TU Muenchen
 
 Mixfix declarations, infixes, binders.
--- a/src/Pure/Syntax/parser.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Syntax/parser.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Syntax/parser.ML
-    ID:         $Id$
     Author:     Carsten Clasohm, Sonia Mahjoub, and Markus Wenzel, TU Muenchen
 
 General context-free parser for the inner syntax of terms, types, etc.
--- a/src/Pure/Syntax/printer.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Syntax/printer.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Syntax/printer.ML
-    ID:         $Id$
     Author:     Tobias Nipkow and Markus Wenzel, TU Muenchen
 
 Pretty printing of asts, terms, types and print (ast) translation.
--- a/src/Pure/Syntax/simple_syntax.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Syntax/simple_syntax.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Syntax/simple_syntax.ML
-    ID:         $Id$
     Author:     Makarius
 
 Simple syntax for types and terms --- for bootstrapping Pure.
--- a/src/Pure/Syntax/syn_ext.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Syntax/syn_ext.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Syntax/syn_ext.ML
-    ID:         $Id$
     Author:     Markus Wenzel and Carsten Clasohm, TU Muenchen
 
 Syntax extension (internal interface).
@@ -379,7 +378,7 @@
 
 fun stamp_trfun s (c, f) = (c, (f, s));
 fun mk_trfun tr = stamp_trfun (stamp ()) tr;
-fun eq_trfun ((_, s1:stamp), (_, s2)) = s1 = s2;
+fun eq_trfun ((_, s1: stamp), (_, s2)) = s1 = s2;
 
 
 (* token translations *)
@@ -387,7 +386,7 @@
 fun tokentrans_mode m trs = map (fn (s, f) => (m, s, f)) trs;
 
 val standard_token_classes =
-  ["class", "tfree", "tvar", "free", "bound", "var", "num", "float", "xnum", "xstr"];
+  ["class", "tfree", "tvar", "free", "bound", "var", "numeral", "inner_string"];
 
 val standard_token_markers = map (fn s => "_" ^ s) standard_token_classes;
 
--- a/src/Pure/Syntax/syn_trans.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Syntax/syn_trans.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Syntax/syn_trans.ML
-    ID:         $Id$
     Author:     Tobias Nipkow and Markus Wenzel, TU Muenchen
 
 Syntax translation functions.
@@ -264,7 +263,7 @@
   let
     val vars = vars_of tm;
     val body = body_of tm;
-    val rev_new_vars = rename_wrt_term body vars;
+    val rev_new_vars = Term.rename_wrt_term body vars;
     fun subst (x, T) b =
       if can Name.dest_internal x andalso not (Term.loose_bvar1 (b, 0))
       then (Const ("_idtdummy", T), incr_boundvars ~1 b)
@@ -302,7 +301,7 @@
     (strip_abss strip_abs_vars strip_abs_body (eta_contr tm));
 
 fun atomic_abs_tr' (x, T, t) =
-  let val [xT] = rename_wrt_term t [(x, T)]
+  let val [xT] = Term.rename_wrt_term t [(x, T)]
   in (mark_boundT xT, subst_bound (mark_bound (fst xT), t)) end;
 
 fun abs_ast_tr' (*"_abs"*) asts =
--- a/src/Pure/Syntax/type_ext.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Syntax/type_ext.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Syntax/type_ext.ML
-    ID:         $Id$
     Author:     Tobias Nipkow and Markus Wenzel, TU Muenchen
 
 Utilities for input and output of types.  Also the concrete syntax of
--- a/src/Pure/Thy/latex.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Thy/latex.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -174,7 +174,7 @@
 fun latex_markup (s, _) =
   if s = Markup.keywordN then ("\\isakeyword{", "}")
   else if s = Markup.commandN then ("\\isacommand{", "}")
-  else ("", "");
+  else Markup.no_output;
 
 fun latex_indent "" _ = ""
   | latex_indent s _ = enclose "\\isaindent{" "}" s;
--- a/src/Pure/Thy/present.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Thy/present.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -467,7 +467,7 @@
             val _ = add_file (Path.append html_prefix base_html,
               HTML.ml_file (Url.File base) (File.read path));
             in (Url.File base_html, Url.File raw_path, loadit) end
-      | NONE => error ("Browser info: expected to find ML file" ^ quote (Path.implode raw_path)));
+      | NONE => error ("Browser info: expected to find ML file " ^ quote (Path.implode raw_path)));
 
     val files_html = map prep_file files;
 
--- a/src/Pure/Thy/thy_edit.ML	Mon Jan 19 20:37:08 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,194 +0,0 @@
-(*  Title:      Pure/Thy/thy_edit.ML
-    ID:         $Id$
-    Author:     Makarius
-
-Basic editor operations on theory sources.
-*)
-
-signature THY_EDIT =
-sig
-  val token_source: Scan.lexicon * Scan.lexicon -> Position.T -> (string, 'a) Source.source ->
-    (OuterLex.token, (SymbolPos.T, Position.T * (Symbol.symbol, (string, 'a)
-      Source.source) Source.source) Source.source) Source.source
-  val parse_tokens: Scan.lexicon * Scan.lexicon -> Position.T -> string -> OuterLex.token list
-  val present_token: OuterLex.token -> output
-  val report_token: OuterLex.token -> unit
-  datatype span_kind = Command of string | Ignored | Malformed
-  type span
-  val span_kind: span -> span_kind
-  val span_content: span -> OuterLex.token list
-  val span_range: span -> Position.range
-  val span_source: (OuterLex.token, 'a) Source.source ->
-    (span, (OuterLex.token, 'a) Source.source) Source.source
-  val parse_spans: Scan.lexicon * Scan.lexicon -> Position.T -> string -> span list
-  val present_span: span -> output
-  val report_span: span -> unit
-  val unit_source: (span, 'a) Source.source ->
-    (span * span list * bool, (span, 'a) Source.source) Source.source
-end;
-
-structure ThyEdit: THY_EDIT =
-struct
-
-structure K = OuterKeyword;
-structure T = OuterLex;
-structure P = OuterParse;
-
-
-(** tokens **)
-
-(* parse *)
-
-fun token_source lexs pos src =
-  Symbol.source {do_recover = true} src
-  |> T.source {do_recover = SOME false} (K lexs) pos;
-
-fun parse_tokens lexs pos str =
-  Source.of_string str
-  |> token_source lexs pos
-  |> Source.exhaust;
-
-
-(* present *)
-
-local
-
-val token_kind_markup =
- fn T.Command       => (Markup.commandN, [])
-  | T.Keyword       => (Markup.keywordN, [])
-  | T.Ident         => Markup.ident
-  | T.LongIdent     => Markup.ident
-  | T.SymIdent      => Markup.ident
-  | T.Var           => Markup.ident
-  | T.TypeIdent     => Markup.ident
-  | T.TypeVar       => Markup.ident
-  | T.Nat           => Markup.ident
-  | T.String        => Markup.string
-  | T.AltString     => Markup.altstring
-  | T.Verbatim      => Markup.verbatim
-  | T.Space         => Markup.none
-  | T.Comment       => Markup.comment
-  | T.InternalValue => Markup.none
-  | T.Malformed     => Markup.malformed
-  | T.Error _       => Markup.malformed
-  | T.Sync          => Markup.control
-  | T.EOF           => Markup.control;
-
-in
-
-fun present_token tok =
-  Markup.enclose (token_kind_markup (T.kind_of tok)) (Output.output (T.unparse tok));
-
-fun report_token tok =
-  Position.report (token_kind_markup (T.kind_of tok)) (T.position_of tok);
-
-end;
-
-
-
-(** spans **)
-
-(* type span *)
-
-datatype span_kind = Command of string | Ignored | Malformed;
-datatype span = Span of span_kind * OuterLex.token list;
-
-fun span_kind (Span (k, _)) = k;
-fun span_content (Span (_, toks)) = toks;
-
-fun span_range span =
-  (case span_content span of
-    [] => (Position.none, Position.none)
-  | toks =>
-      let
-        val start_pos = T.position_of (hd toks);
-        val end_pos = T.end_position_of (List.last toks);
-      in (start_pos, end_pos) end);
-
-
-(* parse *)
-
-local
-
-val is_whitespace = T.is_kind T.Space orf T.is_kind T.Comment;
-
-val body = Scan.unless (Scan.many is_whitespace -- Scan.ahead (P.command || P.eof)) P.not_eof;
-
-val span =
-  Scan.ahead P.command -- P.not_eof -- Scan.repeat body
-    >> (fn ((name, c), bs) => Span (Command name, c :: bs)) ||
-  Scan.many1 is_whitespace >> (fn toks => Span (Ignored, toks)) ||
-  Scan.repeat1 body >> (fn toks => Span (Malformed, toks));
-
-in
-
-fun span_source src = Source.source T.stopper (Scan.bulk span) NONE src;
-
-end;
-
-fun parse_spans lexs pos str =
-  Source.of_string str
-  |> token_source lexs pos
-  |> span_source
-  |> Source.exhaust;
-
-
-(* present *)
-
-local
-
-fun kind_markup (Command name) = Markup.command_span name
-  | kind_markup Ignored = Markup.ignored_span
-  | kind_markup Malformed = Markup.malformed_span;
-
-in
-
-fun present_span span =
-  Markup.enclose (kind_markup (span_kind span)) (implode (map present_token (span_content span)));
-
-fun report_span span =
-  Position.report (kind_markup (span_kind span)) (Position.encode_range (span_range span));
-
-end;
-
-
-
-(** units: commands with proof **)
-
-(* scanning spans *)
-
-val eof = Span (Command "", []);
-
-fun is_eof (Span (Command "", _)) = true
-  | is_eof _ = false;
-
-val not_eof = not o is_eof;
-
-val stopper = Scan.stopper (K eof) is_eof;
-
-
-(* unit_source *)
-
-local
-
-fun command_with pred = Scan.one (fn (Span (Command name, _)) => pred name | _ => false);
-
-val proof = Scan.pass 1 (Scan.repeat (Scan.depend (fn d =>
-  if d <= 0 then Scan.fail
-  else
-    command_with K.is_qed_global >> pair ~1 ||
-    command_with K.is_proof_goal >> pair (d + 1) ||
-    (if d = 0 then Scan.fail else command_with K.is_qed >> pair (d - 1)) ||
-    Scan.unless (command_with K.is_theory) (Scan.one not_eof) >> pair d)) -- Scan.state);
-
-val unit =
-  command_with K.is_theory_goal -- proof >> (fn (a, (bs, d)) => (a, bs, d >= 0)) ||
-  Scan.one not_eof >> (fn a => (a, [], true));
-
-in
-
-fun unit_source src = Source.source stopper (Scan.bulk unit) NONE src;
-
-end;
-
-end;
--- a/src/Pure/Thy/thy_info.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Thy/thy_info.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Thy/thy_info.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Main part of theory loader database, including handling of theory and
@@ -24,6 +23,8 @@
   val get_parents: string -> string list
   val touch_thy: string -> unit
   val touch_child_thys: string -> unit
+  val remove_thy: string -> unit
+  val kill_thy: string -> unit
   val provide_file: Path.T -> string -> unit
   val load_file: bool -> Path.T -> unit
   val exec_file: bool -> Path.T -> Context.generic -> Context.generic
@@ -32,8 +33,6 @@
   val use_thys: string list -> unit
   val use_thy: string -> unit
   val time_use_thy: string -> unit
-  val remove_thy: string -> unit
-  val kill_thy: string -> unit
   val thy_ord: theory * theory -> order
   val begin_theory: string -> string list -> (Path.T * bool) list -> bool -> theory
   val end_theory: theory -> unit
@@ -232,6 +231,35 @@
 end;
 
 
+(* manage pending proofs *)
+
+fun join_thy name =
+  (case lookup_theory name of
+    NONE => []
+  | SOME thy => PureThy.join_proofs thy);
+
+fun cancel_thy name =
+  (case lookup_theory name of
+    NONE => ()
+  | SOME thy => PureThy.cancel_proofs thy);
+
+
+(* remove theory *)
+
+fun remove_thy name =
+  if is_finished name then error (loader_msg "cannot remove finished theory" [name])
+  else
+    let
+      val succs = thy_graph Graph.all_succs [name];
+      val _ = List.app cancel_thy succs;
+      val _ = priority (loader_msg "removing" succs);
+      val _ = CRITICAL (fn () =>
+        (List.app (perform Remove) succs; change_thys (Graph.del_nodes succs)));
+    in () end;
+
+val kill_thy = if_known_thy remove_thy;
+
+
 (* load_file *)
 
 local
@@ -300,50 +328,62 @@
     val _ = CRITICAL (fn () =>
       change_deps name (Option.map (fn {master, text, parents, files, ...} =>
         make_deps upd_time master text parents files)));
-    val _ = OuterSyntax.load_thy name pos text (time orelse ! Output.timing);
-  in
-    CRITICAL (fn () =>
-     (change_deps name
-        (Option.map (fn {update_time, master, parents, files, ...} =>
-          make_deps update_time master [] parents files));
-      perform Update name))
-  end;
+    val after_load = OuterSyntax.load_thy name pos text (time orelse ! Output.timing);
+    val _ =
+      CRITICAL (fn () =>
+       (change_deps name
+          (Option.map (fn {update_time, master, parents, files, ...} =>
+            make_deps update_time master [] parents files));
+        perform Update name));
+  in after_load end;
 
 
 (* scheduling loader tasks *)
 
-datatype task = Task of (unit -> unit) | Finished | Running;
+datatype task = Task of (unit -> unit -> unit) | Finished | Running;
 fun task_finished Finished = true | task_finished _ = false;
 
 local
 
-fun schedule_seq tasks =
-  Graph.topological_order tasks
-  |> List.app (fn name => (case Graph.get_node tasks name of Task body => body () | _ => ()));
-
 fun schedule_futures task_graph =
   let
     val tasks = Graph.topological_order task_graph |> map_filter (fn name =>
       (case Graph.get_node task_graph name of Task body => SOME (name, body) | _ => NONE));
 
-    fun future (name, body) tab =
+    val par_proofs = ! parallel_proofs;
+
+    fun fork (name, body) tab =
       let
         val deps = Graph.imm_preds task_graph name
           |> map_filter (fn parent =>
             (case Symtab.lookup tab parent of SOME future => SOME (parent, future) | NONE => NONE));
         fun failed (parent, future) = if can Future.join future then NONE else SOME parent;
 
-        val x = Future.fork_deps (map #2 deps) (fn () =>
+        val future = Future.fork_deps (map #2 deps) (fn () =>
           (case map_filter failed deps of
             [] => body ()
           | bad => error (loader_msg
               ("failed to load " ^ quote name ^ " (unresolved " ^ commas_quote bad ^ ")") [])));
+        val future' =
+          if par_proofs then future
+          else Future.map (fn after_load => (after_load (); fn () => ())) future;
+      in Symtab.update (name, future') tab end;
 
-      in (x, Symtab.update (name, x) tab) end;
+    val futures = fold fork tasks Symtab.empty;
 
-    val thy_results = Future.join_results (#1 (fold_map future tasks Symtab.empty));
-    val proof_results = PureThy.join_proofs (map_filter (try get_theory o #1) tasks);
-  in ignore (Exn.release_all (thy_results @ proof_results)) end;
+    val exns = rev tasks |> maps (fn (name, _) =>
+      let
+        val after_load = Future.join (the (Symtab.lookup futures name));
+        val proof_exns = join_thy name;
+        val _ = null proof_exns orelse raise Exn.EXCEPTIONS proof_exns;
+        val _ = after_load ();
+      in [] end handle exn => (kill_thy name; [exn]));
+
+  in ignore (Exn.release_all (map Exn.Exn (rev exns))) end;
+
+fun schedule_seq tasks =
+  Graph.topological_order tasks
+  |> List.app (fn name => (case Graph.get_node tasks name of Task body => body () () | _ => ()));
 
 in
 
@@ -464,19 +504,6 @@
 end;
 
 
-(* remove theory *)
-
-fun remove_thy name =
-  if is_finished name then error (loader_msg "cannot remove finished theory" [name])
-  else
-    let val succs = thy_graph Graph.all_succs [name] in
-      priority (loader_msg "removing" succs);
-      CRITICAL (fn () => (List.app (perform Remove) succs; change_thys (Graph.del_nodes succs)))
-    end;
-
-val kill_thy = if_known_thy remove_thy;
-
-
 (* update_time *)
 
 structure UpdateTime = TheoryDataFun
--- a/src/Pure/Thy/thy_load.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Thy/thy_load.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Thy/thy_load.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Theory loader primitives, including load path.
@@ -22,9 +21,11 @@
   val ml_exts: string list
   val ml_path: string -> Path.T
   val thy_path: string -> Path.T
+  val split_thy_path: Path.T -> Path.T * string
   val check_file: Path.T -> Path.T -> (Path.T * File.ident) option
   val check_ml: Path.T -> Path.T -> (Path.T * File.ident) option
   val check_thy: Path.T -> string -> Path.T * File.ident
+  val check_name: string -> string -> unit
   val deps_thy: Path.T -> string ->
    {master: Path.T * File.ident, text: string list, imports: string list, uses: Path.T list}
   val load_ml: Path.T -> Path.T -> Path.T * File.ident
@@ -62,6 +63,11 @@
 val ml_path = Path.ext "ML" o Path.basic;
 val thy_path = Path.ext "thy" o Path.basic;
 
+fun split_thy_path path =
+  (case Path.split_ext path of
+    (path', "thy") => (Path.dir path', Path.implode (Path.base path'))
+  | _ => error ("Bad theory file specification " ^ Path.implode path));
+
 
 (* check files *)
 
@@ -90,6 +96,11 @@
     | SOME thy_id => thy_id)
   end;
 
+fun check_name name name' =
+  if name = name' then ()
+  else error ("Filename " ^ quote (Path.implode (thy_path name)) ^
+    " does not agree with theory name " ^ quote name');
+
 
 (* theory deps *)
 
@@ -99,9 +110,7 @@
     val text = explode (File.read path);
     val (name', imports, uses) =
       ThyHeader.read (Path.position path) (Source.of_list_limited 8000 text);
-    val _ = name = name' orelse
-      error ("Filename " ^ quote (Path.implode (Path.base path)) ^
-        " does not agree with theory name " ^ quote name');
+    val _ = check_name name name';
     val uses = map (Path.explode o #1) uses;
   in {master = master, text = text, imports = imports, uses = uses} end;
 
--- a/src/Pure/Thy/thy_output.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Thy/thy_output.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -11,6 +11,7 @@
   val quotes: bool ref
   val indent: int ref
   val source: bool ref
+  val break: bool ref
   val add_commands: (string * (Args.src -> Toplevel.node option -> string)) list -> unit
   val add_options: (string * (string -> (unit -> string) -> unit -> string)) list -> unit
   val defined_command: string -> bool
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Thy/thy_syntax.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -0,0 +1,193 @@
+(*  Title:      Pure/Thy/thy_syntax.ML
+    Author:     Makarius
+
+Superficial theory syntax: tokens and spans.
+*)
+
+signature THY_SYNTAX =
+sig
+  val token_source: Scan.lexicon * Scan.lexicon -> Position.T -> (string, 'a) Source.source ->
+    (OuterLex.token, (SymbolPos.T, Position.T * (Symbol.symbol, (string, 'a)
+      Source.source) Source.source) Source.source) Source.source
+  val parse_tokens: Scan.lexicon * Scan.lexicon -> Position.T -> string -> OuterLex.token list
+  val present_token: OuterLex.token -> output
+  val report_token: OuterLex.token -> unit
+  datatype span_kind = Command of string | Ignored | Malformed
+  type span
+  val span_kind: span -> span_kind
+  val span_content: span -> OuterLex.token list
+  val span_range: span -> Position.range
+  val span_source: (OuterLex.token, 'a) Source.source ->
+    (span, (OuterLex.token, 'a) Source.source) Source.source
+  val parse_spans: Scan.lexicon * Scan.lexicon -> Position.T -> string -> span list
+  val present_span: span -> output
+  val report_span: span -> unit
+  val unit_source: (span, 'a) Source.source ->
+    (span * span list * bool, (span, 'a) Source.source) Source.source
+end;
+
+structure ThySyntax: THY_SYNTAX =
+struct
+
+structure K = OuterKeyword;
+structure T = OuterLex;
+structure P = OuterParse;
+
+
+(** tokens **)
+
+(* parse *)
+
+fun token_source lexs pos src =
+  Symbol.source {do_recover = true} src
+  |> T.source {do_recover = SOME false} (K lexs) pos;
+
+fun parse_tokens lexs pos str =
+  Source.of_string str
+  |> token_source lexs pos
+  |> Source.exhaust;
+
+
+(* present *)
+
+local
+
+val token_kind_markup =
+ fn T.Command       => (Markup.commandN, [])
+  | T.Keyword       => (Markup.keywordN, [])
+  | T.Ident         => Markup.ident
+  | T.LongIdent     => Markup.ident
+  | T.SymIdent      => Markup.ident
+  | T.Var           => Markup.var
+  | T.TypeIdent     => Markup.tfree
+  | T.TypeVar       => Markup.tvar
+  | T.Nat           => Markup.ident
+  | T.String        => Markup.string
+  | T.AltString     => Markup.altstring
+  | T.Verbatim      => Markup.verbatim
+  | T.Space         => Markup.none
+  | T.Comment       => Markup.comment
+  | T.InternalValue => Markup.none
+  | T.Malformed     => Markup.malformed
+  | T.Error _       => Markup.malformed
+  | T.Sync          => Markup.control
+  | T.EOF           => Markup.control;
+
+in
+
+fun present_token tok =
+  Markup.enclose (token_kind_markup (T.kind_of tok)) (Output.output (T.unparse tok));
+
+fun report_token tok =
+  Position.report (token_kind_markup (T.kind_of tok)) (T.position_of tok);
+
+end;
+
+
+
+(** spans **)
+
+(* type span *)
+
+datatype span_kind = Command of string | Ignored | Malformed;
+datatype span = Span of span_kind * OuterLex.token list;
+
+fun span_kind (Span (k, _)) = k;
+fun span_content (Span (_, toks)) = toks;
+
+fun span_range span =
+  (case span_content span of
+    [] => (Position.none, Position.none)
+  | toks =>
+      let
+        val start_pos = T.position_of (hd toks);
+        val end_pos = T.end_position_of (List.last toks);
+      in (start_pos, end_pos) end);
+
+
+(* parse *)
+
+local
+
+val is_whitespace = T.is_kind T.Space orf T.is_kind T.Comment;
+
+val body = Scan.unless (Scan.many is_whitespace -- Scan.ahead (P.command || P.eof)) P.not_eof;
+
+val span =
+  Scan.ahead P.command -- P.not_eof -- Scan.repeat body
+    >> (fn ((name, c), bs) => Span (Command name, c :: bs)) ||
+  Scan.many1 is_whitespace >> (fn toks => Span (Ignored, toks)) ||
+  Scan.repeat1 body >> (fn toks => Span (Malformed, toks));
+
+in
+
+fun span_source src = Source.source T.stopper (Scan.bulk span) NONE src;
+
+end;
+
+fun parse_spans lexs pos str =
+  Source.of_string str
+  |> token_source lexs pos
+  |> span_source
+  |> Source.exhaust;
+
+
+(* present *)
+
+local
+
+fun kind_markup (Command name) = Markup.command_span name
+  | kind_markup Ignored = Markup.ignored_span
+  | kind_markup Malformed = Markup.malformed_span;
+
+in
+
+fun present_span span =
+  Markup.enclose (kind_markup (span_kind span)) (implode (map present_token (span_content span)));
+
+fun report_span span =
+  Position.report (kind_markup (span_kind span)) (Position.encode_range (span_range span));
+
+end;
+
+
+
+(** units: commands with proof **)
+
+(* scanning spans *)
+
+val eof = Span (Command "", []);
+
+fun is_eof (Span (Command "", _)) = true
+  | is_eof _ = false;
+
+val not_eof = not o is_eof;
+
+val stopper = Scan.stopper (K eof) is_eof;
+
+
+(* unit_source *)
+
+local
+
+fun command_with pred = Scan.one (fn (Span (Command name, _)) => pred name | _ => false);
+
+val proof = Scan.pass 1 (Scan.repeat (Scan.depend (fn d =>
+  if d <= 0 then Scan.fail
+  else
+    command_with K.is_qed_global >> pair ~1 ||
+    command_with K.is_proof_goal >> pair (d + 1) ||
+    (if d = 0 then Scan.fail else command_with K.is_qed >> pair (d - 1)) ||
+    Scan.unless (command_with K.is_theory) (Scan.one not_eof) >> pair d)) -- Scan.state);
+
+val unit =
+  command_with K.is_theory_goal -- proof >> (fn (a, (bs, d)) => (a, bs, d >= 0)) ||
+  Scan.one not_eof >> (fn a => (a, [], true));
+
+in
+
+fun unit_source src = Source.source stopper (Scan.bulk unit) NONE src;
+
+end;
+
+end;
--- a/src/Pure/Tools/invoke.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Tools/invoke.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -6,9 +6,9 @@
 
 signature INVOKE =
 sig
-  val invoke: string * Attrib.src list -> Locale.expr -> string option list ->
+  val invoke: string * Attrib.src list -> Old_Locale.expr -> string option list ->
     (Binding.T * string option * mixfix) list -> bool -> Proof.state -> Proof.state
-  val invoke_i: string * attribute list -> Locale.expr -> term option list ->
+  val invoke_i: string * attribute list -> Old_Locale.expr -> term option list ->
     (Binding.T * typ option * mixfix) list -> bool -> Proof.state -> Proof.state
 end;
 
@@ -88,7 +88,7 @@
           |> (snd o ProofContext.note_thmss_i "" notes)
           |> ProofContext.restore_naming ctxt
         end) #>
-      Proof.put_facts NONE #> Seq.single;
+      Proof.put_facts NONE;
   in
     state'
     |> Proof.chain_facts chain_facts
@@ -104,8 +104,8 @@
 in
 
 fun invoke x =
-  gen_invoke Attrib.attribute Locale.read_expr Syntax.parse_term ProofContext.add_fixes x;
-fun invoke_i x = gen_invoke (K I) Locale.cert_expr (K I) ProofContext.add_fixes_i x;
+  gen_invoke Attrib.attribute Old_Locale.read_expr Syntax.parse_term ProofContext.add_fixes x;
+fun invoke_i x = gen_invoke (K I) Old_Locale.cert_expr (K I) ProofContext.add_fixes_i x;
 
 end;
 
--- a/src/Pure/Tools/isabelle_process.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Tools/isabelle_process.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,15 +1,14 @@
 (*  Title:      Pure/Tools/isabelle_process.ML
-    ID:         $Id$
     Author:     Makarius
 
 Isabelle process wrapper -- interaction via external program.
 
 General format of process output:
 
-  (a) unmarked stdout/stderr, no line structure (output should be
+  (1) unmarked stdout/stderr, no line structure (output should be
   processed immediately as it arrives);
 
-  (b) properly marked-up messages, e.g. for writeln channel
+  (2) properly marked-up messages, e.g. for writeln channel
 
   "\002A" ^ props ^ "\002,\n" ^ text ^ "\002.\n"
 
@@ -17,13 +16,14 @@
   each, and the remaining text is any number of lines (output is
   supposed to be processed in one piece);
 
-  (c) special init message holds "pid" and "session" property.
+  (3) special init message holds "pid" and "session" property;
+
+  (4) message content is encoded in YXML format.
 *)
 
 signature ISABELLE_PROCESS =
 sig
   val isabelle_processN: string
-  val xmlN: string
   val init: string -> unit
 end;
 
@@ -33,7 +33,6 @@
 (* print modes *)
 
 val isabelle_processN = "isabelle_process";
-val xmlN = "XML";
 
 val _ = Output.add_mode isabelle_processN Output.default_output Output.default_escape;
 val _ = Markup.add_mode isabelle_processN YXML.output_markup;
@@ -56,41 +55,32 @@
   let val clean = clean_string [Symbol.STX, "\n", "\r"]
   in implode (map (fn (x, y) => clean x ^ "=" ^ clean y ^ special_sep ^ "\n") props) end;
 
-fun message_text class ts =
-  let
-    val doc = XML.Elem (Markup.messageN, [(Markup.classN, class)], ts);
-    val msg =
-      if print_mode_active xmlN then XML.header ^ XML.string_of doc
-      else YXML.string_of doc;
-  in clean_string [Symbol.STX] msg end;
-
-fun message_pos ts = get_first get_pos ts
-and get_pos (elem as XML.Elem (name, atts, ts)) =
-      if name = Markup.positionN then SOME (Position.of_properties atts)
-      else message_pos ts
-  | get_pos _ = NONE;
+fun message_pos trees = trees |> get_first
+  (fn XML.Elem (name, atts, ts) =>
+        if name = Markup.positionN then SOME (Position.of_properties atts)
+        else message_pos ts
+    | _ => NONE);
 
 fun output out_stream s = NAMED_CRITICAL "IO" (fn () =>
   (TextIO.output (out_stream, s); TextIO.output (out_stream, "\n")));
 
 in
 
-fun message _ _ _ "" = ()
-  | message out_stream ch class body =
+fun message _ _ "" = ()
+  | message out_stream ch body =
       let
-        val (txt, pos) =
-          let val ts = YXML.parse_body body
-          in (message_text class ts, the_default Position.none (message_pos ts)) end;
+        val pos = the_default Position.none (message_pos (YXML.parse_body body));
         val props =
           Position.properties_of (Position.thread_data ())
           |> Position.default_properties pos;
+        val txt = clean_string [Symbol.STX] body;
       in output out_stream (special ch ^ message_props props ^ txt ^ special_end) end;
 
 fun init_message out_stream =
   let
     val pid = (Markup.pidN, process_id ());
     val session = (Markup.sessionN, List.last (Session.id ()) handle List.Empty => "unknown");
-    val text = message_text Markup.initN [XML.Text (Session.welcome ())];
+    val text = Session.welcome ();
   in output out_stream (special "A" ^ message_props [pid, session] ^ text ^ special_end) end;
 
 end;
@@ -123,13 +113,13 @@
     val _ = SimpleThread.fork false (auto_flush out_stream);
     val _ = SimpleThread.fork false (auto_flush TextIO.stdErr);
   in
-    Output.status_fn   := message out_stream "B" Markup.statusN;
-    Output.writeln_fn  := message out_stream "C" Markup.writelnN;
-    Output.priority_fn := message out_stream "D" Markup.priorityN;
-    Output.tracing_fn  := message out_stream "E" Markup.tracingN;
-    Output.warning_fn  := message out_stream "F" Markup.warningN;
-    Output.error_fn    := message out_stream "G" Markup.errorN;
-    Output.debug_fn    := message out_stream "H" Markup.debugN;
+    Output.status_fn   := message out_stream "B";
+    Output.writeln_fn  := message out_stream "C";
+    Output.priority_fn := message out_stream "D";
+    Output.tracing_fn  := message out_stream "E";
+    Output.warning_fn  := message out_stream "F";
+    Output.error_fn    := message out_stream "G";
+    Output.debug_fn    := message out_stream "H";
     Output.prompt_fn   := ignore;
     out_stream
   end;
--- a/src/Pure/Tools/isabelle_process.scala	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Tools/isabelle_process.scala	Mon Jan 19 21:20:18 2009 +0100
@@ -50,6 +50,21 @@
       ('2' : Int) -> Kind.STDOUT,
       ('3' : Int) -> Kind.SIGNAL,
       ('4' : Int) -> Kind.EXIT)
+    // message markup
+    val markup = Map(
+      Kind.INIT -> Markup.INIT,
+      Kind.STATUS -> Markup.STATUS,
+      Kind.WRITELN -> Markup.WRITELN,
+      Kind.PRIORITY -> Markup.PRIORITY,
+      Kind.TRACING -> Markup.TRACING,
+      Kind.WARNING -> Markup.WARNING,
+      Kind.ERROR -> Markup.ERROR,
+      Kind.DEBUG -> Markup.DEBUG,
+      Kind.SYSTEM -> Markup.SYSTEM,
+      Kind.STDIN -> Markup.STDIN,
+      Kind.STDOUT -> Markup.STDOUT,
+      Kind.SIGNAL -> Markup.SIGNAL,
+      Kind.EXIT -> Markup.EXIT)
     //}}}
     def is_raw(kind: Value) =
       kind == STDOUT
@@ -67,7 +82,10 @@
 
   class Result(val kind: Kind.Value, val props: Properties, val result: String) {
     override def toString = {
-      val res = XML.content(YXML.parse_failsafe(result)).mkString
+      val trees = YXML.parse_body_failsafe(result)
+      val res =
+        if (kind == Kind.STATUS) trees.map(_.toString).mkString
+        else trees.flatMap(XML.content(_).mkString).mkString
       if (props == null) kind.toString + " [[" + res + "]]"
       else kind.toString + " " + props.toString + " [[" + res + "]]"
     }
@@ -76,6 +94,10 @@
     def is_system = Kind.is_system(kind)
   }
 
+  def parse_message(kind: Kind.Value, result: String) = {
+    XML.Elem(Markup.MESSAGE, List((Markup.CLASS, Kind.markup(kind))),
+      YXML.parse_body_failsafe(result))
+  }
 }
 
 
--- a/src/Pure/Tools/isabelle_syntax.scala	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/Tools/isabelle_syntax.scala	Mon Jan 19 21:20:18 2009 +0100
@@ -13,7 +13,8 @@
 
   /* string token */
 
-  def append_string(str: String, result: StringBuilder) = {
+  def append_string(str: String, result: StringBuilder)
+  {
     result.append("\"")
     for (c <- str) {
       if (c < 32 || c == '\\' || c == '\"') {
@@ -27,16 +28,41 @@
     result.append("\"")
   }
 
-  def encode_string(str: String) = {
+  def encode_string(str: String) =
+  {
     val result = new StringBuilder(str.length + 20)
     append_string(str, result)
     result.toString
   }
 
 
+  /* list */
+
+  def append_list[A](append_elem: (A, StringBuilder) => Unit, body: Iterable[A],
+    result: StringBuilder)
+  {
+    result.append("(")
+    val elems = body.elements
+    if (elems.hasNext) append_elem(elems.next, result)
+    while (elems.hasNext) {
+      result.append(", ")
+      append_elem(elems.next, result)
+    }
+    result.append(")")
+  }
+
+  def encode_list[A](append_elem: (A, StringBuilder) => Unit, elems: Iterable[A]) =
+  {
+    val result = new StringBuilder(20)
+    append_list(append_elem, elems, result)
+    result.toString
+  }
+
+
   /* properties */
 
-  def append_properties(props: Properties, result: StringBuilder) = {
+  def append_properties(props: Properties, result: StringBuilder)
+  {
     result.append("(")
     val names = props.propertyNames.asInstanceOf[Enumeration[String]]
     while (names.hasMoreElements) {
--- a/src/Pure/axclass.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/axclass.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -295,7 +295,7 @@
   in
     thy
     |> Sign.primitive_classrel (c1, c2)
-    |> put_classrel ((c1, c2), Drule.unconstrainTs th)
+    |> put_classrel ((c1, c2), Thm.close_derivation (Drule.unconstrainTs th))
     |> perhaps complete_arities
   end;
 
--- a/src/Pure/codegen.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/codegen.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/codegen.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Generic code generator.
@@ -278,7 +277,7 @@
 
 fun preprocess_term thy t =
   let
-    val x = Free (Name.variant (add_term_names (t, [])) "x", fastype_of t);
+    val x = Free (Name.variant (OldTerm.add_term_names (t, [])) "x", fastype_of t);
     (* fake definition *)
     val eq = setmp quick_and_dirty true (SkipProof.make_thm thy)
       (Logic.mk_equals (x, t));
@@ -460,7 +459,7 @@
 
 fun rename_terms ts =
   let
-    val names = List.foldr add_term_names
+    val names = List.foldr OldTerm.add_term_names
       (map (fst o fst) (rev (fold Term.add_vars ts []))) ts;
     val reserved = filter ML_Syntax.is_reserved names;
     val (illegal, alt_names) = split_list (map_filter (fn s =>
@@ -485,7 +484,7 @@
 (**** retrieve definition of constant ****)
 
 fun is_instance T1 T2 =
-  Type.raw_instance (T1, if null (typ_tfrees T2) then T2 else Logic.varifyT T2);
+  Type.raw_instance (T1, if null (OldTerm.typ_tfrees T2) then T2 else Logic.varifyT T2);
 
 fun get_assoc_code thy (s, T) = Option.map snd (find_first (fn ((s', T'), _) =>
   s = s' andalso is_instance T T') (#consts (CodegenData.get thy)));
@@ -598,8 +597,8 @@
      else Pretty.block (separate (Pretty.brk 1) (p :: ps));
 
 fun new_names t xs = Name.variant_list
-  (map (fst o fst o dest_Var) (term_vars t) union
-    add_term_names (t, ML_Syntax.reserved_names)) (map mk_id xs);
+  (map (fst o fst o dest_Var) (OldTerm.term_vars t) union
+    OldTerm.add_term_names (t, ML_Syntax.reserved_names)) (map mk_id xs);
 
 fun new_name t x = hd (new_names t [x]);
 
@@ -917,9 +916,9 @@
     val ctxt = ProofContext.init thy;
     val e =
       let
-        val _ = (null (term_tvars t) andalso null (term_tfrees t)) orelse
+        val _ = (null (Term.add_tvars t []) andalso null (Term.add_tfrees t [])) orelse
           error "Term to be evaluated contains type variables";
-        val _ = (null (term_vars t) andalso null (term_frees t)) orelse
+        val _ = (null (Term.add_vars t []) andalso null (Term.add_frees t [])) orelse
           error "Term to be evaluated contains variables";
         val (code, gr) = setmp mode ["term_of"]
           (generate_code_i thy [] "Generated")
--- a/src/Pure/context.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/context.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -132,7 +132,15 @@
 
 val copy_data = Datatab.map' invoke_copy;
 val extend_data = Datatab.map' invoke_extend;
-fun merge_data pp = Datatab.join (invoke_merge pp) o pairself extend_data;
+
+fun merge_data pp (data1, data2) =
+  Datatab.keys (Datatab.merge (K true) (data1, data2))
+  |> Par_List.map (fn k =>
+    (case (Datatab.lookup data1 k, Datatab.lookup data2 k) of
+      (SOME x, NONE) => (k, invoke_extend k x)
+    | (NONE, SOME y) => (k, invoke_extend k y)
+    | (SOME x, SOME y) => (k, invoke_merge pp k (invoke_extend k x, invoke_extend k y))))
+  |> Datatab.make;
 
 end;
 
--- a/src/Pure/drule.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/drule.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,7 +1,5 @@
 (*  Title:      Pure/drule.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1993  University of Cambridge
 
 Derived rules and other operations on theorems.
 *)
@@ -112,6 +110,7 @@
   val sort_triv: theory -> typ * sort -> thm list
   val unconstrainTs: thm -> thm
   val with_subgoal: int -> (thm -> thm) -> thm -> thm
+  val comp_no_flatten: thm * int -> int -> thm -> thm
   val rename_bvars: (string * string) list -> thm -> thm
   val rename_bvars': string option list -> thm -> thm
   val incr_type_indexes: int -> thm -> thm
@@ -340,7 +339,7 @@
      val thy = Thm.theory_of_thm fth
      val {prop, tpairs, ...} = rep_thm fth
  in
-   case List.foldr add_term_vars [] (prop :: Thm.terms_of_tpairs tpairs) of
+   case List.foldr OldTerm.add_term_vars [] (prop :: Thm.terms_of_tpairs tpairs) of
        [] => (fth, fn i => fn x => x)   (*No vars: nothing to do!*)
      | vars =>
          let fun newName (Var(ix,_)) = (ix, gensym (string_of_indexname ix))
@@ -363,13 +362,13 @@
      val thy = Thm.theory_of_thm fth
      val {prop, tpairs, ...} = rep_thm fth
  in
-   case List.foldr add_term_vars [] (prop :: Thm.terms_of_tpairs tpairs) of
+   case List.foldr OldTerm.add_term_vars [] (prop :: Thm.terms_of_tpairs tpairs) of
        [] => (fth, fn x => x)
      | vars =>
          let fun newName (Var(ix,_), (pairs,used)) =
                    let val v = Name.variant used (string_of_indexname ix)
                    in  ((ix,v)::pairs, v::used)  end;
-             val (alist, _) = List.foldr newName ([], Library.foldr add_term_names
+             val (alist, _) = List.foldr newName ([], Library.foldr OldTerm.add_term_names
                (prop :: Thm.terms_of_tpairs tpairs, [])) vars
              fun mk_inst (Var(v,T)) =
                  (cterm_of thy (Var(v,T)),
@@ -807,8 +806,8 @@
 
 (*Increment the indexes of only the type variables*)
 fun incr_type_indexes inc th =
-  let val tvs = term_tvars (prop_of th)
-      and thy = theory_of_thm th
+  let val tvs = OldTerm.term_tvars (prop_of th)
+      and thy = Thm.theory_of_thm th
       fun inc_tvar ((a,i),s) = pairself (ctyp_of thy) (TVar ((a,i),s), TVar ((a,i+inc),s))
   in Thm.instantiate (map inc_tvar tvs, []) th end;
 
@@ -820,6 +819,14 @@
 fun th1 INCR_COMP th2 = incr_indexes th2 th1 COMP th2;
 fun th1 COMP_INCR th2 = th1 COMP incr_indexes th1 th2;
 
+fun comp_no_flatten (th, n) i rule =
+  (case distinct Thm.eq_thm (Seq.list_of
+      (Thm.compose_no_flatten false (th, n) i (incr_indexes th rule))) of
+    [th'] => th'
+  | [] => raise THM ("comp_no_flatten", i, [th, rule])
+  | _ => raise THM ("comp_no_flatten: unique result expected", i, [th, rule]));
+
+
 
 (*** Instantiate theorem th, reading instantiations in theory thy ****)
 
--- a/src/Pure/envir.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/envir.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,7 +1,5 @@
 (*  Title:      Pure/envir.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1988  University of Cambridge
 
 Environments.  The type of a term variable / sort of a type variable is
 part of its name. The lookup function must apply type substitutions,
@@ -118,7 +116,7 @@
 fun vupdate ((aU as (a, U), t), env as Envir {iTs, ...}) = case t of
       Var (nT as (name', T)) =>
         if a = name' then env     (*cycle!*)
-        else if Term.indexname_ord (a, name') = LESS then
+        else if TermOrd.indexname_ord (a, name') = LESS then
            (case lookup (env, nT) of  (*if already assigned, chase*)
                 NONE => update ((nT, Var (a, T)), env)
               | SOME u => vupdate ((aU, u), env))
--- a/src/Pure/facts.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/facts.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/facts.ML
-    ID:         $Id$
     Author:     Makarius
 
 Environment of named facts, optionally indexed by proposition.
@@ -166,7 +165,7 @@
 
 (* indexed props *)
 
-val prop_ord = Term.term_ord o pairself Thm.full_prop_of;
+val prop_ord = TermOrd.term_ord o pairself Thm.full_prop_of;
 
 fun props (Facts {props, ...}) = sort_distinct prop_ord (Net.content props);
 fun could_unify (Facts {props, ...}) = Net.unify_term props;
--- a/src/Pure/goal.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/goal.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,12 +1,12 @@
 (*  Title:      Pure/goal.ML
-    ID:         $Id$
-    Author:     Makarius and Lawrence C Paulson
+    Author:     Makarius
 
 Goals in tactical theorem proving.
 *)
 
 signature BASIC_GOAL =
 sig
+  val parallel_proofs: bool ref
   val SELECT_GOAL: tactic -> int -> tactic
   val CONJUNCTS: tactic -> int -> tactic
   val PRECISE_CONJUNCTS: int -> tactic -> int -> tactic
@@ -20,6 +20,7 @@
   val conclude: thm -> thm
   val finish: thm -> thm
   val norm_result: thm -> thm
+  val future_enabled: unit -> bool
   val future_result: Proof.context -> thm future -> term -> thm
   val prove_internal: cterm list -> cterm -> (thm list -> tactic) -> thm
   val prove_multi: Proof.context -> string list -> term list -> term list ->
@@ -58,18 +59,14 @@
   --- (protect)
   #C
 *)
-fun protect th = th COMP_INCR Drule.protectI;
+fun protect th = Drule.comp_no_flatten (th, 0) 1 Drule.protectI;
 
 (*
   A ==> ... ==> #C
   ---------------- (conclude)
   A ==> ... ==> C
 *)
-fun conclude th =
-  (case SINGLE (Thm.compose_no_flatten false (th, Thm.nprems_of th) 1)
-      (Drule.incr_indexes th Drule.protectD) of
-    SOME th' => th'
-  | NONE => raise THM ("Failed to conclude goal", 0, [th]));
+fun conclude th = Drule.comp_no_flatten (th, Thm.nprems_of th) 1 Drule.protectD;
 
 (*
   #C
@@ -96,6 +93,14 @@
   #> Drule.zero_var_indexes;
 
 
+(* future_enabled *)
+
+val parallel_proofs = ref true;
+
+fun future_enabled () =
+  Future.enabled () andalso ! parallel_proofs andalso is_some (Future.thread_data ());
+
+
 (* future_result *)
 
 fun future_result ctxt result prop =
@@ -177,9 +182,9 @@
             else err ("Proved a different theorem: " ^ string_of_term (Thm.prop_of res))
           end);
     val res =
-      if immediate orelse #maxidx (Thm.rep_cterm stmt) >= 0 orelse not (Future.enabled ())
+      if immediate orelse #maxidx (Thm.rep_cterm stmt) >= 0 orelse not (future_enabled ())
       then result ()
-      else future_result ctxt' (Future.fork_pri 1 result) (Thm.term_of stmt);
+      else future_result ctxt' (Future.fork_pri ~1 result) (Thm.term_of stmt);
   in
     Conjunction.elim_balanced (length props) res
     |> map (Assumption.export false ctxt' ctxt)
--- a/src/Pure/logic.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/logic.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,7 +1,6 @@
 (*  Title:      Pure/logic.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   Cambridge University 1992
+    Author:     Makarius
 
 Abstract syntax operations of the Pure meta-logic.
 *)
@@ -515,7 +514,7 @@
 (*reverses parameters for substitution*)
 fun goal_params st i =
   let val gi = get_goal st i
-      val rfrees = map Free (rename_wrt_term gi (strip_params gi))
+      val rfrees = map Free (Term.rename_wrt_term gi (strip_params gi))
   in (gi, rfrees) end;
 
 fun concl_of_goal st i =
--- a/src/Pure/meta_simplifier.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/meta_simplifier.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,6 +1,5 @@
 (*  Title:      Pure/meta_simplifier.ML
-    ID:         $Id$
-    Author:     Tobias Nipkow and Stefan Berghofer
+    Author:     Tobias Nipkow and Stefan Berghofer, TU Muenchen
 
 Meta-level Simplification.
 *)
@@ -788,7 +787,7 @@
   mk_eq_True = K NONE,
   reorient = default_reorient};
 
-val empty_ss = init_ss basic_mk_rews Term.termless (K (K no_tac)) ([], []);
+val empty_ss = init_ss basic_mk_rews TermOrd.termless (K (K no_tac)) ([], []);
 
 
 (* merge *)  (*NOTE: ignores some fields of 2nd simpset*)
--- a/src/Pure/more_thm.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/more_thm.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/more_thm.ML
-    ID:         $Id$
     Author:     Makarius
 
 Further operations on type ctyp/cterm/thm, outside the inference kernel.
@@ -128,12 +127,12 @@
     val {shyps = shyps1, hyps = hyps1, tpairs = tpairs1, prop = prop1, ...} = Thm.rep_thm th1;
     val {shyps = shyps2, hyps = hyps2, tpairs = tpairs2, prop = prop2, ...} = Thm.rep_thm th2;
   in
-    (case Term.fast_term_ord (prop1, prop2) of
+    (case TermOrd.fast_term_ord (prop1, prop2) of
       EQUAL =>
-        (case list_ord (prod_ord Term.fast_term_ord Term.fast_term_ord) (tpairs1, tpairs2) of
+        (case list_ord (prod_ord TermOrd.fast_term_ord TermOrd.fast_term_ord) (tpairs1, tpairs2) of
           EQUAL =>
-            (case list_ord Term.fast_term_ord (hyps1, hyps2) of
-              EQUAL => list_ord Term.sort_ord (shyps1, shyps2)
+            (case list_ord TermOrd.fast_term_ord (hyps1, hyps2) of
+              EQUAL => list_ord TermOrd.sort_ord (shyps1, shyps2)
             | ord => ord)
         | ord => ord)
     | ord => ord)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/old_term.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -0,0 +1,94 @@
+(*  Title:      Pure/old_term.ML
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+
+Some old-style term operations.
+*)
+
+signature OLD_TERM =
+sig
+  val it_term_types: (typ * 'a -> 'a) -> term * 'a -> 'a
+  val add_term_names: term * string list -> string list
+  val add_typ_tvars: typ * (indexname * sort) list -> (indexname * sort) list
+  val add_typ_tfree_names: typ * string list -> string list
+  val add_typ_tfrees: typ * (string * sort) list -> (string * sort) list
+  val add_term_tvars: term * (indexname * sort) list -> (indexname * sort) list
+  val add_term_tfrees: term * (string * sort) list -> (string * sort) list
+  val add_term_tfree_names: term * string list -> string list
+  val typ_tfrees: typ -> (string * sort) list
+  val typ_tvars: typ -> (indexname * sort) list
+  val term_tfrees: term -> (string * sort) list
+  val term_tvars: term -> (indexname * sort) list
+  val add_term_vars: term * term list -> term list
+  val term_vars: term -> term list
+  val add_term_frees: term * term list -> term list
+  val term_frees: term -> term list
+end;
+
+structure OldTerm: OLD_TERM =
+struct
+
+(*iterate a function over all types in a term*)
+fun it_term_types f =
+let fun iter(Const(_,T), a) = f(T,a)
+      | iter(Free(_,T), a) = f(T,a)
+      | iter(Var(_,T), a) = f(T,a)
+      | iter(Abs(_,T,t), a) = iter(t,f(T,a))
+      | iter(f$u, a) = iter(f, iter(u, a))
+      | iter(Bound _, a) = a
+in iter end
+
+(*Accumulates the names in the term, suppressing duplicates.
+  Includes Frees and Consts.  For choosing unambiguous bound var names.*)
+fun add_term_names (Const(a,_), bs) = insert (op =) (NameSpace.base a) bs
+  | add_term_names (Free(a,_), bs) = insert (op =) a bs
+  | add_term_names (f$u, bs) = add_term_names (f, add_term_names(u, bs))
+  | add_term_names (Abs(_,_,t), bs) = add_term_names(t,bs)
+  | add_term_names (_, bs) = bs;
+
+(*Accumulates the TVars in a type, suppressing duplicates.*)
+fun add_typ_tvars(Type(_,Ts),vs) = List.foldr add_typ_tvars vs Ts
+  | add_typ_tvars(TFree(_),vs) = vs
+  | add_typ_tvars(TVar(v),vs) = insert (op =) v vs;
+
+(*Accumulates the TFrees in a type, suppressing duplicates.*)
+fun add_typ_tfree_names(Type(_,Ts),fs) = List.foldr add_typ_tfree_names fs Ts
+  | add_typ_tfree_names(TFree(f,_),fs) = insert (op =) f fs
+  | add_typ_tfree_names(TVar(_),fs) = fs;
+
+fun add_typ_tfrees(Type(_,Ts),fs) = List.foldr add_typ_tfrees fs Ts
+  | add_typ_tfrees(TFree(f),fs) = insert (op =) f fs
+  | add_typ_tfrees(TVar(_),fs) = fs;
+
+(*Accumulates the TVars in a term, suppressing duplicates.*)
+val add_term_tvars = it_term_types add_typ_tvars;
+
+(*Accumulates the TFrees in a term, suppressing duplicates.*)
+val add_term_tfrees = it_term_types add_typ_tfrees;
+val add_term_tfree_names = it_term_types add_typ_tfree_names;
+
+(*Non-list versions*)
+fun typ_tfrees T = add_typ_tfrees(T,[]);
+fun typ_tvars T = add_typ_tvars(T,[]);
+fun term_tfrees t = add_term_tfrees(t,[]);
+fun term_tvars t = add_term_tvars(t,[]);
+
+
+(*Accumulates the Vars in the term, suppressing duplicates.*)
+fun add_term_vars (t, vars: term list) = case t of
+    Var   _ => OrdList.insert TermOrd.term_ord t vars
+  | Abs (_,_,body) => add_term_vars(body,vars)
+  | f$t =>  add_term_vars (f, add_term_vars(t, vars))
+  | _ => vars;
+
+fun term_vars t = add_term_vars(t,[]);
+
+(*Accumulates the Frees in the term, suppressing duplicates.*)
+fun add_term_frees (t, frees: term list) = case t of
+    Free   _ => OrdList.insert TermOrd.term_ord t frees
+  | Abs (_,_,body) => add_term_frees(body,frees)
+  | f$t =>  add_term_frees (f, add_term_frees(t, frees))
+  | _ => frees;
+
+fun term_frees t = add_term_frees(t,[]);
+
+end;
--- a/src/Pure/pattern.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/pattern.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,6 +1,5 @@
 (*  Title:      Pure/pattern.ML
-    ID:         $Id$
-    Author:     Tobias Nipkow, Christine Heinzelmann, and Stefan Berghofer
+    Author:     Tobias Nipkow, Christine Heinzelmann, and Stefan Berghofer, TU Muenchen
 
 Unification of Higher-Order Patterns.
 
@@ -223,7 +222,7 @@
                      fun lam(is) = mkabs(binders,is,app(H,map (idx is) ks));
                  in Envir.update (((G, Gty), lam js), Envir.update (((F, Fty), lam is), env'))
                  end;
-  in if Term.indexname_ord (G,F) = LESS then ff(F,Fty,is,G,Gty,js) else ff(G,Gty,js,F,Fty,is) end
+  in if TermOrd.indexname_ord (G,F) = LESS then ff(F,Fty,is,G,Gty,js) else ff(G,Gty,js,F,Fty,is) end
 
 fun unify_types thy (T,U) (env as Envir.Envir{asol,iTs,maxidx}) =
   if T=U then env
--- a/src/Pure/proofterm.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/proofterm.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/proofterm.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 LF style proof terms.
@@ -189,7 +188,7 @@
 
 (* proof body *)
 
-val oracle_ord = prod_ord fast_string_ord Term.fast_term_ord;
+val oracle_ord = prod_ord fast_string_ord TermOrd.fast_term_ord;
 fun thm_ord ((i, _): pthm, (j, _)) = int_ord (j, i);
 
 fun make_body prf =
@@ -412,15 +411,15 @@
 fun del_conflicting_tvars envT T = TermSubst.instantiateT
   (map_filter (fn ixnS as (_, S) =>
      (Type.lookup envT ixnS; NONE) handle TYPE _ =>
-        SOME (ixnS, TFree ("'dummy", S))) (typ_tvars T)) T;
+        SOME (ixnS, TFree ("'dummy", S))) (OldTerm.typ_tvars T)) T;
 
 fun del_conflicting_vars env t = TermSubst.instantiate
   (map_filter (fn ixnS as (_, S) =>
      (Type.lookup (type_env env) ixnS; NONE) handle TYPE _ =>
-        SOME (ixnS, TFree ("'dummy", S))) (term_tvars t),
+        SOME (ixnS, TFree ("'dummy", S))) (OldTerm.term_tvars t),
    map_filter (fn Var (ixnT as (_, T)) =>
      (Envir.lookup (env, ixnT); NONE) handle TYPE _ =>
-        SOME (ixnT, Free ("dummy", T))) (term_vars t)) t;
+        SOME (ixnT, Free ("dummy", T))) (OldTerm.term_vars t)) t;
 
 fun norm_proof env =
   let
@@ -545,15 +544,13 @@
   let
     val (fs, Tfs, vs, Tvs) = fold_proof_terms
       (fn t => fn (fs, Tfs, vs, Tvs) =>
-         (add_term_frees (t, fs), add_term_tfree_names (t, Tfs),
-          add_term_vars (t, vs), add_term_tvar_ixns (t, Tvs)))
+         (Term.add_free_names t fs, Term.add_tfree_names t Tfs,
+          Term.add_var_names t vs, Term.add_tvar_names t Tvs))
       (fn T => fn (fs, Tfs, vs, Tvs) =>
-         (fs, add_typ_tfree_names (T, Tfs),
-          vs, add_typ_ixns (Tvs, T)))
+         (fs, Term.add_tfree_namesT T Tfs,
+          vs, Term.add_tvar_namesT T Tvs))
       prf ([], [], [], []);
-    val fs' = map (fst o dest_Free) fs;
-    val vs' = map (fst o dest_Var) vs;
-    val names = vs' ~~ Name.variant_list fs' (map fst vs');
+    val names = vs ~~ Name.variant_list fs (map fst vs);
     val names' = Tvs ~~ Name.variant_list Tfs (map fst Tvs);
     val rnames = map swap names;
     val rnames' = map swap names';
@@ -591,8 +588,9 @@
   let
     val fs = Term.fold_types (Term.fold_atyps
       (fn TFree v => if member (op =) fixed v then I else insert (op =) v | _ => I)) t [];
-    val ixns = add_term_tvar_ixns (t, []);
-    val fmap = fs ~~ Name.variant_list (map #1 ixns) (map fst fs);
+    val used = Name.context
+      |> fold_types (fold_atyps (fn TVar ((a, _), _) => Name.declare a | _ => I)) t;
+    val fmap = fs ~~ (#1 (Name.variants (map fst fs) used));
     fun thaw (f as (a, S)) =
       (case AList.lookup (op =) fmap f of
         NONE => TFree f
@@ -614,8 +612,8 @@
 
 fun freezeT t prf =
   let
-    val used = it_term_types add_typ_tfree_names (t, [])
-    and tvars = map #1 (it_term_types add_typ_tvars (t, []));
+    val used = OldTerm.it_term_types OldTerm.add_typ_tfree_names (t, [])
+    and tvars = map #1 (OldTerm.it_term_types OldTerm.add_typ_tvars (t, []));
     val (alist, _) = List.foldr new_name ([], used) tvars;
   in
     (case alist of
--- a/src/Pure/pure_thy.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/pure_thy.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/pure_thy.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Theorem storage.  Pure theory syntax and logical content.
@@ -11,7 +10,8 @@
   val intern_fact: theory -> xstring -> string
   val defined_fact: theory -> string -> bool
   val hide_fact: bool -> string -> theory -> theory
-  val join_proofs: theory list -> unit Exn.result list
+  val join_proofs: theory -> exn list
+  val cancel_proofs: theory -> unit
   val get_fact: Context.generic -> theory -> Facts.ref -> thm list
   val get_thms: theory -> xstring -> thm list
   val get_thm: theory -> xstring -> thm
@@ -59,11 +59,11 @@
 
 structure FactsData = TheoryDataFun
 (
-  type T = Facts.T * unit lazy list;
-  val empty = (Facts.empty, []);
+  type T = Facts.T * (unit lazy list * Task_Queue.group list);
+  val empty = (Facts.empty, ([], []));
   val copy = I;
-  fun extend (facts, _) = (facts, []);
-  fun merge _ ((facts1, _), (facts2, _)) = (Facts.merge (facts1, facts2), []);
+  fun extend (facts, _) = (facts, ([], []));
+  fun merge _ ((facts1, _), (facts2, _)) = (Facts.merge (facts1, facts2), ([], []));
 );
 
 
@@ -81,7 +81,10 @@
 
 fun lazy_proof thm = Lazy.lazy (fn () => Thm.join_proof thm);
 
-val join_proofs = maps (fn thy => map (Exn.capture Lazy.force) (rev (#2 (FactsData.get thy))));
+fun join_proofs thy =
+  map_filter (Exn.get_exn o Lazy.force_result) (rev (#1 (#2 (FactsData.get thy))));
+
+fun cancel_proofs thy = List.app Future.cancel_group (#2 (#2 (FactsData.get thy)));
 
 
 
@@ -142,7 +145,8 @@
 (* enter_thms *)
 
 fun enter_proofs (thy, thms) =
-  (FactsData.map (apsnd (fold (cons o lazy_proof) thms)) thy, thms);
+  (FactsData.map (apsnd (fn (proofs, groups) =>
+    (fold (cons o lazy_proof) thms proofs, fold Thm.pending_groups thms groups))) thy, thms);
 
 fun enter_thms pre_name post_name app_att (b, thms) thy =
   if Binding.is_empty b
--- a/src/Pure/search.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/search.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/search.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson and Norbert Voelker
 
 Search tacticals.
@@ -42,10 +41,8 @@
 (** Instantiation of heaps for best-first search **)
 
 (*total ordering on theorems, allowing duplicates to be found*)
-structure ThmHeap =
-  HeapFun (type elem = int * thm
-    val ord = Library.prod_ord Library.int_ord
-      (Term.term_ord o Library.pairself (#prop o Thm.rep_thm)));
+structure ThmHeap = HeapFun(type elem = int * thm
+  val ord = prod_ord int_ord (TermOrd.term_ord o pairself Thm.prop_of));
 
 
 structure Search : SEARCH =
--- a/src/Pure/sorts.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/sorts.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/sorts.ML
-    ID:         $Id$
     Author:     Markus Wenzel and Stefan Berghofer, TU Muenchen
 
 The order-sorted algebra of type classes.
@@ -74,13 +73,13 @@
 
 (** ordered lists of sorts **)
 
-val make = OrdList.make Term.sort_ord;
-val op subset = OrdList.subset Term.sort_ord;
-val op union = OrdList.union Term.sort_ord;
-val subtract = OrdList.subtract Term.sort_ord;
+val make = OrdList.make TermOrd.sort_ord;
+val op subset = OrdList.subset TermOrd.sort_ord;
+val op union = OrdList.union TermOrd.sort_ord;
+val subtract = OrdList.subtract TermOrd.sort_ord;
 
-val remove_sort = OrdList.remove Term.sort_ord;
-val insert_sort = OrdList.insert Term.sort_ord;
+val remove_sort = OrdList.remove TermOrd.sort_ord;
+val insert_sort = OrdList.insert TermOrd.sort_ord;
 
 fun insert_typ (TFree (_, S)) Ss = insert_sort S Ss
   | insert_typ (TVar (_, S)) Ss = insert_sort S Ss
--- a/src/Pure/tactic.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/tactic.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,9 +1,7 @@
 (*  Title:      Pure/tactic.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1991  University of Cambridge
 
-Basic tactics.
+Fundamental tactics.
 *)
 
 signature BASIC_TACTIC =
@@ -192,7 +190,7 @@
 (*Determine print names of goal parameters (reversed)*)
 fun innermost_params i st =
   let val (_, _, Bi, _) = dest_state (st, i)
-  in rename_wrt_term Bi (Logic.strip_params Bi) end;
+  in Term.rename_wrt_term Bi (Logic.strip_params Bi) end;
 
 (*params of subgoal i as they are printed*)
 fun params_of_state i st = rev (innermost_params i st);
--- a/src/Pure/tctical.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/tctical.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,7 +1,5 @@
 (*  Title:      Pure/tctical.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1993  University of Cambridge
 
 Tacticals.
 *)
@@ -404,19 +402,18 @@
   (*Left-to-right replacements: ctpairs = [...,(vi,ti),...].
     Instantiates distinct free variables by terms of same type.*)
   fun free_instantiate ctpairs =
-      forall_elim_list (map snd ctpairs) o forall_intr_list (map fst ctpairs);
+    forall_elim_list (map snd ctpairs) o forall_intr_list (map fst ctpairs);
 
-  fun free_of s ((a,i), T) =
-        Free(s ^ (case i of 0 => a | _ => a ^ "_" ^ string_of_int i),
-             T)
+  fun free_of s ((a, i), T) =
+    Free (s ^ (case i of 0 => a | _ => a ^ "_" ^ string_of_int i), T)
 
-  fun mk_inst (var as Var(v,T))  = (var,  free_of "METAHYP1_" (v,T))
+  fun mk_inst v = (Var v, free_of "METAHYP1_" v)
 in
 
 (*Common code for METAHYPS and metahyps_thms*)
 fun metahyps_split_prem prem =
   let (*find all vars in the hyps -- should find tvars also!*)
-      val hyps_vars = List.foldr add_term_vars [] (Logic.strip_assums_hyp prem)
+      val hyps_vars = fold Term.add_vars (Logic.strip_assums_hyp prem) []
       val insts = map mk_inst hyps_vars
       (*replace the hyps_vars by Frees*)
       val prem' = subst_atomic insts prem
@@ -434,20 +431,18 @@
       val cparams = map cterm fparams
       fun swap_ctpair (t,u) = (cterm u, cterm t)
       (*Subgoal variables: make Free; lift type over params*)
-      fun mk_subgoal_inst concl_vars (var as Var(v,T)) =
-          if member (op =) concl_vars var
-          then (var, true, free_of "METAHYP2_" (v,T))
-          else (var, false,
-                free_of "METAHYP2_" (v, map #2 params --->T))
+      fun mk_subgoal_inst concl_vars (v, T) =
+          if member (op =) concl_vars (v, T)
+          then ((v, T), true, free_of "METAHYP2_" (v, T))
+          else ((v, T), false, free_of "METAHYP2_" (v, map #2 params ---> T))
       (*Instantiate subgoal vars by Free applied to params*)
-      fun mk_ctpair (t,in_concl,u) =
-          if in_concl then (cterm t,  cterm u)
-          else (cterm t,  cterm (list_comb (u,fparams)))
+      fun mk_ctpair (v, in_concl, u) =
+          if in_concl then (cterm (Var v), cterm u)
+          else (cterm (Var v), cterm (list_comb (u, fparams)))
       (*Restore Vars with higher type and index*)
-      fun mk_subgoal_swap_ctpair
-                (t as Var((a,i),_), in_concl, u as Free(_,U)) =
-          if in_concl then (cterm u, cterm t)
-          else (cterm u, cterm(Var((a, i+maxidx), U)))
+      fun mk_subgoal_swap_ctpair (((a, i), T), in_concl, u as Free (_, U)) =
+          if in_concl then (cterm u, cterm (Var ((a, i), T)))
+          else (cterm u, cterm (Var ((a, i + maxidx), U)))
       (*Embed B in the original context of params and hyps*)
       fun embed B = list_all_free (params, Logic.list_implies (hyps, B))
       (*Strip the context using elimination rules*)
@@ -456,8 +451,8 @@
       fun relift st =
         let val prop = Thm.prop_of st
             val subgoal_vars = (*Vars introduced in the subgoals*)
-                  List.foldr add_term_vars [] (Logic.strip_imp_prems prop)
-            and concl_vars = add_term_vars (Logic.strip_imp_concl prop, [])
+              fold Term.add_vars (Logic.strip_imp_prems prop) []
+            and concl_vars = Term.add_vars (Logic.strip_imp_concl prop) []
             val subgoal_insts = map (mk_subgoal_inst concl_vars) subgoal_vars
             val st' = Thm.instantiate ([], map mk_ctpair subgoal_insts) st
             val emBs = map (cterm o embed) (prems_of st')
@@ -490,9 +485,8 @@
   let
     fun typed ty = " has type: " ^ Syntax.string_of_typ_global thy ty;
     fun find_vars thy (Const (c, ty)) =
-        (case Term.typ_tvars ty
-         of [] => I
-          | _ => insert (op =) (c ^ typed ty))
+          if null (Term.add_tvarsT ty []) then I
+          else insert (op =) (c ^ typed ty)
       | find_vars thy (Var (xi, ty)) = insert (op =) (Term.string_of_vname xi ^ typed ty)
       | find_vars _ (Free _) = I
       | find_vars _ (Bound _) = I
--- a/src/Pure/term.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/term.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,12 +1,11 @@
 (*  Title:      Pure/term.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   Cambridge University 1992
+    Author:     Makarius
 
 Simply typed lambda-calculus: types, terms, and basic operations.
 *)
 
-infix 9  $;
+infix 9 $;
 infixr 5 -->;
 infixr --->;
 infix aconv;
@@ -75,12 +74,7 @@
   val fold_term_types: (term -> typ -> 'a -> 'a) -> term -> 'a -> 'a
   val fold_types: (typ -> 'a -> 'a) -> term -> 'a -> 'a
   val burrow_types: (typ list -> typ list) -> term list -> term list
-  val it_term_types: (typ * 'a -> 'a) -> term * 'a -> 'a
-  val add_term_names: term * string list -> string list
   val aconv: term * term -> bool
-  structure Vartab: TABLE
-  structure Typtab: TABLE
-  structure Termtab: TABLE
   val propT: typ
   val strip_all_body: term -> term
   val strip_all_vars: term -> (string * typ) list
@@ -94,8 +88,6 @@
   val subst_bound: term * term -> term
   val betapply: term * term -> term
   val betapplys: term * term list -> term
-  val eq_ix: indexname * indexname -> bool
-  val could_unify: term * term -> bool
   val subst_free: (term * term) list -> term -> term
   val abstract_over: term * term -> term
   val lambda: term -> term -> term
@@ -115,31 +107,10 @@
   val maxidx_of_typ: typ -> int
   val maxidx_of_typs: typ list -> int
   val maxidx_of_term: term -> int
-  val add_term_consts: term * string list -> string list
-  val term_consts: term -> string list
   val exists_subtype: (typ -> bool) -> typ -> bool
   val exists_type: (typ -> bool) -> term -> bool
   val exists_subterm: (term -> bool) -> term -> bool
   val exists_Const: (string * typ -> bool) -> term -> bool
-  val add_term_free_names: term * string list -> string list
-  val add_typ_tvars: typ * (indexname * sort) list -> (indexname * sort) list
-  val add_typ_tfree_names: typ * string list -> string list
-  val add_typ_tfrees: typ * (string * sort) list -> (string * sort) list
-  val add_typ_varnames: typ * string list -> string list
-  val add_term_tvars: term * (indexname * sort) list -> (indexname * sort) list
-  val add_term_tfrees: term * (string * sort) list -> (string * sort) list
-  val add_term_tfree_names: term * string list -> string list
-  val typ_tfrees: typ -> (string * sort) list
-  val typ_tvars: typ -> (indexname * sort) list
-  val term_tfrees: term -> (string * sort) list
-  val term_tvars: term -> (indexname * sort) list
-  val add_typ_ixns: indexname list * typ -> indexname list
-  val add_term_tvar_ixns: term * indexname list -> indexname list
-  val add_term_vars: term * term list -> term list
-  val term_vars: term -> term list
-  val add_term_frees: term * term list -> term list
-  val term_frees: term -> term list
-  val rename_wrt_term: term -> (string * 'a) list -> (string * 'a) list
   val show_question_marks: bool ref
 end;
 
@@ -151,41 +122,40 @@
   val a_itselfT: typ
   val all: typ -> term
   val argument_type_of: term -> int -> typ
-  val head_name_of: term -> string
+  val add_tvar_namesT: typ -> indexname list -> indexname list
+  val add_tvar_names: term -> indexname list -> indexname list
   val add_tvarsT: typ -> (indexname * sort) list -> (indexname * sort) list
   val add_tvars: term -> (indexname * sort) list -> (indexname * sort) list
+  val add_var_names: term -> indexname list -> indexname list
   val add_vars: term -> (indexname * typ) list -> (indexname * typ) list
-  val add_varnames: term -> indexname list -> indexname list
+  val add_tfree_namesT: typ -> string list -> string list
+  val add_tfree_names: term -> string list -> string list
   val add_tfreesT: typ -> (string * sort) list -> (string * sort) list
   val add_tfrees: term -> (string * sort) list -> (string * sort) list
+  val add_free_names: term -> string list -> string list
   val add_frees: term -> (string * typ) list -> (string * typ) list
+  val add_const_names: term -> string list -> string list
+  val add_consts: term -> (string * typ) list -> (string * typ) list
   val hidden_polymorphism: term -> (indexname * sort) list
+  val declare_typ_names: typ -> Name.context -> Name.context
+  val declare_term_names: term -> Name.context -> Name.context
+  val declare_term_frees: term -> Name.context -> Name.context
+  val variant_frees: term -> (string * 'a) list -> (string * 'a) list
+  val rename_wrt_term: term -> (string * 'a) list -> (string * 'a) list
+  val eq_ix: indexname * indexname -> bool
+  val eq_tvar: (indexname * sort) * (indexname * sort) -> bool
+  val eq_var: (indexname * typ) * (indexname * typ) -> bool
+  val could_unify: term * term -> bool
   val strip_abs_eta: int -> term -> (string * typ) list * term
-  val fast_indexname_ord: indexname * indexname -> order
-  val indexname_ord: indexname * indexname -> order
-  val sort_ord: sort * sort -> order
-  val typ_ord: typ * typ -> order
-  val fast_term_ord: term * term -> order
-  val term_ord: term * term -> order
-  val hd_ord: term * term -> order
-  val termless: term * term -> bool
-  val term_lpo: (term -> int) -> term * term -> order
   val match_bvars: (term * term) * (string * string) list -> (string * string) list
   val map_abs_vars: (string -> string) -> term -> term
   val rename_abs: term -> term -> term -> term option
-  val eq_tvar: (indexname * sort) * (indexname * sort) -> bool
-  val eq_var: (indexname * typ) * (indexname * typ) -> bool
-  val tvar_ord: (indexname * sort) * (indexname * sort) -> order
-  val var_ord: (indexname * typ) * (indexname * typ) -> order
   val close_schematic_term: term -> term
   val maxidx_typ: typ -> int -> int
   val maxidx_typs: typ list -> int -> int
   val maxidx_term: term -> int -> int
   val has_abs: term -> bool
   val dest_abs: string * typ * term -> string * term
-  val declare_typ_names: typ -> Name.context -> Name.context
-  val declare_term_names: term -> Name.context -> Name.context
-  val variant_frees: term -> (string * 'a) list -> (string * 'a) list
   val dummy_patternN: string
   val dummy_pattern: typ -> term
   val is_dummy_pattern: term -> bool
@@ -413,16 +383,6 @@
 fun head_of (f$t) = head_of f
   | head_of u = u;
 
-fun head_name_of tm =
-  (case head_of tm of
-    t as Const (c, _) =>
-      if NameSpace.is_qualified c then c
-      else raise TERM ("Malformed constant name", [t])
-  | t as Free (x, _) =>
-      if not (NameSpace.is_qualified x) then x
-      else raise TERM ("Malformed fixed variable name", [t])
-  | t => raise TERM ("No fixed head of term", [t]));
-
 (*number of atoms and abstractions in a term*)
 fun size_of_term tm =
   let
@@ -451,29 +411,16 @@
       | map_aux (t $ u) = map_aux t $ map_aux u;
   in map_aux end;
 
-(* iterate a function over all types in a term *)
-fun it_term_types f =
-let fun iter(Const(_,T), a) = f(T,a)
-      | iter(Free(_,T), a) = f(T,a)
-      | iter(Var(_,T), a) = f(T,a)
-      | iter(Abs(_,T,t), a) = iter(t,f(T,a))
-      | iter(f$u, a) = iter(f, iter(u, a))
-      | iter(Bound _, a) = a
-in iter end
-
 
 (* fold types and terms *)
 
-(*fold atoms of type*)
 fun fold_atyps f (Type (_, Ts)) = fold (fold_atyps f) Ts
   | fold_atyps f T = f T;
 
-(*fold atoms of term*)
 fun fold_aterms f (t $ u) = fold_aterms f t #> fold_aterms f u
   | fold_aterms f (Abs (_, _, t)) = fold_aterms f t
   | fold_aterms f a = f a;
 
-(*fold types of term*)
 fun fold_term_types f (t as Const (_, T)) = f t T
   | fold_term_types f (t as Free (_, T)) = f t T
   | fold_term_types f (t as Var (_, T)) = f t T
@@ -504,13 +451,20 @@
   in ts' end;
 
 (*collect variables*)
+val add_tvar_namesT = fold_atyps (fn TVar (xi, _) => insert (op =) xi | _ => I);
+val add_tvar_names = fold_types add_tvar_namesT;
 val add_tvarsT = fold_atyps (fn TVar v => insert (op =) v | _ => I);
 val add_tvars = fold_types add_tvarsT;
+val add_var_names = fold_aterms (fn Var (xi, _) => insert (op =) xi | _ => I);
 val add_vars = fold_aterms (fn Var v => insert (op =) v | _ => I);
-val add_varnames = fold_aterms (fn Var (xi, _) => insert (op =) xi | _ => I);
+val add_tfree_namesT = fold_atyps (fn TFree (xi, _) => insert (op =) xi | _ => I);
+val add_tfree_names = fold_types add_tfree_namesT;
 val add_tfreesT = fold_atyps (fn TFree v => insert (op =) v | _ => I);
 val add_tfrees = fold_types add_tfreesT;
+val add_free_names = fold_aterms (fn Free (x, _) => insert (op =) x | _ => I);
 val add_frees = fold_aterms (fn Free v => insert (op =) v | _ => I);
+val add_const_names = fold_aterms (fn Const (c, _) => insert (op =) c | _ => I);
+val add_consts = fold_aterms (fn Const c => insert (op =) c | _ => I);
 
 (*extra type variables in a term, not covered by its type*)
 fun hidden_polymorphism t =
@@ -522,10 +476,37 @@
   in extra_tvars end;
 
 
+(* renaming variables *)
 
-(** Comparing terms, types, sorts etc. **)
+val declare_typ_names = fold_atyps (fn TFree (a, _) => Name.declare a | _ => I);
+
+fun declare_term_names tm =
+  fold_aterms
+    (fn Const (a, _) => Name.declare (NameSpace.base a)
+      | Free (a, _) => Name.declare a
+      | _ => I) tm #>
+  fold_types declare_typ_names tm;
+
+val declare_term_frees = fold_aterms (fn Free (x, _) => Name.declare x | _ => I);
+
+fun variant_frees t frees =
+  fst (Name.variants (map fst frees) (declare_term_names t Name.context)) ~~ map snd frees;
 
-(* alpha equivalence -- tuned for equalities *)
+fun rename_wrt_term t frees = rev (variant_frees t frees);  (*reversed result!*)
+
+
+
+(** Comparing terms **)
+
+(* variables *)
+
+fun eq_ix ((x, i): indexname, (y, j)) = i = j andalso x = y;
+
+fun eq_tvar ((xi, S: sort), (xi', S')) = eq_ix (xi, xi') andalso S = S';
+fun eq_var ((xi, T: typ), (xi', T')) = eq_ix (xi, xi') andalso T = T';
+
+
+(* alpha equivalence *)
 
 fun tm1 aconv tm2 =
   pointer_eq (tm1, tm2) orelse
@@ -535,184 +516,24 @@
     | (a1, a2) => a1 = a2);
 
 
-(* fast syntactic ordering -- tuned for inequalities *)
-
-fun fast_indexname_ord ((x, i), (y, j)) =
-  (case int_ord (i, j) of EQUAL => fast_string_ord (x, y) | ord => ord);
-
-fun sort_ord SS =
-  if pointer_eq SS then EQUAL
-  else dict_ord fast_string_ord SS;
-
-local
-
-fun cons_nr (TVar _) = 0
-  | cons_nr (TFree _) = 1
-  | cons_nr (Type _) = 2;
-
-in
-
-fun typ_ord TU =
-  if pointer_eq TU then EQUAL
-  else
-    (case TU of
-      (Type (a, Ts), Type (b, Us)) =>
-        (case fast_string_ord (a, b) of EQUAL => dict_ord typ_ord (Ts, Us) | ord => ord)
-    | (TFree (a, S), TFree (b, S')) =>
-        (case fast_string_ord (a, b) of EQUAL => sort_ord (S, S') | ord => ord)
-    | (TVar (xi, S), TVar (yj, S')) =>
-        (case fast_indexname_ord (xi, yj) of EQUAL => sort_ord (S, S') | ord => ord)
-    | (T, U) => int_ord (cons_nr T, cons_nr U));
-
-end;
-
-local
-
-fun cons_nr (Const _) = 0
-  | cons_nr (Free _) = 1
-  | cons_nr (Var _) = 2
-  | cons_nr (Bound _) = 3
-  | cons_nr (Abs _) = 4
-  | cons_nr (_ $ _) = 5;
-
-fun struct_ord (Abs (_, _, t), Abs (_, _, u)) = struct_ord (t, u)
-  | struct_ord (t1 $ t2, u1 $ u2) =
-      (case struct_ord (t1, u1) of EQUAL => struct_ord (t2, u2) | ord => ord)
-  | struct_ord (t, u) = int_ord (cons_nr t, cons_nr u);
-
-fun atoms_ord (Abs (_, _, t), Abs (_, _, u)) = atoms_ord (t, u)
-  | atoms_ord (t1 $ t2, u1 $ u2) =
-      (case atoms_ord (t1, u1) of EQUAL => atoms_ord (t2, u2) | ord => ord)
-  | atoms_ord (Const (a, _), Const (b, _)) = fast_string_ord (a, b)
-  | atoms_ord (Free (x, _), Free (y, _)) = fast_string_ord (x, y)
-  | atoms_ord (Var (xi, _), Var (yj, _)) = fast_indexname_ord (xi, yj)
-  | atoms_ord (Bound i, Bound j) = int_ord (i, j)
-  | atoms_ord _ = sys_error "atoms_ord";
-
-fun types_ord (Abs (_, T, t), Abs (_, U, u)) =
-      (case typ_ord (T, U) of EQUAL => types_ord (t, u) | ord => ord)
-  | types_ord (t1 $ t2, u1 $ u2) =
-      (case types_ord (t1, u1) of EQUAL => types_ord (t2, u2) | ord => ord)
-  | types_ord (Const (_, T), Const (_, U)) = typ_ord (T, U)
-  | types_ord (Free (_, T), Free (_, U)) = typ_ord (T, U)
-  | types_ord (Var (_, T), Var (_, U)) = typ_ord (T, U)
-  | types_ord (Bound _, Bound _) = EQUAL
-  | types_ord _ = sys_error "types_ord";
-
-in
-
-fun fast_term_ord tu =
-  if pointer_eq tu then EQUAL
-  else
-    (case struct_ord tu of
-      EQUAL => (case atoms_ord tu of EQUAL => types_ord tu | ord => ord)
-    | ord => ord);
-
-structure Vartab = TableFun(type key = indexname val ord = fast_indexname_ord);
-structure Typtab = TableFun(type key = typ val ord = typ_ord);
-structure Termtab = TableFun(type key = term val ord = fast_term_ord);
-
-end;
-
-
-(* term_ord *)
-
-(*a linear well-founded AC-compatible ordering for terms:
-  s < t <=> 1. size(s) < size(t) or
-            2. size(s) = size(t) and s=f(...) and t=g(...) and f<g or
-            3. size(s) = size(t) and s=f(s1..sn) and t=f(t1..tn) and
-               (s1..sn) < (t1..tn) (lexicographically)*)
+(*A fast unification filter: true unless the two terms cannot be unified.
+  Terms must be NORMAL.  Treats all Vars as distinct. *)
+fun could_unify (t, u) =
+  let
+    fun matchrands (f $ t) (g $ u) = could_unify (t, u) andalso matchrands f g
+      | matchrands _ _ = true;
+  in
+    case (head_of t, head_of u) of
+      (_, Var _) => true
+    | (Var _, _) => true
+    | (Const (a, _), Const (b, _)) => a = b andalso matchrands t u
+    | (Free (a, _), Free (b, _)) => a = b andalso matchrands t u
+    | (Bound i, Bound j) => i = j andalso matchrands t u
+    | (Abs _, _) => true   (*because of possible eta equality*)
+    | (_, Abs _) => true
+    | _ => false
+  end;
 
-fun indexname_ord ((x, i), (y, j)) =
-  (case int_ord (i, j) of EQUAL => string_ord (x, y) | ord => ord);
-
-local
-
-fun hd_depth (t $ _, n) = hd_depth (t, n + 1)
-  | hd_depth p = p;
-
-fun dest_hd (Const (a, T)) = (((a, 0), T), 0)
-  | dest_hd (Free (a, T)) = (((a, 0), T), 1)
-  | dest_hd (Var v) = (v, 2)
-  | dest_hd (Bound i) = ((("", i), dummyT), 3)
-  | dest_hd (Abs (_, T, _)) = ((("", 0), T), 4);
-
-in
-
-fun term_ord tu =
-  if pointer_eq tu then EQUAL
-  else
-    (case tu of
-      (Abs (_, T, t), Abs(_, U, u)) =>
-        (case term_ord (t, u) of EQUAL => typ_ord (T, U) | ord => ord)
-    | (t, u) =>
-        (case int_ord (size_of_term t, size_of_term u) of
-          EQUAL =>
-            (case prod_ord hd_ord int_ord (hd_depth (t, 0), hd_depth (u, 0)) of
-              EQUAL => args_ord (t, u) | ord => ord)
-        | ord => ord))
-and hd_ord (f, g) =
-  prod_ord (prod_ord indexname_ord typ_ord) int_ord (dest_hd f, dest_hd g)
-and args_ord (f $ t, g $ u) =
-      (case args_ord (f, g) of EQUAL => term_ord (t, u) | ord => ord)
-  | args_ord _ = EQUAL;
-
-fun termless tu = (term_ord tu = LESS);
-
-end;
-
-
-(** Lexicographic path order on terms **)
-
-(*
-  See Baader & Nipkow, Term rewriting, CUP 1998.
-  Without variables.  Const, Var, Bound, Free and Abs are treated all as
-  constants.
-
-  f_ord maps terms to integers and serves two purposes:
-  - Predicate on constant symbols.  Those that are not recognised by f_ord
-    must be mapped to ~1.
-  - Order on the recognised symbols.  These must be mapped to distinct
-    integers >= 0.
-  The argument of f_ord is never an application.
-*)
-
-local
-
-fun unrecognized (Const (a, T)) = ((1, ((a, 0), T)), 0)
-  | unrecognized (Free (a, T)) = ((1, ((a, 0), T)), 0)
-  | unrecognized (Var v) = ((1, v), 1)
-  | unrecognized (Bound i) = ((1, (("", i), dummyT)), 2)
-  | unrecognized (Abs (_, T, _)) = ((1, (("", 0), T)), 3);
-
-fun dest_hd f_ord t =
-      let val ord = f_ord t in
-        if ord = ~1 then unrecognized t else ((0, (("", ord), fastype_of t)), 0)
-      end
-
-fun term_lpo f_ord (s, t) =
-  let val (f, ss) = strip_comb s and (g, ts) = strip_comb t in
-    if forall (fn si => term_lpo f_ord (si, t) = LESS) ss
-    then case hd_ord f_ord (f, g) of
-        GREATER =>
-          if forall (fn ti => term_lpo f_ord (s, ti) = GREATER) ts
-          then GREATER else LESS
-      | EQUAL =>
-          if forall (fn ti => term_lpo f_ord (s, ti) = GREATER) ts
-          then list_ord (term_lpo f_ord) (ss, ts)
-          else LESS
-      | LESS => LESS
-    else GREATER
-  end
-and hd_ord f_ord (f, g) = case (f, g) of
-    (Abs (_, T, t), Abs (_, U, u)) =>
-      (case term_lpo f_ord (t, u) of EQUAL => typ_ord (T, U) | ord => ord)
-  | (_, _) => prod_ord (prod_ord int_ord
-                  (prod_ord indexname_ord typ_ord)) int_ord
-                (dest_hd f_ord f, dest_hd f_ord g)
-in
-val term_lpo = term_lpo
-end;
 
 
 (** Connectives of higher order logic **)
@@ -842,7 +663,7 @@
   of bound variables and implicit eta-expansion*)
 fun strip_abs_eta k t =
   let
-    val used = fold_aterms (fn Free (v, _) => Name.declare v | _ => I) t Name.context;
+    val used = fold_aterms declare_term_frees t Name.context;
     fun strip_abs t (0, used) = (([], t), (0, used))
       | strip_abs (Abs (v, T, t)) (k, used) =
           let
@@ -863,35 +684,8 @@
   in (vs1 @ vs2, t'') end;
 
 
-(* comparing variables *)
-
-fun eq_ix ((x, i): indexname, (y, j)) = i = j andalso x = y;
-
-fun eq_tvar ((xi, S: sort), (xi', S')) = eq_ix (xi, xi') andalso S = S';
-fun eq_var ((xi, T: typ), (xi', T')) = eq_ix (xi, xi') andalso T = T';
-
-val tvar_ord = prod_ord indexname_ord sort_ord;
-val var_ord = prod_ord indexname_ord typ_ord;
-
-
-(*A fast unification filter: true unless the two terms cannot be unified.
-  Terms must be NORMAL.  Treats all Vars as distinct. *)
-fun could_unify (t,u) =
-  let fun matchrands (f$t, g$u) = could_unify(t,u) andalso  matchrands(f,g)
-        | matchrands _ = true
-  in case (head_of t , head_of u) of
-        (_, Var _) => true
-      | (Var _, _) => true
-      | (Const(a,_), Const(b,_)) =>  a=b andalso matchrands(t,u)
-      | (Free(a,_), Free(b,_)) =>  a=b andalso matchrands(t,u)
-      | (Bound i, Bound j) =>  i=j andalso matchrands(t,u)
-      | (Abs _, _) =>  true   (*because of possible eta equality*)
-      | (_, Abs _) =>  true
-      | _ => false
-  end;
-
 (*Substitute new for free occurrences of old in a term*)
-fun subst_free [] = (fn t=>t)
+fun subst_free [] = I
   | subst_free pairs =
       let fun substf u =
             case AList.lookup (op aconv) pairs u of
@@ -1011,11 +805,11 @@
 (** Identifying first-order terms **)
 
 (*Differs from proofterm/is_fun in its treatment of TVar*)
-fun is_funtype (Type("fun",[_,_])) = true
+fun is_funtype (Type ("fun", [_, _])) = true
   | is_funtype _ = false;
 
 (*Argument Ts is a reverse list of binder types, needed if term t contains Bound vars*)
-fun has_not_funtype Ts t = not (is_funtype (fastype_of1 (Ts,t)));
+fun has_not_funtype Ts t = not (is_funtype (fastype_of1 (Ts, t)));
 
 (*First order means in all terms of the form f(t1,...,tn) no argument has a
   function type. The supplied quantifiers are excluded: their argument always
@@ -1057,7 +851,7 @@
 
 
 
-(**** Syntax-related declarations ****)
+(** misc syntax operations **)
 
 (* substructure *)
 
@@ -1086,109 +880,13 @@
       | _ => false);
   in ex end;
 
+fun exists_Const P = exists_subterm (fn Const c => P c | _ => false);
+
 fun has_abs (Abs _) = true
   | has_abs (t $ u) = has_abs t orelse has_abs u
   | has_abs _ = false;
 
 
-
-(** Consts etc. **)
-
-fun add_term_consts (Const (c, _), cs) = insert (op =) c cs
-  | add_term_consts (t $ u, cs) = add_term_consts (t, add_term_consts (u, cs))
-  | add_term_consts (Abs (_, _, t), cs) = add_term_consts (t, cs)
-  | add_term_consts (_, cs) = cs;
-
-fun term_consts t = add_term_consts(t,[]);
-
-fun exists_Const P = exists_subterm (fn Const c => P c | _ => false);
-
-
-(** TFrees and TVars **)
-
-(*Accumulates the names of Frees in the term, suppressing duplicates.*)
-fun add_term_free_names (Free(a,_), bs) = insert (op =) a bs
-  | add_term_free_names (f$u, bs) = add_term_free_names (f, add_term_free_names(u, bs))
-  | add_term_free_names (Abs(_,_,t), bs) = add_term_free_names(t,bs)
-  | add_term_free_names (_, bs) = bs;
-
-(*Accumulates the names in the term, suppressing duplicates.
-  Includes Frees and Consts.  For choosing unambiguous bound var names.*)
-fun add_term_names (Const(a,_), bs) = insert (op =) (NameSpace.base a) bs
-  | add_term_names (Free(a,_), bs) = insert (op =) a bs
-  | add_term_names (f$u, bs) = add_term_names (f, add_term_names(u, bs))
-  | add_term_names (Abs(_,_,t), bs) = add_term_names(t,bs)
-  | add_term_names (_, bs) = bs;
-
-(*Accumulates the TVars in a type, suppressing duplicates. *)
-fun add_typ_tvars(Type(_,Ts),vs) = List.foldr add_typ_tvars vs Ts
-  | add_typ_tvars(TFree(_),vs) = vs
-  | add_typ_tvars(TVar(v),vs) = insert (op =) v vs;
-
-(*Accumulates the TFrees in a type, suppressing duplicates. *)
-fun add_typ_tfree_names(Type(_,Ts),fs) = List.foldr add_typ_tfree_names fs Ts
-  | add_typ_tfree_names(TFree(f,_),fs) = insert (op =) f fs
-  | add_typ_tfree_names(TVar(_),fs) = fs;
-
-fun add_typ_tfrees(Type(_,Ts),fs) = List.foldr add_typ_tfrees fs Ts
-  | add_typ_tfrees(TFree(f),fs) = insert (op =) f fs
-  | add_typ_tfrees(TVar(_),fs) = fs;
-
-fun add_typ_varnames(Type(_,Ts),nms) = List.foldr add_typ_varnames nms Ts
-  | add_typ_varnames(TFree(nm,_),nms) = insert (op =) nm nms
-  | add_typ_varnames(TVar((nm,_),_),nms) = insert (op =) nm nms;
-
-(*Accumulates the TVars in a term, suppressing duplicates. *)
-val add_term_tvars = it_term_types add_typ_tvars;
-
-(*Accumulates the TFrees in a term, suppressing duplicates. *)
-val add_term_tfrees = it_term_types add_typ_tfrees;
-val add_term_tfree_names = it_term_types add_typ_tfree_names;
-
-(*Non-list versions*)
-fun typ_tfrees T = add_typ_tfrees(T,[]);
-fun typ_tvars T = add_typ_tvars(T,[]);
-fun term_tfrees t = add_term_tfrees(t,[]);
-fun term_tvars t = add_term_tvars(t,[]);
-
-(*special code to enforce left-to-right collection of TVar-indexnames*)
-
-fun add_typ_ixns(ixns,Type(_,Ts)) = Library.foldl add_typ_ixns (ixns,Ts)
-  | add_typ_ixns(ixns,TVar(ixn,_)) = if member (op =) ixns ixn then ixns
-                                     else ixns@[ixn]
-  | add_typ_ixns(ixns,TFree(_)) = ixns;
-
-fun add_term_tvar_ixns(Const(_,T),ixns) = add_typ_ixns(ixns,T)
-  | add_term_tvar_ixns(Free(_,T),ixns) = add_typ_ixns(ixns,T)
-  | add_term_tvar_ixns(Var(_,T),ixns) = add_typ_ixns(ixns,T)
-  | add_term_tvar_ixns(Bound _,ixns) = ixns
-  | add_term_tvar_ixns(Abs(_,T,t),ixns) =
-      add_term_tvar_ixns(t,add_typ_ixns(ixns,T))
-  | add_term_tvar_ixns(f$t,ixns) =
-      add_term_tvar_ixns(t,add_term_tvar_ixns(f,ixns));
-
-
-(** Frees and Vars **)
-
-(*Accumulates the Vars in the term, suppressing duplicates*)
-fun add_term_vars (t, vars: term list) = case t of
-    Var   _ => OrdList.insert term_ord t vars
-  | Abs (_,_,body) => add_term_vars(body,vars)
-  | f$t =>  add_term_vars (f, add_term_vars(t, vars))
-  | _ => vars;
-
-fun term_vars t = add_term_vars(t,[]);
-
-(*Accumulates the Frees in the term, suppressing duplicates*)
-fun add_term_frees (t, frees: term list) = case t of
-    Free   _ => OrdList.insert term_ord t frees
-  | Abs (_,_,body) => add_term_frees(body,frees)
-  | f$t =>  add_term_frees (f, add_term_frees(t, frees))
-  | _ => frees;
-
-fun term_frees t = add_term_frees(t,[]);
-
-
 (* dest abstraction *)
 
 fun dest_abs (x, T, body) =
@@ -1203,23 +901,6 @@
   end;
 
 
-(* renaming variables *)
-
-val declare_typ_names = fold_atyps (fn TFree (a, _) => Name.declare a | _ => I);
-
-fun declare_term_names tm =
-  fold_aterms
-    (fn Const (a, _) => Name.declare (NameSpace.base a)
-      | Free (a, _) => Name.declare a
-      | _ => I) tm #>
-  fold_types declare_typ_names tm;
-
-fun variant_frees t frees =
-  fst (Name.variants (map fst frees) (declare_term_names t Name.context)) ~~ map snd frees;
-
-fun rename_wrt_term t frees = rev (variant_frees t frees);  (*reversed result!*)
-
-
 (* dummy patterns *)
 
 val dummy_patternN = "dummy_pattern";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/term_ord.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -0,0 +1,209 @@
+(*  Title:      Pure/term_ord.ML
+    Author:     Tobias Nipkow and Makarius, TU Muenchen
+
+Term orderings.
+*)
+
+signature TERM_ORD =
+sig
+  val fast_indexname_ord: indexname * indexname -> order
+  val sort_ord: sort * sort -> order
+  val typ_ord: typ * typ -> order
+  val fast_term_ord: term * term -> order
+  val indexname_ord: indexname * indexname -> order
+  val tvar_ord: (indexname * sort) * (indexname * sort) -> order
+  val var_ord: (indexname * typ) * (indexname * typ) -> order
+  val term_ord: term * term -> order
+  val hd_ord: term * term -> order
+  val termless: term * term -> bool
+  val term_lpo: (term -> int) -> term * term -> order
+end;
+
+structure TermOrd: TERM_ORD =
+struct
+
+(* fast syntactic ordering -- tuned for inequalities *)
+
+fun fast_indexname_ord ((x, i), (y, j)) =
+  (case int_ord (i, j) of EQUAL => fast_string_ord (x, y) | ord => ord);
+
+fun sort_ord SS =
+  if pointer_eq SS then EQUAL
+  else dict_ord fast_string_ord SS;
+
+local
+
+fun cons_nr (TVar _) = 0
+  | cons_nr (TFree _) = 1
+  | cons_nr (Type _) = 2;
+
+in
+
+fun typ_ord TU =
+  if pointer_eq TU then EQUAL
+  else
+    (case TU of
+      (Type (a, Ts), Type (b, Us)) =>
+        (case fast_string_ord (a, b) of EQUAL => dict_ord typ_ord (Ts, Us) | ord => ord)
+    | (TFree (a, S), TFree (b, S')) =>
+        (case fast_string_ord (a, b) of EQUAL => sort_ord (S, S') | ord => ord)
+    | (TVar (xi, S), TVar (yj, S')) =>
+        (case fast_indexname_ord (xi, yj) of EQUAL => sort_ord (S, S') | ord => ord)
+    | (T, U) => int_ord (cons_nr T, cons_nr U));
+
+end;
+
+local
+
+fun cons_nr (Const _) = 0
+  | cons_nr (Free _) = 1
+  | cons_nr (Var _) = 2
+  | cons_nr (Bound _) = 3
+  | cons_nr (Abs _) = 4
+  | cons_nr (_ $ _) = 5;
+
+fun struct_ord (Abs (_, _, t), Abs (_, _, u)) = struct_ord (t, u)
+  | struct_ord (t1 $ t2, u1 $ u2) =
+      (case struct_ord (t1, u1) of EQUAL => struct_ord (t2, u2) | ord => ord)
+  | struct_ord (t, u) = int_ord (cons_nr t, cons_nr u);
+
+fun atoms_ord (Abs (_, _, t), Abs (_, _, u)) = atoms_ord (t, u)
+  | atoms_ord (t1 $ t2, u1 $ u2) =
+      (case atoms_ord (t1, u1) of EQUAL => atoms_ord (t2, u2) | ord => ord)
+  | atoms_ord (Const (a, _), Const (b, _)) = fast_string_ord (a, b)
+  | atoms_ord (Free (x, _), Free (y, _)) = fast_string_ord (x, y)
+  | atoms_ord (Var (xi, _), Var (yj, _)) = fast_indexname_ord (xi, yj)
+  | atoms_ord (Bound i, Bound j) = int_ord (i, j)
+  | atoms_ord _ = sys_error "atoms_ord";
+
+fun types_ord (Abs (_, T, t), Abs (_, U, u)) =
+      (case typ_ord (T, U) of EQUAL => types_ord (t, u) | ord => ord)
+  | types_ord (t1 $ t2, u1 $ u2) =
+      (case types_ord (t1, u1) of EQUAL => types_ord (t2, u2) | ord => ord)
+  | types_ord (Const (_, T), Const (_, U)) = typ_ord (T, U)
+  | types_ord (Free (_, T), Free (_, U)) = typ_ord (T, U)
+  | types_ord (Var (_, T), Var (_, U)) = typ_ord (T, U)
+  | types_ord (Bound _, Bound _) = EQUAL
+  | types_ord _ = sys_error "types_ord";
+
+in
+
+fun fast_term_ord tu =
+  if pointer_eq tu then EQUAL
+  else
+    (case struct_ord tu of
+      EQUAL => (case atoms_ord tu of EQUAL => types_ord tu | ord => ord)
+    | ord => ord);
+
+end;
+
+
+(* term_ord *)
+
+(*a linear well-founded AC-compatible ordering for terms:
+  s < t <=> 1. size(s) < size(t) or
+            2. size(s) = size(t) and s=f(...) and t=g(...) and f<g or
+            3. size(s) = size(t) and s=f(s1..sn) and t=f(t1..tn) and
+               (s1..sn) < (t1..tn) (lexicographically)*)
+
+fun indexname_ord ((x, i), (y, j)) =
+  (case int_ord (i, j) of EQUAL => string_ord (x, y) | ord => ord);
+
+val tvar_ord = prod_ord indexname_ord sort_ord;
+val var_ord = prod_ord indexname_ord typ_ord;
+
+
+local
+
+fun hd_depth (t $ _, n) = hd_depth (t, n + 1)
+  | hd_depth p = p;
+
+fun dest_hd (Const (a, T)) = (((a, 0), T), 0)
+  | dest_hd (Free (a, T)) = (((a, 0), T), 1)
+  | dest_hd (Var v) = (v, 2)
+  | dest_hd (Bound i) = ((("", i), dummyT), 3)
+  | dest_hd (Abs (_, T, _)) = ((("", 0), T), 4);
+
+in
+
+fun term_ord tu =
+  if pointer_eq tu then EQUAL
+  else
+    (case tu of
+      (Abs (_, T, t), Abs(_, U, u)) =>
+        (case term_ord (t, u) of EQUAL => typ_ord (T, U) | ord => ord)
+    | (t, u) =>
+        (case int_ord (size_of_term t, size_of_term u) of
+          EQUAL =>
+            (case prod_ord hd_ord int_ord (hd_depth (t, 0), hd_depth (u, 0)) of
+              EQUAL => args_ord (t, u) | ord => ord)
+        | ord => ord))
+and hd_ord (f, g) =
+  prod_ord (prod_ord indexname_ord typ_ord) int_ord (dest_hd f, dest_hd g)
+and args_ord (f $ t, g $ u) =
+      (case args_ord (f, g) of EQUAL => term_ord (t, u) | ord => ord)
+  | args_ord _ = EQUAL;
+
+fun termless tu = (term_ord tu = LESS);
+
+end;
+
+
+(* Lexicographic path order on terms *)
+
+(*
+  See Baader & Nipkow, Term rewriting, CUP 1998.
+  Without variables.  Const, Var, Bound, Free and Abs are treated all as
+  constants.
+
+  f_ord maps terms to integers and serves two purposes:
+  - Predicate on constant symbols.  Those that are not recognised by f_ord
+    must be mapped to ~1.
+  - Order on the recognised symbols.  These must be mapped to distinct
+    integers >= 0.
+  The argument of f_ord is never an application.
+*)
+
+local
+
+fun unrecognized (Const (a, T)) = ((1, ((a, 0), T)), 0)
+  | unrecognized (Free (a, T)) = ((1, ((a, 0), T)), 0)
+  | unrecognized (Var v) = ((1, v), 1)
+  | unrecognized (Bound i) = ((1, (("", i), dummyT)), 2)
+  | unrecognized (Abs (_, T, _)) = ((1, (("", 0), T)), 3);
+
+fun dest_hd f_ord t =
+  let val ord = f_ord t
+  in if ord = ~1 then unrecognized t else ((0, (("", ord), fastype_of t)), 0) end;
+
+fun term_lpo f_ord (s, t) =
+  let val (f, ss) = strip_comb s and (g, ts) = strip_comb t in
+    if forall (fn si => term_lpo f_ord (si, t) = LESS) ss
+    then case hd_ord f_ord (f, g) of
+        GREATER =>
+          if forall (fn ti => term_lpo f_ord (s, ti) = GREATER) ts
+          then GREATER else LESS
+      | EQUAL =>
+          if forall (fn ti => term_lpo f_ord (s, ti) = GREATER) ts
+          then list_ord (term_lpo f_ord) (ss, ts)
+          else LESS
+      | LESS => LESS
+    else GREATER
+  end
+and hd_ord f_ord (f, g) = case (f, g) of
+    (Abs (_, T, t), Abs (_, U, u)) =>
+      (case term_lpo f_ord (t, u) of EQUAL => typ_ord (T, U) | ord => ord)
+  | (_, _) => prod_ord (prod_ord int_ord
+                  (prod_ord indexname_ord typ_ord)) int_ord
+                (dest_hd f_ord f, dest_hd f_ord g);
+
+in
+val term_lpo = term_lpo
+end;
+
+
+end;
+
+structure Vartab = TableFun(type key = indexname val ord = TermOrd.fast_indexname_ord);
+structure Typtab = TableFun(type key = typ val ord = TermOrd.typ_ord);
+structure Termtab = TableFun(type key = term val ord = TermOrd.fast_term_ord);
--- a/src/Pure/term_subst.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/term_subst.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/term_subst.ML
-    ID:         $Id$
     Author:     Makarius
 
 Efficient term substitution -- avoids copying.
@@ -163,9 +162,9 @@
 
 fun zero_var_indexes_inst ts =
   let
-    val tvars = sort Term.tvar_ord (fold Term.add_tvars ts []);
+    val tvars = sort TermOrd.tvar_ord (fold Term.add_tvars ts []);
     val instT = map (apsnd TVar) (zero_var_inst tvars);
-    val vars = sort Term.var_ord (map (apsnd (instantiateT instT)) (fold Term.add_vars ts []));
+    val vars = sort TermOrd.var_ord (map (apsnd (instantiateT instT)) (fold Term.add_vars ts []));
     val inst = map (apsnd Var) (zero_var_inst vars);
   in (instT, inst) end;
 
--- a/src/Pure/thm.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/thm.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,7 +1,6 @@
 (*  Title:      Pure/thm.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1994  University of Cambridge
+    Author:     Makarius
 
 The very core of Isabelle's Meta Logic: certified types and terms,
 derivations, theorems, framework rules (including lifting and
@@ -147,6 +146,7 @@
   val varifyT': (string * sort) list -> thm -> ((string * sort) * indexname) list * thm
   val freezeT: thm -> thm
   val future: thm future -> cterm -> thm
+  val pending_groups: thm -> Task_Queue.group list -> Task_Queue.group list
   val proof_body_of: thm -> proof_body
   val proof_of: thm -> proof
   val join_proof: thm -> unit
@@ -380,9 +380,9 @@
 
 fun full_prop_of (Thm (_, {tpairs, prop, ...})) = attach_tpairs tpairs prop;
 
-val union_hyps = OrdList.union Term.fast_term_ord;
-val insert_hyps = OrdList.insert Term.fast_term_ord;
-val remove_hyps = OrdList.remove Term.fast_term_ord;
+val union_hyps = OrdList.union TermOrd.fast_term_ord;
+val insert_hyps = OrdList.insert TermOrd.fast_term_ord;
+val remove_hyps = OrdList.remove TermOrd.fast_term_ord;
 
 
 (* merge theories of cterms/thms -- trivial absorption only *)
@@ -1155,7 +1155,7 @@
 (* Replace all TFrees not fixed or in the hyps by new TVars *)
 fun varifyT' fixed (Thm (der, {thy_ref, maxidx, shyps, hyps, tpairs, prop, ...})) =
   let
-    val tfrees = List.foldr add_term_tfrees fixed hyps;
+    val tfrees = fold Term.add_tfrees hyps fixed;
     val prop1 = attach_tpairs tpairs prop;
     val (al, prop2) = Type.varify tfrees prop1;
     val (ts, prop3) = Logic.strip_prems (length tpairs, [], prop2);
@@ -1561,9 +1561,9 @@
 (*Quick test whether rule is resolvable with the subgoal with hyps Hs
   and conclusion B.  If eres_flg then checks 1st premise of rule also*)
 fun could_bires (Hs, B, eres_flg, rule) =
-    let fun could_reshyp (A1::_) = exists (fn H => could_unify (A1, H)) Hs
+    let fun could_reshyp (A1::_) = exists (fn H => Term.could_unify (A1, H)) Hs
           | could_reshyp [] = false;  (*no premise -- illegal*)
-    in  could_unify(concl_of rule, B) andalso
+    in  Term.could_unify(concl_of rule, B) andalso
         (not eres_flg  orelse  could_reshyp (prems_of rule))
     end;
 
@@ -1613,7 +1613,7 @@
     val _ = T <> propT andalso raise CTERM ("future: prop expected", [ct]);
 
     val i = serial ();
-    val future = future_thm |> Future.map (future_result i thy sorts prop o norm_proof);
+    val future = future_thm |> Future.map (future_result i thy sorts prop);
     val promise = (i, future);
   in
     Thm (make_deriv i [promise] [promise] [] [] (Pt.promise_proof thy i prop),
@@ -1627,7 +1627,13 @@
   end;
 
 
-(* fulfilled proof *)
+(* pending task groups *)
+
+fun pending_groups (Thm (Deriv {open_promises, ...}, _)) =
+  fold (insert Task_Queue.eq_group o Future.group_of o #2) open_promises;
+
+
+(* fulfilled proofs *)
 
 fun raw_proof_of (Thm (Deriv {body, ...}, _)) = Proofterm.proof_of body;
 
@@ -1696,7 +1702,7 @@
   val empty = NameSpace.empty_table;
   val copy = I;
   val extend = I;
-  fun merge _ oracles = NameSpace.merge_tables (op =) oracles
+  fun merge _ oracles : T = NameSpace.merge_tables (op =) oracles
     handle Symtab.DUP dup => err_dup_ora dup;
 );
 
--- a/src/Pure/type.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/type.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/type.ML
-    ID:         $Id$
     Author:     Tobias Nipkow, Lawrence C Paulson, and Markus Wenzel
 
 Type signatures and certified types, special treatment of type vars,
@@ -266,9 +265,9 @@
 (* no_tvars *)
 
 fun no_tvars T =
-  (case typ_tvars T of [] => T
+  (case Term.add_tvarsT T [] of [] => T
   | vs => raise TYPE ("Illegal schematic type variable(s): " ^
-      commas_quote (map (Term.string_of_vname o #1) vs), [T], []));
+      commas_quote (map (Term.string_of_vname o #1) (rev vs)), [T], []));
 
 
 (* varify *)
@@ -277,8 +276,9 @@
   let
     val fs = Term.fold_types (Term.fold_atyps
       (fn TFree v => if member (op =) fixed v then I else insert (op =) v | _ => I)) t [];
-    val ixns = add_term_tvar_ixns (t, []);
-    val fmap = fs ~~ map (rpair 0) (Name.variant_list (map #1 ixns) (map fst fs))
+    val used = Name.context
+      |> fold_types (fold_atyps (fn TVar ((a, _), _) => Name.declare a | _ => I)) t;
+    val fmap = fs ~~ map (rpair 0) (#1 (Name.variants (map fst fs) used));
     fun thaw (f as (a, S)) =
       (case AList.lookup (op =) fmap f of
         NONE => TFree f
@@ -307,8 +307,8 @@
 (*this sort of code could replace unvarifyT*)
 fun freeze_thaw_type T =
   let
-    val used = add_typ_tfree_names (T, [])
-    and tvars = map #1 (add_typ_tvars (T, []));
+    val used = OldTerm.add_typ_tfree_names (T, [])
+    and tvars = map #1 (OldTerm.add_typ_tvars (T, []));
     val (alist, _) = List.foldr new_name ([], used) tvars;
   in (map_type_tvar (freeze_one alist) T, map_type_tfree (thaw_one (map swap alist))) end;
 
@@ -316,8 +316,8 @@
 
 fun freeze_thaw t =
   let
-    val used = it_term_types add_typ_tfree_names (t, [])
-    and tvars = map #1 (it_term_types add_typ_tvars (t, []));
+    val used = OldTerm.it_term_types OldTerm.add_typ_tfree_names (t, [])
+    and tvars = map #1 (OldTerm.it_term_types OldTerm.add_typ_tvars (t, []));
     val (alist, _) = List.foldr new_name ([], used) tvars;
   in
     (case alist of
@@ -400,7 +400,7 @@
     fun occ (Type (_, Ts)) = exists occ Ts
       | occ (TFree _) = false
       | occ (TVar (w, S)) =
-          eq_ix (v, w) orelse
+          Term.eq_ix (v, w) orelse
             (case lookup tye (w, S) of
               NONE => false
             | SOME U => occ U);
@@ -437,7 +437,7 @@
     fun unif (ty1, ty2) tye =
       (case (devar tye ty1, devar tye ty2) of
         (T as TVar (v, S1), U as TVar (w, S2)) =>
-          if eq_ix (v, w) then
+          if Term.eq_ix (v, w) then
             if S1 = S2 then tye else tvar_clash v S1 S2
           else if Sorts.sort_le classes (S1, S2) then
             Vartab.update_new (w, (S2, T)) tye
@@ -465,7 +465,7 @@
 fun raw_unify (ty1, ty2) tye =
   (case (devar tye ty1, devar tye ty2) of
     (T as TVar (v, S1), U as TVar (w, S2)) =>
-      if eq_ix (v, w) then
+      if Term.eq_ix (v, w) then
         if S1 = S2 then tye else tvar_clash v S1 S2
       else Vartab.update_new (w, (S2, T)) tye
   | (TVar (v, S), T) =>
--- a/src/Pure/unify.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/unify.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/unify.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   Cambridge University 1992
 
@@ -104,7 +103,7 @@
   | occur (Free _)  = false
   | occur (Var (w, T))  =
       if member (op =) (!seen) w then false
-      else if eq_ix(v,w) then true
+      else if Term.eq_ix (v, w) then true
         (*no need to lookup: v has no assignment*)
       else (seen := w:: !seen;
             case Envir.lookup (env, (w, T)) of
@@ -167,7 +166,7 @@
   | occur (Free _)  = NoOcc
   | occur (Var (w, T))  =
       if member (op =) (!seen) w then NoOcc
-      else if eq_ix(v,w) then Rigid
+      else if Term.eq_ix (v, w) then Rigid
       else (seen := w:: !seen;
             case Envir.lookup (env, (w, T)) of
           NONE    => NoOcc
@@ -176,7 +175,7 @@
   | occur (t as f$_) =  (*switch to nonrigid search?*)
      (case head_of_in (env,f) of
         Var (w,_) => (*w is not assigned*)
-    if eq_ix(v,w) then Rigid
+    if Term.eq_ix (v, w) then Rigid
     else  nonrigid t
       | Abs(_,_,body) => nonrigid t (*not in normal form*)
       | _ => occomb t)
@@ -541,10 +540,10 @@
   let val t = Envir.norm_term env t0  and  u = Envir.norm_term env u0
   in  case  (head_of t, head_of u) of
       (Var(v,T), Var(w,U)) =>  (*Check for identical variables...*)
-  if eq_ix(v,w) then     (*...occur check would falsely return true!*)
+  if Term.eq_ix (v, w) then     (*...occur check would falsely return true!*)
       if T=U then (env, add_tpair (rbinder, (t,u), tpairs))
       else raise TERM ("add_ffpair: Var name confusion", [t,u])
-  else if Term.indexname_ord (v, w) = LESS then (*prefer to update the LARGER variable*)
+  else if TermOrd.indexname_ord (v, w) = LESS then (*prefer to update the LARGER variable*)
        clean_ffpair thy ((rbinder, u, t), (env,tpairs))
         else clean_ffpair thy ((rbinder, t, u), (env,tpairs))
     | _ => raise TERM ("add_ffpair: Vars expected", [t,u])
--- a/src/Pure/variable.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Pure/variable.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/variable.ML
-    ID:         $Id$
     Author:     Makarius
 
 Fixed type/term variables and polymorphic term abbreviations.
@@ -179,12 +178,12 @@
 (* names *)
 
 fun declare_type_names t =
-  map_names (fold_types (fold_atyps (fn TFree (a, _) => Name.declare a | _ => I)) t) #>
+  map_names (fold_types (fold_atyps Term.declare_typ_names) t) #>
   map_maxidx (fold_types Term.maxidx_typ t);
 
 fun declare_names t =
   declare_type_names t #>
-  map_names (fold_aterms (fn Free (x, _) => Name.declare x | _ => I) t) #>
+  map_names (fold_aterms Term.declare_term_frees t) #>
   map_maxidx (Term.maxidx_term t);
 
 
--- a/src/Sequents/modal.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Sequents/modal.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,9 +1,8 @@
 (*  Title:      Sequents/modal.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1992  University of Cambridge
 
-Simple modal reasoner
+Simple modal reasoner.
 *)
 
 
@@ -39,7 +38,7 @@
   Assumes each formula in seqc is surrounded by sequence variables
   -- checks that each concl formula looks like some subgoal formula.*)
 fun could_res (seqp,seqc) =
-      forall (fn Qc => exists (fn Qp => could_unify (Qp,Qc)) 
+      forall (fn Qc => exists (fn Qp => Term.could_unify (Qp,Qc)) 
                               (forms_of_seq seqp))
              (forms_of_seq seqc);
 
--- a/src/Sequents/prover.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Sequents/prover.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,9 +1,8 @@
-(*  Title:      Sequents/prover
-    ID:         $Id$
+(*  Title:      Sequents/prover.ML
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1992  University of Cambridge
 
-Simple classical reasoner for the sequent calculus, based on "theorem packs"
+Simple classical reasoner for the sequent calculus, based on "theorem packs".
 *)
 
 
@@ -65,7 +64,7 @@
   -- checks that each concl formula looks like some subgoal formula.
   It SHOULD check order as well, using recursion rather than forall/exists*)
 fun could_res (seqp,seqc) =
-      forall (fn Qc => exists (fn Qp => could_unify (Qp,Qc)) 
+      forall (fn Qc => exists (fn Qp => Term.could_unify (Qp,Qc)) 
                               (forms_of_seq seqp))
              (forms_of_seq seqc);
 
@@ -78,7 +77,7 @@
 	  could_res (leftp,leftc) andalso could_res (rightp,rightc)
     | (_ $ Abs(_,_,leftp) $ rightp,
        _ $ Abs(_,_,leftc) $ rightc) =>
-	  could_res (leftp,leftc)  andalso  could_unify (rightp,rightc)
+	  could_res (leftp,leftc)  andalso  Term.could_unify (rightp,rightc)
     | _ => false;
 
 
--- a/src/Tools/Compute_Oracle/compute.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Tools/Compute_Oracle/compute.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Tools/Compute_Oracle/compute.ML
-    ID:         $Id$
     Author:     Steven Obua
 *)
 
@@ -168,7 +167,7 @@
   | machine_of_prog (ProgHaskell _) = HASKELL
   | machine_of_prog (ProgSML _) = SML
 
-structure Sorttab = TableFun(type key = sort val ord = Term.sort_ord)
+structure Sorttab = TableFun(type key = sort val ord = TermOrd.sort_ord)
 
 type naming = int -> string
 
--- a/src/Tools/Compute_Oracle/linker.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Tools/Compute_Oracle/linker.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Tools/Compute_Oracle/linker.ML
-    ID:         $Id$
     Author:     Steven Obua
 
 Linker.ML solves the problem that the computing oracle does not
@@ -51,7 +50,7 @@
   | constant_of _ = raise Link "constant_of"
 
 fun bool_ord (x,y) = if x then (if y then EQUAL else GREATER) else (if y then LESS else EQUAL)
-fun constant_ord (Constant (x1,x2,x3), Constant (y1,y2,y3)) = (prod_ord (prod_ord bool_ord fast_string_ord) Term.typ_ord) (((x1,x2),x3), ((y1,y2),y3))
+fun constant_ord (Constant (x1,x2,x3), Constant (y1,y2,y3)) = (prod_ord (prod_ord bool_ord fast_string_ord) TermOrd.typ_ord) (((x1,x2),x3), ((y1,y2),y3))
 fun constant_modty_ord (Constant (x1,x2,_), Constant (y1,y2,_)) = (prod_ord bool_ord fast_string_ord) ((x1,x2), (y1,y2))
 
 
@@ -71,7 +70,7 @@
     handle Type.TYPE_MATCH => NONE
 
 fun subst_ord (A:Type.tyenv, B:Type.tyenv) =
-    (list_ord (prod_ord Term.fast_indexname_ord (prod_ord Term.sort_ord Term.typ_ord))) (Vartab.dest A, Vartab.dest B)
+    (list_ord (prod_ord TermOrd.fast_indexname_ord (prod_ord TermOrd.sort_ord TermOrd.typ_ord))) (Vartab.dest A, Vartab.dest B)
 
 structure Substtab = TableFun(type key = Type.tyenv val ord = subst_ord);
 
@@ -82,7 +81,7 @@
 
 datatype instances = Instances of unit ConsttabModTy.table * Type.tyenv Consttab.table Consttab.table * constant list list * unit Substtab.table
 
-fun is_polymorphic (Constant (_, _, ty)) = not (null (typ_tvars ty))
+fun is_polymorphic (Constant (_, _, ty)) = not (null (Term.add_tvarsT ty []))
 
 fun distinct_constants cs =
     Consttab.keys (fold (fn c => Consttab.update (c, ())) cs Consttab.empty)
@@ -91,7 +90,7 @@
     let
         val cs = distinct_constants (filter is_polymorphic cs)
         val old_cs = cs
-(*      fun collect_tvars ty tab = fold (fn v => fn tab => Typtab.update (TVar v, ()) tab) (typ_tvars ty) tab
+(*      fun collect_tvars ty tab = fold (fn v => fn tab => Typtab.update (TVar v, ()) tab) (OldTerm.typ_tvars ty) tab
         val tvars_count = length (Typtab.keys (fold (fn c => fn tab => collect_tvars (typ_of_constant c) tab) cs Typtab.empty))
         fun tvars_of ty = collect_tvars ty Typtab.empty
         val cs = map (fn c => (c, tvars_of (typ_of_constant c))) cs
@@ -263,10 +262,10 @@
 let
     val (left, right) = collect_consts_of_thm th
     val polycs = filter Linker.is_polymorphic left
-    val tytab = fold (fn p => fn tab => fold (fn n => fn tab => Typtab.update (TVar n, ()) tab) (typ_tvars (Linker.typ_of_constant p)) tab) polycs Typtab.empty
+    val tytab = fold (fn p => fn tab => fold (fn n => fn tab => Typtab.update (TVar n, ()) tab) (OldTerm.typ_tvars (Linker.typ_of_constant p)) tab) polycs Typtab.empty
     fun check_const (c::cs) cs' =
         let
-            val tvars = typ_tvars (Linker.typ_of_constant c)
+            val tvars = OldTerm.typ_tvars (Linker.typ_of_constant c)
             val wrong = fold (fn n => fn wrong => wrong orelse is_none (Typtab.lookup tytab (TVar n))) tvars false
         in
             if wrong then raise PCompute "right hand side of theorem contains type variables which do not occur on the left hand side"
@@ -379,7 +378,7 @@
         ComputeThm (hyps, shyps, prop)
     end
 
-val cthm_ord' = prod_ord (prod_ord (list_ord Term.term_ord) (list_ord Term.sort_ord)) Term.term_ord
+val cthm_ord' = prod_ord (prod_ord (list_ord TermOrd.term_ord) (list_ord TermOrd.sort_ord)) TermOrd.term_ord
 
 fun cthm_ord (ComputeThm (h1, sh1, p1), ComputeThm (h2, sh2, p2)) = cthm_ord' (((h1,sh1), p1), ((h2, sh2), p2))
 
--- a/src/Tools/IsaPlanner/isand.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Tools/IsaPlanner/isand.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Tools/IsaPlanner/isand.ML
-    ID:		$Id$
     Author:     Lucas Dixon, University of Edinburgh
 
 Natural Deduction tools.
@@ -187,8 +186,8 @@
 (* change type-vars to fresh type frees *)
 fun fix_tvars_to_tfrees_in_terms names ts = 
     let 
-      val tfree_names = map fst (List.foldr Term.add_term_tfrees [] ts);
-      val tvars = List.foldr Term.add_term_tvars [] ts;
+      val tfree_names = map fst (List.foldr OldTerm.add_term_tfrees [] ts);
+      val tvars = List.foldr OldTerm.add_term_tvars [] ts;
       val (names',renamings) = 
           List.foldr (fn (tv as ((n,i),s),(Ns,Rs)) => 
                          let val n2 = Name.variant Ns n in 
@@ -213,15 +212,15 @@
 fun unfix_tfrees ns th = 
     let 
       val varfiytfrees = map (Term.dest_TFree o Thm.typ_of) ns
-      val skiptfrees = subtract (op =) varfiytfrees (Term.add_term_tfrees (Thm.prop_of th,[]));
+      val skiptfrees = subtract (op =) varfiytfrees (OldTerm.add_term_tfrees (Thm.prop_of th,[]));
     in #2 (Thm.varifyT' skiptfrees th) end;
 
 (* change schematic/meta vars to fresh free vars, avoiding name clashes 
    with "names" *)
 fun fix_vars_to_frees_in_terms names ts = 
     let 
-      val vars = map Term.dest_Var (List.foldr Term.add_term_vars [] ts);
-      val Ns = List.foldr Term.add_term_names names ts;
+      val vars = map Term.dest_Var (List.foldr OldTerm.add_term_vars [] ts);
+      val Ns = List.foldr OldTerm.add_term_names names ts;
       val (_,renamings) = 
           Library.foldl (fn ((Ns,Rs),v as ((n,i),ty)) => 
                     let val n2 = Name.variant Ns n in
@@ -246,7 +245,7 @@
       val ctypify = Thm.ctyp_of sgn
       val tpairs = Thm.terms_of_tpairs (Thm.tpairs_of th);
       val prop = (Thm.prop_of th);
-      val tvars = List.foldr Term.add_term_tvars [] (prop :: tpairs);
+      val tvars = List.foldr OldTerm.add_term_tvars [] (prop :: tpairs);
       val ctyfixes = 
           map_filter 
             (fn (v as ((s,i),ty)) => 
@@ -259,7 +258,7 @@
       val ctermify = Thm.cterm_of sgn
       val tpairs = Thm.terms_of_tpairs (Thm.tpairs_of th);
       val prop = (Thm.prop_of th);
-      val vars = map Term.dest_Var (List.foldr Term.add_term_vars 
+      val vars = map Term.dest_Var (List.foldr OldTerm.add_term_vars 
                                                [] (prop :: tpairs));
       val cfixes = 
           map_filter 
@@ -360,7 +359,7 @@
       val sgn = Thm.theory_of_thm th;
       val ctermify = Thm.cterm_of sgn;
 
-      val allfrees = map Term.dest_Free (Term.term_frees (Thm.prop_of th))
+      val allfrees = map Term.dest_Free (OldTerm.term_frees (Thm.prop_of th))
       val cfrees = map (ctermify o Free o lookupfree allfrees) vs
 
       val sgs = prems_of th;
@@ -420,8 +419,8 @@
     let
       val t = Term.strip_all_body alledt;
       val alls = rev (Term.strip_all_vars alledt);
-      val varnames = map (fst o fst o Term.dest_Var) (Term.term_vars t)
-      val names = Term.add_term_names (t,varnames);
+      val varnames = map (fst o fst o Term.dest_Var) (OldTerm.term_vars t)
+      val names = OldTerm.add_term_names (t,varnames);
       val fvs = map Free 
                     (Name.variant_list names (map fst alls)
                        ~~ (map snd alls));
@@ -429,8 +428,8 @@
 
 fun fix_alls_term i t = 
     let 
-      val varnames = map (fst o fst o Term.dest_Var) (Term.term_vars t)
-      val names = Term.add_term_names (t,varnames);
+      val varnames = map (fst o fst o Term.dest_Var) (OldTerm.term_vars t)
+      val names = OldTerm.add_term_names (t,varnames);
       val gt = Logic.get_goal t i;
       val body = Term.strip_all_body gt;
       val alls = rev (Term.strip_all_vars gt);
--- a/src/Tools/IsaPlanner/rw_inst.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Tools/IsaPlanner/rw_inst.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Tools/IsaPlanner/rw_inst.ML
-    ID:         $Id$
     Author:     Lucas Dixon, University of Edinburgh
 
 Rewriting using a conditional meta-equality theorem which supports
@@ -72,7 +71,7 @@
       val (tpairl,tpairr) = Library.split_list (#tpairs rep)
       val prop = #prop rep
     in
-      List.foldr Term.add_term_names [] (prop :: (tpairl @ (tpairr @ hyps)))
+      List.foldr OldTerm.add_term_names [] (prop :: (tpairl @ (tpairr @ hyps)))
     end;
 
 (* Given a list of variables that were bound, and a that has been
@@ -137,15 +136,15 @@
 fun mk_renamings tgt rule_inst = 
     let
       val rule_conds = Thm.prems_of rule_inst
-      val names = foldr Term.add_term_names [] (tgt :: rule_conds);
+      val names = foldr OldTerm.add_term_names [] (tgt :: rule_conds);
       val (conds_tyvs,cond_vs) = 
           Library.foldl (fn ((tyvs, vs), t) => 
                     (Library.union
-                       (Term.term_tvars t, tyvs),
+                       (OldTerm.term_tvars t, tyvs),
                      Library.union 
-                       (map Term.dest_Var (Term.term_vars t), vs))) 
+                       (map Term.dest_Var (OldTerm.term_vars t), vs))) 
                 (([],[]), rule_conds);
-      val termvars = map Term.dest_Var (Term.term_vars tgt); 
+      val termvars = map Term.dest_Var (OldTerm.term_vars tgt); 
       val vars_to_fix = Library.union (termvars, cond_vs);
       val (renamings, names2) = 
           foldr (fn (((n,i),ty), (vs, names')) => 
@@ -168,8 +167,8 @@
       val ignore_ixs = map fst ignore_insts;
       val (tvars, tfrees) = 
             foldr (fn (t, (varixs, tfrees)) => 
-                      (Term.add_term_tvars (t,varixs),
-                       Term.add_term_tfrees (t,tfrees)))
+                      (OldTerm.add_term_tvars (t,varixs),
+                       OldTerm.add_term_tfrees (t,tfrees)))
                   ([],[]) ts;
         val unfixed_tvars = 
             List.filter (fn (ix,s) => not (member (op =) ignore_ixs ix)) tvars;
--- a/src/Tools/code/code_ml.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Tools/code/code_ml.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -954,7 +954,7 @@
 fun eval eval'' term_of reff thy ct args =
   let
     val ctxt = ProofContext.init thy;
-    val _ = if null (term_frees (term_of ct)) then () else error ("Term "
+    val _ = if null (Term.add_frees (term_of ct) []) then () else error ("Term "
       ^ quote (Syntax.string_of_term_global thy (term_of ct))
       ^ " to be evaluated contains free variables");
     fun eval' naming program ((vs, ty), t) deps =
--- a/src/Tools/code/code_thingol.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Tools/code/code_thingol.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Tools/code/code_thingol.ML
-    ID:         $Id$
     Author:     Florian Haftmann, TU Muenchen
 
 Intermediate language ("Thin-gol") representing executable code.
@@ -620,7 +619,7 @@
     fun stmt_fun ((vs, raw_ty), raw_thms) =
       let
         val ty = Logic.unvarifyT raw_ty;
-        val thms = if (null o Term.typ_tfrees) ty orelse (null o fst o strip_type) ty
+        val thms = if null (Term.add_tfreesT ty []) orelse (null o fst o strip_type) ty
           then raw_thms
           else (map o apfst) (Code_Unit.expand_eta thy 1) raw_thms;
       in
@@ -708,8 +707,7 @@
         let
           val k = length ts;
           val tys = (curry Library.take (i - k) o curry Library.drop k o fst o strip_type) ty;
-          val ctxt = (fold o fold_aterms)
-            (fn Free (v, _) => Name.declare v | _ => I) ts Name.context;
+          val ctxt = (fold o fold_aterms) Term.declare_term_frees ts Name.context;
           val vs = Name.names ctxt "a" tys;
         in
           fold_map (translate_typ thy algbr funcgr) tys
--- a/src/Tools/induct.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Tools/induct.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -440,7 +440,7 @@
     val thy = ProofContext.theory_of ctxt;
     val maxidx = Thm.maxidx_of st;
     val goal = Thm.term_of (Thm.cprem_of st i);  (*exception Subscript*)
-    val params = rev (rename_wrt_term goal (Logic.strip_params goal));
+    val params = rev (Term.rename_wrt_term goal (Logic.strip_params goal));
   in
     if not (null params) then
       (warning ("Cannot determine rule instantiation due to pending parameter(s): " ^
--- a/src/Tools/nbe.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Tools/nbe.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Tools/nbe.ML
-    ID:         $Id$
     Authors:    Klaus Aehlig, LMU Muenchen; Tobias Nipkow, Florian Haftmann, TU Muenchen
 
 Normalization by evaluation, based on generic code generator.
@@ -448,7 +447,7 @@
       singleton (TypeInfer.infer_types (Syntax.pp_global thy) (Sign.tsig_of thy) I
         (try (Type.strip_sorts o Sign.the_const_type thy)) (K NONE) Name.context 0)
       (TypeInfer.constrain ty t);
-    fun check_tvars t = if null (Term.term_tvars t) then t else
+    fun check_tvars t = if null (Term.add_tvars t []) then t else
       error ("Illegal schematic type variables in normalized term: "
         ^ setmp show_types true (Syntax.string_of_term_global thy) t);
     val string_of_term = setmp show_types true (Syntax.string_of_term_global thy);
--- a/src/Tools/quickcheck.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Tools/quickcheck.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -70,11 +70,11 @@
 
 fun prep_test_term t =
   let
-    val _ = (null (term_tvars t) andalso null (term_tfrees t)) orelse
+    val _ = (null (Term.add_tvars t []) andalso null (Term.add_tfrees t [])) orelse
       error "Term to be tested contains type variables";
-    val _ = null (term_vars t) orelse
+    val _ = null (Term.add_vars t []) orelse
       error "Term to be tested contains schematic variables";
-    val frees = map dest_Free (term_frees t);
+    val frees = map dest_Free (OldTerm.term_frees t);
   in (map fst frees, list_abs_free (frees, t)) end
 
 fun test_term ctxt quiet generator_name size i t =
--- a/src/Tools/value.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/Tools/value.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -20,7 +20,7 @@
   val empty = [];
   val copy = I;
   val extend = I;
-  fun merge pp = AList.merge (op =) (K true);
+  fun merge _ data = AList.merge (op =) (K true) data;
 )
 
 val add_evaluator = Evaluator.map o AList.update (op =);
--- a/src/ZF/Tools/inductive_package.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/ZF/Tools/inductive_package.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,7 +1,5 @@
 (*  Title:      ZF/Tools/inductive_package.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1994  University of Cambridge
 
 Fixedpoint definition module -- for Inductive/Coinductive Definitions
 
@@ -101,7 +99,7 @@
                Syntax.string_of_term ctxt t);
 
   (*** Construct the fixedpoint definition ***)
-  val mk_variant = Name.variant (foldr add_term_names [] intr_tms);
+  val mk_variant = Name.variant (foldr OldTerm.add_term_names [] intr_tms);
 
   val z' = mk_variant"z" and X' = mk_variant"X" and w' = mk_variant"w";
 
@@ -113,7 +111,7 @@
   fun fp_part intr = (*quantify over rule's free vars except parameters*)
     let val prems = map dest_tprop (Logic.strip_imp_prems intr)
         val dummy = List.app (fn rec_hd => List.app (chk_prem rec_hd) prems) rec_hds
-        val exfrees = term_frees intr \\ rec_params
+        val exfrees = OldTerm.term_frees intr \\ rec_params
         val zeq = FOLogic.mk_eq (Free(z',iT), #1 (rule_concl intr))
     in foldr FOLogic.mk_exists
              (BalancedTree.make FOLogic.mk_conj (zeq::prems)) exfrees
@@ -304,7 +302,7 @@
 
      (*Make a premise of the induction rule.*)
      fun induct_prem ind_alist intr =
-       let val quantfrees = map dest_Free (term_frees intr \\ rec_params)
+       let val quantfrees = map dest_Free (OldTerm.term_frees intr \\ rec_params)
            val iprems = foldr (add_induct_prem ind_alist) []
                               (Logic.strip_imp_prems intr)
            val (t,X) = Ind_Syntax.rule_concl intr
--- a/src/ZF/Tools/primrec_package.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/ZF/Tools/primrec_package.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,10 +1,9 @@
 (*  Title:      ZF/Tools/primrec_package.ML
-    ID:         $Id$
-    Author:     Stefan Berghofer and Norbert Voelker
-    Copyright   1998  TU Muenchen
-    ZF version by Lawrence C Paulson (Cambridge)
+    Author:     Norbert Voelker, FernUni Hagen
+    Author:     Stefan Berghofer, TU Muenchen
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 
-Package for defining functions on datatypes by primitive recursion
+Package for defining functions on datatypes by primitive recursion.
 *)
 
 signature PRIMREC_PACKAGE =
@@ -33,7 +32,7 @@
 fun process_eqn thy (eq, rec_fn_opt) =
   let
     val (lhs, rhs) =
-        if null (term_vars eq) then
+        if null (Term.add_vars eq []) then
             dest_eqn eq handle TERM _ => raise RecError "not a proper equation"
         else raise RecError "illegal schematic variable(s)";
 
@@ -66,7 +65,7 @@
   in
     if has_duplicates (op =) lfrees then
       raise RecError "repeated variable name in pattern"
-    else if not ((map dest_Free (term_frees rhs)) subset lfrees) then
+    else if not ((map dest_Free (OldTerm.term_frees rhs)) subset lfrees) then
       raise RecError "extra variables on rhs"
     else if length middle > 1 then
       raise RecError "more than one non-variable in pattern"
--- a/src/ZF/arith_data.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/ZF/arith_data.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,7 +1,5 @@
 (*  Title:      ZF/arith_data.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   2000  University of Cambridge
 
 Arithmetic simplification: cancellation of common terms
 *)
@@ -106,7 +104,7 @@
 
 (*Dummy version: the "coefficient" is always 1.
   In the result, the factors are sorted terms*)
-fun dest_coeff t = (1, mk_prod (sort Term.term_ord (dest_prod t)));
+fun dest_coeff t = (1, mk_prod (sort TermOrd.term_ord (dest_prod t)));
 
 (*Find first coefficient-term THAT MATCHES u*)
 fun find_first_coeff past u [] = raise TERM("find_first_coeff", [])
--- a/src/ZF/int_arith.ML	Mon Jan 19 20:37:08 2009 +0100
+++ b/src/ZF/int_arith.ML	Mon Jan 19 21:20:18 2009 +0100
@@ -1,7 +1,5 @@
 (*  Title:      ZF/int_arith.ML
-    ID:         $Id$
     Author:     Larry Paulson
-    Copyright   2000  University of Cambridge
 
 Simprocs for linear arithmetic.
 *)
@@ -72,7 +70,7 @@
 (*Express t as a product of (possibly) a numeral with other sorted terms*)
 fun dest_coeff sign (Const (@{const_name "zminus"}, _) $ t) = dest_coeff (~sign) t
   | dest_coeff sign t =
-    let val ts = sort Term.term_ord (dest_prod t)
+    let val ts = sort TermOrd.term_ord (dest_prod t)
         val (n, ts') = find_first_numeral [] ts
                           handle TERM _ => (1, ts)
     in (sign*n, mk_prod ts') end;