Automated merge with ssh://paulson@atbroy100.informatik.tu-muenchen.de//home/isabelle-repository/repos/isabelle
authorpaulson
Thu, 29 Jan 2009 12:24:00 +0000
changeset 29685 aba49b4fe959
parent 29675 fa6f988f1c50 (diff)
parent 29684 40bf9f4e7a4e (current diff)
child 29686 4cd2874eb5ff
Automated merge with ssh://paulson@atbroy100.informatik.tu-muenchen.de//home/isabelle-repository/repos/isabelle
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Admin/churn	Thu Jan 29 12:24:00 2009 +0000
@@ -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/settings/at-mac-poly-5.1-para	Thu Jan 29 12:05:19 2009 +0000
+++ b/Admin/isatest/settings/at-mac-poly-5.1-para	Thu Jan 29 12:24:00 2009 +0000
@@ -25,4 +25,4 @@
 
 ISABELLE_USEDIR_OPTIONS="-i false -d false -M 4"
 
-HOL_USEDIR_OPTIONS="-p 2"
+HOL_USEDIR_OPTIONS="-p 2 -Q false"
--- a/Admin/isatest/settings/at-poly-5.1-para-e	Thu Jan 29 12:05:19 2009 +0000
+++ b/Admin/isatest/settings/at-poly-5.1-para-e	Thu Jan 29 12:24:00 2009 +0000
@@ -24,4 +24,4 @@
 
 ISABELLE_USEDIR_OPTIONS="-i true -d pdf -v true -M 20"
 
-HOL_USEDIR_OPTIONS="-p 2"
+HOL_USEDIR_OPTIONS="-p 2 -Q false"
--- a/Admin/isatest/settings/at64-poly-5.1-para	Thu Jan 29 12:05:19 2009 +0000
+++ b/Admin/isatest/settings/at64-poly-5.1-para	Thu Jan 29 12:24:00 2009 +0000
@@ -24,4 +24,4 @@
 
 ISABELLE_USEDIR_OPTIONS="-i true -d pdf -v true -M 2"
 
-HOL_USEDIR_OPTIONS="-p 2"
+HOL_USEDIR_OPTIONS="-p 2 -Q false"
--- a/Admin/makedist	Thu Jan 29 12:05:19 2009 +0000
+++ b/Admin/makedist	Thu Jan 29 12:24:00 2009 +0000
@@ -1,12 +1,10 @@
 #!/usr/bin/env bash
 #
-# $Id$
-#
 # makedist -- make Isabelle source distribution
 
 ## global settings
 
-REPOS="http://isabelle.in.tum.de/repos/isabelle"
+REPOS="https://isabelle.in.tum.de/repos/isabelle"
 
 DISTPREFIX=${DISTPREFIX:-~/tmp/isadist}
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Admin/user-aliases	Thu Jan 29 12:24:00 2009 +0000
@@ -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/NEWS	Thu Jan 29 12:05:19 2009 +0000
+++ b/NEWS	Thu Jan 29 12:24:00 2009 +0000
@@ -66,13 +66,19 @@
 
 *** Pure ***
 
-* Type Binding.T gradually replaces formerly used type bstring for names
+* Class declaration: sc. "base sort" must not be given in import list
+any longer but is inferred from the specification.  Particularly in HOL,
+write
+
+    class foo = ...     instead of      class foo = type + ...
+
+* Type binding gradually replaces formerly used type bstring for names
 to be bound.  Name space interface for declarations has been simplified:
 
   NameSpace.declare: NameSpace.naming
-    -> Binding.T -> NameSpace.T -> string * NameSpace.T
+    -> binding -> NameSpace.T -> string * NameSpace.T
   NameSpace.bind: NameSpace.naming
-    -> Binding.T * 'a -> 'a NameSpace.table -> string * 'a NameSpace.table
+    -> binding * 'a -> 'a NameSpace.table -> string * 'a NameSpace.table
          (*exception Symtab.DUP*)
 
 See further modules src/Pure/General/binding.ML and
@@ -187,6 +193,10 @@
 
 *** HOL ***
 
+* Theory "Reflection" now resides in HOL/Library.
+
+* Entry point to Word library now simply named "Word".  INCOMPATIBILITY.
+
 * Made source layout more coherent with logical distribution
 structure:
 
@@ -304,6 +314,11 @@
 Occasionally this is more convenient than the old fold combinator which is
 now defined in terms of the new one and renamed to fold_image.
 
+* HOL/Ring_and_Field and HOL/OrderedGroup: The lemmas "group_simps" and
+"ring_simps" have been replaced by "algebra_simps" (which can be extended with
+further lemmas!). At the moment both still exist but the former will disappear
+at some point.
+
 * HOL/Ring_and_Field and HOL/Divides: Definition of "op dvd" has been
 moved to separate class dvd in Ring_and_Field; a couple of lemmas on
 dvd has been generalized to class comm_semiring_1.  Likewise a bunch
--- a/README_REPOSITORY	Thu Jan 29 12:05:19 2009 +0000
+++ b/README_REPOSITORY	Thu Jan 29 12:24:00 2009 +0000
@@ -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/contrib/SystemOnTPTP/remote	Thu Jan 29 12:05:19 2009 +0000
+++ b/contrib/SystemOnTPTP/remote	Thu Jan 29 12:24:00 2009 +0000
@@ -3,93 +3,137 @@
 # Wrapper for custom remote provers on SystemOnTPTP
 # Author: Fabian Immler, TU Muenchen
 #
-# Similar to the vampire wrapper, but compatible provers can be passed in the
-# command line, with %s for the problemfile e.g. 
-#
-# ./remote Vampire---9.0 jumpirefix --output_syntax tptp --mode casc -t 3600 %s
-# ./remote Vampire---10.0 drakosha.pl 60 %s
-# ./remote SPASS---3.01 SPASS -Auto -SOS=1 -PGiven=0 -PProblem=0 -Splits=0 -FullRed=0 -DocProof %s
-# ./remote Metis---2.1 metis --show proof --show saturation %s
-# ./remote SNARK---20080805r005 run-snark %s
 
 use warnings;
 use strict;
-
 use Getopt::Std;
 use HTTP::Request::Common;
-use LWP::UserAgent;
+use LWP;
 
-# address of proof-server
 my $SystemOnTPTPFormReplyURL = "http://www.cs.miami.edu/~tptp/cgi-bin/SystemOnTPTPFormReply";
 
-if(scalar(@ARGV) < 2) {
-    print "prover and command missing";
-    exit -1;
-}
-my $prover = shift(@ARGV);
-my $command = shift(@ARGV);
-
-# pass arguments
-my $options = "";
-while(scalar(@ARGV)>1){
-	$options.=" ".shift(@ARGV);
-}
-# last argument is problemfile to be uploaded
-my $problem = [shift(@ARGV)];
-
-# fill in form
+# default parameters
 my %URLParameters = (
     "NoHTML" => 1,
     "QuietFlag" => "-q01",
     "X2TPTP" => "-S",
     "SubmitButton" => "RunSelectedSystems",
     "ProblemSource" => "UPLOAD",
-    "UPLOADProblem" => $problem,
-    "System___$prover" => "$prover",
-    "Format___$prover" => "tptp",
-    "Command___$prover" => "$command $options %s"
-);
+    );
+
+# check connection
+my $TestAgent = LWP::UserAgent->new;
+$TestAgent->timeout(5);
+my $TestRequest = GET($SystemOnTPTPFormReplyURL);
+my $TestResponse = $TestAgent->request($TestRequest);
+if(! $TestResponse->is_success) {
+  print "HTTP-Error: " . $TestResponse->message . "\n";
+  exit(-1);
+}
+
+#----Get format and transform options if specified
+my %Options;
+getopts("hws:t:c:",\%Options);
+
+#----Usage
+if (exists($Options{'h'})) {
+    print("Usage: remote <options> [<File name>]\n");
+    print("    <options> are ...\n");
+    print("    -h            - print this help\n");
+    print("    -w            - list available ATP systems\n");
+    print("    -s<system>    - specified system to use\n");
+    print("    -t<timelimit> - CPU time limit for system\n");
+    print("    -c<command>   - custom command for system\n");
+    print("    <File name>   - TPTP problem file\n");
+    exit(0);
+}
+
+#----What systems flag
+if (exists($Options{'w'})) {
+    $URLParameters{"SubmitButton"} = "ListSystems";
+    delete($URLParameters{"ProblemSource"});
+}
+#----Selected system
+my $System;
+if (exists($Options{'s'})) {
+    $System = $Options{'s'};
+} else {
+    # use Vampire as default
+    $System = "Vampire---9.0";
+}
+$URLParameters{"System___$System"} = $System;
+
+#----Time limit
+if (exists($Options{'t'})) {
+    $URLParameters{"TimeLimit___$System"} = $Options{'t'};
+}
+#----Custom command
+if (exists($Options{'c'})) {
+    $URLParameters{"Command___$System"} = $Options{'c'};
+}
+
+#----Get single file name
+if (exists($URLParameters{"ProblemSource"})) {
+    if (scalar(@ARGV) >= 1) {
+        $URLParameters{"UPLOADProblem"} = [shift(@ARGV)];
+    } else {
+die("Missing problem file");
+    }
+}
 
 # Query Server
 my $Agent = LWP::UserAgent->new;
+if (exists($Options{'t'})) {
+  # give server more time to respond
+  $Agent->timeout($Options{'t'} * 2 + 10);
+}
 my $Request = POST($SystemOnTPTPFormReplyURL,
 	Content_Type => 'form-data',Content => \%URLParameters);
 my $Response = $Agent->request($Request);
-  
-#catch errors, let isabelle/watcher know
-if($Response->is_success && $Response->content !~ /NO SOLUTION OUTPUT BY SYSTEM/
-&& $Response->content =~ m/%\s*Result\s*:\s*Unsatisfiable.*?\n%\s*Output\s*:\s*.*?Refutation.*?\n/){
-    # convert to isabelle-friendly format
-    my @lines = split( /%\s*Result\s*:\s*Unsatisfiable.*?\n%\s*Output\s*:\s*.*?Refutation.*?\n/, $Response->content);
-    @lines = split( /\n/, $lines[1]);    my $extract = "";
-    my $inproof = 0 > 1;
-    my $ende = 0 > 1;
+
+#catch errors / failure
+if(! $Response->is_success){
+  	print "HTTP-Error: " . $Response->message . "\n";
+    exit(-1);
+} elsif (exists($Options{'w'})) {
+  print $Response->content;
+  exit (0);
+} elsif ($Response->content =~ /NO SOLUTION OUTPUT BY SYSTEM/){
+  if ($Response->content =~ /%\s*Result\s*:(.*)\n%\s*Output\s*:(.*)\n%/) {
+    print "No Solution Output\nResult: $1\nOutput: $2\n";
+  } else {
+    print "No Solution Output\n";
+  }
+  exit(-1);
+} elsif ($Response->content =~ /ERROR: Could not form TPTP format derivation/) {
+  print "Could not form TPTP format derivation\n";
+  exit(-1);
+} elsif ($Response->content =~ /WARNING: (\S*) does not exist/) {
+  print "Specified System $1 does not exist\n";
+  exit(-1);
+} elsif ($Response->content =~ /^\s*$/) {
+  print "Empty response (specified bad system? Inappropriate problem file format?)\n";
+  exit(-1);
+} elsif ($Response->content !~ /%\s*Result\s*:(.*)\n%\s*Output\s*:(.*)\n%/) {
+  print "Bad response: \n".$Response->content;
+  exit(-1);
+} else {
+    my @lines = split( /\n/, $Response->content);
+    my $extract = "";
     foreach my $line (@lines){
-        if(! $ende){
-            #ignore comments
-            if(! $inproof){
-                if ($line !~ /^%/ && !($line eq "")) {
-                    $extract .= "$line";
-                    $inproof = 1;
-                }
-            } else {
-                if ($line !~ /^%/) {
-                    $extract .= "$line";
-                } else {
-                    $ende = 1;
-                }
-            }
+        #ignore comments
+        if ($line !~ /^%/ && !($line eq "")) {
+            $extract .= "$line";
         }
     }
-    # insert newlines after '.'
+    # insert newlines after ').'
     $extract =~ s/\s//g;
     $extract =~ s/\)\.cnf/\)\.\ncnf/g;
+    
+    # orientation for res_reconstruct.ML
     print "# SZS output start CNFRefutation.\n";
     print "$extract\n";
     print "# SZS output end CNFRefutation.\n";
-} else {
-	print "HTTP-Request: " . $Response->message;
-	print "\nCANNOT PROVE: \n";
-  print $Response->content;
+    exit(0);
 }
 
--- a/contrib/SystemOnTPTP/vampire	Thu Jan 29 12:05:19 2009 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,60 +0,0 @@
-#!/usr/bin/env perl
-#
-# Vampire Wrapper for SystemOnTPTP
-# Author: Fabian Immler, TU Muenchen
-#
-# - querys a Vampire theorem prover on SystemOnTPTP
-#   (http://www.cs.miami.edu/~tptp/cgi-bin/SystemOnTPTP)
-# - behaves like a local Vampire
-# => can be used for Isabelle when Vampire is not available (e.g. on a Mac)
-#
-
-use warnings;
-use strict;
-
-use Getopt::Std;
-use HTTP::Request::Common;
-use LWP::UserAgent;
-
-# address of proof-server
-my $SystemOnTPTPFormReplyURL = "http://www.cs.miami.edu/~tptp/cgi-bin/SystemOnTPTPFormReply";
-
-#name of prover and its executable on the server, e.g.
-# Vampire---9.0
-# jumpirefix
-my $prover = "Vampire---9.0";
-my $command = "jumpirefix";
-
-# pass arguments
-my $options = "";
-while(scalar(@ARGV)>1){
-	$options.=" ".shift(@ARGV);
-}
-# last argument is problemfile to be uploaded
-my $problem = [shift(@ARGV)];
-
-# fill in form
-my %URLParameters = (
-    "NoHTML" => 1,
-    "QuietFlag" => "-q01",
-    "SubmitButton" => "RunSelectedSystems",
-    "ProblemSource" => "UPLOAD",
-    "UPLOADProblem" => $problem,
-    "System___$prover" => "$prover",
-    "Format___$prover" => "tptp",
-    "Command___$prover" => "$command $options %s"
-);
-
-# Query Server
-my $Agent = LWP::UserAgent->new;
-my $Request = POST($SystemOnTPTPFormReplyURL,
-	Content_Type => 'form-data',Content => \%URLParameters);
-my $Response = $Agent->request($Request);
-    
-#catch errors, let isabelle/watcher know
-if($Response->is_success){
-	print $Response->content;
-} else {
-	print $Response->message;
-	print "\nCANNOT PROVE\n";
-}
--- a/doc-src/IsarAdvanced/Classes/Thy/Classes.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/doc-src/IsarAdvanced/Classes/Thy/Classes.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -368,14 +368,14 @@
 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':    (* FIXME proper prefix? *)
+interpretation %quote idem_class:
   idem "f \<Colon> (\<alpha>\<Colon>idem) \<Rightarrow> \<alpha>"
 proof qed (rule idem)
 
@@ -459,7 +459,7 @@
   of monoids for lists:
 *}
 
-class_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"
 
-class_interpretation %quote list_monoid: monoid [append "[]"] where
+interpretation %quote list_monoid!: monoid append "[]" where
   "monoid.pow_nat append [] = replicate"
 proof -
-  class_interpret monoid [append "[]"] ..
+  interpret monoid append "[]" ..
   show "monoid.pow_nat append [] = replicate"
   proof
     fix n
--- a/doc-src/IsarAdvanced/Classes/Thy/document/Classes.tex	Thu Jan 29 12:05:19 2009 +0000
+++ b/doc-src/IsarAdvanced/Classes/Thy/document/Classes.tex	Thu Jan 29 12:24:00 2009 +0000
@@ -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}%
@@ -681,7 +711,7 @@
 %
 \isatagquote
 \isacommand{interpretation}\isamarkupfalse%
-\ idem{\isacharunderscore}class{\isacharprime}{\isacharcolon}\ \ \ \ \isanewline
+\ idem{\isacharunderscore}class{\isacharcolon}\isanewline
 \ \ idem\ {\isachardoublequoteopen}f\ {\isasymColon}\ {\isacharparenleft}{\isasymalpha}{\isasymColon}idem{\isacharparenright}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\isanewline
 \isacommand{proof}\isamarkupfalse%
 \ \isacommand{qed}\isamarkupfalse%
@@ -843,8 +873,8 @@
 \endisadelimquote
 %
 \isatagquote
-\isacommand{class{\isacharunderscore}interpretation}\isamarkupfalse%
-\ list{\isacharunderscore}monoid{\isacharcolon}\ monoid\ {\isacharbrackleft}append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}{\isacharbrackright}\isanewline
+\isacommand{interpretation}\isamarkupfalse%
+\ list{\isacharunderscore}monoid{\isacharbang}{\isacharcolon}\ monoid\ append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline
 \ \ \isacommand{proof}\isamarkupfalse%
 \ \isacommand{qed}\isamarkupfalse%
 \ auto%
@@ -874,13 +904,13 @@
 \ \ {\isachardoublequoteopen}replicate\ {\isadigit{0}}\ {\isacharunderscore}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline
 \ \ {\isacharbar}\ {\isachardoublequoteopen}replicate\ {\isacharparenleft}Suc\ n{\isacharparenright}\ xs\ {\isacharequal}\ xs\ {\isacharat}\ replicate\ n\ xs{\isachardoublequoteclose}\isanewline
 \isanewline
-\isacommand{class{\isacharunderscore}interpretation}\isamarkupfalse%
-\ list{\isacharunderscore}monoid{\isacharcolon}\ monoid\ {\isacharbrackleft}append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}{\isacharbrackright}\ \isakeyword{where}\isanewline
+\isacommand{interpretation}\isamarkupfalse%
+\ 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{class{\isacharunderscore}interpret}\isamarkupfalse%
-\ monoid\ {\isacharbrackleft}append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}{\isacharbrackright}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
+\ \ \isacommand{interpret}\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
--- a/doc-src/IsarAdvanced/Codegen/Thy/Introduction.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/doc-src/IsarAdvanced/Codegen/Thy/Introduction.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/doc-src/IsarAdvanced/Codegen/Thy/ML.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -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.
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/doc-src/IsarAdvanced/Codegen/Thy/Program.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/doc-src/IsarAdvanced/Codegen/Thy/Setup.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -4,7 +4,7 @@
 begin
 
 ML {* no_document use_thys
-  ["Efficient_Nat", "Code_Char_chr", "Product_ord", "Imperative_HOL",
+  ["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/Introduction.tex	Thu Jan 29 12:05:19 2009 +0000
+++ b/doc-src/IsarAdvanced/Codegen/Thy/document/Introduction.tex	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/doc-src/IsarAdvanced/Codegen/Thy/document/ML.tex	Thu Jan 29 12:24:00 2009 +0000
@@ -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.
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/doc-src/IsarAdvanced/Codegen/Thy/document/Program.tex	Thu Jan 29 12:24:00 2009 +0000
@@ -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%
@@ -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%
 %
@@ -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}
@@ -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.
@@ -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
@@ -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
--- a/doc-src/IsarImplementation/Thy/ML.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/doc-src/IsarImplementation/Thy/ML.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -291,7 +291,7 @@
   in particular files \emph{Pure/library.ML} and \emph{Pure/General/*.ML}.
 *}
 
-section {* Linear transformations *}
+section {* Linear transformations \label{sec:ML-linear-trans} *}
 
 text %mlref {*
   \begin{mldecls}
@@ -317,9 +317,9 @@
   a theory by constant declararion and primitive definitions:
 
   \smallskip\begin{mldecls}
-  @{ML "Sign.declare_const: Properties.T -> (Binding.T * typ) * mixfix
+  @{ML "Sign.declare_const: Properties.T -> (binding * typ) * mixfix
   -> theory -> term * theory"} \\
-  @{ML "Thm.add_def: bool -> bool -> bstring * term -> theory -> thm * theory"}
+  @{ML "Thm.add_def: bool -> bool -> binding * term -> theory -> thm * theory"}
   \end{mldecls}
 
   \noindent Written with naive application, an addition of constant
@@ -328,7 +328,7 @@
 
   \smallskip\begin{mldecls}
   @{ML "(fn (t, thy) => Thm.add_def false false
-  (\"bar_def\", Logic.mk_equals (t, @{term \"%x. x\"})) thy)
+  (Binding.name \"bar_def\", Logic.mk_equals (t, @{term \"%x. x\"})) thy)
     (Sign.declare_const []
       ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn) thy)"}
   \end{mldecls}
@@ -347,7 +347,7 @@
 |> Sign.declare_const [] ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
 |> (fn (t, thy) => thy
 |> Thm.add_def false false
-     (\"bar_def\", Logic.mk_equals (t, @{term \"%x. x\"})))"}
+     (Binding.name \"bar_def\", Logic.mk_equals (t, @{term \"%x. x\"})))"}
   \end{mldecls}
 *}
 
@@ -370,7 +370,7 @@
 @{ML "thy
 |> Sign.declare_const [] ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
 |-> (fn t => Thm.add_def false false
-      (\"bar_def\", Logic.mk_equals (t, @{term \"%x. x\"})))
+      (Binding.name \"bar_def\", Logic.mk_equals (t, @{term \"%x. x\"})))
 "}
   \end{mldecls}
 
@@ -380,7 +380,7 @@
 @{ML "thy
 |> Sign.declare_const [] ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
 |>> (fn t => Logic.mk_equals (t, @{term \"%x. x\"}))
-|-> (fn def => Thm.add_def false false (\"bar_def\", def))
+|-> (fn def => Thm.add_def false false (Binding.name \"bar_def\", def))
 "}
   \end{mldecls}
 
@@ -392,7 +392,7 @@
 |> Sign.declare_const [] ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
 ||> Sign.add_path \"foobar\"
 |-> (fn t => Thm.add_def false false
-      (\"bar_def\", Logic.mk_equals (t, @{term \"%x. x\"})))
+      (Binding.name \"bar_def\", Logic.mk_equals (t, @{term \"%x. x\"})))
 ||> Sign.restore_naming thy
 "}
   \end{mldecls}
@@ -404,7 +404,7 @@
 |> Sign.declare_const [] ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
 ||>> Sign.declare_const [] ((Binding.name \"foobar\", @{typ \"foo => foo\"}), NoSyn)
 |-> (fn (t1, t2) => Thm.add_def false false
-      (\"bar_def\", Logic.mk_equals (t1, t2)))
+      (Binding.name \"bar_def\", Logic.mk_equals (t1, t2)))
 "}
   \end{mldecls}
 *}
@@ -451,7 +451,7 @@
        ((Binding.name const, @{typ \"foo => foo\"}), NoSyn)) consts
   |>> map (fn t => Logic.mk_equals (t, @{term \"%x. x\"}))
   |-> (fn defs => fold_map (fn def =>
-       Thm.add_def false false (\"\", def)) defs)
+       Thm.add_def false false (Binding.empty, def)) defs)
 end"}
   \end{mldecls}
 *}
--- a/doc-src/IsarImplementation/Thy/document/ML.tex	Thu Jan 29 12:05:19 2009 +0000
+++ b/doc-src/IsarImplementation/Thy/document/ML.tex	Thu Jan 29 12:24:00 2009 +0000
@@ -319,7 +319,7 @@
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
-\isamarkupsection{Linear transformations%
+\isamarkupsection{Linear transformations \label{sec:ML-linear-trans}%
 }
 \isamarkuptrue%
 %
@@ -366,9 +366,9 @@
   a theory by constant declararion and primitive definitions:
 
   \smallskip\begin{mldecls}
-  \verb|Sign.declare_const: Properties.T -> (Binding.T * typ) * mixfix|\isasep\isanewline%
+  \verb|Sign.declare_const: Properties.T -> (binding * typ) * mixfix|\isasep\isanewline%
 \verb|  -> theory -> term * theory| \\
-  \verb|Thm.add_def: bool -> bool -> bstring * term -> theory -> thm * theory|
+  \verb|Thm.add_def: bool -> bool -> binding * term -> theory -> thm * theory|
   \end{mldecls}
 
   \noindent Written with naive application, an addition of constant
@@ -377,7 +377,7 @@
 
   \smallskip\begin{mldecls}
   \verb|(fn (t, thy) => Thm.add_def false false|\isasep\isanewline%
-\verb|  ("bar_def", Logic.mk_equals (t, @{term "%x. x"})) thy)|\isasep\isanewline%
+\verb|  (Binding.name "bar_def", Logic.mk_equals (t, @{term "%x. x"})) thy)|\isasep\isanewline%
 \verb|    (Sign.declare_const []|\isasep\isanewline%
 \verb|      ((Binding.name "bar", @{typ "foo => foo"}), NoSyn) thy)|
   \end{mldecls}
@@ -397,7 +397,7 @@
 \verb||\verb,|,\verb|> Sign.declare_const [] ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
 \verb||\verb,|,\verb|> (fn (t, thy) => thy|\isasep\isanewline%
 \verb||\verb,|,\verb|> Thm.add_def false false|\isasep\isanewline%
-\verb|     ("bar_def", Logic.mk_equals (t, @{term "%x. x"})))|
+\verb|     (Binding.name "bar_def", Logic.mk_equals (t, @{term "%x. x"})))|
   \end{mldecls}%
 \end{isamarkuptext}%
 \isamarkuptrue%
@@ -435,7 +435,7 @@
 \verb|thy|\isasep\isanewline%
 \verb||\verb,|,\verb|> Sign.declare_const [] ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
 \verb||\verb,|,\verb|-> (fn t => Thm.add_def false false|\isasep\isanewline%
-\verb|      ("bar_def", Logic.mk_equals (t, @{term "%x. x"})))|\isasep\isanewline%
+\verb|      (Binding.name "bar_def", Logic.mk_equals (t, @{term "%x. x"})))|\isasep\isanewline%
 
   \end{mldecls}
 
@@ -445,7 +445,7 @@
 \verb|thy|\isasep\isanewline%
 \verb||\verb,|,\verb|> Sign.declare_const [] ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
 \verb||\verb,|,\verb|>> (fn t => Logic.mk_equals (t, @{term "%x. x"}))|\isasep\isanewline%
-\verb||\verb,|,\verb|-> (fn def => Thm.add_def false false ("bar_def", def))|\isasep\isanewline%
+\verb||\verb,|,\verb|-> (fn def => Thm.add_def false false (Binding.name "bar_def", def))|\isasep\isanewline%
 
   \end{mldecls}
 
@@ -457,7 +457,7 @@
 \verb||\verb,|,\verb|> Sign.declare_const [] ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
 \verb||\verb,|,\verb||\verb,|,\verb|> Sign.add_path "foobar"|\isasep\isanewline%
 \verb||\verb,|,\verb|-> (fn t => Thm.add_def false false|\isasep\isanewline%
-\verb|      ("bar_def", Logic.mk_equals (t, @{term "%x. x"})))|\isasep\isanewline%
+\verb|      (Binding.name "bar_def", Logic.mk_equals (t, @{term "%x. x"})))|\isasep\isanewline%
 \verb||\verb,|,\verb||\verb,|,\verb|> Sign.restore_naming thy|\isasep\isanewline%
 
   \end{mldecls}
@@ -469,7 +469,7 @@
 \verb||\verb,|,\verb|> Sign.declare_const [] ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
 \verb||\verb,|,\verb||\verb,|,\verb|>> Sign.declare_const [] ((Binding.name "foobar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
 \verb||\verb,|,\verb|-> (fn (t1, t2) => Thm.add_def false false|\isasep\isanewline%
-\verb|      ("bar_def", Logic.mk_equals (t1, t2)))|\isasep\isanewline%
+\verb|      (Binding.name "bar_def", Logic.mk_equals (t1, t2)))|\isasep\isanewline%
 
   \end{mldecls}%
 \end{isamarkuptext}%
@@ -531,7 +531,7 @@
 \verb|       ((Binding.name const, @{typ "foo => foo"}), NoSyn)) consts|\isasep\isanewline%
 \verb|  |\verb,|,\verb|>> map (fn t => Logic.mk_equals (t, @{term "%x. x"}))|\isasep\isanewline%
 \verb|  |\verb,|,\verb|-> (fn defs => fold_map (fn def =>|\isasep\isanewline%
-\verb|       Thm.add_def false false ("", def)) defs)|\isasep\isanewline%
+\verb|       Thm.add_def false false (Binding.empty, def)) defs)|\isasep\isanewline%
 \verb|end|
   \end{mldecls}%
 \end{isamarkuptext}%
--- a/doc-src/IsarImplementation/Thy/document/logic.tex	Thu Jan 29 12:05:19 2009 +0000
+++ b/doc-src/IsarImplementation/Thy/document/logic.tex	Thu Jan 29 12:24:00 2009 +0000
@@ -325,9 +325,9 @@
   \indexml{fastype\_of}\verb|fastype_of: term -> typ| \\
   \indexml{lambda}\verb|lambda: term -> term -> term| \\
   \indexml{betapply}\verb|betapply: term * term -> term| \\
-  \indexml{Sign.declare\_const}\verb|Sign.declare_const: Properties.T -> (Binding.T * typ) * mixfix ->|\isasep\isanewline%
+  \indexml{Sign.declare\_const}\verb|Sign.declare_const: Properties.T -> (binding * typ) * mixfix ->|\isasep\isanewline%
 \verb|  theory -> term * theory| \\
-  \indexml{Sign.add\_abbrev}\verb|Sign.add_abbrev: string -> Properties.T -> Binding.T * term ->|\isasep\isanewline%
+  \indexml{Sign.add\_abbrev}\verb|Sign.add_abbrev: string -> Properties.T -> binding * term ->|\isasep\isanewline%
 \verb|  theory -> (term * term) * theory| \\
   \indexml{Sign.const\_typargs}\verb|Sign.const_typargs: theory -> string * typ -> typ list| \\
   \indexml{Sign.const\_instance}\verb|Sign.const_instance: theory -> string * typ list -> typ| \\
@@ -594,9 +594,9 @@
 \verb|  -> (string * ('a -> thm)) * theory| \\
   \end{mldecls}
   \begin{mldecls}
-  \indexml{Theory.add\_axioms\_i}\verb|Theory.add_axioms_i: (string * term) list -> theory -> theory| \\
+  \indexml{Theory.add\_axioms\_i}\verb|Theory.add_axioms_i: (binding * term) list -> theory -> theory| \\
   \indexml{Theory.add\_deps}\verb|Theory.add_deps: string -> string * typ -> (string * typ) list -> theory -> theory| \\
-  \indexml{Theory.add\_defs\_i}\verb|Theory.add_defs_i: bool -> bool -> (bstring * term) list -> theory -> theory| \\
+  \indexml{Theory.add\_defs\_i}\verb|Theory.add_defs_i: bool -> bool -> (binding * term) list -> theory -> theory| \\
   \end{mldecls}
 
   \begin{description}
--- a/doc-src/IsarImplementation/Thy/document/prelim.tex	Thu Jan 29 12:05:19 2009 +0000
+++ b/doc-src/IsarImplementation/Thy/document/prelim.tex	Thu Jan 29 12:24:00 2009 +0000
@@ -816,13 +816,13 @@
   \indexmltype{NameSpace.naming}\verb|type NameSpace.naming| \\
   \indexml{NameSpace.default\_naming}\verb|NameSpace.default_naming: NameSpace.naming| \\
   \indexml{NameSpace.add\_path}\verb|NameSpace.add_path: string -> NameSpace.naming -> NameSpace.naming| \\
-  \indexml{NameSpace.full\_name}\verb|NameSpace.full_name: NameSpace.naming -> Binding.T -> string| \\
+  \indexml{NameSpace.full\_name}\verb|NameSpace.full_name: NameSpace.naming -> binding -> string| \\
   \end{mldecls}
   \begin{mldecls}
   \indexmltype{NameSpace.T}\verb|type NameSpace.T| \\
   \indexml{NameSpace.empty}\verb|NameSpace.empty: NameSpace.T| \\
   \indexml{NameSpace.merge}\verb|NameSpace.merge: NameSpace.T * NameSpace.T -> NameSpace.T| \\
-  \indexml{NameSpace.declare}\verb|NameSpace.declare: NameSpace.naming -> Binding.T -> NameSpace.T -> string * NameSpace.T| \\
+  \indexml{NameSpace.declare}\verb|NameSpace.declare: NameSpace.naming -> binding -> NameSpace.T -> string * NameSpace.T| \\
   \indexml{NameSpace.intern}\verb|NameSpace.intern: NameSpace.T -> string -> string| \\
   \indexml{NameSpace.extern}\verb|NameSpace.extern: NameSpace.T -> string -> string| \\
   \end{mldecls}
--- a/doc-src/IsarImplementation/Thy/logic.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/doc-src/IsarImplementation/Thy/logic.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -323,9 +323,9 @@
   @{index_ML fastype_of: "term -> typ"} \\
   @{index_ML lambda: "term -> term -> term"} \\
   @{index_ML betapply: "term * term -> term"} \\
-  @{index_ML Sign.declare_const: "Properties.T -> (Binding.T * typ) * mixfix ->
+  @{index_ML Sign.declare_const: "Properties.T -> (binding * typ) * mixfix ->
   theory -> term * theory"} \\
-  @{index_ML Sign.add_abbrev: "string -> Properties.T -> Binding.T * term ->
+  @{index_ML Sign.add_abbrev: "string -> Properties.T -> binding * term ->
   theory -> (term * term) * theory"} \\
   @{index_ML Sign.const_typargs: "theory -> string * typ -> typ list"} \\
   @{index_ML Sign.const_instance: "theory -> string * typ list -> typ"} \\
@@ -596,9 +596,9 @@
   -> (string * ('a -> thm)) * theory"} \\
   \end{mldecls}
   \begin{mldecls}
-  @{index_ML Theory.add_axioms_i: "(string * term) list -> theory -> theory"} \\
+  @{index_ML Theory.add_axioms_i: "(binding * term) list -> theory -> theory"} \\
   @{index_ML Theory.add_deps: "string -> string * typ -> (string * typ) list -> theory -> theory"} \\
-  @{index_ML Theory.add_defs_i: "bool -> bool -> (bstring * term) list -> theory -> theory"} \\
+  @{index_ML Theory.add_defs_i: "bool -> bool -> (binding * term) list -> theory -> theory"} \\
   \end{mldecls}
 
   \begin{description}
--- a/doc-src/IsarImplementation/Thy/prelim.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/doc-src/IsarImplementation/Thy/prelim.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -707,13 +707,13 @@
   @{index_ML_type NameSpace.naming} \\
   @{index_ML NameSpace.default_naming: NameSpace.naming} \\
   @{index_ML NameSpace.add_path: "string -> NameSpace.naming -> NameSpace.naming"} \\
-  @{index_ML NameSpace.full_name: "NameSpace.naming -> Binding.T -> string"} \\
+  @{index_ML NameSpace.full_name: "NameSpace.naming -> binding -> string"} \\
   \end{mldecls}
   \begin{mldecls}
   @{index_ML_type NameSpace.T} \\
   @{index_ML NameSpace.empty: NameSpace.T} \\
   @{index_ML NameSpace.merge: "NameSpace.T * NameSpace.T -> NameSpace.T"} \\
-  @{index_ML NameSpace.declare: "NameSpace.naming -> Binding.T -> NameSpace.T -> string * NameSpace.T"} \\
+  @{index_ML NameSpace.declare: "NameSpace.naming -> binding -> NameSpace.T -> string * NameSpace.T"} \\
   @{index_ML NameSpace.intern: "NameSpace.T -> string -> string"} \\
   @{index_ML NameSpace.extern: "NameSpace.T -> string -> string"} \\
   \end{mldecls}
--- a/doc-src/IsarRef/Thy/HOL_Specific.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/doc-src/IsarRef/Thy/HOL_Specific.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/doc-src/IsarRef/Thy/Spec.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,3 @@
-(* $Id$ *)
-
 theory Spec
 imports Main
 begin
@@ -438,7 +436,6 @@
   \begin{matharray}{rcl}
     @{command_def "interpretation"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
     @{command_def "interpret"} & : & @{text "proof(state) | proof(chain \<rightarrow> proof(prove)"} \\
-    @{command_def "print_interps"}@{text "\<^sup>*"} & : &  @{text "context \<rightarrow>"} \\
   \end{matharray}
 
   \indexouternonterm{interp}
@@ -447,8 +444,6 @@
     ;
     'interpret' interp
     ;
-    'print\_interps' '!'? name
-    ;
     instantiation: ('[' (inst+) ']')?
     ;
     interp: (name ':')? \\ (contextexpr instantiation |
@@ -533,13 +528,6 @@
   interprets @{text expr} in the proof context and is otherwise
   similar to interpretation in theories.
 
-  \item @{command "print_interps"}~@{text loc} prints the
-  interpretations of a particular locale @{text loc} that are active
-  in the current context, either theory or proof context.  The
-  exclamation point argument triggers printing of \emph{witness}
-  theorems justifying interpretations.  These are normally omitted
-  from the output.
-  
   \end{description}
 
   \begin{warn}
--- a/doc-src/IsarRef/Thy/document/HOL_Specific.tex	Thu Jan 29 12:05:19 2009 +0000
+++ b/doc-src/IsarRef/Thy/document/HOL_Specific.tex	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/doc-src/IsarRef/Thy/document/Spec.tex	Thu Jan 29 12:24:00 2009 +0000
@@ -3,8 +3,6 @@
 \def\isabellecontext{Spec}%
 %
 \isadelimtheory
-\isanewline
-\isanewline
 %
 \endisadelimtheory
 %
@@ -455,7 +453,6 @@
   \begin{matharray}{rcl}
     \indexdef{}{command}{interpretation}\hypertarget{command.interpretation}{\hyperlink{command.interpretation}{\mbox{\isa{\isacommand{interpretation}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ proof{\isacharparenleft}prove{\isacharparenright}{\isachardoublequote}} \\
     \indexdef{}{command}{interpret}\hypertarget{command.interpret}{\hyperlink{command.interpret}{\mbox{\isa{\isacommand{interpret}}}}} & : & \isa{{\isachardoublequote}proof{\isacharparenleft}state{\isacharparenright}\ {\isacharbar}\ proof{\isacharparenleft}chain\ {\isasymrightarrow}\ proof{\isacharparenleft}prove{\isacharparenright}{\isachardoublequote}} \\
-    \indexdef{}{command}{print\_interps}\hypertarget{command.print-interps}{\hyperlink{command.print-interps}{\mbox{\isa{\isacommand{print{\isacharunderscore}interps}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : &  \isa{{\isachardoublequote}context\ {\isasymrightarrow}{\isachardoublequote}} \\
   \end{matharray}
 
   \indexouternonterm{interp}
@@ -464,8 +461,6 @@
     ;
     'interpret' interp
     ;
-    'print\_interps' '!'? name
-    ;
     instantiation: ('[' (inst+) ']')?
     ;
     interp: (name ':')? \\ (contextexpr instantiation |
@@ -545,13 +540,6 @@
   interprets \isa{expr} in the proof context and is otherwise
   similar to interpretation in theories.
 
-  \item \hyperlink{command.print-interps}{\mbox{\isa{\isacommand{print{\isacharunderscore}interps}}}}~\isa{loc} prints the
-  interpretations of a particular locale \isa{loc} that are active
-  in the current context, either theory or proof context.  The
-  exclamation point argument triggers printing of \emph{witness}
-  theorems justifying interpretations.  These are normally omitted
-  from the output.
-  
   \end{description}
 
   \begin{warn}
--- a/doc-src/Locales/Locales/Examples.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/doc-src/Locales/Locales/Examples.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -608,7 +608,7 @@
   and @{text lattice} be placed between @{text partial_order}
   and @{text total_order}, as shown in Figure~\ref{fig:lattices}(b).
   Changes to the locale hierarchy may be declared
-  with the \isakeyword{interpretation} command. *}
+  with the \isakeyword{sublocale} command. *}
 
   sublocale %visible total_order \<subseteq> lattice
 
--- a/doc-src/Locales/Locales/Examples3.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/doc-src/Locales/Locales/Examples3.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -178,8 +178,6 @@
   nat_dvd_join_eq} are named since they are handy in the proof of
   the subsequent interpretation. *}
 
-ML %invisible {* set quick_and_dirty *}
-
 (*
 definition
   is_lcm :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where
@@ -200,8 +198,6 @@
 lemma %invisible gcd_lcm_distr:
   "gcd x (lcm y z) = lcm (gcd x y) (gcd x z)" sorry
 
-ML %invisible {* reset quick_and_dirty *}
-  
 interpretation %visible nat_dvd:
   distrib_lattice "op dvd :: nat \<Rightarrow> nat \<Rightarrow> bool"
   apply unfold_locales
@@ -262,7 +258,7 @@
   preserving maps can be declared in the following way.  *}
 
   locale order_preserving =
-    partial_order + po': partial_order le' for le' (infixl "\<preceq>" 50) +
+    le: partial_order + le': partial_order le' for le' (infixl "\<preceq>" 50) +
     fixes \<phi> :: "'a \<Rightarrow> 'b"
     assumes hom_le: "x \<sqsubseteq> y \<Longrightarrow> \<phi> x \<preceq> \<phi> y"
 
@@ -288,8 +284,7 @@
   obtained by appending the conclusions of the left locale and of the
   right locale.  *}
 
-text {* % FIXME needs update
-  The locale @{text order_preserving} contains theorems for both
+text {* The locale @{text order_preserving} contains theorems for both
   orders @{text \<sqsubseteq>} and @{text \<preceq>}.  How can one refer to a theorem for
   a particular order, @{text \<sqsubseteq>} or @{text \<preceq>}?  Names in locales are
   qualified by the locale parameters.  More precisely, a name is
@@ -298,8 +293,8 @@
 
 context %invisible order_preserving begin
 
-text {* % FIXME needs update?
-  @{thm [source] less_le_trans}: @{thm less_le_trans}
+text {*
+  @{thm [source] le.less_le_trans}: @{thm le.less_le_trans}
 
   @{thm [source] hom_le}: @{thm hom_le}
   *}
@@ -307,12 +302,11 @@
 text {* When renaming a locale, the morphism is also applied
   to the qualifiers.  Hence theorems for the partial order @{text \<preceq>}
   are qualified by @{text le'}.  For example, @{thm [source]
-  po'.less_le_trans}: @{thm [display, indent=2] po'.less_le_trans} *}
+  le'.less_le_trans}: @{thm [display, indent=2] le'.less_le_trans} *}
 
 end %invisible
 
-text {* % FIXME needs update?
-  This example reveals that there is no infix syntax for the strict
+text {* This example reveals that there is no infix syntax for the strict
   version of @{text \<preceq>}!  This can, of course, not be introduced
   automatically, but it can be declared manually through an abbreviation.
   *}
@@ -321,7 +315,7 @@
     less' (infixl "\<prec>" 50) where "less' \<equiv> partial_order.less le'"
 
 text {* Now the theorem is displayed nicely as
-  @{thm [locale=order_preserving] po'.less_le_trans}.  *}
+  @{thm [locale=order_preserving] le'.less_le_trans}.  *}
 
 text {* Not only names of theorems are qualified.  In fact, all names
   are qualified, in particular names introduced by definitions and
@@ -333,7 +327,7 @@
 text {* Two more locales illustrate working with locale expressions.
   A map @{text \<phi>} is a lattice homomorphism if it preserves meet and join. *}
 
-  locale lattice_hom = lattice + lat'!: lattice le' for le' (infixl "\<preceq>" 50) +
+  locale lattice_hom = le: lattice + le': lattice le' for le' (infixl "\<preceq>" 50) +
     fixes \<phi>
     assumes hom_meet:
 	"\<phi> (lattice.meet le x y) = lattice.meet le' (\<phi> x) (\<phi> y)"
@@ -341,9 +335,9 @@
 	"\<phi> (lattice.join le x y) = lattice.join le' (\<phi> x) (\<phi> y)"
 
   abbreviation (in lattice_hom)
-    meet' (infixl "\<sqinter>''" 50) where "meet' \<equiv> lat'.meet"
+    meet' (infixl "\<sqinter>''" 50) where "meet' \<equiv> le'.meet"
   abbreviation (in lattice_hom)
-    join' (infixl "\<squnion>''" 50) where "join' \<equiv> lat'.join"
+    join' (infixl "\<squnion>''" 50) where "join' \<equiv> le'.join"
 
 text {* A homomorphism is an endomorphism if both orders coincide. *}
 
@@ -400,17 +394,17 @@
   sublocale lattice_hom \<subseteq> order_preserving proof unfold_locales
     fix x y
     assume "x \<sqsubseteq> y"
-    then have "y = (x \<squnion> y)" by (simp add: join_connection)
+    then have "y = (x \<squnion> y)" by (simp add: le.join_connection)
     then have "\<phi> y = (\<phi> x \<squnion>' \<phi> y)" by (simp add: hom_join [symmetric])
-    then show "\<phi> x \<preceq> \<phi> y" by (simp add: lat'.join_connection)
+    then show "\<phi> x \<preceq> \<phi> y" by (simp add: le'.join_connection)
   qed
 
 text {* Theorems and other declarations --- syntax, in particular ---
   from the locale @{text order_preserving} are now active in @{text
   lattice_hom}, for example
 
-  @{thm [locale=lattice_hom, source] lat'.less_le_trans}:
-  @{thm [locale=lattice_hom] lat'.less_le_trans}
+  @{thm [locale=lattice_hom, source] le'.less_le_trans}:
+  @{thm [locale=lattice_hom] le'.less_le_trans}
   *}
 
 
@@ -450,7 +444,9 @@
 
   \textit{attr-name} & ::=
   & \textit{name} $|$ \textit{attribute} $|$
-    \textit{name} \textit{attribute} \\[2ex]
+    \textit{name} \textit{attribute} \\
+  \textit{qualifier} & ::=
+  & \textit{name} [``\textbf{!}''] \\[2ex]
 
   \multicolumn{3}{l}{Context Elements} \\
 
@@ -490,19 +486,23 @@
 
   \multicolumn{3}{l}{Locale Expressions} \\
 
-  \textit{rename} & ::=
-  & \textit{name} [ \textit{mixfix} ] $|$ ``\textbf{\_}'' \\
-  \textit{expr}  & ::= 
-  & \textit{renamed-expr} ( ``\textbf{+}'' \textit{renamed-expr} )$^*$ \\
-  \textit{renamed-expr} & ::=
-  & ( \textit{qualified-name} $|$
-    ``\textbf{(}'' \textit{expr} ``\textbf{)}'' ) \textit{rename}$^*$ \\[2ex]
+  \textit{pos-insts} & ::=
+  & ( \textit{term} $|$ ``\textbf{\_}'' )$^*$ \\
+  \textit{named-insts} & ::=
+  & \textbf{where} \textit{name} ``\textbf{=}'' \textit{term}
+  ( \textbf{and} \textit{name} ``\textbf{=}'' \textit{term} )$^*$ \\
+  \textit{instance} & ::=
+  & [ \textit{qualifier} \textbf{:} ]
+    \textit{qualified-name} ( \textit{pos-insts} $|$ \textit{named-inst} ) \\
+  \textit{expression}  & ::= 
+  & \textit{instance} ( ``\textbf{+}'' \textit{instance} )$^*$
+    [ \textbf{for} \textit{fixes} ( \textbf{and} \textit{fixes} )$^*$ ] \\[2ex]
 
   \multicolumn{3}{l}{Declaration of Locales} \\
 
   \textit{locale} & ::=
   & \textit{element}$^+$ \\
-  & | & \textit{locale-expr} [ ``\textbf{+}'' \textit{element}$^+$ ] \\
+  & | & \textit{expression} [ ``\textbf{+}'' \textit{element}$^+$ ] \\
   \textit{toplevel} & ::=
   & \textbf{locale} \textit{name} [ ``\textbf{=}''
     \textit{locale} ] \\[2ex]
@@ -511,19 +511,17 @@
 
   \textit{equation} & ::= & [ \textit{attr-name} ``\textbf{:}'' ]
     \textit{prop} \\
-  \textit{insts} & ::= & [ ``\textbf{[}'' \textit{term}$^+$
-    ``\textbf{]}'' ] \\
-  & & [ \textbf{where} \textit{equation} ( \textbf{and}
-    \textit{equation} )$^*$ ] \\
+  \textit{equations} & ::= &  \textbf{where} \textit{equation} ( \textbf{and}
+    \textit{equation} )$^*$  \\
   \textit{toplevel} & ::=
-  & \textbf{interpretation} \textit{name} ( ``$<$'' $|$
-    ``$\subseteq$'' ) \textit{expr} \textit{proof} \\
+  & \textbf{sublocale} \textit{name} ( ``$<$'' $|$
+    ``$\subseteq$'' ) \textit{expression} \textit{proof} \\
   & |
-  & \textbf{interpretation} [ \textit{attr-name} ``\textbf{:}'' ]
-    \textit{expr} \textit{insts} \textit{proof} \\
+  & \textbf{interpretation}
+    \textit{expression} [ \textit{equations} ] \textit{proof} \\
   & |
-  & \textbf{interpret} [ \textit{attr-name} ``\textbf{:}'' ]
-    \textit{expr} \textit{insts} \textit{proof} \\[2ex]
+  & \textbf{interpret}
+    \textit{expression} \textit{proof} \\[2ex]
 
   \multicolumn{3}{l}{Diagnostics} \\
 
@@ -533,7 +531,7 @@
 \end{tabular}
 \end{center}
 \hrule
-\caption{Syntax of Locale Commands.}
+\caption{Syntax of Locale Commands (abridged).}
 \label{tab:commands}
 \end{table}
   *}
--- a/doc-src/Locales/Locales/ROOT.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/doc-src/Locales/Locales/ROOT.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,4 +1,4 @@
 no_document use_thy "GCD";
 use_thy "Examples1";
 use_thy "Examples2";
-use_thy "Examples3";
+setmp_noncritical quick_and_dirty true use_thy "Examples3";
--- a/doc-src/Locales/Locales/document/Examples.tex	Thu Jan 29 12:05:19 2009 +0000
+++ b/doc-src/Locales/Locales/document/Examples.tex	Thu Jan 29 12:24:00 2009 +0000
@@ -1213,7 +1213,7 @@
   and \isa{lattice} be placed between \isa{partial{\isacharunderscore}order}
   and \isa{total{\isacharunderscore}order}, as shown in Figure~\ref{fig:lattices}(b).
   Changes to the locale hierarchy may be declared
-  with the \isakeyword{interpretation} command.%
+  with the \isakeyword{sublocale} command.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
--- a/doc-src/Locales/Locales/document/Examples3.tex	Thu Jan 29 12:05:19 2009 +0000
+++ b/doc-src/Locales/Locales/document/Examples3.tex	Thu Jan 29 12:24:00 2009 +0000
@@ -362,18 +362,10 @@
 \endisadeliminvisible
 %
 \isataginvisible
-\isacommand{ML}\isamarkupfalse%
-\ {\isacharverbatimopen}\ set\ quick{\isacharunderscore}and{\isacharunderscore}dirty\ {\isacharverbatimclose}\isanewline
-\isanewline
-\isanewline
-\isanewline
 \isacommand{lemma}\isamarkupfalse%
 \ gcd{\isacharunderscore}lcm{\isacharunderscore}distr{\isacharcolon}\isanewline
 \ \ {\isachardoublequoteopen}gcd\ x\ {\isacharparenleft}lcm\ y\ z{\isacharparenright}\ {\isacharequal}\ lcm\ {\isacharparenleft}gcd\ x\ y{\isacharparenright}\ {\isacharparenleft}gcd\ x\ z{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{sorry}\isamarkupfalse%
-\isanewline
-\isanewline
-\isacommand{ML}\isamarkupfalse%
-\ {\isacharverbatimopen}\ reset\ quick{\isacharunderscore}and{\isacharunderscore}dirty\ {\isacharverbatimclose}%
+%
 \endisataginvisible
 {\isafoldinvisible}%
 %
@@ -383,7 +375,7 @@
 \isanewline
 %
 \isadelimvisible
-\ \ \isanewline
+\isanewline
 %
 \endisadelimvisible
 %
@@ -476,7 +468,7 @@
 \isamarkuptrue%
 \ \ \isacommand{locale}\isamarkupfalse%
 \ order{\isacharunderscore}preserving\ {\isacharequal}\isanewline
-\ \ \ \ partial{\isacharunderscore}order\ {\isacharplus}\ po{\isacharprime}{\isacharcolon}\ partial{\isacharunderscore}order\ le{\isacharprime}\ \isakeyword{for}\ le{\isacharprime}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasympreceq}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ {\isacharplus}\isanewline
+\ \ \ \ le{\isacharcolon}\ partial{\isacharunderscore}order\ {\isacharplus}\ le{\isacharprime}{\isacharcolon}\ partial{\isacharunderscore}order\ le{\isacharprime}\ \isakeyword{for}\ le{\isacharprime}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasympreceq}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ {\isacharplus}\isanewline
 \ \ \ \ \isakeyword{fixes}\ {\isasymphi}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b{\isachardoublequoteclose}\isanewline
 \ \ \ \ \isakeyword{assumes}\ hom{\isacharunderscore}le{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymsqsubseteq}\ y\ {\isasymLongrightarrow}\ {\isasymphi}\ x\ {\isasympreceq}\ {\isasymphi}\ y{\isachardoublequoteclose}%
 \begin{isamarkuptext}%
@@ -505,8 +497,7 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-% FIXME needs update
-  The locale \isa{order{\isacharunderscore}preserving} contains theorems for both
+The locale \isa{order{\isacharunderscore}preserving} contains theorems for both
   orders \isa{{\isasymsqsubseteq}} and \isa{{\isasympreceq}}.  How can one refer to a theorem for
   a particular order, \isa{{\isasymsqsubseteq}} or \isa{{\isasympreceq}}?  Names in locales are
   qualified by the locale parameters.  More precisely, a name is
@@ -530,8 +521,7 @@
 \endisadeliminvisible
 %
 \begin{isamarkuptext}%
-% FIXME needs update?
-  \isa{less{\isacharunderscore}le{\isacharunderscore}trans}: \isa{{\isasymlbrakk}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasympreceq}\ {\isacharquery}x\ {\isacharquery}y{\isacharsemicolon}\ {\isacharquery}y\ {\isasympreceq}\ {\isacharquery}z{\isasymrbrakk}\ {\isasymLongrightarrow}\ partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasympreceq}\ {\isacharquery}x\ {\isacharquery}z}
+\isa{le{\isachardot}less{\isacharunderscore}le{\isacharunderscore}trans}: \isa{{\isasymlbrakk}{\isacharquery}x\ {\isasymsqsubset}\ {\isacharquery}y{\isacharsemicolon}\ {\isacharquery}y\ {\isasymsqsubseteq}\ {\isacharquery}z{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}x\ {\isasymsqsubset}\ {\isacharquery}z}
 
   \isa{hom{\isacharunderscore}le}: \isa{{\isacharquery}x\ {\isasymsqsubseteq}\ {\isacharquery}y\ {\isasymLongrightarrow}\ {\isasymphi}\ {\isacharquery}x\ {\isasympreceq}\ {\isasymphi}\ {\isacharquery}y}%
 \end{isamarkuptext}%
@@ -540,7 +530,7 @@
 \begin{isamarkuptext}%
 When renaming a locale, the morphism is also applied
   to the qualifiers.  Hence theorems for the partial order \isa{{\isasympreceq}}
-  are qualified by \isa{le{\isacharprime}}.  For example, \isa{po{\isacharprime}{\isachardot}less{\isacharunderscore}le{\isacharunderscore}trans}: \begin{isabelle}%
+  are qualified by \isa{le{\isacharprime}}.  For example, \isa{le{\isacharprime}{\isachardot}less{\isacharunderscore}le{\isacharunderscore}trans}: \begin{isabelle}%
 \ \ {\isasymlbrakk}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasympreceq}\ {\isacharquery}x\ {\isacharquery}y{\isacharsemicolon}\ {\isacharquery}y\ {\isasympreceq}\ {\isacharquery}z{\isasymrbrakk}\isanewline
 \isaindent{\ \ }{\isasymLongrightarrow}\ partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasympreceq}\ {\isacharquery}x\ {\isacharquery}z%
 \end{isabelle}%
@@ -562,8 +552,7 @@
 \endisadeliminvisible
 %
 \begin{isamarkuptext}%
-% FIXME needs update?
-  This example reveals that there is no infix syntax for the strict
+This example reveals that there is no infix syntax for the strict
   version of \isa{{\isasympreceq}}!  This can, of course, not be introduced
   automatically, but it can be declared manually through an abbreviation.%
 \end{isamarkuptext}%
@@ -592,7 +581,7 @@
 \end{isamarkuptext}%
 \isamarkuptrue%
 \ \ \isacommand{locale}\isamarkupfalse%
-\ lattice{\isacharunderscore}hom\ {\isacharequal}\ lattice\ {\isacharplus}\ lat{\isacharprime}{\isacharbang}{\isacharcolon}\ lattice\ le{\isacharprime}\ \isakeyword{for}\ le{\isacharprime}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasympreceq}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ {\isacharplus}\isanewline
+\ lattice{\isacharunderscore}hom\ {\isacharequal}\ le{\isacharcolon}\ lattice\ {\isacharplus}\ le{\isacharprime}{\isacharcolon}\ lattice\ le{\isacharprime}\ \isakeyword{for}\ le{\isacharprime}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasympreceq}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ {\isacharplus}\isanewline
 \ \ \ \ \isakeyword{fixes}\ {\isasymphi}\isanewline
 \ \ \ \ \isakeyword{assumes}\ hom{\isacharunderscore}meet{\isacharcolon}\isanewline
 \ \ \ \ \ \ \ \ {\isachardoublequoteopen}{\isasymphi}\ {\isacharparenleft}lattice{\isachardot}meet\ le\ x\ y{\isacharparenright}\ {\isacharequal}\ lattice{\isachardot}meet\ le{\isacharprime}\ {\isacharparenleft}{\isasymphi}\ x{\isacharparenright}\ {\isacharparenleft}{\isasymphi}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
@@ -601,10 +590,10 @@
 \isanewline
 \ \ \isacommand{abbreviation}\isamarkupfalse%
 \ {\isacharparenleft}\isakeyword{in}\ lattice{\isacharunderscore}hom{\isacharparenright}\isanewline
-\ \ \ \ meet{\isacharprime}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymsqinter}{\isacharprime}{\isacharprime}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ \isakeyword{where}\ {\isachardoublequoteopen}meet{\isacharprime}\ {\isasymequiv}\ lat{\isacharprime}{\isachardot}meet{\isachardoublequoteclose}\isanewline
+\ \ \ \ meet{\isacharprime}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymsqinter}{\isacharprime}{\isacharprime}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ \isakeyword{where}\ {\isachardoublequoteopen}meet{\isacharprime}\ {\isasymequiv}\ le{\isacharprime}{\isachardot}meet{\isachardoublequoteclose}\isanewline
 \ \ \isacommand{abbreviation}\isamarkupfalse%
 \ {\isacharparenleft}\isakeyword{in}\ lattice{\isacharunderscore}hom{\isacharparenright}\isanewline
-\ \ \ \ join{\isacharprime}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymsqunion}{\isacharprime}{\isacharprime}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ \isakeyword{where}\ {\isachardoublequoteopen}join{\isacharprime}\ {\isasymequiv}\ lat{\isacharprime}{\isachardot}join{\isachardoublequoteclose}%
+\ \ \ \ join{\isacharprime}\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymsqunion}{\isacharprime}{\isacharprime}{\isachardoublequoteclose}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\ \isakeyword{where}\ {\isachardoublequoteopen}join{\isacharprime}\ {\isasymequiv}\ le{\isacharprime}{\isachardot}join{\isachardoublequoteclose}%
 \begin{isamarkuptext}%
 A homomorphism is an endomorphism if both orders coincide.%
 \end{isamarkuptext}%
@@ -678,7 +667,7 @@
 \ \ \ \ \isacommand{then}\isamarkupfalse%
 \ \isacommand{have}\isamarkupfalse%
 \ {\isachardoublequoteopen}y\ {\isacharequal}\ {\isacharparenleft}x\ {\isasymsqunion}\ y{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
-\ {\isacharparenleft}simp\ add{\isacharcolon}\ join{\isacharunderscore}connection{\isacharparenright}\isanewline
+\ {\isacharparenleft}simp\ add{\isacharcolon}\ le{\isachardot}join{\isacharunderscore}connection{\isacharparenright}\isanewline
 \ \ \ \ \isacommand{then}\isamarkupfalse%
 \ \isacommand{have}\isamarkupfalse%
 \ {\isachardoublequoteopen}{\isasymphi}\ y\ {\isacharequal}\ {\isacharparenleft}{\isasymphi}\ x\ {\isasymsqunion}{\isacharprime}\ {\isasymphi}\ y{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
@@ -686,7 +675,7 @@
 \ \ \ \ \isacommand{then}\isamarkupfalse%
 \ \isacommand{show}\isamarkupfalse%
 \ {\isachardoublequoteopen}{\isasymphi}\ x\ {\isasympreceq}\ {\isasymphi}\ y{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
-\ {\isacharparenleft}simp\ add{\isacharcolon}\ lat{\isacharprime}{\isachardot}join{\isacharunderscore}connection{\isacharparenright}\isanewline
+\ {\isacharparenleft}simp\ add{\isacharcolon}\ le{\isacharprime}{\isachardot}join{\isacharunderscore}connection{\isacharparenright}\isanewline
 \ \ \isacommand{qed}\isamarkupfalse%
 %
 \endisatagproof
@@ -700,7 +689,7 @@
 Theorems and other declarations --- syntax, in particular ---
   from the locale \isa{order{\isacharunderscore}preserving} are now active in \isa{lattice{\isacharunderscore}hom}, for example
 
-  \isa{lat{\isacharprime}{\isachardot}less{\isacharunderscore}le{\isacharunderscore}trans}:
+  \isa{le{\isacharprime}{\isachardot}less{\isacharunderscore}le{\isacharunderscore}trans}:
   \isa{{\isasymlbrakk}{\isacharquery}x\ {\isasymprec}\ {\isacharquery}y{\isacharsemicolon}\ {\isacharquery}y\ {\isasympreceq}\ {\isacharquery}z{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}x\ {\isasymprec}\ {\isacharquery}z}%
 \end{isamarkuptext}%
 \isamarkuptrue%
@@ -744,7 +733,9 @@
 
   \textit{attr-name} & ::=
   & \textit{name} $|$ \textit{attribute} $|$
-    \textit{name} \textit{attribute} \\[2ex]
+    \textit{name} \textit{attribute} \\
+  \textit{qualifier} & ::=
+  & \textit{name} [``\textbf{!}''] \\[2ex]
 
   \multicolumn{3}{l}{Context Elements} \\
 
@@ -784,19 +775,23 @@
 
   \multicolumn{3}{l}{Locale Expressions} \\
 
-  \textit{rename} & ::=
-  & \textit{name} [ \textit{mixfix} ] $|$ ``\textbf{\_}'' \\
-  \textit{expr}  & ::= 
-  & \textit{renamed-expr} ( ``\textbf{+}'' \textit{renamed-expr} )$^*$ \\
-  \textit{renamed-expr} & ::=
-  & ( \textit{qualified-name} $|$
-    ``\textbf{(}'' \textit{expr} ``\textbf{)}'' ) \textit{rename}$^*$ \\[2ex]
+  \textit{pos-insts} & ::=
+  & ( \textit{term} $|$ ``\textbf{\_}'' )$^*$ \\
+  \textit{named-insts} & ::=
+  & \textbf{where} \textit{name} ``\textbf{=}'' \textit{term}
+  ( \textbf{and} \textit{name} ``\textbf{=}'' \textit{term} )$^*$ \\
+  \textit{instance} & ::=
+  & [ \textit{qualifier} \textbf{:} ]
+    \textit{qualified-name} ( \textit{pos-insts} $|$ \textit{named-inst} ) \\
+  \textit{expression}  & ::= 
+  & \textit{instance} ( ``\textbf{+}'' \textit{instance} )$^*$
+    [ \textbf{for} \textit{fixes} ( \textbf{and} \textit{fixes} )$^*$ ] \\[2ex]
 
   \multicolumn{3}{l}{Declaration of Locales} \\
 
   \textit{locale} & ::=
   & \textit{element}$^+$ \\
-  & | & \textit{locale-expr} [ ``\textbf{+}'' \textit{element}$^+$ ] \\
+  & | & \textit{expression} [ ``\textbf{+}'' \textit{element}$^+$ ] \\
   \textit{toplevel} & ::=
   & \textbf{locale} \textit{name} [ ``\textbf{=}''
     \textit{locale} ] \\[2ex]
@@ -805,19 +800,17 @@
 
   \textit{equation} & ::= & [ \textit{attr-name} ``\textbf{:}'' ]
     \textit{prop} \\
-  \textit{insts} & ::= & [ ``\textbf{[}'' \textit{term}$^+$
-    ``\textbf{]}'' ] \\
-  & & [ \textbf{where} \textit{equation} ( \textbf{and}
-    \textit{equation} )$^*$ ] \\
+  \textit{equations} & ::= &  \textbf{where} \textit{equation} ( \textbf{and}
+    \textit{equation} )$^*$  \\
   \textit{toplevel} & ::=
-  & \textbf{interpretation} \textit{name} ( ``$<$'' $|$
-    ``$\subseteq$'' ) \textit{expr} \textit{proof} \\
+  & \textbf{sublocale} \textit{name} ( ``$<$'' $|$
+    ``$\subseteq$'' ) \textit{expression} \textit{proof} \\
   & |
-  & \textbf{interpretation} [ \textit{attr-name} ``\textbf{:}'' ]
-    \textit{expr} \textit{insts} \textit{proof} \\
+  & \textbf{interpretation}
+    \textit{expression} [ \textit{equations} ] \textit{proof} \\
   & |
-  & \textbf{interpret} [ \textit{attr-name} ``\textbf{:}'' ]
-    \textit{expr} \textit{insts} \textit{proof} \\[2ex]
+  & \textbf{interpret}
+    \textit{expression} \textit{proof} \\[2ex]
 
   \multicolumn{3}{l}{Diagnostics} \\
 
@@ -827,7 +820,7 @@
 \end{tabular}
 \end{center}
 \hrule
-\caption{Syntax of Locale Commands.}
+\caption{Syntax of Locale Commands (abridged).}
 \label{tab:commands}
 \end{table}%
 \end{isamarkuptext}%
--- a/doc-src/Locales/Locales/document/root.tex	Thu Jan 29 12:05:19 2009 +0000
+++ b/doc-src/Locales/Locales/document/root.tex	Thu Jan 29 12:24:00 2009 +0000
@@ -22,14 +22,17 @@
 
 \begin{document}
 
-\title{Tutorial to Locales and Locale Interpretation}
+\title{Tutorial to Locales and Locale Interpretation \\[1ex]
+  \large 2nd revision, for Isabelle 2009}
 \author{Clemens Ballarin}
 \date{}
 
 \maketitle
 
+%\thispagestyle{myheadings}
+%\markright{Technical Report TUM-I0723, Technische Universit\"at M\"unchen, 2007}
 \thispagestyle{myheadings}
-\markright{Technical Report TUM-I0723, Technische Universit\"at M\"unchen, 2007}
+\markright{This tutorial is currently not consistent.}
 
 \begin{abstract}
   Locales are Isabelle's mechanism to deal with parametric theories.
@@ -40,6 +43,10 @@
 
   This tutorial is intended for locale novices; familiarity with
   Isabelle and Isar is presumed.
+  The 2nd revision accommodates changes introduced by the locales
+  reimplementation for Isabelle 2009.  Most notably, in complex
+  specifications (\emph{locale expressions}) renaming has been
+  generalised to instantiation.
 \end{abstract}
 
 \parindent 0pt\parskip 0.5ex
--- a/doc-src/more_antiquote.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/doc-src/more_antiquote.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -113,13 +113,13 @@
   val parse_const_terms = Scan.repeat1 Args.term
     >> (fn ts => fn thy => map (Code_Unit.check_const thy) ts);
   val parse_consts = Scan.lift (Args.parens (Args.$$$ "consts")) |-- parse_const_terms
-    >> (fn mk_cs => fn thy => fn naming => map (the o Code_Thingol.lookup_const naming) (mk_cs thy));
+    >> (fn mk_cs => fn thy => fn naming => map_filter (Code_Thingol.lookup_const naming) (mk_cs thy));
   val parse_types = Scan.lift (Args.parens (Args.$$$ "types") |-- Scan.repeat1 Args.name)
-    >> (fn tycos => fn thy => fn naming => map (the o Code_Thingol.lookup_tyco naming o Sign.intern_type thy) tycos);
+    >> (fn tycos => fn thy => fn naming => map_filter (Code_Thingol.lookup_tyco naming o Sign.intern_type thy) tycos);
   val parse_classes = Scan.lift (Args.parens (Args.$$$ "classes") |-- Scan.repeat1 Args.name)
-    >> (fn classes => fn thy => fn naming => map (the o Code_Thingol.lookup_class naming o Sign.intern_class thy) classes);
+    >> (fn classes => fn thy => fn naming => map_filter (Code_Thingol.lookup_class naming o Sign.intern_class thy) classes);
   val parse_instances = Scan.lift (Args.parens (Args.$$$ "instances") |-- Scan.repeat1 (Args.name --| Args.$$$ "::" -- Args.name))
-    >> (fn insts => fn thy => fn naming => map (the o Code_Thingol.lookup_instance naming o apsnd (Sign.intern_type thy) o apfst (Sign.intern_class thy) o swap) insts);
+    >> (fn insts => fn thy => fn naming => map_filter (Code_Thingol.lookup_instance naming o apsnd (Sign.intern_type thy) o apfst (Sign.intern_class thy) o swap) insts);
   val parse_names = parse_consts || parse_types || parse_classes || parse_instances; 
 
   fun code_stmts src ctxt ((mk_cs, mk_stmtss), target) =
--- a/etc/isar-keywords-ZF.el	Thu Jan 29 12:05:19 2009 +0000
+++ b/etc/isar-keywords-ZF.el	Thu Jan 29 12:24:00 2009 +0000
@@ -3,14 +3,16 @@
 ;; Generated from Pure + Pure-ProofGeneral + FOL + ZF.
 ;; *** DO NOT EDIT *** DO NOT EDIT *** DO NOT EDIT ***
 ;;
-;; $Id$
-;;
 
 (defconst isar-keywords-major
   '("\\."
     "\\.\\."
     "Isabelle\\.command"
+    "Isar\\.begin_document"
     "Isar\\.command"
+    "Isar\\.define_command"
+    "Isar\\.edit_document"
+    "Isar\\.end_document"
     "Isar\\.insert"
     "Isar\\.remove"
     "ML"
@@ -89,7 +91,6 @@
     "instantiation"
     "interpret"
     "interpretation"
-    "invoke"
     "judgment"
     "kill"
     "kill_thy"
@@ -135,7 +136,6 @@
     "print_drafts"
     "print_facts"
     "print_induct_rules"
-    "print_interps"
     "print_locale"
     "print_locales"
     "print_methods"
@@ -249,7 +249,11 @@
 
 (defconst isar-keywords-control
   '("Isabelle\\.command"
+    "Isar\\.begin_document"
     "Isar\\.command"
+    "Isar\\.define_command"
+    "Isar\\.edit_document"
+    "Isar\\.end_document"
     "Isar\\.insert"
     "Isar\\.remove"
     "ProofGeneral\\.inform_file_processed"
@@ -298,7 +302,6 @@
     "print_drafts"
     "print_facts"
     "print_induct_rules"
-    "print_interps"
     "print_locale"
     "print_locales"
     "print_methods"
@@ -438,8 +441,7 @@
 (defconst isar-keywords-proof-goal
   '("have"
     "hence"
-    "interpret"
-    "invoke"))
+    "interpret"))
 
 (defconst isar-keywords-proof-block
   '("next"
--- a/etc/isar-keywords.el	Thu Jan 29 12:05:19 2009 +0000
+++ b/etc/isar-keywords.el	Thu Jan 29 12:24:00 2009 +0000
@@ -3,14 +3,16 @@
 ;; Generated from Pure + Pure-ProofGeneral + HOL + HOLCF + IOA + HOL-Nominal + HOL-Statespace.
 ;; *** DO NOT EDIT *** DO NOT EDIT *** DO NOT EDIT ***
 ;;
-;; $Id$
-;;
 
 (defconst isar-keywords-major
   '("\\."
     "\\.\\."
     "Isabelle\\.command"
+    "Isar\\.begin_document"
     "Isar\\.command"
+    "Isar\\.define_command"
+    "Isar\\.edit_document"
+    "Isar\\.end_document"
     "Isar\\.insert"
     "Isar\\.remove"
     "ML"
@@ -46,9 +48,6 @@
     "chapter"
     "class"
     "class_deps"
-    "class_interpret"
-    "class_interpretation"
-    "class_locale"
     "classes"
     "classrel"
     "code_abort"
@@ -119,7 +118,6 @@
     "instantiation"
     "interpret"
     "interpretation"
-    "invoke"
     "judgment"
     "kill"
     "kill_thy"
@@ -172,7 +170,6 @@
     "print_drafts"
     "print_facts"
     "print_induct_rules"
-    "print_interps"
     "print_locale"
     "print_locales"
     "print_methods"
@@ -312,7 +309,11 @@
 
 (defconst isar-keywords-control
   '("Isabelle\\.command"
+    "Isar\\.begin_document"
     "Isar\\.command"
+    "Isar\\.define_command"
+    "Isar\\.edit_document"
+    "Isar\\.end_document"
     "Isar\\.insert"
     "Isar\\.remove"
     "ProofGeneral\\.inform_file_processed"
@@ -369,7 +370,6 @@
     "print_drafts"
     "print_facts"
     "print_induct_rules"
-    "print_interps"
     "print_locale"
     "print_locales"
     "print_methods"
@@ -423,7 +423,6 @@
     "axiomatization"
     "axioms"
     "class"
-    "class_locale"
     "classes"
     "classrel"
     "code_abort"
@@ -507,7 +506,6 @@
 
 (defconst isar-keywords-theory-goal
   '("ax_specification"
-    "class_interpretation"
     "corollary"
     "cpodef"
     "function"
@@ -546,11 +544,9 @@
     "subsubsect"))
 
 (defconst isar-keywords-proof-goal
-  '("class_interpret"
-    "have"
+  '("have"
     "hence"
-    "interpret"
-    "invoke"))
+    "interpret"))
 
 (defconst isar-keywords-proof-block
   '("next"
--- a/etc/settings	Thu Jan 29 12:05:19 2009 +0000
+++ b/etc/settings	Thu Jan 29 12:24:00 2009 +0000
@@ -242,7 +242,6 @@
   "$ISABELLE_HOME/contrib/vampire/$ML_PLATFORM" \
   "$ISABELLE_HOME/../vampire/$ML_PLATFORM" \
   "/usr/local/Vampire" \
-  "$ISABELLE_HOME/contrib/SystemOnTPTP" \
   "")
 SPASS_HOME=$(choosefrom \
   "$ISABELLE_HOME/contrib/spass/$ML_PLATFORM/bin" \
--- a/lib/jedit/isabelle.xml	Thu Jan 29 12:05:19 2009 +0000
+++ b/lib/jedit/isabelle.xml	Thu Jan 29 12:24:00 2009 +0000
@@ -2,7 +2,6 @@
 <!DOCTYPE MODE SYSTEM "xmode.dtd">
 <!-- Generated from Pure + HOL + HOLCF + IOA + HOL-Nominal + HOL-Statespace + FOL + ZF. -->
 <!-- *** DO NOT EDIT *** DO NOT EDIT *** DO NOT EDIT *** -->
-<!-- $Id$ -->
 <MODE>
   <PROPS>
     <PROPERTY NAME="commentStart" VALUE="(*"/>
@@ -36,7 +35,11 @@
       <OPERATOR>.</OPERATOR>
       <OPERATOR>..</OPERATOR>
       <INVALID>Isabelle.command</INVALID>
+      <INVALID>Isar.begin_document</INVALID>
       <INVALID>Isar.command</INVALID>
+      <INVALID>Isar.define_command</INVALID>
+      <INVALID>Isar.edit_document</INVALID>
+      <INVALID>Isar.end_document</INVALID>
       <INVALID>Isar.insert</INVALID>
       <INVALID>Isar.remove</INVALID>
       <OPERATOR>ML</OPERATOR>
@@ -171,7 +174,6 @@
       <OPERATOR>interpret</OPERATOR>
       <OPERATOR>interpretation</OPERATOR>
       <KEYWORD4>intros</KEYWORD4>
-      <OPERATOR>invoke</OPERATOR>
       <KEYWORD4>is</KEYWORD4>
       <OPERATOR>judgment</OPERATOR>
       <INVALID>kill</INVALID>
@@ -239,7 +241,6 @@
       <LABEL>print_drafts</LABEL>
       <LABEL>print_facts</LABEL>
       <LABEL>print_induct_rules</LABEL>
-      <LABEL>print_interps</LABEL>
       <LABEL>print_locale</LABEL>
       <LABEL>print_locales</LABEL>
       <LABEL>print_methods</LABEL>
--- a/src/FOL/ex/LocaleTest.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/FOL/ex/LocaleTest.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,16 +1,13 @@
 (*  Title:      FOL/ex/NewLocaleTest.thy
     Author:     Clemens Ballarin, TU Muenchen
 
-Testing environment for locale expressions --- experimental.
+Testing environment for locale expressions.
 *)
 
 theory LocaleTest
 imports FOL
 begin
 
-ML_val {* set Toplevel.debug *}
-
-
 typedecl int arities int :: "term"
 consts plus :: "int => int => int" (infixl "+" 60)
   zero :: int ("0")
--- a/src/HOL/ATP_Linkup.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/ATP_Linkup.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      HOL/ATP_Linkup.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson
     Author:     Jia Meng, NICTA
     Author:     Fabian Immler, TUM
@@ -8,7 +7,7 @@
 header {* The Isabelle-ATP Linkup *}
 
 theory ATP_Linkup
-imports Record Hilbert_Choice
+imports Divides Record Hilbert_Choice Plain
 uses
   "Tools/polyhash.ML"
   "Tools/res_clause.ML"
@@ -112,10 +111,13 @@
 setup {* AtpManager.add_prover "spass_no_tc" (AtpWrapper.spass_opts 40 false) *}
 
 text {* remote provers via SystemOnTPTP *}
-setup {* AtpManager.add_prover "remote_vamp9"
-  (AtpWrapper.remote_prover "Vampire---9.0" "jumpirefix --output_syntax tptp --mode casc -t 3600") *}
-setup {* AtpManager.add_prover "remote_vamp10"
-  (AtpWrapper.remote_prover "Vampire---10.0" "drakosha.pl 60") *}
+setup {* AtpManager.add_prover "remote_vampire"
+  (AtpWrapper.remote_prover "-s Vampire---9.0") *}
+setup {* AtpManager.add_prover "remote_spass"
+  (AtpWrapper.remote_prover "-s SPASS---3.01") *}
+setup {* AtpManager.add_prover "remote_e"
+  (AtpWrapper.remote_prover "-s EP---1.0") *}
+  
 
 
 subsection {* The Metis prover *}
--- a/src/HOL/Algebra/abstract/Field.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Algebra/abstract/Field.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,6 +1,5 @@
 (*
     Properties of abstract class field
-    $Id$
     Author: Clemens Ballarin, started 15 November 1997
 *)
 
@@ -14,7 +13,6 @@
 
 instance field < factorial
   apply intro_classes
-   apply (rule TrueI)
   apply (erule field_fact_prime)
   done
 
--- a/src/HOL/Algebra/abstract/PID.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Algebra/abstract/PID.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,6 +1,5 @@
 (*
     Principle ideal domains
-    $Id$
     Author: Clemens Ballarin, started 5 October 1999
 *)
 
@@ -8,7 +7,6 @@
 
 instance pid < factorial
   apply intro_classes
-   apply (rule TrueI)
   apply (erule pid_irred_imp_prime)
   done
 
--- a/src/HOL/Algebra/abstract/Ring2.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Algebra/abstract/Ring2.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -59,8 +59,9 @@
   Proper definition using divisor chain condition currently not supported.
   factorial_divisor:    "wf {(a, b). a dvd b & ~ (b dvd a)}"
 *)
-  assumes factorial_divisor: "True"
-  and factorial_prime: "irred a ==> prime a"
+  (*assumes factorial_divisor: "True"*)
+  assumes factorial_prime: "irred a ==> prime a"
+
 
 subsection {* Euclidean domains *}
 
@@ -224,10 +225,6 @@
   {* Method.no_args (Method.SIMPLE_METHOD' (full_simp_tac ring_ss)) *}
   {* computes distributive normal form in rings *}
 
-lemmas ring_simps =
-  l_zero r_zero l_neg r_neg minus_minus minus0
-  l_one r_one l_null r_null l_minus r_minus
-
 
 subsection {* Rings and the summation operator *}
 
--- a/src/HOL/Algebra/poly/UnivPoly2.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Algebra/poly/UnivPoly2.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -189,7 +189,7 @@
 lemma coeff_smult [simp]: "coeff (a *s p) n = (a::'a::ring) * coeff p n"
 proof -
   have "!!f. f : UP ==> (%n. a * f n) : UP"
-    by (unfold UP_def) (force simp add: ring_simps)
+    by (unfold UP_def) (force simp add: algebra_simps)
 *)      (* this force step is slow *)
 (*  then show ?thesis
     apply (simp add: coeff_def smult_def Abs_UP_inverse Rep_UP)
@@ -198,7 +198,7 @@
 lemma coeff_smult [simp]: "coeff (a *s p) n = (a::'a::ring) * coeff p n"
 proof -
   have "Rep_UP p : UP ==> (%n. a * Rep_UP p n) : UP"
-    by (unfold UP_def) (force simp add: ring_simps)
+    by (unfold UP_def) (force simp add: algebra_simps)
       (* this force step is slow *)
   then show ?thesis
     by (simp add: coeff_def smult_def Abs_UP_inverse Rep_UP)
@@ -220,7 +220,7 @@
 	fix i
 	assume "max n m < i"
 	with boundn and boundm show "f i + g i = 0"
-          by (fastsimp simp add: ring_simps)
+          by (fastsimp simp add: algebra_simps)
       qed
       then show "(%i. (f i + g i)) : UP"
 	by (unfold UP_def) fast
@@ -251,15 +251,15 @@
 	  have "f i * g (k-i) = 0"
 	  proof cases
 	    assume "n < i"
-	    with `bound n f` show ?thesis by (auto simp add: ring_simps)
+	    with `bound n f` show ?thesis by (auto simp add: algebra_simps)
 	  next
 	    assume "~ (n < i)"
 	    with bound have "m < k-i" by arith
-	    with `bound m g` show ?thesis by (auto simp add: ring_simps)
+	    with `bound m g` show ?thesis by (auto simp add: algebra_simps)
 	  qed
 	}
 	then show "setsum (%i. f i * g (k-i)) {..k} = 0"
-	  by (simp add: ring_simps)
+	  by (simp add: algebra_simps)
       qed
       then show "(%n. setsum (%i. f i * g (n-i)) {..n}) : UP"
 	by (unfold UP_def) fast
@@ -270,7 +270,7 @@
 qed
 
 lemma coeff_uminus [simp]: "coeff (-p) n = (-coeff p n::'a::ring)"
-by (unfold up_uminus_def) (simp add: ring_simps)
+by (unfold up_uminus_def) (simp add: algebra_simps)
 
 (* Other lemmas *)
 
@@ -388,7 +388,7 @@
   proof (cases k)
     case 0 then show ?thesis by simp ring
   next
-    case Suc then show ?thesis by (simp add: ring_simps) ring
+    case Suc then show ?thesis by (simp add: algebra_simps) ring
   qed
   then show "coeff (monom a 0 * p) k = coeff (a *s p) k" by ring
 qed
--- a/src/HOL/Code_Eval.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Code_Eval.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Code_Eval.thy
-    ID:         $Id$
     Author:     Florian Haftmann, TU Muenchen
 *)
 
@@ -24,7 +23,7 @@
 code_datatype Const App
 
 class term_of = typerep +
-  fixes term_of :: "'a \<Rightarrow> term"
+  fixes term_of :: "'a::{} \<Rightarrow> term"
 
 lemma term_of_anything: "term_of x \<equiv> t"
   by (rule eq_reflection) (cases "term_of x", cases t, simp)
--- a/src/HOL/Complex.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Complex.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -237,9 +237,9 @@
   show "scaleR 1 x = x"
     by (simp add: expand_complex_eq)
   show "scaleR a x * y = scaleR a (x * y)"
-    by (simp add: expand_complex_eq ring_simps)
+    by (simp add: expand_complex_eq algebra_simps)
   show "x * scaleR a y = scaleR a (x * y)"
-    by (simp add: expand_complex_eq ring_simps)
+    by (simp add: expand_complex_eq algebra_simps)
 qed
 
 end
@@ -312,7 +312,7 @@
        (simp add: power_mult_distrib right_distrib [symmetric] real_sqrt_mult)
   show "norm (x * y) = norm x * norm y"
     by (induct x, induct y)
-       (simp add: real_sqrt_mult [symmetric] power2_eq_square ring_simps)
+       (simp add: real_sqrt_mult [symmetric] power2_eq_square algebra_simps)
   show "sgn x = x /\<^sub>R cmod x" by(simp add: complex_sgn_def)
 qed
 
--- a/src/HOL/Datatype.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Datatype.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Datatype.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Author:     Stefan Berghofer and Markus Wenzel, TU Muenchen
 
@@ -9,7 +8,7 @@
 header {* Analogues of the Cartesian Product and Disjoint Sum for Datatypes *}
 
 theory Datatype
-imports Nat Relation
+imports Nat Product_Type
 begin
 
 typedef (Node)
@@ -510,15 +509,6 @@
 lemmas dsum_subset_Sigma = subset_trans [OF dsum_mono dsum_Sigma, standard]
 
 
-(*** Domain ***)
-
-lemma Domain_dprod [simp]: "Domain (dprod r s) = uprod (Domain r) (Domain s)"
-by auto
-
-lemma Domain_dsum [simp]: "Domain (dsum r s) = usum (Domain r) (Domain s)"
-by auto
-
-
 text {* hides popular names *}
 hide (open) type node item
 hide (open) const Push Node Atom Leaf Numb Lim Split Case
--- a/src/HOL/Dense_Linear_Order.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Dense_Linear_Order.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,10 +1,12 @@
-(* Author: Amine Chaieb, TU Muenchen *)
+(*  Title       : HOL/Dense_Linear_Order.thy
+    Author      : Amine Chaieb, TU Muenchen
+*)
 
 header {* Dense linear order without endpoints
   and a quantifier elimination procedure in Ferrante and Rackoff style *}
 
 theory Dense_Linear_Order
-imports Plain Groebner_Basis
+imports Plain Groebner_Basis Main
 uses
   "~~/src/HOL/Tools/Qelim/langford_data.ML"
   "~~/src/HOL/Tools/Qelim/ferrante_rackoff_data.ML"
@@ -301,7 +303,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 +313,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 +362,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 +409,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)
@@ -552,7 +554,7 @@
 lemma neg_prod_lt:"(c\<Colon>'a\<Colon>ordered_field) < 0 \<Longrightarrow> ((c*x < 0) == (x > 0))"
 proof-
   assume H: "c < 0"
-  have "c*x < 0 = (0/c < x)" by (simp only: neg_divide_less_eq[OF H] ring_simps)
+  have "c*x < 0 = (0/c < x)" by (simp only: neg_divide_less_eq[OF H] algebra_simps)
   also have "\<dots> = (0 < x)" by simp
   finally show  "(c*x < 0) == (x > 0)" by simp
 qed
@@ -560,7 +562,7 @@
 lemma pos_prod_lt:"(c\<Colon>'a\<Colon>ordered_field) > 0 \<Longrightarrow> ((c*x < 0) == (x < 0))"
 proof-
   assume H: "c > 0"
-  hence "c*x < 0 = (0/c > x)" by (simp only: pos_less_divide_eq[OF H] ring_simps)
+  hence "c*x < 0 = (0/c > x)" by (simp only: pos_less_divide_eq[OF H] algebra_simps)
   also have "\<dots> = (0 > x)" by simp
   finally show  "(c*x < 0) == (x < 0)" by simp
 qed
@@ -569,7 +571,7 @@
 proof-
   assume H: "c < 0"
   have "c*x + t< 0 = (c*x < -t)" by (subst less_iff_diff_less_0 [of "c*x" "-t"], simp)
-  also have "\<dots> = (-t/c < x)" by (simp only: neg_divide_less_eq[OF H] ring_simps)
+  also have "\<dots> = (-t/c < x)" by (simp only: neg_divide_less_eq[OF H] algebra_simps)
   also have "\<dots> = ((- 1/c)*t < x)" by simp
   finally show  "(c*x + t < 0) == (x > (- 1/c)*t)" by simp
 qed
@@ -578,7 +580,7 @@
 proof-
   assume H: "c > 0"
   have "c*x + t< 0 = (c*x < -t)"  by (subst less_iff_diff_less_0 [of "c*x" "-t"], simp)
-  also have "\<dots> = (-t/c > x)" by (simp only: pos_less_divide_eq[OF H] ring_simps)
+  also have "\<dots> = (-t/c > x)" by (simp only: pos_less_divide_eq[OF H] algebra_simps)
   also have "\<dots> = ((- 1/c)*t > x)" by simp
   finally show  "(c*x + t < 0) == (x < (- 1/c)*t)" by simp
 qed
@@ -589,7 +591,7 @@
 lemma neg_prod_le:"(c\<Colon>'a\<Colon>ordered_field) < 0 \<Longrightarrow> ((c*x <= 0) == (x >= 0))"
 proof-
   assume H: "c < 0"
-  have "c*x <= 0 = (0/c <= x)" by (simp only: neg_divide_le_eq[OF H] ring_simps)
+  have "c*x <= 0 = (0/c <= x)" by (simp only: neg_divide_le_eq[OF H] algebra_simps)
   also have "\<dots> = (0 <= x)" by simp
   finally show  "(c*x <= 0) == (x >= 0)" by simp
 qed
@@ -597,7 +599,7 @@
 lemma pos_prod_le:"(c\<Colon>'a\<Colon>ordered_field) > 0 \<Longrightarrow> ((c*x <= 0) == (x <= 0))"
 proof-
   assume H: "c > 0"
-  hence "c*x <= 0 = (0/c >= x)" by (simp only: pos_le_divide_eq[OF H] ring_simps)
+  hence "c*x <= 0 = (0/c >= x)" by (simp only: pos_le_divide_eq[OF H] algebra_simps)
   also have "\<dots> = (0 >= x)" by simp
   finally show  "(c*x <= 0) == (x <= 0)" by simp
 qed
@@ -606,7 +608,7 @@
 proof-
   assume H: "c < 0"
   have "c*x + t <= 0 = (c*x <= -t)"  by (subst le_iff_diff_le_0 [of "c*x" "-t"], simp)
-  also have "\<dots> = (-t/c <= x)" by (simp only: neg_divide_le_eq[OF H] ring_simps)
+  also have "\<dots> = (-t/c <= x)" by (simp only: neg_divide_le_eq[OF H] algebra_simps)
   also have "\<dots> = ((- 1/c)*t <= x)" by simp
   finally show  "(c*x + t <= 0) == (x >= (- 1/c)*t)" by simp
 qed
@@ -615,7 +617,7 @@
 proof-
   assume H: "c > 0"
   have "c*x + t <= 0 = (c*x <= -t)" by (subst le_iff_diff_le_0 [of "c*x" "-t"], simp)
-  also have "\<dots> = (-t/c >= x)" by (simp only: pos_le_divide_eq[OF H] ring_simps)
+  also have "\<dots> = (-t/c >= x)" by (simp only: pos_le_divide_eq[OF H] algebra_simps)
   also have "\<dots> = ((- 1/c)*t >= x)" by simp
   finally show  "(c*x + t <= 0) == (x <= (- 1/c)*t)" by simp
 qed
@@ -628,16 +630,16 @@
 proof-
   assume H: "c \<noteq> 0"
   have "c*x + t = 0 = (c*x = -t)" by (subst eq_iff_diff_eq_0 [of "c*x" "-t"], simp)
-  also have "\<dots> = (x = -t/c)" by (simp only: nonzero_eq_divide_eq[OF H] ring_simps)
+  also have "\<dots> = (x = -t/c)" by (simp only: nonzero_eq_divide_eq[OF H] algebra_simps)
   finally show  "(c*x + t = 0) == (x = (- 1/c)*t)" by simp
 qed
 lemma sum_eq:"((x::'a::pordered_ab_group_add) + t = 0) == (x = - t)"
   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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Deriv.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -141,7 +141,7 @@
 lemma inverse_diff_inverse:
   "\<lbrakk>(a::'a::division_ring) \<noteq> 0; b \<noteq> 0\<rbrakk>
    \<Longrightarrow> inverse a - inverse b = - (inverse a * (a - b) * inverse b)"
-by (simp add: ring_simps)
+by (simp add: algebra_simps)
 
 lemma DERIV_inverse_lemma:
   "\<lbrakk>a \<noteq> 0; b \<noteq> (0::'a::real_normed_field)\<rbrakk>
@@ -206,7 +206,7 @@
 case (Suc k)
   from DERIV_mult' [OF f Suc] show ?case
     apply (simp only: of_nat_Suc ring_distribs mult_1_left)
-    apply (simp only: power_Suc right_distrib mult_ac add_ac)
+    apply (simp only: power_Suc algebra_simps)
     done
 qed
 
@@ -724,7 +724,7 @@
       from isCont_bounded [OF le this]
       obtain k where k: "!!x. a \<le> x & x \<le> b --> inverse (M - f x) \<le> k" by auto
       have Minv: "!!x. a \<le> x & x \<le> b --> 0 < inverse (M - f (x))"
-        by (simp add: M3 compare_rls)
+        by (simp add: M3 algebra_simps)
       have "!!x. a \<le> x & x \<le> b --> inverse (M - f x) < k+1" using k
         by (auto intro: order_le_less_trans [of _ k])
       with Minv
@@ -1398,7 +1398,7 @@
     have "?h b - ?h a =
          ((f b)*(g b) - (f a)*(g b) - (g b)*(f b) + (g a)*(f b)) -
           ((f b)*(g a) - (f a)*(g a) - (g b)*(f a) + (g a)*(f a))"
-      by (simp add: mult_ac add_ac right_diff_distrib)
+      by (simp add: algebra_simps)
     hence "?h b - ?h a = 0" by auto
   }
   ultimately have "(b - a) * (g'c * (f b - f a) - f'c * (g b - g a)) = 0" by auto
@@ -1427,7 +1427,7 @@
 
 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)
+  apply (simp add: pderiv_pCons coeff_pCons algebra_simps split: nat.split)
   done
 
 lemma pderiv_eq_0_iff: "pderiv p = 0 \<longleftrightarrow> degree p = 0"
@@ -1451,21 +1451,21 @@
 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)
+by (rule poly_ext, simp add: coeff_pderiv algebra_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)
+by (rule poly_ext, simp add: coeff_pderiv algebra_simps)
 
 lemma pderiv_smult: "pderiv (smult a p) = smult a (pderiv p)"
-by (rule poly_ext, simp add: coeff_pderiv ring_simps)
+by (rule poly_ext, simp add: coeff_pderiv algebra_simps)
 
 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)
+apply (simp add: pderiv_add pderiv_smult pderiv_pCons algebra_simps)
 done
 
 lemma pderiv_power_Suc:
@@ -1475,7 +1475,7 @@
 apply (subst power_Suc)
 apply (subst pderiv_mult)
 apply (erule ssubst)
-apply (simp add: smult_add_left ring_simps)
+apply (simp add: smult_add_left algebra_simps)
 done
 
 lemma DERIV_cmult2: "DERIV f x :> D ==> DERIV (%x. (f x) * c :: real) x :> D * c"
--- a/src/HOL/Divides.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Divides.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -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"
@@ -38,10 +38,10 @@
   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)
+by (simp add: mod_div_equality)
 
 lemma div_mod_equality2: "(b * (a div b) + a mod b) + c = a + c"
-  by (simp add: mod_div_equality2)
+by (simp add: mod_div_equality2)
 
 lemma mod_by_0 [simp]: "a mod 0 = a"
   using mod_div_equality [of a zero] by simp
@@ -63,7 +63,7 @@
     by (simp add: mod_div_equality)
   also from False div_mult_self1 [of b a c] have
     "\<dots> = (c + a div b) * b + (a + c * b) mod b"
-      by (simp add: left_distrib add_ac)
+      by (simp add: algebra_simps)
   finally have "a = a div b * b + (a + c * b) mod b"
     by (simp add: add_commute [of a] add_assoc left_distrib)
   then have "a div b * b + (a + c * b) mod b = a div b * b + a mod b"
@@ -72,7 +72,7 @@
 qed
 
 lemma mod_mult_self2 [simp]: "(a + b * c) mod b = a mod b"
-  by (simp add: mult_commute [of b])
+by (simp add: mult_commute [of b])
 
 lemma div_mult_self1_is_id [simp]: "b \<noteq> 0 \<Longrightarrow> b * a div b = a"
   using div_mult_self2 [of b 0 a] by simp
@@ -217,7 +217,7 @@
   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)
+    by (simp only: algebra_simps)
   also have "\<dots> = (a mod c * b) mod c"
     by (rule mod_mult_self1)
   finally show ?thesis .
@@ -228,7 +228,7 @@
   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)
+    by (simp only: algebra_simps)
   also have "\<dots> = (a * (b mod c)) mod c"
     by (rule mod_mult_self1)
   finally show ?thesis .
@@ -552,7 +552,7 @@
 
 lemma divmod_if [code]: "divmod m n = (if n = 0 \<or> m < n then (0, m) else
   let (q, r) = divmod (m - n) n in (Suc q, r))"
-  by (simp add: divmod_zero divmod_base divmod_step)
+by (simp add: divmod_zero divmod_base divmod_step)
     (simp add: divmod_div_mod)
 
 code_modulename SML
@@ -568,16 +568,16 @@
 subsubsection {* Quotient *}
 
 lemma div_geq: "0 < n \<Longrightarrow>  \<not> m < n \<Longrightarrow> m div n = Suc ((m - n) div n)"
-  by (simp add: le_div_geq linorder_not_less)
+by (simp add: le_div_geq linorder_not_less)
 
 lemma div_if: "0 < n \<Longrightarrow> m div n = (if m < n then 0 else Suc ((m - n) div n))"
-  by (simp add: div_geq)
+by (simp add: div_geq)
 
 lemma div_mult_self_is_m [simp]: "0<n ==> (m*n) div n = (m::nat)"
-  by simp
+by simp
 
 lemma div_mult_self1_is_m [simp]: "0<n ==> (n*m) div n = (m::nat)"
-  by simp
+by simp
 
 
 subsubsection {* Remainder *}
@@ -597,13 +597,13 @@
 qed
 
 lemma mod_geq: "\<not> m < (n\<Colon>nat) \<Longrightarrow> m mod n = (m - n) mod n"
-  by (simp add: le_mod_geq linorder_not_less)
+by (simp add: le_mod_geq linorder_not_less)
 
 lemma mod_if: "m mod (n\<Colon>nat) = (if m < n then m else (m - n) mod n)"
-  by (simp add: le_mod_geq)
+by (simp add: le_mod_geq)
 
 lemma mod_1 [simp]: "m mod Suc 0 = 0"
-  by (induct m) (simp_all add: mod_geq)
+by (induct m) (simp_all add: mod_geq)
 
 lemma mod_mult_distrib: "(m mod n) * (k\<Colon>nat) = (m * k) mod (n * k)"
   apply (cases "n = 0", simp)
@@ -614,11 +614,11 @@
   done
 
 lemma mod_mult_distrib2: "(k::nat) * (m mod n) = (k*m) mod (k*n)"
-  by (simp add: mult_commute [of k] mod_mult_distrib)
+by (simp add: mult_commute [of k] mod_mult_distrib)
 
 (* a simple rearrangement of mod_div_equality: *)
 lemma mult_div_cancel: "(n::nat) * (m div n) = m - (m mod n)"
-  by (cut_tac a = m and b = n in mod_div_equality2, arith)
+by (cut_tac a = m and b = n in mod_div_equality2, arith)
 
 lemma mod_le_divisor[simp]: "0 < n \<Longrightarrow> m mod n \<le> (n::nat)"
   apply (drule mod_less_divisor [where m = m])
@@ -630,7 +630,7 @@
 lemma divmod_rel_mult1_eq:
   "[| divmod_rel b c q r; c > 0 |]
    ==> divmod_rel (a*b) c (a*q + a*r div c) (a*r mod c)"
-by (auto simp add: split_ifs mult_ac divmod_rel_def add_mult_distrib2)
+by (auto simp add: split_ifs divmod_rel_def algebra_simps)
 
 lemma div_mult1_eq: "(a*b) div c = a*(b div c) + a*(b mod c) div (c::nat)"
 apply (cases "c = 0", simp)
@@ -638,19 +638,19 @@
 done
 
 lemma mod_mult1_eq: "(a*b) mod c = a*(b mod c) mod (c::nat)"
-  by (rule mod_mult_right_eq)
+by (rule mod_mult_right_eq)
 
 lemma mod_mult1_eq': "(a*b) mod (c::nat) = ((a mod c) * b) mod c"
-  by (rule mod_mult_left_eq)
+by (rule mod_mult_left_eq)
 
 lemma mod_mult_distrib_mod:
   "(a*b) mod (c::nat) = ((a mod c) * (b mod c)) mod c"
-  by (rule mod_mult_eq)
+by (rule mod_mult_eq)
 
 lemma divmod_rel_add1_eq:
   "[| divmod_rel a c aq ar; divmod_rel b c bq br;  c > 0 |]
    ==> divmod_rel (a + b) c (aq + bq + (ar+br) div c) ((ar + br) mod c)"
-by (auto simp add: split_ifs mult_ac divmod_rel_def add_mult_distrib2)
+by (auto simp add: split_ifs divmod_rel_def algebra_simps)
 
 (*NOT suitable for rewriting: the RHS has an instance of the LHS*)
 lemma div_add1_eq:
@@ -660,7 +660,7 @@
 done
 
 lemma mod_add1_eq: "(a+b) mod (c::nat) = (a mod c + b mod c) mod c"
-  by (rule mod_add_eq)
+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)
@@ -671,7 +671,7 @@
 
 lemma divmod_rel_mult2_eq: "[| divmod_rel a b q r;  0 < b;  0 < c |]
       ==> divmod_rel a (b*c) (q div c) (b*(q mod c) + r)"
-  by (auto simp add: mult_ac divmod_rel_def add_mult_distrib2 [symmetric] mod_lemma)
+by (auto simp add: mult_ac divmod_rel_def add_mult_distrib2 [symmetric] mod_lemma)
 
 lemma div_mult2_eq: "a div (b*c) = (a div b) div (c::nat)"
   apply (cases "b = 0", simp)
@@ -690,7 +690,7 @@
 
 lemma div_mult_mult_lemma:
     "[| (0::nat) < b;  0 < c |] ==> (c*a) div (c*b) = a div b"
-  by (auto simp add: div_mult2_eq)
+by (auto simp add: div_mult2_eq)
 
 lemma div_mult_mult1 [simp]: "(0::nat) < c ==> (c*a) div (c*b) = a div b"
   apply (cases "b = 0")
@@ -706,7 +706,7 @@
 subsubsection{*Further Facts about Quotient and Remainder*}
 
 lemma div_1 [simp]: "m div Suc 0 = m"
-  by (induct m) (simp_all add: div_geq)
+by (induct m) (simp_all add: div_geq)
 
 
 (* Monotonicity of div in first argument *)
@@ -780,10 +780,10 @@
 done
 
 lemma nat_mod_div_trivial: "m mod n div n = (0 :: nat)"
-  by simp
+by simp
 
 lemma nat_mod_mod_trivial: "m mod n mod n = (m mod n :: nat)"
-  by simp
+by simp
 
 
 subsubsection {* The Divides Relation *}
@@ -792,7 +792,7 @@
   unfolding dvd_def by simp
 
 lemma dvd_1_iff_1 [simp]: "(m dvd Suc 0) = (m = Suc 0)"
-  by (simp add: dvd_def)
+by (simp add: dvd_def)
 
 lemma dvd_anti_sym: "[| m dvd n; n dvd m |] ==> m = (n::nat)"
   unfolding dvd_def
@@ -800,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)"
@@ -813,7 +813,7 @@
   done
 
 lemma dvd_diffD1: "[| k dvd m-n; k dvd m; n\<le>m |] ==> k dvd (n::nat)"
-  by (drule_tac m = m in dvd_diff, auto)
+by (drule_tac m = m in dvd_diff, auto)
 
 lemma dvd_reduce: "(k dvd n + k) = (k dvd (n::nat))"
   apply (rule iffI)
@@ -839,7 +839,7 @@
   done
 
 lemma dvd_mod_iff: "k dvd n ==> ((k::nat) dvd m mod n) = (k dvd m)"
-  by (blast intro: dvd_mod_imp_dvd dvd_mod)
+by (blast intro: dvd_mod_imp_dvd dvd_mod)
 
 lemma dvd_mult_cancel: "!!k::nat. [| k*m dvd k*n; 0<k |] ==> m dvd n"
   unfolding dvd_def
@@ -894,7 +894,7 @@
   done
 
 lemma mod_eq_0_iff: "(m mod d = 0) = (\<exists>q::nat. m = d*q)"
-  by (auto simp add: dvd_eq_mod_eq_0 [symmetric] dvd_def)
+by (auto simp add: dvd_eq_mod_eq_0 [symmetric] dvd_def)
 
 lemmas mod_eq_0D [dest!] = mod_eq_0_iff [THEN iffD1]
 
--- a/src/HOL/Equiv_Relations.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Equiv_Relations.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,12 +1,11 @@
-(*  ID:         $Id$
-    Authors:    Lawrence C Paulson, Cambridge University Computer Laboratory
+(*  Authors:    Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1996  University of Cambridge
 *)
 
 header {* Equivalence Relations in Higher-Order Set Theory *}
 
 theory Equiv_Relations
-imports Finite_Set Relation
+imports Finite_Set Relation Plain
 begin
 
 subsection {* Equivalence relations *}
--- a/src/HOL/Finite_Set.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Finite_Set.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Finite_Set.thy
-    ID:         $Id$
     Author:     Tobias Nipkow, Lawrence C Paulson and Markus Wenzel
                 with contributions by Jeremy Avigad
 *)
@@ -7,7 +6,7 @@
 header {* Finite sets *}
 
 theory Finite_Set
-imports Datatype Divides Transitive_Closure
+imports Nat Product_Type Power
 begin
 
 subsection {* Definition and basic properties *}
@@ -381,46 +380,6 @@
 by(blast intro: finite_subset[OF subset_Pow_Union])
 
 
-lemma finite_converse [iff]: "finite (r^-1) = finite r"
-  apply (subgoal_tac "r^-1 = (%(x,y). (y,x))`r")
-   apply simp
-   apply (rule iffI)
-    apply (erule finite_imageD [unfolded inj_on_def])
-    apply (simp split add: split_split)
-   apply (erule finite_imageI)
-  apply (simp add: converse_def image_def, auto)
-  apply (rule bexI)
-   prefer 2 apply assumption
-  apply simp
-  done
-
-
-text {* \paragraph{Finiteness of transitive closure} (Thanks to Sidi
-Ehmety) *}
-
-lemma finite_Field: "finite r ==> finite (Field r)"
-  -- {* A finite relation has a finite field (@{text "= domain \<union> range"}. *}
-  apply (induct set: finite)
-   apply (auto simp add: Field_def Domain_insert Range_insert)
-  done
-
-lemma trancl_subset_Field2: "r^+ <= Field r \<times> Field r"
-  apply clarify
-  apply (erule trancl_induct)
-   apply (auto simp add: Field_def)
-  done
-
-lemma finite_trancl: "finite (r^+) = finite r"
-  apply auto
-   prefer 2
-   apply (rule trancl_subset_Field2 [THEN finite_subset])
-   apply (rule finite_SigmaI)
-    prefer 3
-    apply (blast intro: r_into_trancl' finite_subset)
-   apply (auto simp add: finite_Field)
-  done
-
-
 subsection {* Class @{text finite}  *}
 
 setup {* Sign.add_path "finite" *} -- {*FIXME: name tweaking*}
@@ -472,9 +431,6 @@
 instance "+" :: (finite, finite) finite
   by default (simp only: UNIV_Plus_UNIV [symmetric] finite_Plus finite)
 
-instance option :: (finite) finite
-  by default (simp add: insert_None_conv_UNIV [symmetric])
-
 
 subsection {* A fold functional for finite sets *}
 
@@ -873,7 +829,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"
@@ -941,6 +897,46 @@
      "inj_on f B ==> setsum f B = setsum id (f ` B)"
 by (auto simp add: setsum_reindex)
 
+lemma setsum_reindex_nonzero: 
+  assumes fS: "finite S"
+  and nz: "\<And> x y. x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> x \<noteq> y \<Longrightarrow> f x = f y \<Longrightarrow> h (f x) = 0"
+  shows "setsum h (f ` S) = setsum (h o f) S"
+using nz
+proof(induct rule: finite_induct[OF fS])
+  case 1 thus ?case by simp
+next
+  case (2 x F) 
+  {assume fxF: "f x \<in> f ` F" hence "\<exists>y \<in> F . f y = f x" by auto
+    then obtain y where y: "y \<in> F" "f x = f y" by auto 
+    from "2.hyps" y have xy: "x \<noteq> y" by auto
+    
+    from "2.prems"[of x y] "2.hyps" xy y have h0: "h (f x) = 0" by simp
+    have "setsum h (f ` insert x F) = setsum h (f ` F)" using fxF by auto
+    also have "\<dots> = setsum (h o f) (insert x F)" 
+      unfolding setsum_insert[OF `finite F` `x\<notin>F`]
+      using h0 
+      apply simp
+      apply (rule "2.hyps"(3))
+      apply (rule_tac y="y" in  "2.prems")
+      apply simp_all
+      done
+    finally have ?case .}
+  moreover
+  {assume fxF: "f x \<notin> f ` F"
+    have "setsum h (f ` insert x F) = h (f x) + setsum h (f ` F)" 
+      using fxF "2.hyps" by simp 
+    also have "\<dots> = setsum (h o f) (insert x F)"
+      unfolding setsum_insert[OF `finite F` `x\<notin>F`]
+      apply simp
+      apply (rule cong[OF refl[of "op + (h (f x))"]])
+      apply (rule "2.hyps"(3))
+      apply (rule_tac y="y" in  "2.prems")
+      apply simp_all
+      done
+    finally have ?case .}
+  ultimately show ?case by blast
+qed
+
 lemma setsum_cong:
   "A = B ==> (!!x. x:B ==> f x = g x) ==> setsum f A = setsum g B"
 by(fastsimp simp: setsum_def intro: comm_monoid_add.fold_image_cong)
@@ -958,6 +954,7 @@
     ==> setsum h B = setsum g A"
 by (simp add: setsum_reindex cong: setsum_cong)
 
+
 lemma setsum_0[simp]: "setsum (%i. 0) A = 0"
 apply (clarsimp simp: setsum_def)
 apply (erule finite_induct, auto)
@@ -975,6 +972,73 @@
   ==> A Int B = {} ==> setsum g (A Un B) = setsum g A + setsum g B"
 by (subst setsum_Un_Int [symmetric], auto)
 
+lemma setsum_mono_zero_left: 
+  assumes fT: "finite T" and ST: "S \<subseteq> T"
+  and z: "\<forall>i \<in> T - S. f i = 0"
+  shows "setsum f S = setsum f T"
+proof-
+  have eq: "T = S \<union> (T - S)" using ST by blast
+  have d: "S \<inter> (T - S) = {}" using ST by blast
+  from fT ST have f: "finite S" "finite (T - S)" by (auto intro: finite_subset)
+  show ?thesis 
+  by (simp add: setsum_Un_disjoint[OF f d, unfolded eq[symmetric]] setsum_0'[OF z])
+qed
+
+lemma setsum_mono_zero_right: 
+  assumes fT: "finite T" and ST: "S \<subseteq> T"
+  and z: "\<forall>i \<in> T - S. f i = 0"
+  shows "setsum f T = setsum f S"
+using setsum_mono_zero_left[OF fT ST z] by simp
+
+lemma setsum_mono_zero_cong_left: 
+  assumes fT: "finite T" and ST: "S \<subseteq> T"
+  and z: "\<forall>i \<in> T - S. g i = 0"
+  and fg: "\<And>x. x \<in> S \<Longrightarrow> f x = g x"
+  shows "setsum f S = setsum g T"
+proof-
+  have eq: "T = S \<union> (T - S)" using ST by blast
+  have d: "S \<inter> (T - S) = {}" using ST by blast
+  from fT ST have f: "finite S" "finite (T - S)" by (auto intro: finite_subset)
+  show ?thesis 
+    using fg by (simp add: setsum_Un_disjoint[OF f d, unfolded eq[symmetric]] setsum_0'[OF z])
+qed
+
+lemma setsum_mono_zero_cong_right: 
+  assumes fT: "finite T" and ST: "S \<subseteq> T"
+  and z: "\<forall>i \<in> T - S. f i = 0"
+  and fg: "\<And>x. x \<in> S \<Longrightarrow> f x = g x"
+  shows "setsum f T = setsum g S"
+using setsum_mono_zero_cong_left[OF fT ST z] fg[symmetric] by auto 
+
+lemma setsum_delta: 
+  assumes fS: "finite S"
+  shows "setsum (\<lambda>k. if k=a then b k else 0) S = (if a \<in> S then b a else 0)"
+proof-
+  let ?f = "(\<lambda>k. if k=a then b k else 0)"
+  {assume a: "a \<notin> S"
+    hence "\<forall> k\<in> S. ?f k = 0" by simp
+    hence ?thesis  using a by simp}
+  moreover 
+  {assume a: "a \<in> S"
+    let ?A = "S - {a}"
+    let ?B = "{a}"
+    have eq: "S = ?A \<union> ?B" using a by blast 
+    have dj: "?A \<inter> ?B = {}" by simp
+    from fS have fAB: "finite ?A" "finite ?B" by auto  
+    have "setsum ?f S = setsum ?f ?A + setsum ?f ?B"
+      using setsum_Un_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]]
+      by simp
+    then have ?thesis  using a by simp}
+  ultimately show ?thesis by blast
+qed
+lemma setsum_delta': 
+  assumes fS: "finite S" shows 
+  "setsum (\<lambda>k. if a = k then b k else 0) S = 
+     (if a\<in> S then b a else 0)"
+  using setsum_delta[OF fS, of a b, symmetric] 
+  by (auto intro: setsum_cong)
+
+
 (*But we can't get rid of finite I. If infinite, although the rhs is 0, 
   the lhs need not be, since UNION I A could still be finite.*)
 lemma setsum_UN_disjoint:
@@ -1033,12 +1097,12 @@
 lemma setsum_Un_nat: "finite A ==> finite B ==>
   (setsum f (A Un B) :: nat) = setsum f A + setsum f B - setsum f (A Int B)"
   -- {* For the natural numbers, we have subtraction. *}
-by (subst setsum_Un_Int [symmetric], auto simp add: ring_simps)
+by (subst setsum_Un_Int [symmetric], auto simp add: algebra_simps)
 
 lemma setsum_Un: "finite A ==> finite B ==>
   (setsum f (A Un B) :: 'a :: ab_group_add) =
    setsum f A + setsum f B - setsum f (A Int B)"
-by (subst setsum_Un_Int [symmetric], auto simp add: ring_simps)
+by (subst setsum_Un_Int [symmetric], auto simp add: algebra_simps)
 
 lemma setsum_diff1_nat: "(setsum f (A - {a}) :: nat) =
   (if a:A then setsum f A - f a else setsum f A)"
@@ -1409,6 +1473,18 @@
     B = f ` A ==> g = h \<circ> f ==> setprod h B = setprod g A"
 by (frule setprod_reindex, simp)
 
+lemma strong_setprod_reindex_cong: assumes i: "inj_on f A"
+  and B: "B = f ` A" and eq: "\<And>x. x \<in> A \<Longrightarrow> g x = (h \<circ> f) x"
+  shows "setprod h B = setprod g A"
+proof-
+    have "setprod h B = setprod (h o f) A"
+      by (simp add: B setprod_reindex[OF i, of h])
+    then show ?thesis apply simp
+      apply (rule setprod_cong)
+      apply simp
+      by (erule eq[symmetric])
+qed
+
 
 lemma setprod_1: "setprod (%i. 1) A = 1"
 apply (case_tac "finite A")
@@ -1429,6 +1505,37 @@
   ==> A Int B = {} ==> setprod g (A Un B) = setprod g A * setprod g B"
 by (subst setprod_Un_Int [symmetric], auto)
 
+lemma setprod_delta: 
+  assumes fS: "finite S"
+  shows "setprod (\<lambda>k. if k=a then b k else 1) S = (if a \<in> S then b a else 1)"
+proof-
+  let ?f = "(\<lambda>k. if k=a then b k else 1)"
+  {assume a: "a \<notin> S"
+    hence "\<forall> k\<in> S. ?f k = 1" by simp
+    hence ?thesis  using a by (simp add: setprod_1 cong add: setprod_cong) }
+  moreover 
+  {assume a: "a \<in> S"
+    let ?A = "S - {a}"
+    let ?B = "{a}"
+    have eq: "S = ?A \<union> ?B" using a by blast 
+    have dj: "?A \<inter> ?B = {}" by simp
+    from fS have fAB: "finite ?A" "finite ?B" by auto  
+    have fA1: "setprod ?f ?A = 1" apply (rule setprod_1') by auto
+    have "setprod ?f ?A * setprod ?f ?B = setprod ?f S"
+      using setprod_Un_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]]
+      by simp
+    then have ?thesis  using a by (simp add: fA1 cong add: setprod_cong cong del: if_weak_cong)}
+  ultimately show ?thesis by blast
+qed
+
+lemma setprod_delta': 
+  assumes fS: "finite S" shows 
+  "setprod (\<lambda>k. if a = k then b k else 1) S = 
+     (if a\<in> S then b a else 1)"
+  using setprod_delta[OF fS, of a b, symmetric] 
+  by (auto intro: setprod_cong)
+
+
 lemma setprod_UN_disjoint:
     "finite I ==> (ALL i:I. finite (A i)) ==>
         (ALL i:I. ALL j:I. i \<noteq> j --> A i Int A j = {}) ==>
@@ -1711,7 +1818,7 @@
 lemma setsum_constant [simp]: "(\<Sum>x \<in> A. y) = of_nat(card A) * y"
 apply (cases "finite A")
 apply (erule finite_induct)
-apply (auto simp add: ring_simps)
+apply (auto simp add: algebra_simps)
 done
 
 lemma setprod_constant: "finite A ==> (\<Prod>x\<in> A. (y::'a::{recpower, comm_monoid_mult})) = y^(card A)"
@@ -1719,6 +1826,34 @@
 apply (auto simp add: power_Suc)
 done
 
+lemma setprod_gen_delta:
+  assumes fS: "finite S"
+  shows "setprod (\<lambda>k. if k=a then b k else c) S = (if a \<in> S then (b a ::'a::{comm_monoid_mult, recpower}) * c^ (card S - 1) else c^ card S)"
+proof-
+  let ?f = "(\<lambda>k. if k=a then b k else c)"
+  {assume a: "a \<notin> S"
+    hence "\<forall> k\<in> S. ?f k = c" by simp
+    hence ?thesis  using a setprod_constant[OF fS, of c] by (simp add: setprod_1 cong add: setprod_cong) }
+  moreover 
+  {assume a: "a \<in> S"
+    let ?A = "S - {a}"
+    let ?B = "{a}"
+    have eq: "S = ?A \<union> ?B" using a by blast 
+    have dj: "?A \<inter> ?B = {}" by simp
+    from fS have fAB: "finite ?A" "finite ?B" by auto  
+    have fA0:"setprod ?f ?A = setprod (\<lambda>i. c) ?A"
+      apply (rule setprod_cong) by auto
+    have cA: "card ?A = card S - 1" using fS a by auto
+    have fA1: "setprod ?f ?A = c ^ card ?A"  unfolding fA0 apply (rule setprod_constant) using fS by auto
+    have "setprod ?f ?A * setprod ?f ?B = setprod ?f S"
+      using setprod_Un_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]]
+      by simp
+    then have ?thesis using a cA
+      by (simp add: fA1 ring_simps cong add: setprod_cong cong del: if_weak_cong)}
+  ultimately show ?thesis by blast
+qed
+
+
 lemma setsum_bounded:
   assumes le: "\<And>i. i\<in>A \<Longrightarrow> f i \<le> (K::'a::{semiring_1, pordered_ab_semigroup_add})"
   shows "setsum f A \<le> of_nat(card A) * K"
@@ -1760,7 +1895,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 +2333,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,23 +2348,23 @@
   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
     thus ?case
     proof
       assume "a = x" thus ?thesis using insert
-        by (simp add: mult_ac_idem)
+        by (simp add: mult_ac)
     next
       assume "a \<in> F"
       hence bel: "fold1 inf F \<le> a" by (rule insert)
       have "inf (fold1 inf (insert x F)) a = inf x (inf (fold1 inf F) a)"
-        using insert by (simp add: mult_ac_idem)
+        using insert by (simp add: mult_ac)
       also have "inf (fold1 inf F) a = fold1 inf F"
         using bel by (auto intro: antisym)
       also have "inf x \<dots> = fold1 inf (insert x F)"
-        using insert by (simp add: mult_ac_idem)
+        using insert by (simp add: mult_ac)
       finally have aux: "inf (fold1 inf (insert x F)) a = fold1 inf (insert x F)" .
       moreover have "inf (fold1 inf (insert x F)) a \<le> a" by simp
       ultimately show ?thesis by simp
@@ -2288,7 +2423,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 +2438,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 +2468,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 +2492,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 +2521,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 +2532,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 +2581,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 +2592,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 +2603,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 +2616,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 +2650,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 +2659,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 +2668,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 +2677,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 +2686,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 +2696,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 +2707,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 +2718,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 +2728,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 +2737,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 +2746,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 +2755,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 +2764,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 +2774,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 +2784,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 +2795,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 +2805,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 +2816,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 +2826,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 +2836,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/FunDef.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/FunDef.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      HOL/FunDef.thy
-    ID:         $Id$
     Author:     Alexander Krauss, TU Muenchen
 *)
 
--- a/src/HOL/Fundamental_Theorem_Algebra.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Fundamental_Theorem_Algebra.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -45,7 +45,7 @@
     have th2: "2 *(y * sqrt ((sqrt (x * x + y * y) - x) * (sqrt (x * x + y * y) + x) / 4)) / \<bar>y\<bar> = y"
       using iffD2[OF real_sqrt_pow2_iff sum_power2_ge_zero[of x y]] y0
       unfolding power2_eq_square 
-      by (simp add: ring_simps real_sqrt_divide sqrt4)
+      by (simp add: algebra_simps real_sqrt_divide sqrt4)
      from y0 xy have ?thesis  apply (simp add: csqrt_def power2_eq_square)
        apply (simp add: real_sqrt_sum_squares_mult_ge_zero[of x y] real_sqrt_pow2[OF th(1)[of x y], unfolded power2_eq_square] real_sqrt_pow2[OF th(2)[of x y], unfolded power2_eq_square] real_sqrt_mult[symmetric])
       using th1 th2  ..}
@@ -109,7 +109,7 @@
 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)
+apply (simp add: offset_poly_pCons algebra_simps)
 done
 
 lemma offset_poly_eq_0_lemma: "smult c p + pCons a p = 0 \<Longrightarrow> p = 0"
@@ -135,15 +135,15 @@
 done
 
 definition
-  "plength p = (if p = 0 then 0 else Suc (degree p))"
+  "psize p = (if p = 0 then 0 else Suc (degree p))"
 
-lemma plength_eq_0_iff [simp]: "plength p = 0 \<longleftrightarrow> p = 0"
-  unfolding plength_def by simp
+lemma psize_eq_0_iff [simp]: "psize p = 0 \<longleftrightarrow> p = 0"
+  unfolding psize_def by simp
 
-lemma poly_offset: "\<exists> q. plength q = plength p \<and> (\<forall>x. poly q (x::complex) = poly p (a + x))"
+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 "plength (offset_poly p a) = plength p"
-    unfolding plength_def
+  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)
@@ -350,7 +350,7 @@
   from md z have xy: "x^2 + y^2 = 1" by (simp add: cmod_def)
   {assume C: "cmod (z + 1) \<ge> 1" "cmod (z - 1) \<ge> 1" "cmod (z + ii) \<ge> 1" "cmod (z - ii) \<ge> 1"
     from C z xy have "2*x \<le> 1" "2*x \<ge> -1" "2*y \<le> 1" "2*y \<ge> -1"
-      by (simp_all add: cmod_def power2_eq_square ring_simps)
+      by (simp_all add: cmod_def power2_eq_square algebra_simps)
     hence "abs (2*x) \<le> 1" "abs (2*y) \<le> 1" by simp_all
     hence "(abs (2 * x))^2 <= 1^2" "(abs (2 * y)) ^2 <= 1^2"
       by - (rule power_mono, simp, simp)+
@@ -391,9 +391,9 @@
     1" 
       apply (simp add: power2_eq_square norm_mult[symmetric] norm_inverse[symmetric])
       using right_inverse[OF b']
-      by (simp add: power2_eq_square[symmetric] power_inverse[symmetric] ring_simps)
+      by (simp add: power2_eq_square[symmetric] power_inverse[symmetric] algebra_simps)
     have th0: "cmod (complex_of_real (cmod b) / b) = 1"
-      apply (simp add: complex_Re_mult cmod_def power2_eq_square Re_complex_of_real Im_complex_of_real divide_inverse ring_simps )
+      apply (simp add: complex_Re_mult cmod_def power2_eq_square Re_complex_of_real Im_complex_of_real divide_inverse algebra_simps )
       by (simp add: real_sqrt_mult[symmetric] th0)        
     from o have "\<exists>m. n = Suc (2*m)" by presburger+
     then obtain m where m: "n = Suc (2*m)" by blast
@@ -667,10 +667,10 @@
       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 (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)
+	unfolding norm_mult by (simp add: algebra_simps)
       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) 
+	by (simp add: diff_le_eq algebra_simps) 
       from th1 th2 have "d \<le> cmod (poly (pCons a (pCons c cs)) z)"  by arith}
     hence ?case by blast}
   moreover
@@ -685,7 +685,7 @@
       have ath: "\<And>mzh mazh ma. mzh <= mazh + ma ==> abs(d) + ma <= mzh ==> d <= mazh" by arith
       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)
+	by (simp add: algebra_simps)
       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}
@@ -719,8 +719,8 @@
 text{* Constant function (non-syntactic characterization). *}
 definition "constant f = (\<forall>x y. f x = f y)"
 
-lemma nonconstant_length: "\<not> (constant (poly p)) \<Longrightarrow> plength p \<ge> 2"
-  unfolding constant_def plength_def
+lemma nonconstant_length: "\<not> (constant (poly p)) \<Longrightarrow> psize p \<ge> 2"
+  unfolding constant_def psize_def
   apply (induct p, auto)
   done
  
@@ -733,9 +733,9 @@
 
 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 (plength q + k) = plength p \<and> 
+  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 plength_def
+unfolding psize_def
 using nz
 proof(induct p)
   case 0 thus ?case by simp
@@ -761,7 +761,7 @@
 lemma poly_decompose:
   assumes nc: "~constant(poly p)"
   shows "\<exists>k a q. a\<noteq>(0::'a::{recpower,idom}) \<and> k\<noteq>0 \<and>
-               plength q + k + 1 = plength p \<and> 
+               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)
@@ -781,7 +781,7 @@
     apply simp
     apply (rule_tac x="q" in exI)
     apply (auto simp add: power_Suc)
-    apply (auto simp add: plength_def split: if_splits)
+    apply (auto simp add: psize_def split: if_splits)
     done
 qed
 
@@ -791,10 +791,10 @@
   assumes nc: "~constant(poly p)"
   shows "\<exists>z::complex. poly p z = 0"
 using nc
-proof(induct n\<equiv> "plength p" arbitrary: p rule: nat_less_induct)
+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 = plength p \<longrightarrow> (\<exists>(z::complex). poly p z = 0)" and nc: "\<not> constant ?p" and n: "n = plength 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)
@@ -804,7 +804,7 @@
   moreover
   {assume pc0: "?p c \<noteq> 0"
     from poly_offset[of p c] obtain q where
-      q: "plength q = plength 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
@@ -823,8 +823,8 @@
     have qr: "\<forall>z. poly q z = poly (smult (inverse ?a0) q) z * ?a0"
       by simp
     let ?r = "smult (inverse ?a0) q"
-    have lgqr: "plength q = plength ?r"
-      using a00 unfolding plength_def degree_def
+    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
@@ -844,13 +844,13 @@
       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" "plength s + k + 1 = plength ?r" 
+      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=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)}
+	  using kas(4)[rule_format, of w] s0 r01 by (simp add: algebra_simps)}
       note hth = this [symmetric]
 	from reduce_poly_simple[OF kas(1,2)] 
       have "\<exists>w. cmod (poly ?r w) < 1" unfolding hth by blast}
@@ -861,12 +861,12 @@
 	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(1) kas(2) have th02: "k+1 = plength (pCons 1 (monom a (k - 1)))"
-	by (simp add: plength_def degree_monom_eq)
+      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_pCons poly_monom
-	using kas(2) by (cases k, auto simp add: ring_simps)
+	using kas(2) by (cases k, auto simp add: algebra_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)
@@ -879,7 +879,7 @@
 	t: "t > 0" "t < 1" "t < inverse (cmod w ^ (k + 1) * m)" by blast
       let ?ct = "complex_of_real t"
       let ?w = "?ct * w"
-      have "1 + ?w^k * (a + ?w * poly s ?w) = 1 + ?ct^k * (w^k * a) + ?w^k * ?w * poly s ?w" using kas(1) by (simp add: ring_simps power_mult_distrib)
+      have "1 + ?w^k * (a + ?w * poly s ?w) = 1 + ?ct^k * (w^k * a) + ?w^k * ?w * poly s ?w" using kas(1) by (simp add: algebra_simps power_mult_distrib)
       also have "\<dots> = complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w"
 	unfolding wm1 by (simp)
       finally have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) = cmod (complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w)" 
@@ -898,7 +898,7 @@
       have th30: "t^k * (t* (cmod w ^ (k + 1) * m)) < t^k * 1" 
 	apply - apply (rule mult_strict_left_mono) by simp_all
       have "cmod (?w^k * ?w * poly s ?w) = t^k * (t* (cmod w ^ (k+1) * cmod (poly s ?w)))"  using w0 t(1)
-	by (simp add: ring_simps power_mult_distrib norm_of_real norm_power norm_mult)
+	by (simp add: algebra_simps power_mult_distrib norm_of_real norm_power norm_mult)
       then have "cmod (?w^k * ?w * poly s ?w) \<le> t^k * (t* (cmod w ^ (k + 1) * m))"
 	using t(1,2) m(2)[rule_format, OF tw] w0
 	apply (simp only: )
@@ -970,6 +970,7 @@
 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"
@@ -1307,14 +1308,14 @@
   {assume l: ?lhs
     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)
+       using u qrp' [symmetric] t by (simp add: algebra_simps mult_smult_right)
      then have ?rhs ..}
   moreover
   {assume r: ?rhs
     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)
+      by (simp add: algebra_simps mult_smult_right smult_smult)
     hence ?lhs ..}
   ultimately have "?lhs = ?rhs" by blast }
 thus "?lhs \<equiv> ?rhs"  by - (atomize(full), blast) 
@@ -1352,11 +1353,11 @@
 lemma basic_cqe_conv3:
   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 ^ (plength p)))"
+  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-
-  from l have dp:"degree (pCons a p) = plength p" by (simp add: plength_def)
+  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 ^ (plength p)))"
+  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
--- a/src/HOL/GCD.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/GCD.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      HOL/GCD.thy
-    ID:         $Id$
     Author:     Christophe Tabacznyj and Lawrence C Paulson
     Copyright   1996  University of Cambridge
 *)
@@ -7,7 +6,7 @@
 header {* The Greatest Common Divisor *}
 
 theory GCD
-imports Plain Presburger
+imports Plain Presburger Main
 begin
 
 text {*
--- a/src/HOL/Groebner_Basis.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Groebner_Basis.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -165,7 +165,7 @@
 
 interpretation class_semiring!: gb_semiring
     "op +" "op *" "op ^" "0::'a::{comm_semiring_1, recpower}" "1"
-  proof qed (auto simp add: ring_simps power_Suc)
+  proof qed (auto simp add: algebra_simps power_Suc)
 
 lemmas nat_arith =
   add_nat_number_of
@@ -345,13 +345,13 @@
 
 interpretation class_ringb!: ringb
   "op +" "op *" "op ^" "0::'a::{idom,recpower,number_ring}" "1" "op -" "uminus"
-proof(unfold_locales, simp add: ring_simps power_Suc, auto)
+proof(unfold_locales, simp add: algebra_simps power_Suc, auto)
   fix w x y z ::"'a::{idom,recpower,number_ring}"
   assume p: "w * y + x * z = w * z + x * y" and ynz: "y \<noteq> z"
   hence ynz': "y - z \<noteq> 0" by simp
   from p have "w * y + x* z - w*z - x*y = 0" by simp
-  hence "w* (y - z) - x * (y - z) = 0" by (simp add: ring_simps)
-  hence "(y - z) * (w - x) = 0" by (simp add: ring_simps)
+  hence "w* (y - z) - x * (y - z) = 0" by (simp add: algebra_simps)
+  hence "(y - z) * (w - x) = 0" by (simp add: algebra_simps)
   with  no_zero_divirors_neq0 [OF ynz']
   have "w - x = 0" by blast
   thus "w = x"  by simp
@@ -361,20 +361,20 @@
 
 interpretation natgb!: semiringb
   "op +" "op *" "op ^" "0::nat" "1"
-proof (unfold_locales, simp add: ring_simps power_Suc)
+proof (unfold_locales, simp add: algebra_simps power_Suc)
   fix w x y z ::"nat"
   { assume p: "w * y + x * z = w * z + x * y" and ynz: "y \<noteq> z"
     hence "y < z \<or> y > z" by arith
     moreover {
       assume lt:"y <z" hence "\<exists>k. z = y + k \<and> k > 0" by (rule_tac x="z - y" in exI, auto)
       then obtain k where kp: "k>0" and yz:"z = y + k" by blast
-      from p have "(w * y + x *y) + x*k = (w * y + x*y) + w*k" by (simp add: yz ring_simps)
+      from p have "(w * y + x *y) + x*k = (w * y + x*y) + w*k" by (simp add: yz algebra_simps)
       hence "x*k = w*k" by simp
       hence "w = x" using kp by (simp add: mult_cancel2) }
     moreover {
       assume lt: "y >z" hence "\<exists>k. y = z + k \<and> k>0" by (rule_tac x="y - z" in exI, auto)
       then obtain k where kp: "k>0" and yz:"y = z + k" by blast
-      from p have "(w * z + x *z) + w*k = (w * z + x*z) + x*k" by (simp add: yz ring_simps)
+      from p have "(w * z + x *z) + w*k = (w * z + x*z) + x*k" by (simp add: yz algebra_simps)
       hence "w*k = x*k" by simp
       hence "w = x" using kp by (simp add: mult_cancel2)}
     ultimately have "w=x" by blast }
--- a/src/HOL/HOL.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/HOL.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -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 *}
@@ -208,34 +208,34 @@
 
 subsubsection {* Generic classes and algebraic operations *}
 
-class default = type +
+class default =
   fixes default :: 'a
 
-class zero = type + 
+class zero = 
   fixes zero :: 'a  ("0")
 
-class one = type +
+class one =
   fixes one  :: 'a  ("1")
 
 hide (open) const zero one
 
-class plus = type +
+class plus =
   fixes plus :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixl "+" 65)
 
-class minus = type +
+class minus =
   fixes minus :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixl "-" 65)
 
-class uminus = type +
+class uminus =
   fixes uminus :: "'a \<Rightarrow> 'a"  ("- _" [81] 80)
 
-class times = type +
+class times =
   fixes times :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixl "*" 70)
 
-class inverse = type +
+class inverse =
   fixes inverse :: "'a \<Rightarrow> 'a"
     and divide :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixl "'/" 70)
 
-class abs = type +
+class abs =
   fixes abs :: "'a \<Rightarrow> 'a"
 begin
 
@@ -247,10 +247,10 @@
 
 end
 
-class sgn = type +
+class sgn =
   fixes sgn :: "'a \<Rightarrow> 'a"
 
-class ord = type +
+class ord =
   fixes less_eq :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
     and less :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
 begin
@@ -1675,7 +1675,7 @@
 
 text {* Equality *}
 
-class eq = type +
+class eq =
   fixes eq :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
   assumes eq_equals: "eq x y \<longleftrightarrow> x = y"
 begin
@@ -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/Hilbert_Choice.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Hilbert_Choice.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Hilbert_Choice.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson
     Copyright   2001  University of Cambridge
 *)
@@ -7,7 +6,7 @@
 header {* Hilbert's Epsilon-Operator and the Axiom of Choice *}
 
 theory Hilbert_Choice
-imports Nat Wellfounded
+imports Nat Wellfounded Plain
 uses ("Tools/meson.ML") ("Tools/specification_package.ML")
 begin
 
--- a/src/HOL/Import/hol4rews.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Import/hol4rews.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -390,7 +390,7 @@
         val thm2 = standard thm1;
       in
         thy
-        |> PureThy.store_thm (bthm, thm2)
+        |> PureThy.store_thm (Binding.name bthm, thm2)
         |> snd
         |> add_hol4_theorem bthy bthm hth
       end;
--- a/src/HOL/Import/proof_kernel.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Import/proof_kernel.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1928,7 +1928,7 @@
                        Replaying _ => thy
                      | _ => (ImportRecorder.add_consts [(constname, ctype, csyn)]; Sign.add_consts_i [(constname,ctype,csyn)] thy)
         val eq = mk_defeq constname rhs' thy1
-        val (thms, thy2) = PureThy.add_defs false [((thmname,eq),[])] thy1
+        val (thms, thy2) = PureThy.add_defs false [((Binding.name thmname,eq),[])] thy1
         val _ = ImportRecorder.add_defs thmname eq
         val def_thm = hd thms
         val thm' = def_thm RS meta_eq_to_obj_eq_thm
--- a/src/HOL/Import/replay.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Import/replay.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -340,7 +340,7 @@
 	  | delta (Hol_move (fullname, moved_thmname)) thy = 
 	    add_hol4_move fullname moved_thmname thy
 	  | delta (Defs (thmname, eq)) thy =
-	    snd (PureThy.add_defs false [((thmname, eq), [])] thy)
+	    snd (PureThy.add_defs false [((Binding.name thmname, eq), [])] thy)
 	  | delta (Hol_theorem (thyname, thmname, th)) thy =
 	    add_hol4_theorem thyname thmname ([], th_of thy th) thy
 	  | delta (Typedef (thmname, typ, c, repabs, th)) thy = 
--- a/src/HOL/Int.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Int.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Int.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
                 Tobias Nipkow, Florian Haftmann, TU Muenchen
     Copyright   1994  University of Cambridge
@@ -157,13 +156,13 @@
   show "i - j = i + - j"
     by (simp add: diff_int_def)
   show "(i * j) * k = i * (j * k)"
-    by (cases i, cases j, cases k) (simp add: mult ring_simps)
+    by (cases i, cases j, cases k) (simp add: mult algebra_simps)
   show "i * j = j * i"
-    by (cases i, cases j) (simp add: mult ring_simps)
+    by (cases i, cases j) (simp add: mult algebra_simps)
   show "1 * i = i"
     by (cases i) (simp add: One_int_def mult)
   show "(i + j) * k = i * k + j * k"
-    by (cases i, cases j, cases k) (simp add: add mult ring_simps)
+    by (cases i, cases j, cases k) (simp add: add mult algebra_simps)
   show "0 \<noteq> (1::int)"
     by (simp add: Zero_int_def One_int_def)
 qed
@@ -301,36 +300,35 @@
 lemma of_int: "of_int (Abs_Integ (intrel `` {(i,j)})) = of_nat i - of_nat j"
 proof -
   have "(\<lambda>(i,j). { of_nat i - (of_nat j :: 'a) }) respects intrel"
-    by (simp add: congruent_def compare_rls of_nat_add [symmetric]
+    by (simp add: congruent_def algebra_simps of_nat_add [symmetric]
             del: of_nat_add) 
   thus ?thesis
     by (simp add: of_int_def UN_equiv_class [OF equiv_intrel])
 qed
 
 lemma of_int_0 [simp]: "of_int 0 = 0"
-  by (simp add: of_int Zero_int_def)
+by (simp add: of_int Zero_int_def)
 
 lemma of_int_1 [simp]: "of_int 1 = 1"
-  by (simp add: of_int One_int_def)
+by (simp add: of_int One_int_def)
 
 lemma of_int_add [simp]: "of_int (w+z) = of_int w + of_int z"
-  by (cases w, cases z, simp add: compare_rls of_int OrderedGroup.compare_rls add)
+by (cases w, cases z, simp add: algebra_simps of_int add)
 
 lemma of_int_minus [simp]: "of_int (-z) = - (of_int z)"
-  by (cases z, simp add: compare_rls of_int minus)
+by (cases z, simp add: algebra_simps of_int minus)
 
 lemma of_int_diff [simp]: "of_int (w - z) = of_int w - of_int z"
-  by (simp add: OrderedGroup.diff_minus diff_minus)
+by (simp add: OrderedGroup.diff_minus diff_minus)
 
 lemma of_int_mult [simp]: "of_int (w*z) = of_int w * of_int z"
 apply (cases w, cases z)
-apply (simp add: compare_rls of_int left_diff_distrib right_diff_distrib
-                 mult add_ac of_nat_mult)
+apply (simp add: algebra_simps of_int mult of_nat_mult)
 done
 
 text{*Collapse nested embeddings*}
 lemma of_int_of_nat_eq [simp]: "of_int (of_nat n) = of_nat n"
-  by (induct n) auto
+by (induct n) auto
 
 end
 
@@ -339,7 +337,7 @@
 
 lemma of_int_le_iff [simp]:
   "of_int w \<le> of_int z \<longleftrightarrow> w \<le> z"
-  by (cases w, cases z, simp add: of_int le minus compare_rls of_nat_add [symmetric] del: of_nat_add)
+by (cases w, cases z, simp add: of_int le minus algebra_simps of_nat_add [symmetric] del: of_nat_add)
 
 text{*Special cases where either operand is zero*}
 lemmas of_int_0_le_iff [simp] = of_int_le_iff [of 0, simplified]
@@ -511,7 +509,7 @@
   also have "\<dots> = (\<exists>n. z - w = of_nat n)"
     by (auto elim: zero_le_imp_eq_int)
   also have "\<dots> = (\<exists>n. z = w + of_nat n)"
-    by (simp only: group_simps)
+    by (simp only: algebra_simps)
   finally show ?thesis .
 qed
 
@@ -599,7 +597,7 @@
   Bit1 :: "int \<Rightarrow> int" where
   [code del]: "Bit1 k = 1 + k + k"
 
-class number = type + -- {* for numeric types: nat, int, real, \dots *}
+class number = -- {* for numeric types: nat, int, real, \dots *}
   fixes number_of :: "int \<Rightarrow> 'a"
 
 use "Tools/numeral.ML"
@@ -829,7 +827,7 @@
 next
   case (neg n)
   thus ?thesis by (simp del: of_nat_Suc of_nat_add
-    add: compare_rls of_nat_1 [symmetric] of_nat_add [symmetric])
+    add: algebra_simps of_nat_1 [symmetric] of_nat_add [symmetric])
 qed
 
 lemma bin_less_0_simps:
@@ -1067,16 +1065,16 @@
 lemma eq_number_of_eq:
   "((number_of x::'a::number_ring) = number_of y) =
    iszero (number_of (x + uminus y) :: 'a)"
-  unfolding iszero_def number_of_add number_of_minus
-  by (simp add: compare_rls)
+unfolding iszero_def number_of_add number_of_minus
+by (simp add: algebra_simps)
 
 lemma iszero_number_of_Pls:
   "iszero ((number_of Pls)::'a::number_ring)"
-  unfolding iszero_def numeral_0_eq_0 ..
+unfolding iszero_def numeral_0_eq_0 ..
 
 lemma nonzero_number_of_Min:
   "~ iszero ((number_of Min)::'a::number_ring)"
-  unfolding iszero_def numeral_m1_eq_minus_1 by simp
+unfolding iszero_def numeral_m1_eq_minus_1 by simp
 
 
 subsubsection {* Comparisons, for Ordered Rings *}
@@ -1162,7 +1160,7 @@
 next
   case (neg n)
   thus ?thesis by (simp del: of_nat_Suc of_nat_add
-    add: compare_rls of_nat_1 [symmetric] of_nat_add [symmetric])
+    add: algebra_simps of_nat_1 [symmetric] of_nat_add [symmetric])
 qed
 
 text {* Less-Than or Equals *}
@@ -1249,9 +1247,7 @@
 lemma add_number_of_diff2 [simp]:
   "number_of v + (c - number_of w) =
    number_of (v + uminus w) + (c::'a::number_ring)"
-apply (subst diff_number_of_eq [symmetric])
-apply (simp only: compare_rls)
-done
+by (simp add: algebra_simps diff_number_of_eq [symmetric])
 
 
 subsection {* The Set of Integers *}
--- a/src/HOL/IntDiv.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/IntDiv.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,82 +1,69 @@
 (*  Title:      HOL/IntDiv.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1999  University of Cambridge
 
 *)
 
-header{*The Division Operators div and mod; the Divides Relation dvd*}
+header{* The Division Operators div and mod *}
 
 theory IntDiv
 imports Int Divides FunDef
 begin
 
-constdefs
-  quorem :: "(int*int) * (int*int) => bool"
+definition divmod_rel :: "int \<Rightarrow> int \<Rightarrow> int \<times> int \<Rightarrow> bool" where
     --{*definition of quotient and remainder*}
-    [code]: "quorem == %((a,b), (q,r)).
-                      a = b*q + r &
-                      (if 0 < b then 0\<le>r & r<b else b<r & r \<le> 0)"
+    [code]: "divmod_rel a b = (\<lambda>(q, r). a = b * q + r \<and>
+               (if 0 < b then 0 \<le> r \<and> r < b else b < r \<and> r \<le> 0))"
 
-  adjust :: "[int, int*int] => int*int"
+definition adjust :: "int \<Rightarrow> int \<times> int \<Rightarrow> int \<times> int" where
     --{*for the division algorithm*}
-    [code]: "adjust b == %(q,r). if 0 \<le> r-b then (2*q + 1, r-b)
-                         else (2*q, r)"
+    [code]: "adjust b = (\<lambda>(q, r). if 0 \<le> r - b then (2 * q + 1, r - b)
+                         else (2 * q, r))"
 
 text{*algorithm for the case @{text "a\<ge>0, b>0"}*}
-function
-  posDivAlg :: "int \<Rightarrow> int \<Rightarrow> int \<times> int"
-where
-  "posDivAlg a b =
-     (if (a<b | b\<le>0) then (0,a)
-        else adjust b (posDivAlg a (2*b)))"
+function posDivAlg :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where
+  "posDivAlg a b = (if a < b \<or>  b \<le> 0 then (0, a)
+     else adjust b (posDivAlg a (2 * b)))"
 by auto
-termination by (relation "measure (%(a,b). nat(a - b + 1))") auto
+termination by (relation "measure (\<lambda>(a, b). nat (a - b + 1))") auto
 
 text{*algorithm for the case @{text "a<0, b>0"}*}
-function
-  negDivAlg :: "int \<Rightarrow> int \<Rightarrow> int \<times> int"
-where
-  "negDivAlg a b  =
-     (if (0\<le>a+b | b\<le>0) then (-1,a+b)
-      else adjust b (negDivAlg a (2*b)))"
+function negDivAlg :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where
+  "negDivAlg a b = (if 0 \<le>a + b \<or> b \<le> 0  then (-1, a + b)
+     else adjust b (negDivAlg a (2 * b)))"
 by auto
-termination by (relation "measure (%(a,b). nat(- a - b))") auto
+termination by (relation "measure (\<lambda>(a, b). nat (- a - b))") auto
 
 text{*algorithm for the general case @{term "b\<noteq>0"}*}
-constdefs
-  negateSnd :: "int*int => int*int"
-    [code]: "negateSnd == %(q,r). (q,-r)"
+definition negateSnd :: "int \<times> int \<Rightarrow> int \<times> int" where
+  [code inline]: "negateSnd = apsnd uminus"
 
-definition
-  divAlg :: "int \<times> int \<Rightarrow> int \<times> int"
+definition divmod :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where
     --{*The full division algorithm considers all possible signs for a, b
        including the special case @{text "a=0, b<0"} because 
        @{term negDivAlg} requires @{term "a<0"}.*}
-where
-  "divAlg = (\<lambda>(a, b). (if 0\<le>a then
-                  if 0\<le>b then posDivAlg a b
-                  else if a=0 then (0, 0)
+  "divmod a b = (if 0 \<le> a then if 0 \<le> b then posDivAlg a b
+                  else if a = 0 then (0, 0)
                        else negateSnd (negDivAlg (-a) (-b))
                else 
-                  if 0<b then negDivAlg a b
-                  else negateSnd (posDivAlg (-a) (-b))))"
+                  if 0 < b then negDivAlg a b
+                  else negateSnd (posDivAlg (-a) (-b)))"
 
 instantiation int :: Divides.div
 begin
 
 definition
-  div_def: "a div b = fst (divAlg (a, b))"
+  div_def: "a div b = fst (divmod a b)"
 
 definition
-  mod_def: "a mod b = snd (divAlg (a, b))"
+  mod_def: "a mod b = snd (divmod a b)"
 
 instance ..
 
 end
 
-lemma divAlg_mod_div:
-  "divAlg (p, q) = (p div q, p mod q)"
+lemma divmod_mod_div:
+  "divmod p q = (p div q, p mod q)"
   by (auto simp add: div_def mod_def)
 
 text{*
@@ -97,7 +84,7 @@
 
     fun negateSnd (q,r:int) = (q,~r);
 
-    fun divAlg (a,b) = if 0\<le>a then 
+    fun divmod (a,b) = if 0\<le>a then 
 			  if b>0 then posDivAlg (a,b) 
 			   else if a=0 then (0,0)
 				else negateSnd (negDivAlg (~a,~b))
@@ -131,9 +118,9 @@
     auto)
 
 lemma unique_quotient:
-     "[| quorem ((a,b), (q,r));  quorem ((a,b), (q',r'));  b \<noteq> 0 |]  
+     "[| divmod_rel a b (q, r); divmod_rel a b (q', r');  b \<noteq> 0 |]  
       ==> q = q'"
-apply (simp add: quorem_def linorder_neq_iff split: split_if_asm)
+apply (simp add: divmod_rel_def linorder_neq_iff split: split_if_asm)
 apply (blast intro: order_antisym
              dest: order_eq_refl [THEN unique_quotient_lemma] 
              order_eq_refl [THEN unique_quotient_lemma_neg] sym)+
@@ -141,10 +128,10 @@
 
 
 lemma unique_remainder:
-     "[| quorem ((a,b), (q,r));  quorem ((a,b), (q',r'));  b \<noteq> 0 |]  
+     "[| divmod_rel a b (q, r); divmod_rel a b (q', r');  b \<noteq> 0 |]  
       ==> r = r'"
 apply (subgoal_tac "q = q'")
- apply (simp add: quorem_def)
+ apply (simp add: divmod_rel_def)
 apply (blast intro: unique_quotient)
 done
 
@@ -171,10 +158,10 @@
 text{*Correctness of @{term posDivAlg}: it computes quotients correctly*}
 theorem posDivAlg_correct:
   assumes "0 \<le> a" and "0 < b"
-  shows "quorem ((a, b), posDivAlg a b)"
+  shows "divmod_rel a b (posDivAlg a b)"
 using prems apply (induct a b rule: posDivAlg.induct)
 apply auto
-apply (simp add: quorem_def)
+apply (simp add: divmod_rel_def)
 apply (subst posDivAlg_eqn, simp add: right_distrib)
 apply (case_tac "a < b")
 apply simp_all
@@ -200,10 +187,10 @@
   It doesn't work if a=0 because the 0/b equals 0, not -1*)
 lemma negDivAlg_correct:
   assumes "a < 0" and "b > 0"
-  shows "quorem ((a, b), negDivAlg a b)"
+  shows "divmod_rel a b (negDivAlg a b)"
 using prems apply (induct a b rule: negDivAlg.induct)
 apply (auto simp add: linorder_not_le)
-apply (simp add: quorem_def)
+apply (simp add: divmod_rel_def)
 apply (subst negDivAlg_eqn, assumption)
 apply (case_tac "a + b < (0\<Colon>int)")
 apply simp_all
@@ -215,8 +202,8 @@
 subsection{*Existence Shown by Proving the Division Algorithm to be Correct*}
 
 (*the case a=0*)
-lemma quorem_0: "b \<noteq> 0 ==> quorem ((0,b), (0,0))"
-by (auto simp add: quorem_def linorder_neq_iff)
+lemma divmod_rel_0: "b \<noteq> 0 ==> divmod_rel 0 b (0, 0)"
+by (auto simp add: divmod_rel_def linorder_neq_iff)
 
 lemma posDivAlg_0 [simp]: "posDivAlg 0 b = (0, 0)"
 by (subst posDivAlg.simps, auto)
@@ -227,26 +214,26 @@
 lemma negateSnd_eq [simp]: "negateSnd(q,r) = (q,-r)"
 by (simp add: negateSnd_def)
 
-lemma quorem_neg: "quorem ((-a,-b), qr) ==> quorem ((a,b), negateSnd qr)"
-by (auto simp add: split_ifs quorem_def)
+lemma divmod_rel_neg: "divmod_rel (-a) (-b) qr ==> divmod_rel a b (negateSnd qr)"
+by (auto simp add: split_ifs divmod_rel_def)
 
-lemma divAlg_correct: "b \<noteq> 0 ==> quorem ((a,b), divAlg (a, b))"
-by (force simp add: linorder_neq_iff quorem_0 divAlg_def quorem_neg
+lemma divmod_correct: "b \<noteq> 0 ==> divmod_rel a b (divmod a b)"
+by (force simp add: linorder_neq_iff divmod_rel_0 divmod_def divmod_rel_neg
                     posDivAlg_correct negDivAlg_correct)
 
 text{*Arbitrary definitions for division by zero.  Useful to simplify 
     certain equations.*}
 
 lemma DIVISION_BY_ZERO [simp]: "a div (0::int) = 0 & a mod (0::int) = a"
-by (simp add: div_def mod_def divAlg_def posDivAlg.simps)  
+by (simp add: div_def mod_def divmod_def posDivAlg.simps)  
 
 
 text{*Basic laws about division and remainder*}
 
 lemma zmod_zdiv_equality: "(a::int) = b * (a div b) + (a mod b)"
 apply (case_tac "b = 0", simp)
-apply (cut_tac a = a and b = b in divAlg_correct)
-apply (auto simp add: quorem_def div_def mod_def)
+apply (cut_tac a = a and b = b in divmod_correct)
+apply (auto simp add: divmod_rel_def div_def mod_def)
 done
 
 lemma zdiv_zmod_equality: "(b * (a div b) + (a mod b)) + k = (a::int)+k"
@@ -288,16 +275,16 @@
 *}
 
 lemma pos_mod_conj : "(0::int) < b ==> 0 \<le> a mod b & a mod b < b"
-apply (cut_tac a = a and b = b in divAlg_correct)
-apply (auto simp add: quorem_def mod_def)
+apply (cut_tac a = a and b = b in divmod_correct)
+apply (auto simp add: divmod_rel_def mod_def)
 done
 
 lemmas pos_mod_sign  [simp] = pos_mod_conj [THEN conjunct1, standard]
    and pos_mod_bound [simp] = pos_mod_conj [THEN conjunct2, standard]
 
 lemma neg_mod_conj : "b < (0::int) ==> a mod b \<le> 0 & b < a mod b"
-apply (cut_tac a = a and b = b in divAlg_correct)
-apply (auto simp add: quorem_def div_def mod_def)
+apply (cut_tac a = a and b = b in divmod_correct)
+apply (auto simp add: divmod_rel_def div_def mod_def)
 done
 
 lemmas neg_mod_sign  [simp] = neg_mod_conj [THEN conjunct1, standard]
@@ -307,47 +294,47 @@
 
 subsection{*General Properties of div and mod*}
 
-lemma quorem_div_mod: "b \<noteq> 0 ==> quorem ((a, b), (a div b, a mod b))"
+lemma divmod_rel_div_mod: "b \<noteq> 0 ==> divmod_rel a b (a div b, a mod b)"
 apply (cut_tac a = a and b = b in zmod_zdiv_equality)
-apply (force simp add: quorem_def linorder_neq_iff)
+apply (force simp add: divmod_rel_def linorder_neq_iff)
 done
 
-lemma quorem_div: "[| quorem((a,b),(q,r));  b \<noteq> 0 |] ==> a div b = q"
-by (simp add: quorem_div_mod [THEN unique_quotient])
+lemma divmod_rel_div: "[| divmod_rel a b (q, r);  b \<noteq> 0 |] ==> a div b = q"
+by (simp add: divmod_rel_div_mod [THEN unique_quotient])
 
-lemma quorem_mod: "[| quorem((a,b),(q,r));  b \<noteq> 0 |] ==> a mod b = r"
-by (simp add: quorem_div_mod [THEN unique_remainder])
+lemma divmod_rel_mod: "[| divmod_rel a b (q, r);  b \<noteq> 0 |] ==> a mod b = r"
+by (simp add: divmod_rel_div_mod [THEN unique_remainder])
 
 lemma div_pos_pos_trivial: "[| (0::int) \<le> a;  a < b |] ==> a div b = 0"
-apply (rule quorem_div)
-apply (auto simp add: quorem_def)
+apply (rule divmod_rel_div)
+apply (auto simp add: divmod_rel_def)
 done
 
 lemma div_neg_neg_trivial: "[| a \<le> (0::int);  b < a |] ==> a div b = 0"
-apply (rule quorem_div)
-apply (auto simp add: quorem_def)
+apply (rule divmod_rel_div)
+apply (auto simp add: divmod_rel_def)
 done
 
 lemma div_pos_neg_trivial: "[| (0::int) < a;  a+b \<le> 0 |] ==> a div b = -1"
-apply (rule quorem_div)
-apply (auto simp add: quorem_def)
+apply (rule divmod_rel_div)
+apply (auto simp add: divmod_rel_def)
 done
 
 (*There is no div_neg_pos_trivial because  0 div b = 0 would supersede it*)
 
 lemma mod_pos_pos_trivial: "[| (0::int) \<le> a;  a < b |] ==> a mod b = a"
-apply (rule_tac q = 0 in quorem_mod)
-apply (auto simp add: quorem_def)
+apply (rule_tac q = 0 in divmod_rel_mod)
+apply (auto simp add: divmod_rel_def)
 done
 
 lemma mod_neg_neg_trivial: "[| a \<le> (0::int);  b < a |] ==> a mod b = a"
-apply (rule_tac q = 0 in quorem_mod)
-apply (auto simp add: quorem_def)
+apply (rule_tac q = 0 in divmod_rel_mod)
+apply (auto simp add: divmod_rel_def)
 done
 
 lemma mod_pos_neg_trivial: "[| (0::int) < a;  a+b \<le> 0 |] ==> a mod b = a+b"
-apply (rule_tac q = "-1" in quorem_mod)
-apply (auto simp add: quorem_def)
+apply (rule_tac q = "-1" in divmod_rel_mod)
+apply (auto simp add: divmod_rel_def)
 done
 
 text{*There is no @{text mod_neg_pos_trivial}.*}
@@ -356,15 +343,15 @@
 (*Simpler laws such as -a div b = -(a div b) FAIL, but see just below*)
 lemma zdiv_zminus_zminus [simp]: "(-a) div (-b) = a div (b::int)"
 apply (case_tac "b = 0", simp)
-apply (simp add: quorem_div_mod [THEN quorem_neg, simplified, 
-                                 THEN quorem_div, THEN sym])
+apply (simp add: divmod_rel_div_mod [THEN divmod_rel_neg, simplified, 
+                                 THEN divmod_rel_div, THEN sym])
 
 done
 
 (*Simpler laws such as -a mod b = -(a mod b) FAIL, but see just below*)
 lemma zmod_zminus_zminus [simp]: "(-a) mod (-b) = - (a mod (b::int))"
 apply (case_tac "b = 0", simp)
-apply (subst quorem_div_mod [THEN quorem_neg, simplified, THEN quorem_mod],
+apply (subst divmod_rel_div_mod [THEN divmod_rel_neg, simplified, THEN divmod_rel_mod],
        auto)
 done
 
@@ -372,22 +359,22 @@
 subsection{*Laws for div and mod with Unary Minus*}
 
 lemma zminus1_lemma:
-     "quorem((a,b),(q,r))  
-      ==> quorem ((-a,b), (if r=0 then -q else -q - 1),  
-                          (if r=0 then 0 else b-r))"
-by (force simp add: split_ifs quorem_def linorder_neq_iff right_diff_distrib)
+     "divmod_rel a b (q, r)
+      ==> divmod_rel (-a) b (if r=0 then -q else -q - 1,  
+                          if r=0 then 0 else b-r)"
+by (force simp add: split_ifs divmod_rel_def linorder_neq_iff right_diff_distrib)
 
 
 lemma zdiv_zminus1_eq_if:
      "b \<noteq> (0::int)  
       ==> (-a) div b =  
           (if a mod b = 0 then - (a div b) else  - (a div b) - 1)"
-by (blast intro: quorem_div_mod [THEN zminus1_lemma, THEN quorem_div])
+by (blast intro: divmod_rel_div_mod [THEN zminus1_lemma, THEN divmod_rel_div])
 
 lemma zmod_zminus1_eq_if:
      "(-a::int) mod b = (if a mod b = 0 then 0 else  b - (a mod b))"
 apply (case_tac "b = 0", simp)
-apply (blast intro: quorem_div_mod [THEN zminus1_lemma, THEN quorem_mod])
+apply (blast intro: divmod_rel_div_mod [THEN zminus1_lemma, THEN divmod_rel_mod])
 done
 
 lemma zdiv_zminus2: "a div (-b) = (-a::int) div b"
@@ -420,91 +407,91 @@
 apply (simp add: right_diff_distrib)
 done
 
-lemma self_quotient: "[| quorem((a,a),(q,r));  a \<noteq> (0::int) |] ==> q = 1"
-apply (simp add: split_ifs quorem_def linorder_neq_iff)
+lemma self_quotient: "[| divmod_rel a a (q, r);  a \<noteq> (0::int) |] ==> q = 1"
+apply (simp add: split_ifs divmod_rel_def linorder_neq_iff)
 apply (rule order_antisym, safe, simp_all)
 apply (rule_tac [3] a = "-a" and r = "-r" in self_quotient_aux1)
 apply (rule_tac a = "-a" and r = "-r" in self_quotient_aux2)
 apply (force intro: self_quotient_aux1 self_quotient_aux2 simp add: add_commute)+
 done
 
-lemma self_remainder: "[| quorem((a,a),(q,r));  a \<noteq> (0::int) |] ==> r = 0"
+lemma self_remainder: "[| divmod_rel a a (q, r);  a \<noteq> (0::int) |] ==> r = 0"
 apply (frule self_quotient, assumption)
-apply (simp add: quorem_def)
+apply (simp add: divmod_rel_def)
 done
 
 lemma zdiv_self [simp]: "a \<noteq> 0 ==> a div a = (1::int)"
-by (simp add: quorem_div_mod [THEN self_quotient])
+by (simp add: divmod_rel_div_mod [THEN self_quotient])
 
 (*Here we have 0 mod 0 = 0, also assumed by Knuth (who puts m mod 0 = 0) *)
 lemma zmod_self [simp]: "a mod a = (0::int)"
 apply (case_tac "a = 0", simp)
-apply (simp add: quorem_div_mod [THEN self_remainder])
+apply (simp add: divmod_rel_div_mod [THEN self_remainder])
 done
 
 
 subsection{*Computation of Division and Remainder*}
 
 lemma zdiv_zero [simp]: "(0::int) div b = 0"
-by (simp add: div_def divAlg_def)
+by (simp add: div_def divmod_def)
 
 lemma div_eq_minus1: "(0::int) < b ==> -1 div b = -1"
-by (simp add: div_def divAlg_def)
+by (simp add: div_def divmod_def)
 
 lemma zmod_zero [simp]: "(0::int) mod b = 0"
-by (simp add: mod_def divAlg_def)
+by (simp add: mod_def divmod_def)
 
 lemma zdiv_minus1: "(0::int) < b ==> -1 div b = -1"
-by (simp add: div_def divAlg_def)
+by (simp add: div_def divmod_def)
 
 lemma zmod_minus1: "(0::int) < b ==> -1 mod b = b - 1"
-by (simp add: mod_def divAlg_def)
+by (simp add: mod_def divmod_def)
 
 text{*a positive, b positive *}
 
 lemma div_pos_pos: "[| 0 < a;  0 \<le> b |] ==> a div b = fst (posDivAlg a b)"
-by (simp add: div_def divAlg_def)
+by (simp add: div_def divmod_def)
 
 lemma mod_pos_pos: "[| 0 < a;  0 \<le> b |] ==> a mod b = snd (posDivAlg a b)"
-by (simp add: mod_def divAlg_def)
+by (simp add: mod_def divmod_def)
 
 text{*a negative, b positive *}
 
 lemma div_neg_pos: "[| a < 0;  0 < b |] ==> a div b = fst (negDivAlg a b)"
-by (simp add: div_def divAlg_def)
+by (simp add: div_def divmod_def)
 
 lemma mod_neg_pos: "[| a < 0;  0 < b |] ==> a mod b = snd (negDivAlg a b)"
-by (simp add: mod_def divAlg_def)
+by (simp add: mod_def divmod_def)
 
 text{*a positive, b negative *}
 
 lemma div_pos_neg:
      "[| 0 < a;  b < 0 |] ==> a div b = fst (negateSnd (negDivAlg (-a) (-b)))"
-by (simp add: div_def divAlg_def)
+by (simp add: div_def divmod_def)
 
 lemma mod_pos_neg:
      "[| 0 < a;  b < 0 |] ==> a mod b = snd (negateSnd (negDivAlg (-a) (-b)))"
-by (simp add: mod_def divAlg_def)
+by (simp add: mod_def divmod_def)
 
 text{*a negative, b negative *}
 
 lemma div_neg_neg:
      "[| a < 0;  b \<le> 0 |] ==> a div b = fst (negateSnd (posDivAlg (-a) (-b)))"
-by (simp add: div_def divAlg_def)
+by (simp add: div_def divmod_def)
 
 lemma mod_neg_neg:
      "[| a < 0;  b \<le> 0 |] ==> a mod b = snd (negateSnd (posDivAlg (-a) (-b)))"
-by (simp add: mod_def divAlg_def)
+by (simp add: mod_def divmod_def)
 
 text {*Simplify expresions in which div and mod combine numerical constants*}
 
-lemma quoremI:
+lemma divmod_relI:
   "\<lbrakk>a == b * q + r; if 0 < b then 0 \<le> r \<and> r < b else b < r \<and> r \<le> 0\<rbrakk>
-    \<Longrightarrow> quorem ((a, b), (q, r))"
-  unfolding quorem_def by simp
+    \<Longrightarrow> divmod_rel a b (q, r)"
+  unfolding divmod_rel_def by simp
 
-lemmas quorem_div_eq = quoremI [THEN quorem_div, THEN eq_reflection]
-lemmas quorem_mod_eq = quoremI [THEN quorem_mod, THEN eq_reflection]
+lemmas divmod_rel_div_eq = divmod_relI [THEN divmod_rel_div, THEN eq_reflection]
+lemmas divmod_rel_mod_eq = divmod_relI [THEN divmod_rel_mod, THEN eq_reflection]
 lemmas arithmetic_simps =
   arith_simps
   add_special
@@ -548,10 +535,10 @@
 *}
 
 simproc_setup binary_int_div ("number_of m div number_of n :: int") =
-  {* K (divmod_proc (@{thm quorem_div_eq})) *}
+  {* K (divmod_proc (@{thm divmod_rel_div_eq})) *}
 
 simproc_setup binary_int_mod ("number_of m mod number_of n :: int") =
-  {* K (divmod_proc (@{thm quorem_mod_eq})) *}
+  {* K (divmod_proc (@{thm divmod_rel_mod_eq})) *}
 
 (* The following 8 lemmas are made unnecessary by the above simprocs: *)
 
@@ -718,18 +705,18 @@
 text{*proving (a*b) div c = a * (b div c) + a * (b mod c) *}
 
 lemma zmult1_lemma:
-     "[| quorem((b,c),(q,r));  c \<noteq> 0 |]  
-      ==> quorem ((a*b, c), (a*q + a*r div c, a*r mod c))"
-by (force simp add: split_ifs quorem_def linorder_neq_iff right_distrib)
+     "[| divmod_rel b c (q, r);  c \<noteq> 0 |]  
+      ==> divmod_rel (a * b) c (a*q + a*r div c, a*r mod c)"
+by (force simp add: split_ifs divmod_rel_def linorder_neq_iff right_distrib)
 
 lemma zdiv_zmult1_eq: "(a*b) div c = a*(b div c) + a*(b mod c) div (c::int)"
 apply (case_tac "c = 0", simp)
-apply (blast intro: quorem_div_mod [THEN zmult1_lemma, THEN quorem_div])
+apply (blast intro: divmod_rel_div_mod [THEN zmult1_lemma, THEN divmod_rel_div])
 done
 
 lemma zmod_zmult1_eq: "(a*b) mod c = a*(b mod c) mod (c::int)"
 apply (case_tac "c = 0", simp)
-apply (blast intro: quorem_div_mod [THEN zmult1_lemma, THEN quorem_mod])
+apply (blast intro: divmod_rel_div_mod [THEN zmult1_lemma, THEN divmod_rel_mod])
 done
 
 lemma zmod_zmult1_eq': "(a*b) mod (c::int) = ((a mod c) * b) mod c"
@@ -760,20 +747,20 @@
 text{*proving (a+b) div c = a div c + b div c + ((a mod c + b mod c) div c) *}
 
 lemma zadd1_lemma:
-     "[| quorem((a,c),(aq,ar));  quorem((b,c),(bq,br));  c \<noteq> 0 |]  
-      ==> quorem ((a+b, c), (aq + bq + (ar+br) div c, (ar+br) mod c))"
-by (force simp add: split_ifs quorem_def linorder_neq_iff right_distrib)
+     "[| divmod_rel a c (aq, ar);  divmod_rel b c (bq, br);  c \<noteq> 0 |]  
+      ==> divmod_rel (a+b) c (aq + bq + (ar+br) div c, (ar+br) mod c)"
+by (force simp add: split_ifs divmod_rel_def linorder_neq_iff right_distrib)
 
 (*NOT suitable for rewriting: the RHS has an instance of the LHS*)
 lemma zdiv_zadd1_eq:
      "(a+b) div (c::int) = a div c + b div c + ((a mod c + b mod c) div c)"
 apply (case_tac "c = 0", simp)
-apply (blast intro: zadd1_lemma [OF quorem_div_mod quorem_div_mod] quorem_div)
+apply (blast intro: zadd1_lemma [OF divmod_rel_div_mod divmod_rel_div_mod] divmod_rel_div)
 done
 
 lemma zmod_zadd1_eq: "(a+b) mod (c::int) = (a mod c + b mod c) mod c"
 apply (case_tac "c = 0", simp)
-apply (blast intro: zadd1_lemma [OF quorem_div_mod quorem_div_mod] quorem_mod)
+apply (blast intro: zadd1_lemma [OF divmod_rel_div_mod divmod_rel_div_mod] divmod_rel_mod)
 done
 
 instance int :: ring_div
@@ -785,6 +772,33 @@
       by (simp add: zmod_zmult1_eq zmod_zdiv_trivial)
 qed auto
 
+lemma posDivAlg_div_mod:
+  assumes "k \<ge> 0"
+  and "l \<ge> 0"
+  shows "posDivAlg k l = (k div l, k mod l)"
+proof (cases "l = 0")
+  case True then show ?thesis by (simp add: posDivAlg.simps)
+next
+  case False with assms posDivAlg_correct
+    have "divmod_rel k l (fst (posDivAlg k l), snd (posDivAlg k l))"
+    by simp
+  from divmod_rel_div [OF this `l \<noteq> 0`] divmod_rel_mod [OF this `l \<noteq> 0`]
+  show ?thesis by simp
+qed
+
+lemma negDivAlg_div_mod:
+  assumes "k < 0"
+  and "l > 0"
+  shows "negDivAlg k l = (k div l, k mod l)"
+proof -
+  from assms have "l \<noteq> 0" by simp
+  from assms negDivAlg_correct
+    have "divmod_rel k l (fst (negDivAlg k l), snd (negDivAlg k l))"
+    by simp
+  from divmod_rel_div [OF this `l \<noteq> 0`] divmod_rel_mod [OF this `l \<noteq> 0`]
+  show ?thesis by simp
+qed
+
 lemma zdiv_zadd_self1: "a \<noteq> (0::int) ==> (a+b) div a = b div a + 1"
 by (rule div_add_self1) (* already declared [simp] *)
 
@@ -831,12 +845,13 @@
 
 lemma zmult2_lemma_aux1: "[| (0::int) < c;  b < r;  r \<le> 0 |] ==> b*c < b*(q mod c) + r"
 apply (subgoal_tac "b * (c - q mod c) < r * 1")
-apply (simp add: right_diff_distrib)
+ apply (simp add: algebra_simps)
 apply (rule order_le_less_trans)
-apply (erule_tac [2] mult_strict_right_mono)
-apply (rule mult_left_mono_neg)
-apply (auto simp add: compare_rls add_commute [of 1]
-                      add1_zle_eq pos_mod_bound)
+ apply (erule_tac [2] mult_strict_right_mono)
+ apply (rule mult_left_mono_neg)
+  using add1_zle_eq[of "q mod c"]apply(simp add: algebra_simps pos_mod_bound)
+ apply (simp)
+apply (simp)
 done
 
 lemma zmult2_lemma_aux2:
@@ -854,29 +869,30 @@
 
 lemma zmult2_lemma_aux4: "[| (0::int) < c; 0 \<le> r; r < b |] ==> b * (q mod c) + r < b * c"
 apply (subgoal_tac "r * 1 < b * (c - q mod c) ")
-apply (simp add: right_diff_distrib)
+ apply (simp add: right_diff_distrib)
 apply (rule order_less_le_trans)
-apply (erule mult_strict_right_mono)
-apply (rule_tac [2] mult_left_mono)
-apply (auto simp add: compare_rls add_commute [of 1]
-                      add1_zle_eq pos_mod_bound)
+ apply (erule mult_strict_right_mono)
+ apply (rule_tac [2] mult_left_mono)
+  apply simp
+ using add1_zle_eq[of "q mod c"] apply (simp add: algebra_simps pos_mod_bound)
+apply simp
 done
 
-lemma zmult2_lemma: "[| quorem ((a,b), (q,r));  b \<noteq> 0;  0 < c |]  
-      ==> quorem ((a, b*c), (q div c, b*(q mod c) + r))"
-by (auto simp add: mult_ac quorem_def linorder_neq_iff
+lemma zmult2_lemma: "[| divmod_rel a b (q, r);  b \<noteq> 0;  0 < c |]  
+      ==> divmod_rel a (b * c) (q div c, b*(q mod c) + r)"
+by (auto simp add: mult_ac divmod_rel_def linorder_neq_iff
                    zero_less_mult_iff right_distrib [symmetric] 
                    zmult2_lemma_aux1 zmult2_lemma_aux2 zmult2_lemma_aux3 zmult2_lemma_aux4)
 
 lemma zdiv_zmult2_eq: "(0::int) < c ==> a div (b*c) = (a div b) div c"
 apply (case_tac "b = 0", simp)
-apply (force simp add: quorem_div_mod [THEN zmult2_lemma, THEN quorem_div])
+apply (force simp add: divmod_rel_div_mod [THEN zmult2_lemma, THEN divmod_rel_div])
 done
 
 lemma zmod_zmult2_eq:
      "(0::int) < c ==> a mod (b*c) = b*(a div b mod c) + a mod b"
 apply (case_tac "b = 0", simp)
-apply (force simp add: quorem_div_mod [THEN zmult2_lemma, THEN quorem_mod])
+apply (force simp add: divmod_rel_div_mod [THEN zmult2_lemma, THEN divmod_rel_mod])
 done
 
 
@@ -1251,7 +1267,7 @@
 
 lemma zmult_div_cancel: "(n::int) * (m div n) = m - (m mod n)"
   using zmod_zdiv_equality[where a="m" and b="n"]
-  by (simp add: ring_simps)
+  by (simp add: algebra_simps)
 
 lemma zdvd_mult_div_cancel:"(n::int) dvd m \<Longrightarrow> n * (m div n) = m"
 apply (subgoal_tac "m mod n = 0")
@@ -1360,7 +1376,7 @@
 apply (subst split_div, auto)
 apply (subst split_zdiv, auto)
 apply (rule_tac a="int (b * i) + int j" and b="int b" and r="int j" and r'=ja in IntDiv.unique_quotient)
-apply (auto simp add: IntDiv.quorem_def of_nat_mult)
+apply (auto simp add: IntDiv.divmod_rel_def of_nat_mult)
 done
 
 lemma zmod_int: "int (a mod b) = (int a) mod (int b)"
@@ -1368,7 +1384,7 @@
 apply (subst split_zmod, auto)
 apply (rule_tac a="int (b * i) + int j" and b="int b" and q="int i" and q'=ia 
        in unique_remainder)
-apply (auto simp add: IntDiv.quorem_def of_nat_mult)
+apply (auto simp add: IntDiv.divmod_rel_def of_nat_mult)
 done
 
 text{*Suggested by Matthias Daum*}
@@ -1429,7 +1445,7 @@
 lemma of_int_num [code]:
   "of_int k = (if k = 0 then 0 else if k < 0 then
      - of_int (- k) else let
-       (l, m) = divAlg (k, 2);
+       (l, m) = divmod k 2;
        l' = of_int l
      in if m = 0 then l' + l' else l' + l' + 1)"
 proof -
@@ -1457,7 +1473,7 @@
     show "x * of_int 2 = x + x" 
     unfolding int2 of_int_add right_distrib by simp
   qed
-  from aux1 show ?thesis by (auto simp add: divAlg_mod_div Let_def aux2 aux3)
+  from aux1 show ?thesis by (auto simp add: divmod_mod_div Let_def aux2 aux3)
 qed
 
 end
--- a/src/HOL/IsaMakefile	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/IsaMakefile	Thu Jan 29 12:24:00 2009 +0000
@@ -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: \
@@ -66,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
@@ -75,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 \
@@ -131,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 \
@@ -140,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 \
@@ -159,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)
@@ -280,7 +287,6 @@
   GCD.thy \
   Order_Relation.thy \
   Parity.thy \
-  Univ_Poly.thy \
   Lubs.thy \
   Polynomial.thy \
   PReal.thy \
@@ -325,10 +331,11 @@
   Library/Binomial.thy Library/Eval_Witness.thy	\
   Library/Code_Index.thy Library/Code_Char.thy				\
   Library/Code_Char_chr.thy Library/Code_Integer.thy			\
-  Library/Numeral_Type.thy			\
+  Library/Numeral_Type.thy	Library/Reflection.thy		\
   Library/Boolean_Algebra.thy Library/Countable.thy	\
-  Library/RBT.thy		\
-  Library/Enum.thy Library/Float.thy $(SRC)/Tools/float.ML $(SRC)/HOL/Tools/float_arith.ML
+  Library/RBT.thy	Library/Univ_Poly.thy	\
+  Library/Enum.thy Library/Float.thy $(SRC)/Tools/float.ML $(SRC)/HOL/Tools/float_arith.ML \
+  Library/reify_data.ML Library/reflection.ML
 	@cd Library; $(ISABELLE_TOOL) usedir $(OUT)/HOL Library
 
 
@@ -803,14 +810,14 @@
   ex/Lagrange.thy ex/LocaleTest2.thy ex/MT.thy		\
   ex/MergeSort.thy ex/MonoidGroup.thy ex/Multiquote.thy ex/NatSum.thy	\
   ex/Numeral.thy ex/PER.thy ex/PresburgerEx.thy ex/Primrec.thy		\
-  ex/Quickcheck_Examples.thy ex/Reflection.thy ex/reflection_data.ML	\
+  ex/Quickcheck_Examples.thy	\
   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/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/document/root.tex ex/Meson_Test.thy ex/set.thy	\
   ex/svc_funcs.ML ex/svc_test.thy	\
   ex/ImperativeQuicksort.thy	\
   ex/BigO_Complex.thy			\
@@ -962,13 +969,13 @@
 
 HOL-Word: HOL $(OUT)/HOL-Word
 
-$(OUT)/HOL-Word: $(OUT)/HOL Word/ROOT.ML Library/Infinite_Set.thy	\
+$(OUT)/HOL-Word: $(OUT)/HOL Word/ROOT.ML	\
   Library/Boolean_Algebra.thy			\
   Library/Numeral_Type.thy Word/Num_Lemmas.thy Word/TdThs.thy		\
   Word/Size.thy Word/BinGeneral.thy Word/BinOperations.thy		\
   Word/BinBoolList.thy Word/BitSyntax.thy Word/WordDefinition.thy	\
   Word/WordArith.thy Word/WordBitwise.thy Word/WordShift.thy		\
-  Word/WordGenLib.thy Word/WordMain.thy Word/document/root.tex		\
+  Word/WordGenLib.thy Word/Word.thy Word/document/root.tex		\
   Word/document/root.bib
 	@cd Word; $(ISABELLE_TOOL) usedir -b -g true $(OUT)/HOL HOL-Word
 
--- a/src/HOL/Lattices.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Lattices.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Lattices.thy
-    ID:         $Id$
     Author:     Tobias Nipkow
 *)
 
@@ -300,8 +299,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/Abstract_Rat.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Library/Abstract_Rat.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -240,7 +240,7 @@
     by (simp only: of_int_mult[symmetric] of_int_add [symmetric])
   then have "of_int x / of_int d = ?t / of_int d" 
     using cong[OF refl[of ?f] eq] by simp
-  then show ?thesis by (simp add: add_divide_distrib ring_simps prems)
+  then show ?thesis by (simp add: add_divide_distrib algebra_simps prems)
 qed
 
 lemma of_int_div: "(d::int) ~= 0 ==> d dvd n ==>
--- a/src/HOL/Library/BigO.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Library/BigO.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -349,7 +349,7 @@
   apply (drule set_plus_imp_minus)
   apply (rule set_minus_imp_plus)
   apply (drule bigo_mult3 [where g = g and j = g])
-  apply (auto simp add: ring_simps)
+  apply (auto simp add: algebra_simps)
   done
 
 lemma bigo_mult5: "ALL x. f x ~= 0 ==>
@@ -799,14 +799,14 @@
   apply simp
   apply (subst abs_of_nonneg)
   apply (drule_tac x = x in spec) back
-  apply (simp add: compare_rls)
+  apply (simp add: algebra_simps)
   apply (subst diff_minus)+
   apply (rule add_right_mono)
   apply (erule spec)
   apply (rule order_trans) 
   prefer 2
   apply (rule abs_ge_zero)
-  apply (simp add: compare_rls)
+  apply (simp add: algebra_simps)
   done
 
 lemma bigo_lesso3: "f =o g +o O(h) ==>
@@ -823,7 +823,7 @@
   apply simp
   apply (subst abs_of_nonneg)
   apply (drule_tac x = x in spec) back
-  apply (simp add: compare_rls)
+  apply (simp add: algebra_simps)
   apply (subst diff_minus)+
   apply (rule add_left_mono)
   apply (rule le_imp_neg_le)
@@ -831,7 +831,7 @@
   apply (rule order_trans) 
   prefer 2
   apply (rule abs_ge_zero)
-  apply (simp add: compare_rls)
+  apply (simp add: algebra_simps)
   done
 
 lemma bigo_lesso4: "f <o g =o O(k::'a=>'b::ordered_field) ==>
@@ -844,7 +844,7 @@
   apply assumption
   apply (erule bigo_lesseq2) back
   apply (rule allI)
-  apply (auto simp add: func_plus fun_diff_def compare_rls 
+  apply (auto simp add: func_plus fun_diff_def algebra_simps
     split: split_max abs_split)
   done
 
@@ -856,7 +856,7 @@
   apply (rule allI)
   apply (drule_tac x = x in spec)
   apply (subgoal_tac "abs(max (f x - g x) 0) = max (f x - g x) 0")
-  apply (clarsimp simp add: compare_rls add_ac) 
+  apply (clarsimp simp add: algebra_simps) 
   apply (rule abs_of_nonneg)
   apply (rule le_maxI2)
   done
--- a/src/HOL/Library/Boolean_Algebra.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Library/Boolean_Algebra.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,8 +1,5 @@
-(* 
-  ID:     $Id$
-  Author: Brian Huffman
-
-  Boolean algebras as locales.
+(*  Title:      HOL/Library/Boolean_Algebra.thy
+    Author:     Brian Huffman
 *)
 
 header {* Boolean Algebras *}
--- a/src/HOL/Library/Commutative_Ring.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Library/Commutative_Ring.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -173,11 +173,11 @@
 
 text {* mkPinj preserve semantics *}
 lemma mkPinj_ci: "Ipol l (mkPinj a B) = Ipol l (Pinj a B)"
-  by (induct B) (auto simp add: mkPinj_def ring_simps)
+  by (induct B) (auto simp add: mkPinj_def algebra_simps)
 
 text {* mkPX preserves semantics *}
 lemma mkPX_ci: "Ipol l (mkPX A b C) = Ipol l (PX A b C)"
-  by (cases A) (auto simp add: mkPX_def mkPinj_ci power_add ring_simps)
+  by (cases A) (auto simp add: mkPX_def mkPinj_ci power_add algebra_simps)
 
 text {* Correctness theorems for the implemented operations *}
 
@@ -192,13 +192,13 @@
   show ?case
   proof (rule linorder_cases)
     assume "x < y"
-    with 6 show ?case by (simp add: mkPinj_ci ring_simps)
+    with 6 show ?case by (simp add: mkPinj_ci algebra_simps)
   next
     assume "x = y"
     with 6 show ?case by (simp add: mkPinj_ci)
   next
     assume "x > y"
-    with 6 show ?case by (simp add: mkPinj_ci ring_simps)
+    with 6 show ?case by (simp add: mkPinj_ci algebra_simps)
   qed
 next
   case (7 x P Q y R)
@@ -206,7 +206,7 @@
   moreover
   { assume "x = 0" with 7 have ?case by simp }
   moreover
-  { assume "x = 1" with 7 have ?case by (simp add: ring_simps) }
+  { assume "x = 1" with 7 have ?case by (simp add: algebra_simps) }
   moreover
   { assume "x > 1" from 7 have ?case by (cases x) simp_all }
   ultimately show ?case by blast
@@ -225,20 +225,20 @@
   show ?case
   proof (rule linorder_cases)
     assume a: "x < y" hence "EX d. d + x = y" by arith
-    with 9 a show ?case by (auto simp add: mkPX_ci power_add ring_simps)
+    with 9 a show ?case by (auto simp add: mkPX_ci power_add algebra_simps)
   next
     assume a: "y < x" hence "EX d. d + y = x" by arith
-    with 9 a show ?case by (auto simp add: power_add mkPX_ci ring_simps)
+    with 9 a show ?case by (auto simp add: power_add mkPX_ci algebra_simps)
   next
     assume "x = y"
-    with 9 show ?case by (simp add: mkPX_ci ring_simps)
+    with 9 show ?case by (simp add: mkPX_ci algebra_simps)
   qed
-qed (auto simp add: ring_simps)
+qed (auto simp add: algebra_simps)
 
 text {* Multiplication *}
 lemma mul_ci: "Ipol l (P \<otimes> Q) = Ipol l P * Ipol l Q"
   by (induct P Q arbitrary: l rule: mul.induct)
-    (simp_all add: mkPX_ci mkPinj_ci ring_simps add_ci power_add)
+    (simp_all add: mkPX_ci mkPinj_ci algebra_simps add_ci power_add)
 
 text {* Substraction *}
 lemma sub_ci: "Ipol l (P \<ominus> Q) = Ipol l P - Ipol l Q"
@@ -247,7 +247,7 @@
 text {* Square *}
 lemma sqr_ci: "Ipol ls (sqr P) = Ipol ls P * Ipol ls P"
   by (induct P arbitrary: ls)
-    (simp_all add: add_ci mkPinj_ci mkPX_ci mul_ci ring_simps power_add)
+    (simp_all add: add_ci mkPinj_ci mkPX_ci mul_ci algebra_simps power_add)
 
 text {* Power *}
 lemma even_pow:"even n \<Longrightarrow> pow n P = pow (n div 2) (sqr P)"
--- a/src/HOL/Library/Countable.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Library/Countable.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Library/Countable.thy
-    ID:         $Id$
     Author:     Alexander Krauss, TU Muenchen
 *)
 
--- a/src/HOL/Library/Efficient_Nat.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Library/Efficient_Nat.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -56,10 +56,10 @@
   and @{term "op mod \<Colon> nat \<Rightarrow> nat \<Rightarrow> nat"} operations. *}
 
 definition divmod_aux ::  "nat \<Rightarrow> nat \<Rightarrow> nat \<times> nat" where
-  [code del]: "divmod_aux = divmod"
+  [code del]: "divmod_aux = Divides.divmod"
 
 lemma [code]:
-  "divmod n m = (if m = 0 then (0, n) else divmod_aux n m)"
+  "Divides.divmod n m = (if m = 0 then (0, n) else divmod_aux n m)"
   unfolding divmod_aux_def divmod_div_mod by simp
 
 lemma divmod_aux_code [code]:
--- a/src/HOL/Library/Eval_Witness.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Library/Eval_Witness.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -32,7 +32,7 @@
   with the oracle.  
 *}
 
-class ml_equiv = type
+class ml_equiv
 
 text {*
   Instances of @{text "ml_equiv"} should only be declared for those types,
--- a/src/HOL/Library/Float.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Library/Float.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -38,7 +38,7 @@
   show ?thesis
   proof (induct a)
     case (1 n)
-    from pos show ?case by (simp add: ring_simps)
+    from pos show ?case by (simp add: algebra_simps)
   next
     case (2 n)
     show ?case
@@ -59,7 +59,7 @@
     show ?case by simp
   next
     case (Suc m)
-    show ?case by (auto simp add: ring_simps pow2_add1 prems)
+    show ?case by (auto simp add: algebra_simps pow2_add1 prems)
   qed
 next
   case (2 n)
@@ -72,7 +72,7 @@
       apply (subst pow2_neg[of "-1"])
       apply (simp)
       apply (insert pow2_add1[of "-a"])
-      apply (simp add: ring_simps)
+      apply (simp add: algebra_simps)
       apply (subst pow2_neg[of "-a"])
       apply (simp)
       done
@@ -83,7 +83,7 @@
       apply (auto)
       apply (subst pow2_neg[of "a + (-2 - int m)"])
       apply (subst pow2_neg[of "-2 - int m"])
-      apply (auto simp add: ring_simps)
+      apply (auto simp add: algebra_simps)
       apply (subst a)
       apply (subst b)
       apply (simp only: pow2_add1)
@@ -91,13 +91,13 @@
       apply (subst pow2_neg[of "int m + 1"])
       apply auto
       apply (insert prems)
-      apply (auto simp add: ring_simps)
+      apply (auto simp add: algebra_simps)
       done
   qed
 qed
 
 lemma "float (a, e) + float (b, e) = float (a + b, e)"
-by (simp add: float_def ring_simps)
+by (simp add: float_def algebra_simps)
 
 definition
   int_of_real :: "real \<Rightarrow> int" where
@@ -254,7 +254,7 @@
 
 lemma float_transfer_even: "even a \<Longrightarrow> float (a, b) = float (a div 2, b+1)"
   apply (subst float_transfer[where a="a" and b="b" and c="-1", simplified])
-  apply (simp_all add: pow2_def even_def real_is_int_def ring_simps)
+  apply (simp_all add: pow2_def even_def real_is_int_def algebra_simps)
   apply (auto)
 proof -
   fix q::int
@@ -319,7 +319,7 @@
   "float (a1, e1) + float (a2, e2) =
   (if e1<=e2 then float (a1+a2*2^(nat(e2-e1)), e1)
   else float (a1*2^(nat (e1-e2))+a2, e2))"
-  apply (simp add: float_def ring_simps)
+  apply (simp add: float_def algebra_simps)
   apply (auto simp add: pow2_int[symmetric] pow2_add[symmetric])
   done
 
--- a/src/HOL/Library/LaTeXsugar.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Library/LaTeXsugar.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Library/Library.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,4 +1,3 @@
-(* $Id$ *)
 (*<*)
 theory Library
 imports
@@ -36,8 +35,10 @@
   Quicksort
   Quotient
   Ramsey
+  Reflection
   RBT
   State_Monad
+  Univ_Poly
   While_Combinator
   Word
   Zorn
--- a/src/HOL/Library/Multiset.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Library/Multiset.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Library/Nat_Infinity.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -8,17 +8,6 @@
 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 {*
@@ -26,8 +15,6 @@
   infinity.
 *}
 
-end
-
 datatype inat = Fin nat | Infty
 
 notation (xsymbols)
@@ -223,10 +210,10 @@
 end
 
 lemma mult_iSuc: "iSuc m * n = n + m * n"
-  unfolding iSuc_plus_1 by (simp add: ring_simps)
+  unfolding iSuc_plus_1 by (simp add: algebra_simps)
 
 lemma mult_iSuc_right: "m * iSuc n = m + m * n"
-  unfolding iSuc_plus_1 by (simp add: ring_simps)
+  unfolding iSuc_plus_1 by (simp add: algebra_simps)
 
 lemma of_nat_eq_Fin: "of_nat n = Fin n"
   apply (induct n)
--- a/src/HOL/Library/Numeral_Type.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Library/Numeral_Type.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,11 +1,8 @@
-(*
-  ID:     $Id$
-  Author: Brian Huffman
-
-  Numeral Syntax for Types
+(*  Title:      HOL/Library/Numeral_Type.thy
+    Author:     Brian Huffman
 *)
 
-header "Numeral Syntax for Types"
+header {* Numeral Syntax for Types *}
 
 theory Numeral_Type
 imports Plain "~~/src/HOL/Presburger"
--- a/src/HOL/Library/OptionalSugar.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Library/OptionalSugar.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -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/Pocklington.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Library/Pocklington.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -254,7 +254,7 @@
   apply (simp add: nat_mod)
   apply (rule_tac x="q" in exI)
   apply (rule_tac x="q + q" in exI)
-  by (auto simp: ring_simps)
+  by (auto simp: algebra_simps)
 
 lemma cong_to_1: "[a = 1] (mod n) \<longleftrightarrow> a = 0 \<and> n = 1 \<or> (\<exists>m. a = 1 + m * n)"
 proof-
@@ -780,7 +780,7 @@
 	of "(n - 1) div m * m"]
       have yn: "coprime ?y n" by (simp add: coprime_commute) 
       have "?y mod n = (a^m)^((n - 1) div m) mod n" 
-	by (simp add: ring_simps power_mult)
+	by (simp add: algebra_simps power_mult)
       also have "\<dots> = (a^m mod n)^((n - 1) div m) mod n" 
 	using power_mod[of "a^m" n "(n - 1) div m"] by simp
       also have "\<dots> = 1" using m(3)[unfolded modeq_def onen] onen 
@@ -1239,7 +1239,7 @@
   from bqn psp qrn
   have bqn: "a ^ (n - 1) mod n = 1"
     and psp: "list_all (\<lambda>p. prime p \<and> coprime (a^(r *(q div p)) mod n - 1) n) ps"  unfolding arnb[symmetric] power_mod 
-    by (simp_all add: power_mult[symmetric] ring_simps)
+    by (simp_all add: power_mult[symmetric] algebra_simps)
   from n  have n0: "n > 0" by arith
   from mod_div_equality[of "a^(n - 1)" n]
     mod_less_divisor[OF n0, of "a^(n - 1)"]
@@ -1248,7 +1248,7 @@
     apply -
     apply (rule exI[where x="0"])
     apply (rule exI[where x="a^(n - 1) div n"])
-    by (simp add: ring_simps)
+    by (simp add: algebra_simps)
   {fix p assume p: "prime p" "p dvd q"
     from psp psq have pfpsq: "primefact ps q"
       by (auto simp add: primefact_variant list_all_iff)
--- a/src/HOL/Library/Primes.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Library/Primes.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -204,7 +204,7 @@
     from z have z': "?g > 0" by simp
     from bezout_gcd_strong[OF az, of b] 
     obtain x y where xy: "a*x = b*y + ?g" by blast
-    from xy a b have "?g * a'*x = ?g * (b'*y + 1)" by (simp add: ring_simps)
+    from xy a b have "?g * a'*x = ?g * (b'*y + 1)" by (simp add: algebra_simps)
     hence "?g * (a'*x) = ?g * (b'*y + 1)" by (simp add: mult_assoc)
     hence "a'*x = (b'*y + 1)"
       by (simp only: nat_mult_eq_cancel1[OF z']) 
--- a/src/HOL/Library/Quotient.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Library/Quotient.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -21,7 +21,7 @@
  "\<sim> :: 'a => 'a => bool"}.
 *}
 
-class eqv = type +
+class eqv =
   fixes eqv :: "'a \<Rightarrow> 'a \<Rightarrow> bool"    (infixl "\<sim>" 50)
 
 class equiv = eqv +
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Reflection.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -0,0 +1,45 @@
+(*  Title:      HOL/Library/Reflection.thy
+    Author:     Amine Chaieb, TU Muenchen
+*)
+
+header {* Generic reflection and reification *}
+
+theory Reflection
+imports Main
+uses "reify_data.ML" ("reflection.ML")
+begin
+
+setup {* Reify_Data.setup *}
+
+lemma ext2: "(\<forall>x. f x = g x) \<Longrightarrow> f = g"
+  by (blast intro: ext)
+
+use "reflection.ML"
+
+method_setup reify = {* fn src =>
+  Method.syntax (Attrib.thms --
+    Scan.option (Scan.lift (Args.$$$ "(") |-- Args.term --| Scan.lift (Args.$$$ ")") )) src #>
+  (fn ((eqs, to), ctxt) => Method.SIMPLE_METHOD' (Reflection.genreify_tac ctxt (eqs @ (fst (Reify_Data.get ctxt))) to))
+*} "partial automatic reification"
+
+method_setup reflection = {* 
+let 
+  fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ();
+  val onlyN = "only";
+  val rulesN = "rules";
+  val any_keyword = keyword onlyN || keyword rulesN;
+  val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat;
+  val terms = thms >> map (term_of o Drule.dest_term);
+  fun optional scan = Scan.optional scan [];
+in fn src =>
+  Method.syntax (thms -- optional (keyword rulesN |-- thms) -- Scan.option (keyword onlyN |-- Args.term)) src #> 
+    (fn (((eqs,ths),to), ctxt) => 
+      let 
+        val (ceqs,cths) = Reify_Data.get ctxt 
+        val corr_thms = ths@cths
+        val raw_eqs = eqs@ceqs
+      in Method.SIMPLE_METHOD' (Reflection.reflection_tac ctxt corr_thms raw_eqs to) 
+     end) end
+*} "reflection method"
+
+end
--- a/src/HOL/Library/SetsAndFunctions.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Library/SetsAndFunctions.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -101,32 +101,32 @@
 
 instance "fun" :: (type,comm_ring_1)comm_ring_1
   apply default
-   apply (auto simp add: func_plus func_times fun_Compl_def fun_diff_def ext
-     func_one func_zero ring_simps)
+   apply (auto simp add: func_plus func_times fun_Compl_def fun_diff_def
+     func_one func_zero algebra_simps)
   apply (drule fun_cong)
   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)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Univ_Poly.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -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: algebra_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: algebra_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 simp add: expand_fun_eq 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: algebra_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 algebra_simps)
+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 algebra_simps) simp
+  hence "poly ((c#cs) +++ -- (d#ds)) = poly []" by(simp add:expand_fun_eq)
+  hence "c = d \<and> list_all (\<lambda>x. x=0) ((cs +++ -- ds))"
+    unfolding poly_zero by (simp add: poly_minus_def algebra_simps)
+  hence "c = d \<and> (\<forall>x. poly (cs +++ -- ds) x = 0)"
+    unfolding poly_zero[symmetric] by simp 
+  thus ?rhs  by (simp add: poly_minus poly_add algebra_simps expand_fun_eq)
+next
+  assume ?rhs then show ?lhs by(simp add:expand_fun_eq)
+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 algebra_simps)
+   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
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/reflection.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -0,0 +1,327 @@
+(*  Title:      HOL/Library/reflection.ML
+    Author:     Amine Chaieb, TU Muenchen
+
+A trial for automatical reification.
+*)
+
+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
+
+val ext2 = @{thm ext2};
+val nth_Cons_0 = @{thm nth_Cons_0};
+val nth_Cons_Suc = @{thm nth_Cons_Suc};
+
+  (* Make a congruence rule out of a defining equation for the interpretation *)
+  (* th is one defining equation of f, i.e.
+     th is "f (Cp ?t1 ... ?tn) = P(f ?t1, .., f ?tn)" *)
+  (* Cp is a constructor pattern and P is a pattern *)
+
+  (* The result is:
+      [|?A1 = f ?t1 ; .. ; ?An= f ?tn |] ==> P (?A1, .., ?An) = f (Cp ?t1 .. ?tn) *)
+  (*  + the a list of names of the A1 .. An, Those are fresh in the ctxt*)
+
+
+fun mk_congeq ctxt fs th = 
+  let 
+   val (f as Const(fN,fT)) = th |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq 
+     |> fst |> strip_comb |> fst
+   val thy = ProofContext.theory_of ctxt
+   val cert = Thm.cterm_of thy
+   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 => 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 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'
+   val tys = map fastype_of fterms
+   val vs = map Free (xs ~~ tys)
+   val env = fterms ~~ vs
+		    (* FIXME!!!!*)	
+   fun replace_fterms (t as t1 $ t2) =
+       (case AList.lookup (op aconv) env t of
+	    SOME v => v
+	  | NONE => replace_fterms t1 $ replace_fterms t2)
+     | replace_fterms t = (case AList.lookup (op aconv) env t of
+			       SOME v => v
+			     | NONE => t)
+      
+   fun mk_def (Abs(x,xT,t),v) = HOLogic.mk_Trueprop ((HOLogic.all_const xT)$ Abs(x,xT,HOLogic.mk_eq(v$(Bound 0), t)))
+     | mk_def (t, v) = HOLogic.mk_Trueprop (HOLogic.mk_eq (v, t))
+   fun tryext x = (x RS ext2 handle THM _ =>  x)
+   val cong = (Goal.prove ctxt'' [] (map mk_def env)
+			  (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, replace_fterms rhs)))
+			  (fn x => LocalDefs.unfold_tac (#context x) (map tryext (#prems x)) 
+							THEN rtac th' 1)) RS sym
+	      
+   val (cong' :: vars') = 
+       Variable.export ctxt'' ctxt (cong :: map (Drule.mk_term o cert) vs)
+   val vs' = map (fst o fst o Term.dest_Var o Thm.term_of o Drule.dest_term) vars'
+					      
+  in  (vs', cong') end; 
+ (* congs is a list of pairs (P,th) where th is a theorem for *)
+        (* [| f p1 = A1; ...; f pn = An|] ==> f (C p1 .. pn) = P *)
+val FWD = curry (op OF);
+
+ (* da is the decomposition for atoms, ie. it returns ([],g) where g
+ returns the right instance f (AtC n) = t , where AtC is the Atoms
+ constructor and n is the number of the atom corresponding to t *)
+
+(* Generic decomp for reification : matches the actual term with the
+rhs of one cong rule. The result of the matching guides the
+proof synthesis: The matches of the introduced Variables A1 .. An are
+processed recursively
+ The rest is instantiated in the cong rule,i.e. no reification is needed *)
+
+exception REIF of string;
+
+val bds = ref ([]: (typ * ((term list) * (term list))) list);
+
+fun index_of t = 
+ let 
+  val tt = HOLogic.listT (fastype_of t)
+ in 
+  (case AList.lookup Type.could_unify (!bds) tt of
+    NONE => error "index_of : type not found in environements!"
+  | SOME (tbs,tats) =>
+    let
+     val i = find_index_eq t tats
+     val j = find_index_eq t tbs 
+    in (if j= ~1 then 
+	    if i= ~1 
+	    then (bds := AList.update Type.could_unify (tt,(tbs,tats@[t])) (!bds) ; 
+		  length tbs + length tats) 
+	    else i else j)
+    end)
+ end;
+
+fun dest_listT (Type ("List.list", [T])) = T;
+
+fun decomp_genreif da cgns (t,ctxt) =
+ let 
+  val thy = ProofContext.theory_of ctxt 
+  val cert = cterm_of thy
+  fun tryabsdecomp (s,ctxt) = 
+   (case s of 
+     Abs(xn,xT,ta) => 
+     (let
+       val ([xn],ctxt') = Variable.variant_fixes ["x"] ctxt
+       val (xn,ta) = variant_abs (xn,xT,ta)
+       val x = Free(xn,xT)
+       val _ = (case AList.lookup Type.could_unify (!bds) (HOLogic.listT xT)
+		 of NONE => error "tryabsdecomp: Type not found in the Environement"
+		  | SOME (bsT,atsT) => 
+		    (bds := AList.update Type.could_unify (HOLogic.listT xT, ((x::bsT), atsT)) (!bds)))
+      in ([(ta, ctxt')] , 
+	  fn [th] => ((let val (bsT,asT) = the(AList.lookup Type.could_unify (!bds) (HOLogic.listT xT))
+		       in (bds := AList.update Type.could_unify (HOLogic.listT xT,(tl bsT,asT)) (!bds))
+		       end) ; 
+		      hd (Variable.export ctxt' ctxt [(forall_intr (cert x) th) COMP allI])))
+	end)
+    | _ => da (s,ctxt))
+  in 
+  (case cgns of 
+    [] => tryabsdecomp (t,ctxt)
+  | ((vns,cong)::congs) => ((let
+        val cert = cterm_of thy
+	val certy = ctyp_of thy
+        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), Vartab.empty)
+        val (fnvs,invs) = List.partition (fn ((vn,_),_) => vn mem vns) (Vartab.dest tmenv)
+        val (fts,its) = 
+	    (map (snd o snd) fnvs,
+             map (fn ((vn,vi),(tT,t)) => (cert(Var ((vn,vi),tT)), cert t)) invs)
+	val ctyenv = map (fn ((vn,vi),(s,ty)) => (certy (TVar((vn,vi),s)), certy ty)) (Vartab.dest tyenv)
+    in (fts ~~ (replicate (length fts) ctxt), FWD (instantiate (ctyenv, its) cong))
+    end)
+      handle MATCH => decomp_genreif da congs (t,ctxt)))
+  end;
+
+ (* looks for the atoms equation and instantiates it with the right number *)
+
+
+fun mk_decompatom eqs (t,ctxt) =
+let 
+ val tT = fastype_of t
+ fun isat eq = 
+  let 
+   val rhs = eq |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd
+   in exists_Const 
+	  (fn (n,ty) => n="List.nth" 
+			andalso 
+			AList.defined Type.could_unify (!bds) (domain_type ty)) rhs 
+	  andalso Type.could_unify (fastype_of rhs, tT)
+   end
+ fun get_nths t acc = 
+  case t of
+    Const("List.nth",_)$vs$n => insert (fn ((a,_),(b,_)) => a aconv b) (t,(vs,n)) acc
+  | t1$t2 => get_nths t1 (get_nths t2 acc)
+  | Abs(_,_,t') => get_nths t'  acc
+  | _ => acc
+
+ fun 
+   tryeqs [] = error "Can not find the atoms equation"
+ | tryeqs (eq::eqs) = ((
+  let 
+   val rhs = eq |> prop_of |> HOLogic.dest_Trueprop  |> HOLogic.dest_eq |> snd
+   val nths = get_nths rhs []
+   val (vss,ns) = fold_rev (fn (_,(vs,n)) => fn (vss,ns) => 
+                             (insert (op aconv) vs vss, insert (op aconv) n ns)) nths ([],[]) 
+   val (vsns, ctxt') = Variable.variant_fixes (replicate (length vss) "vs") ctxt
+   val (xns, ctxt'') = Variable.variant_fixes (replicate (length nths) "x") ctxt' 
+   val thy = ProofContext.theory_of ctxt''
+   val cert = cterm_of thy
+   val certT = ctyp_of thy
+   val vsns_map = vss ~~ vsns
+   val xns_map = (fst (split_list nths)) ~~ xns
+   val subst = map (fn (nt, xn) => (nt, Var ((xn,0), fastype_of nt))) xns_map
+   val rhs_P = subst_free subst rhs
+   val (tyenv, tmenv) = Pattern.match 
+	                    thy (rhs_P, t)
+	                    (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)) 
+                      (Vartab.dest tyenv)
+   val tml = Vartab.dest tmenv
+   val t's = map (fn xn => snd (valOf (AList.lookup (op =) tml (xn,0)))) xns (* FIXME : Express with sbst*)
+   val subst_ns = map (fn (Const _ $ vs $ n, Var (xn0,T)) => 
+                          (cert n, snd (valOf (AList.lookup (op =) tml xn0)) 
+                             |> (index_of #> HOLogic.mk_nat #> cert))) 
+                      subst
+   val subst_vs = 
+    let 
+     fun ty (Const _ $ (vs as Var (vsn,lT)) $ n, Var (xn0,T)) = (certT T, certT (sbsT T))
+     fun h (Const _ $ (vs as Var (vsn,lT)) $ n, Var (xn0,T)) = 
+      let 
+       val cns = sbst (Const("List.list.Cons", T --> lT --> lT))
+       val lT' = sbsT lT
+       val (bsT,asT) = the (AList.lookup Type.could_unify (!bds) lT)
+       val vsn = valOf (AList.lookup (op =) vsns_map vs)
+       val cvs = cert (fold_rev (fn x => fn xs => cns$x$xs) bsT (Free (vsn, lT')))
+      in (cert vs, cvs) end
+    in map h subst end
+   val cts = map (fn ((vn,vi),(tT,t)) => (cert(Var ((vn,vi),tT)), cert t)) 
+                 (fold (AList.delete (fn (((a: string),_),(b,_)) => a = b)) 
+                       (map (fn n => (n,0)) xns) tml)
+   val substt = 
+    let val ih = Drule.cterm_rule (Thm.instantiate (subst_ty,[]))
+    in map (fn (v,t) => (ih v, ih t)) (subst_ns@subst_vs@cts)  end
+   val th = (instantiate (subst_ty, substt)  eq) RS sym
+  in  hd (Variable.export ctxt'' ctxt [th]) end)
+ handle MATCH => tryeqs eqs)
+in ([], fn _ => tryeqs (filter isat eqs))
+end;
+
+  (* Generic reification procedure: *)
+  (* creates all needed cong rules and then just uses the theorem synthesis *)
+
+  fun mk_congs ctxt raw_eqs = 
+ let
+  val fs = fold_rev (fn eq =>
+		     insert (op =) (eq |> prop_of |> HOLogic.dest_Trueprop 
+			 |> HOLogic.dest_eq |> fst |> strip_comb 
+			 |> fst)) raw_eqs []
+  val tys = fold_rev (fn f => fn ts => (f |> fastype_of |> binder_types |> tl) 
+				    union ts) fs []
+  val _ = bds := AList.make (fn _ => ([],[])) tys
+  val (vs, ctxt') = Variable.variant_fixes (replicate (length tys) "vs") ctxt
+  val thy = ProofContext.theory_of ctxt'
+  val cert = cterm_of thy
+  val vstys = map (fn (t,v) => (t,SOME (cert (Free(v,t))))) 
+		  (tys ~~ vs)
+  val is_Var = can dest_Var
+  fun insteq eq vs = 
+   let
+     val subst = map (fn (v as Var(n,t)) => (cert v, (valOf o valOf) (AList.lookup (op =) vstys t)))  
+  (filter is_Var vs)
+   in Thm.instantiate ([],subst) eq
+   end
+  val eqs = map (fn eq => eq |> prop_of |> HOLogic.dest_Trueprop 
+			     |> HOLogic.dest_eq |> fst |> strip_comb |> snd |> tl
+			     |> (insteq eq)) raw_eqs
+  val (ps,congs) = split_list (map (mk_congeq ctxt' fs) eqs)
+in ps ~~ (Variable.export ctxt' ctxt congs)
+end
+
+fun partition P [] = ([],[])
+  | partition P (x::xs) = 
+     let val (yes,no) = partition P xs
+     in if P x then (x::yes,no) else (yes, x::no) end
+
+fun rearrange congs = 
+let 
+ fun P (_, th) = 
+  let val @{term "Trueprop"}$(Const ("op =",_) $l$_) = concl_of th
+  in can dest_Var l end
+ val (yes,no) = partition P congs 
+ in no @ yes end
+
+
+
+fun genreif ctxt raw_eqs t =
+ let 
+  val congs = rearrange (mk_congs ctxt raw_eqs)
+  val th = divide_and_conquer (decomp_genreif (mk_decompatom raw_eqs) congs) (t,ctxt)
+  fun is_listVar (Var (_,t)) = can dest_listT t
+       | is_listVar _ = false
+  val vars = th |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd
+	       |> strip_comb |> snd |> filter is_listVar
+  val cert = cterm_of (ProofContext.theory_of ctxt)
+  val cvs = map (fn (v as Var(n,t)) => (cert v, the (AList.lookup Type.could_unify (!bds) t) |> snd |> HOLogic.mk_list (dest_listT t) |> cert)) vars
+  val th' = instantiate ([], cvs) th
+  val t' = (fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) th'
+  val th'' = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (t, t')))
+			(fn _ => simp_tac (local_simpset_of ctxt) 1)
+  val _ = bds := []
+in FWD trans [th'',th']
+end
+
+
+fun genreflect ctxt conv corr_thms raw_eqs t =
+let 
+  val reifth = genreif ctxt raw_eqs t
+  fun trytrans [] = error "No suitable correctness theorem found"
+    | trytrans (th::ths) = 
+         (FWD trans [reifth, th RS sym] handle THM _ => trytrans ths)
+  val th = trytrans corr_thms
+  val ft = (Thm.dest_arg1 o Thm.dest_arg o Thm.dest_arg o cprop_of) th
+  val rth = conv ft
+in simplify (HOL_basic_ss addsimps raw_eqs addsimps [nth_Cons_0, nth_Cons_Suc])
+           (simplify (HOL_basic_ss addsimps [rth]) th)
+end
+
+fun genreify_tac ctxt eqs to i = (fn st =>
+  let
+    val P = HOLogic.dest_Trueprop (List.nth (prems_of st, i - 1))
+    val t = (case to of NONE => P | SOME x => x)
+    val th = (genreif ctxt eqs t) RS ssubst
+  in rtac th i st
+  end);
+
+    (* Reflection calls reification and uses the correctness *)
+        (* theorem assumed to be the dead of the list *)
+fun gen_reflection_tac ctxt conv corr_thms raw_eqs to i = (fn st =>
+  let
+    val P = HOLogic.dest_Trueprop (nth (prems_of st) (i - 1));
+    val t = the_default P to;
+    val th = genreflect ctxt conv corr_thms raw_eqs t
+      RS ssubst;
+  in (rtac th i THEN TRY(rtac TrueI i)) st end);
+
+fun reflection_tac ctxt = gen_reflection_tac ctxt Codegen.evaluation_conv;
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/reify_data.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -0,0 +1,39 @@
+(*  Title:      HOL/Library/reflection_data.ML
+    Author:     Amine Chaieb, TU Muenchen
+
+Data for the reification and reflection methods.
+*)
+
+signature REIFY_DATA =
+sig
+  val get: Proof.context -> thm list * thm list
+  val add: attribute
+  val del: attribute
+  val radd: attribute
+  val rdel: attribute
+  val setup: theory -> theory
+end;
+
+structure Reify_Data : REIFY_DATA =
+struct
+
+structure Data = GenericDataFun
+(
+  type T = thm list * thm list;
+  val empty = ([], []);
+  val extend = I;
+  fun merge _ = pairself Thm.merge_thms;
+);
+
+val get = Data.get o Context.Proof;
+
+val add = Thm.declaration_attribute (Data.map o apfst o Thm.add_thm);
+val del = Thm.declaration_attribute (Data.map o apfst o Thm.del_thm);
+val radd = Thm.declaration_attribute (Data.map o apsnd o Thm.add_thm);
+val rdel = Thm.declaration_attribute (Data.map o apsnd o Thm.del_thm);
+
+val setup = Attrib.add_attributes
+  [("reify", Attrib.add_del_args add del, "Reify data"),
+   ("reflection", Attrib.add_del_args radd rdel, "Reflection data")];
+
+end;
--- a/src/HOL/Lim.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Lim.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -52,7 +52,7 @@
 apply (drule_tac r="r" in LIM_D, safe)
 apply (rule_tac x="s" in exI, safe)
 apply (drule_tac x="x + k" in spec)
-apply (simp add: compare_rls)
+apply (simp add: algebra_simps)
 done
 
 lemma LIM_offset_zero: "f -- a --> L \<Longrightarrow> (\<lambda>h. f (a + h)) -- 0 --> L"
--- a/src/HOL/List.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/List.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -547,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 = [])"
@@ -2887,6 +2887,35 @@
 apply (auto simp: sorted_distinct_set_unique)
 done
 
+lemma sorted_take:
+  "sorted xs \<Longrightarrow> sorted (take n xs)"
+proof (induct xs arbitrary: n rule: sorted.induct)
+  case 1 show ?case by simp
+next
+  case 2 show ?case by (cases n) simp_all
+next
+  case (3 x y xs)
+  then have "x \<le> y" by simp
+  show ?case proof (cases n)
+    case 0 then show ?thesis by simp
+  next
+    case (Suc m) 
+    with 3 have "sorted (take m (y # xs))" by simp
+    with Suc  `x \<le> y` show ?thesis by (cases m) simp_all
+  qed
+qed
+
+lemma sorted_drop:
+  "sorted xs \<Longrightarrow> sorted (drop n xs)"
+proof (induct xs arbitrary: n rule: sorted.induct)
+  case 1 show ?case by simp
+next
+  case 2 show ?case by (cases n) simp_all
+next
+  case 3 then show ?case by (cases n) simp_all
+qed
+
+
 end
 
 lemma sorted_upt[simp]: "sorted[i..<j]"
--- a/src/HOL/Ln.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Ln.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -187,7 +187,7 @@
 proof -
   assume a: "0 <= (x::real)" and b: "x < 1"
   have "(1 - x) * (1 + x + x^2) = (1 - x^3)"
-    by (simp add: ring_simps power2_eq_square power3_eq_cube)
+    by (simp add: algebra_simps power2_eq_square power3_eq_cube)
   also have "... <= 1"
     by (auto simp add: a)
   finally have "(1 - x) * (1 + x + x ^ 2) <= 1" .
@@ -325,10 +325,10 @@
     done
   also have "... <= 2 * x^2"
     apply (subgoal_tac "- (-x) - 2 * (-x)^2 <= ln (1 - (-x))")
-    apply (simp add: compare_rls)
+    apply (simp add: algebra_simps)
     apply (rule ln_one_minus_pos_lower_bound)
     apply (insert prems, auto)
-    done 
+    done
   finally show ?thesis .
 qed
 
@@ -375,7 +375,7 @@
   apply simp
   apply (subst ln_div [THEN sym])
   apply arith
-  apply (auto simp add: ring_simps add_frac_eq frac_eq_eq 
+  apply (auto simp add: algebra_simps add_frac_eq frac_eq_eq 
     add_divide_distrib power2_eq_square)
   apply (rule mult_pos_pos, assumption)+
   apply assumption
@@ -394,7 +394,7 @@
     apply auto
     done
   have "x * ln y - x * ln x = x * (ln y - ln x)"
-    by (simp add: ring_simps)
+    by (simp add: algebra_simps)
   also have "... = x * ln(y / x)"
     apply (subst ln_div)
     apply (rule b, rule a, rule refl)
--- a/src/HOL/Lubs.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Lubs.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title       : Lubs.thy
-    ID          : $Id$
     Author      : Jacques D. Fleuriot
     Copyright   : 1998  University of Cambridge
 *)
@@ -7,7 +6,7 @@
 header{*Definitions of Upper Bounds and Least Upper Bounds*}
 
 theory Lubs
-imports Plain
+imports Plain Main
 begin
 
 text{*Thanks to suggestions by James Margetson*}
--- a/src/HOL/Map.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Map.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -503,6 +503,15 @@
 lemma map_add_comm: "dom m1 \<inter> dom m2 = {} \<Longrightarrow> m1++m2 = m2++m1"
 by (rule ext) (force simp: map_add_def dom_def split: option.split)
 
+lemma dom_const [simp]:
+  "dom (\<lambda>x. Some y) = UNIV"
+  by auto
+
+lemma dom_if:
+  "dom (\<lambda>x. if P x then f x else g x) = dom f \<inter> {x. P x} \<union> dom g \<inter> {x. \<not> P x}"
+  by (auto split: if_splits)
+
+
 (* Due to John Matthews - could be rephrased with dom *)
 lemma finite_map_freshness:
   "finite (dom (f :: 'a \<rightharpoonup> 'b)) \<Longrightarrow> \<not> finite (UNIV :: 'a set) \<Longrightarrow>
--- a/src/HOL/Matrix/LP.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Matrix/LP.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -20,7 +20,7 @@
 proof -
   from prems have 1: "y * b <= y * b'" by (simp add: mult_left_mono)
   from prems have 2: "y * (A * x) <= y * b" by (simp add: mult_left_mono) 
-  have 3: "y * (A * x) = c * x + (y * (A - A') + (y * A' - c') + (c'-c)) * x" by (simp add: ring_simps)  
+  have 3: "y * (A * x) = c * x + (y * (A - A') + (y * A' - c') + (c'-c)) * x" by (simp add: algebra_simps)  
   from 1 2 3 have 4: "c * x + (y * (A - A') + (y * A' - c') + (c'-c)) * x <= y * b'" by simp
   have 5: "c * x <= y * b' + abs((y * (A - A') + (y * A' - c') + (c'-c)) * x)"
     by (simp only: 4 estimate_by_abs)  
@@ -32,7 +32,7 @@
     by (simp add: abs_triangle_ineq mult_right_mono)    
   have 9: "(abs (y * (A-A')) + abs (y*A'-c') + abs(c'-c)) * abs x <= (abs y * abs (A-A') + abs (y*A'-c') + abs (c'-c)) * abs x"
     by (simp add: abs_le_mult mult_right_mono)  
-  have 10: "c'-c = -(c-c')" by (simp add: ring_simps)
+  have 10: "c'-c = -(c-c')" by (simp add: algebra_simps)
   have 11: "abs (c'-c) = abs (c-c')" 
     by (subst 10, subst abs_minus_cancel, simp)
   have 12: "(abs y * abs (A-A') + abs (y*A'-c') + abs (c'-c)) * abs x <= (abs y * abs (A-A') + abs (y*A'-c') + \<delta>c) * abs x"
@@ -85,7 +85,7 @@
     apply simp
     done
   then have "a * b = pprt a * pprt b + pprt a * nprt b + nprt a * pprt b + nprt a * nprt b"
-    by (simp add: ring_simps)
+    by (simp add: algebra_simps)
   moreover have "pprt a * pprt b <= pprt a2 * pprt b2"
     by (simp_all add: prems mult_mono)
   moreover have "pprt a * nprt b <= pprt a1 * nprt b2"
@@ -134,10 +134,10 @@
   (is "_ <= _ + ?C")
 proof -
   from prems have "y * (A * x) <= y * b" by (simp add: mult_left_mono) 
-  moreover have "y * (A * x) = c * x + (y * A - c) * x" by (simp add: ring_simps)  
+  moreover have "y * (A * x) = c * x + (y * A - c) * x" by (simp add: algebra_simps)  
   ultimately have "c * x + (y * A - c) * x <= y * b" by simp
   then have "c * x <= y * b - (y * A - c) * x" by (simp add: le_diff_eq)
-  then have cx: "c * x <= y * b + (c - y * A) * x" by (simp add: ring_simps)
+  then have cx: "c * x <= y * b + (c - y * A) * x" by (simp add: algebra_simps)
   have s2: "c - y * A <= c2 - y * A1"
     by (simp add: diff_def prems add_mono mult_left_mono)
   have s1: "c1 - y * A2 <= c - y * A"
--- a/src/HOL/Matrix/Matrix.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Matrix/Matrix.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1573,17 +1573,17 @@
   show "A * B * C = A * (B * C)"
     apply (simp add: times_matrix_def)
     apply (rule mult_matrix_assoc)
-    apply (simp_all add: associative_def ring_simps)
+    apply (simp_all add: associative_def algebra_simps)
     done
   show "(A + B) * C = A * C + B * C"
     apply (simp add: times_matrix_def plus_matrix_def)
     apply (rule l_distributive_matrix[simplified l_distributive_def, THEN spec, THEN spec, THEN spec])
-    apply (simp_all add: associative_def commutative_def ring_simps)
+    apply (simp_all add: associative_def commutative_def algebra_simps)
     done
   show "A * (B + C) = A * B + A * C"
     apply (simp add: times_matrix_def plus_matrix_def)
     apply (rule r_distributive_matrix[simplified r_distributive_def, THEN spec, THEN spec, THEN spec])
-    apply (simp_all add: associative_def commutative_def ring_simps)
+    apply (simp_all add: associative_def commutative_def algebra_simps)
     done
 qed  
 
@@ -1793,7 +1793,7 @@
 by (simp add: scalar_mult_def)
 
 lemma scalar_mult_add: "scalar_mult y (a+b) = (scalar_mult y a) + (scalar_mult y b)"
-by (simp add: scalar_mult_def apply_matrix_add ring_simps)
+by (simp add: scalar_mult_def apply_matrix_add algebra_simps)
 
 lemma Rep_scalar_mult[simp]: "Rep_matrix (scalar_mult y a) j i = y * (Rep_matrix a j i)" 
 by (simp add: scalar_mult_def)
--- a/src/HOL/Matrix/SparseMatrix.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Matrix/SparseMatrix.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -274,7 +274,7 @@
     apply (rule conjI)
     apply (intro strip)
     apply (frule_tac as=brr in sorted_spvec_cons1)
-    apply (simp add: ring_simps sparse_row_matrix_cons)
+    apply (simp add: algebra_simps sparse_row_matrix_cons)
     apply (simplesubst Rep_matrix_zero_imp_mult_zero) 
     apply (simp)
     apply (intro strip)
@@ -296,7 +296,7 @@
     
     apply (intro strip | rule conjI)+
     apply (frule_tac as=arr in sorted_spvec_cons1)
-    apply (simp add: ring_simps)
+    apply (simp add: algebra_simps)
     apply (subst Rep_matrix_zero_imp_mult_zero)
     apply (simp)
     apply (rule disjI2)
@@ -310,7 +310,7 @@
     apply (simp_all)
     apply (frule_tac as=arr in sorted_spvec_cons1)
     apply (frule_tac as=brr in sorted_spvec_cons1)
-    apply (simp add: sparse_row_matrix_cons ring_simps sparse_row_vector_addmult_spvec)
+    apply (simp add: sparse_row_matrix_cons algebra_simps sparse_row_vector_addmult_spvec)
     apply (rule_tac B1 = "sparse_row_matrix brr" in ssubst[OF Rep_matrix_zero_imp_mult_zero])
     apply (auto)
     apply (rule sorted_sparse_row_matrix_zero)
@@ -360,7 +360,7 @@
 lemma sparse_row_mult_spmat[rule_format]: 
   "sorted_spmat A \<longrightarrow> sorted_spvec B \<longrightarrow> sparse_row_matrix (mult_spmat A B) = (sparse_row_matrix A) * (sparse_row_matrix B)"
   apply (induct A)
-  apply (auto simp add: sparse_row_matrix_cons sparse_row_mult_spvec_spmat ring_simps move_matrix_mult)
+  apply (auto simp add: sparse_row_matrix_cons sparse_row_mult_spvec_spmat algebra_simps move_matrix_mult)
   done
 
 lemma sorted_spvec_mult_spmat[rule_format]:
--- a/src/HOL/Matrix/cplex/Cplex.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Matrix/cplex/Cplex.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -34,8 +34,8 @@
   add_spmat (mult_spmat (nprt_spmat s2) (pprt_spmat r1), mult_spmat (nprt_spmat s1) (nprt_spmat r1)))))))"
   apply (simp add: Let_def)
   apply (insert assms)
-  apply (simp add: sparse_row_matrix_op_simps ring_simps)  
-  apply (rule mult_le_dual_prts[where A=A, simplified Let_def ring_simps])
+  apply (simp add: sparse_row_matrix_op_simps algebra_simps)  
+  apply (rule mult_le_dual_prts[where A=A, simplified Let_def algebra_simps])
   apply (auto)
   done
 
--- a/src/HOL/MetisExamples/BT.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/MetisExamples/BT.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/MetisExamples/BigO.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -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
 
@@ -489,7 +486,7 @@
 have 2: "\<And>X3. (0\<Colon>'b) + X3 = X3"
   by (metis diff_eq_eq right_minus_eq)
 have 3: "\<not> (0\<Colon>'b) \<le> f x - lb x"
-  by (metis 1 compare_rls(1))
+  by (metis 1 diff_minus)
 have 4: "\<not> (0\<Colon>'b) + lb x \<le> f x"
   by (metis 3 le_diff_eq)
 show "False"
@@ -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"*}
@@ -1200,13 +1197,13 @@
 apply (erule thin_rl) 
 (*sledgehammer*); 
   apply (case_tac "0 <= f x - k x")
-  apply (simp del: compare_rls diff_minus);
+  apply (simp)
   apply (subst abs_of_nonneg)
   apply (drule_tac x = x in spec) back
 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) ==>
@@ -1219,7 +1216,7 @@
   apply assumption
   apply (erule bigo_lesseq2) back
   apply (rule allI)
-  apply (auto simp add: func_plus fun_diff_def compare_rls 
+  apply (auto simp add: func_plus fun_diff_def algebra_simps
     split: split_max abs_split)
 done
 
--- a/src/HOL/Nat.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Nat.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -425,6 +425,17 @@
 
 end
 
+instantiation nat :: bot
+begin
+
+definition bot_nat :: nat where
+  "bot_nat = 0"
+
+instance proof
+qed (simp add: bot_nat_def)
+
+end
+
 subsubsection {* Introduction properties *}
 
 lemma lessI [iff]: "n < Suc n"
@@ -1254,7 +1265,7 @@
 begin
 
 lemma of_nat_diff: "n \<le> m \<Longrightarrow> of_nat (m - n) = of_nat m - of_nat n"
-  by (simp add: compare_rls of_nat_add [symmetric])
+by (simp add: algebra_simps of_nat_add [symmetric])
 
 end
 
@@ -1515,7 +1526,7 @@
 
 subsection {* size of a datatype value *}
 
-class size = type +
+class size =
   fixes size :: "'a \<Rightarrow> nat" -- {* see further theory @{text Wellfounded} *}
 
 end
--- a/src/HOL/Nominal/Examples/W.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Nominal/Examples/W.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -49,7 +49,7 @@
 
 text {* free type variables *}
 
-class ftv = type +
+class ftv =
   fixes ftv :: "'a \<Rightarrow> tvar list"
 
 instantiation * :: (ftv, ftv) ftv
--- a/src/HOL/Nominal/nominal_atoms.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Nominal/nominal_atoms.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -90,6 +90,9 @@
   let val T = fastype_of x
   in Const ("List.list.Cons", T --> HOLogic.listT T --> HOLogic.listT T) $ x $ xs end;
 
+fun add_thms_string args = PureThy.add_thms ((map o apfst o apfst) Binding.name args);
+fun add_thmss_string args = PureThy.add_thmss ((map o apfst o apfst) Binding.name args);
+
 (* this function sets up all matters related to atom-  *)
 (* kinds; the user specifies a list of atom-kind names *)
 (* atom_decl <ak1> ... <akn>                           *)
@@ -121,7 +124,7 @@
                                           atac 1]
              
               val (inj_thm,thy2) = 
-                   PureThy.add_thms [((ak^"_inj",Goal.prove_global thy1 [] [] stmnt1 proof1), [])] thy1
+                   add_thms_string [((ak^"_inj",Goal.prove_global thy1 [] [] stmnt1 proof1), [])] thy1
               
               (* second statement *)
               val y = Free ("y",ak_type)  
@@ -136,7 +139,7 @@
 
               (* third statement *)
               val (inject_thm,thy3) =
-                  PureThy.add_thms [((ak^"_injection",Goal.prove_global thy2 [] [] stmnt2 proof2), [])] thy2
+                  add_thms_string [((ak^"_injection",Goal.prove_global thy2 [] [] stmnt2 proof2), [])] thy2
   
               val stmnt3 = HOLogic.mk_Trueprop
                            (HOLogic.mk_not
@@ -152,7 +155,7 @@
                                           simp_tac (HOL_basic_ss addsimps simp3) 1]  
            
               val (inf_thm,thy4) =  
-                    PureThy.add_thms [((ak^"_infinite",Goal.prove_global thy3 [] [] stmnt3 proof3), [])] thy3
+                    add_thms_string [((ak^"_infinite",Goal.prove_global thy3 [] [] stmnt3 proof3), [])] thy3
           in 
             ((inj_thm,inject_thm,inf_thm),thy4)
           end) ak_names thy
@@ -186,7 +189,7 @@
         val def2 = Logic.mk_equals (cswap $ ab $ c, cswap_akname $ ab $ c)
       in
         thy |> Sign.add_consts_i [("swap_" ^ ak_name, swapT, NoSyn)] 
-            |> PureThy.add_defs_unchecked true [((name, def2),[])]
+            |> PureThy.add_defs_unchecked true [((Binding.name name, def2),[])]
             |> snd
             |> OldPrimrecPackage.add_primrec_unchecked_i "" [(("", def1),[])]
       end) ak_names_types thy1;
@@ -241,14 +244,14 @@
           val def = Logic.mk_equals
                     (cperm $ pi $ a, if ak_name = ak_name' then cperm_def $ pi $ a else a)
         in
-          PureThy.add_defs_unchecked true [((name, def),[])] thy'
+          PureThy.add_defs_unchecked true [((Binding.name name, def),[])] thy'
         end) ak_names_types thy) ak_names_types thy4;
     
     (* proves that every atom-kind is an instance of at *)
     (* lemma at_<ak>_inst:                              *)
     (* at TYPE(<ak>)                                    *)
     val (prm_cons_thms,thy6) = 
-      thy5 |> PureThy.add_thms (map (fn (ak_name, T) =>
+      thy5 |> add_thms_string (map (fn (ak_name, T) =>
       let
         val ak_name_qu = Sign.full_bname thy5 (ak_name);
         val i_type = Type(ak_name_qu,[]);
@@ -309,7 +312,7 @@
     (* lemma pt_<ak>_inst:                                    *)
     (* pt TYPE('x::pt_<ak>) TYPE(<ak>)                        *)
     val (prm_inst_thms,thy8) = 
-      thy7 |> PureThy.add_thms (map (fn (ak_name, T) =>
+      thy7 |> add_thms_string (map (fn (ak_name, T) =>
       let
         val ak_name_qu = Sign.full_bname thy7 ak_name;
         val pt_name_qu = Sign.full_bname thy7 ("pt_"^ak_name);
@@ -355,7 +358,7 @@
      (* lemma abst_<ak>_inst:                                    *)
      (* fs TYPE('x::pt_<ak>) TYPE (<ak>)                         *)
      val (fs_inst_thms,thy12) = 
-       thy11 |> PureThy.add_thms (map (fn (ak_name, T) =>
+       thy11 |> add_thms_string (map (fn (ak_name, T) =>
        let
          val ak_name_qu = Sign.full_bname thy11 ak_name;
          val fs_name_qu = Sign.full_bname thy11 ("fs_"^ak_name);
@@ -428,7 +431,7 @@
                                         rtac allI 1, rtac allI 1, rtac allI 1,
                                         rtac cp1 1];
            in
-             yield_singleton PureThy.add_thms ((name,
+             yield_singleton add_thms_string ((name,
                Goal.prove_global thy' [] [] statement proof), []) thy'
            end) 
            ak_names_types thy) ak_names_types thy12b;
@@ -460,7 +463,7 @@
 
                  val proof = fn _ => simp_tac simp_s 1;
                in
-                PureThy.add_thms [((name, Goal.prove_global thy' [] [] statement proof), [])] thy'
+                add_thms_string [((name, Goal.prove_global thy' [] [] statement proof), [])] thy'
                end
            else 
             ([],thy')))  (* do nothing branch, if ak_name = ak_name' *) 
@@ -870,114 +873,114 @@
              fun inst_pt_pt_at_cp_dj thms = inst_zip djs (inst_pt_pt_at_cp thms);
            in
             thy32 
-            |>   PureThy.add_thmss [(("alpha", inst_pt_at [abs_fun_eq]),[])]
-            ||>> PureThy.add_thmss [(("alpha'", inst_pt_at [abs_fun_eq']),[])]
-            ||>> PureThy.add_thmss [(("alpha_fresh", inst_pt_at [abs_fun_fresh]),[])]
-            ||>> PureThy.add_thmss [(("alpha_fresh'", inst_pt_at [abs_fun_fresh']),[])]
-            ||>> PureThy.add_thmss [(("perm_swap", inst_pt_at [pt_swap_bij] @ inst_pt_at [pt_swap_bij']),[])]
-            ||>> PureThy.add_thmss 
+            |>   add_thmss_string [(("alpha", inst_pt_at [abs_fun_eq]),[])]
+            ||>> add_thmss_string [(("alpha'", inst_pt_at [abs_fun_eq']),[])]
+            ||>> add_thmss_string [(("alpha_fresh", inst_pt_at [abs_fun_fresh]),[])]
+            ||>> add_thmss_string [(("alpha_fresh'", inst_pt_at [abs_fun_fresh']),[])]
+            ||>> add_thmss_string [(("perm_swap", inst_pt_at [pt_swap_bij] @ inst_pt_at [pt_swap_bij']),[])]
+            ||>> add_thmss_string 
 	      let val thms1 = inst_at at_swap_simps
                   and thms2 = inst_dj [dj_perm_forget]
               in [(("swap_simps", thms1 @ thms2),[])] end 
-            ||>> PureThy.add_thmss 
+            ||>> add_thmss_string 
               let val thms1 = inst_pt_at [pt_pi_rev];
                   val thms2 = inst_pt_at [pt_rev_pi];
               in [(("perm_pi_simp",thms1 @ thms2),[])] end
-            ||>> PureThy.add_thmss [(("perm_fresh_fresh", inst_pt_at [pt_fresh_fresh]),[])]
-            ||>> PureThy.add_thmss [(("perm_bij", inst_pt_at [pt_bij]),[])]
-            ||>> PureThy.add_thmss 
+            ||>> add_thmss_string [(("perm_fresh_fresh", inst_pt_at [pt_fresh_fresh]),[])]
+            ||>> add_thmss_string [(("perm_bij", inst_pt_at [pt_bij]),[])]
+            ||>> add_thmss_string 
               let val thms1 = inst_pt_at [pt_perm_compose];
                   val thms2 = instR cp1 (Library.flat cps');
               in [(("perm_compose",thms1 @ thms2),[])] end
-            ||>> PureThy.add_thmss [(("perm_compose'",inst_pt_at [pt_perm_compose']),[])] 
-            ||>> PureThy.add_thmss [(("perm_app", inst_pt_at [perm_app]),[])]
-            ||>> PureThy.add_thmss [(("supp_atm", (inst_at [at_supp]) @ (inst_dj [dj_supp])),[])]
-            ||>> PureThy.add_thmss [(("exists_fresh", inst_at [at_exists_fresh]),[])]
-            ||>> PureThy.add_thmss [(("exists_fresh'", inst_at [at_exists_fresh']),[])]
-            ||>> PureThy.add_thmss
+            ||>> add_thmss_string [(("perm_compose'",inst_pt_at [pt_perm_compose']),[])] 
+            ||>> add_thmss_string [(("perm_app", inst_pt_at [perm_app]),[])]
+            ||>> add_thmss_string [(("supp_atm", (inst_at [at_supp]) @ (inst_dj [dj_supp])),[])]
+            ||>> add_thmss_string [(("exists_fresh", inst_at [at_exists_fresh]),[])]
+            ||>> add_thmss_string [(("exists_fresh'", inst_at [at_exists_fresh']),[])]
+            ||>> add_thmss_string
               let
                 val thms1 = inst_pt_at [all_eqvt];
                 val thms2 = map (fold_rule [inductive_forall_def]) thms1
               in
                 [(("all_eqvt", thms1 @ thms2), [NominalThmDecls.eqvt_force_add])]
               end
-            ||>> PureThy.add_thmss [(("ex_eqvt", inst_pt_at [ex_eqvt]),[NominalThmDecls.eqvt_force_add])]
-            ||>> PureThy.add_thmss [(("ex1_eqvt", inst_pt_at [ex1_eqvt]),[NominalThmDecls.eqvt_force_add])]
-            ||>> PureThy.add_thmss [(("the_eqvt", inst_pt_at [the_eqvt]),[NominalThmDecls.eqvt_force_add])]
-            ||>> PureThy.add_thmss 
+            ||>> add_thmss_string [(("ex_eqvt", inst_pt_at [ex_eqvt]),[NominalThmDecls.eqvt_force_add])]
+            ||>> add_thmss_string [(("ex1_eqvt", inst_pt_at [ex1_eqvt]),[NominalThmDecls.eqvt_force_add])]
+            ||>> add_thmss_string [(("the_eqvt", inst_pt_at [the_eqvt]),[NominalThmDecls.eqvt_force_add])]
+            ||>> add_thmss_string 
               let val thms1 = inst_at [at_fresh]
                   val thms2 = inst_dj [at_fresh_ineq]
               in [(("fresh_atm", thms1 @ thms2),[])] end
-            ||>> PureThy.add_thmss
+            ||>> add_thmss_string
               let val thms1 = inst_at at_calc
                   and thms2 = inst_dj [dj_perm_forget]
               in [(("calc_atm", thms1 @ thms2),[])] end
-            ||>> PureThy.add_thmss
+            ||>> add_thmss_string
               let val thms1 = inst_pt_at [abs_fun_pi]
                   and thms2 = inst_pt_pt_at_cp [abs_fun_pi_ineq]
               in [(("abs_perm", thms1 @ thms2),[NominalThmDecls.eqvt_add])] end
-            ||>> PureThy.add_thmss
+            ||>> add_thmss_string
               let val thms1 = inst_dj [dj_perm_forget]
                   and thms2 = inst_dj [dj_pp_forget]
               in [(("perm_dj", thms1 @ thms2),[])] end
-            ||>> PureThy.add_thmss
+            ||>> add_thmss_string
               let val thms1 = inst_pt_at_fs [fresh_iff]
                   and thms2 = inst_pt_at [fresh_iff]
                   and thms3 = inst_pt_pt_at_cp_dj [fresh_iff_ineq]
             in [(("abs_fresh", thms1 @ thms2 @ thms3),[])] end
-            ||>> PureThy.add_thmss
+            ||>> add_thmss_string
               let val thms1 = inst_pt_at [abs_fun_supp]
                   and thms2 = inst_pt_at_fs [abs_fun_supp]
                   and thms3 = inst_pt_pt_at_cp_dj [abs_fun_supp_ineq]
               in [(("abs_supp", thms1 @ thms2 @ thms3),[])] end
-            ||>> PureThy.add_thmss
+            ||>> add_thmss_string
               let val thms1 = inst_pt_at [fresh_left]
                   and thms2 = inst_pt_pt_at_cp [fresh_left_ineq]
               in [(("fresh_left", thms1 @ thms2),[])] end
-            ||>> PureThy.add_thmss
+            ||>> add_thmss_string
               let val thms1 = inst_pt_at [fresh_right]
                   and thms2 = inst_pt_pt_at_cp [fresh_right_ineq]
               in [(("fresh_right", thms1 @ thms2),[])] end
-            ||>> PureThy.add_thmss
+            ||>> add_thmss_string
               let val thms1 = inst_pt_at [fresh_bij]
                   and thms2 = inst_pt_pt_at_cp [fresh_bij_ineq]
               in [(("fresh_bij", thms1 @ thms2),[])] end
-            ||>> PureThy.add_thmss
+            ||>> add_thmss_string
               let val thms1 = inst_pt_at fresh_star_bij
                   and thms2 = flat (map (fn ti => inst_pt_pt_at_cp [ti]) fresh_star_bij_ineq);
               in [(("fresh_star_bij", thms1 @ thms2),[])] end
-            ||>> PureThy.add_thmss
+            ||>> add_thmss_string
               let val thms1 = inst_pt_at [fresh_eqvt]
                   and thms2 = inst_pt_pt_at_cp_dj [fresh_eqvt_ineq]
               in [(("fresh_eqvt", thms1 @ thms2),[NominalThmDecls.eqvt_add])] end
-            ||>> PureThy.add_thmss
+            ||>> add_thmss_string
               let val thms1 = inst_pt_at [in_eqvt]
               in [(("in_eqvt", thms1),[NominalThmDecls.eqvt_add])] end
-            ||>> PureThy.add_thmss
+            ||>> add_thmss_string
               let val thms1 = inst_pt_at [eq_eqvt]
               in [(("eq_eqvt", thms1),[NominalThmDecls.eqvt_add])] end
-            ||>> PureThy.add_thmss
+            ||>> add_thmss_string
               let val thms1 = inst_pt_at [set_diff_eqvt]
               in [(("set_diff_eqvt", thms1),[NominalThmDecls.eqvt_add])] end
-            ||>> PureThy.add_thmss
+            ||>> add_thmss_string
               let val thms1 = inst_pt_at [subseteq_eqvt]
               in [(("subseteq_eqvt", thms1),[NominalThmDecls.eqvt_add])] end
-            ||>> PureThy.add_thmss [(("insert_eqvt", inst_pt_at [insert_eqvt]), [NominalThmDecls.eqvt_add])]
-            ||>> PureThy.add_thmss [(("set_eqvt", inst_pt_at [set_eqvt]), [NominalThmDecls.eqvt_add])]
-            ||>> PureThy.add_thmss [(("perm_set_eq", inst_pt_at [perm_set_eq]), [])]
-            ||>> PureThy.add_thmss
+            ||>> add_thmss_string [(("insert_eqvt", inst_pt_at [insert_eqvt]), [NominalThmDecls.eqvt_add])]
+            ||>> add_thmss_string [(("set_eqvt", inst_pt_at [set_eqvt]), [NominalThmDecls.eqvt_add])]
+            ||>> add_thmss_string [(("perm_set_eq", inst_pt_at [perm_set_eq]), [])]
+            ||>> add_thmss_string
               let val thms1 = inst_pt_at [fresh_aux]
                   and thms2 = inst_pt_pt_at_cp_dj [fresh_perm_app_ineq] 
               in  [(("fresh_aux", thms1 @ thms2),[])] end  
-            ||>> PureThy.add_thmss
+            ||>> add_thmss_string
               let val thms1 = inst_pt_at [fresh_perm_app]
                   and thms2 = inst_pt_pt_at_cp_dj [fresh_perm_app_ineq] 
               in  [(("fresh_perm_app", thms1 @ thms2),[])] end 
-            ||>> PureThy.add_thmss
+            ||>> add_thmss_string
               let val thms1 = inst_pt_at [pt_perm_supp]
                   and thms2 = inst_pt_pt_at_cp [pt_perm_supp_ineq] 
               in  [(("supp_eqvt", thms1 @ thms2),[NominalThmDecls.eqvt_add])] end  
-            ||>> PureThy.add_thmss [(("fin_supp",fs_axs),[])]
+            ||>> add_thmss_string [(("fin_supp",fs_axs),[])]
            end;
 
     in 
--- a/src/HOL/Nominal/nominal_induct.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Nominal/nominal_induct.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -6,7 +6,7 @@
 
 structure NominalInduct:
 sig
-  val nominal_induct_tac: Proof.context -> (Binding.T option * term) option list list ->
+  val nominal_induct_tac: Proof.context -> (binding option * term) option list list ->
     (string * typ) list -> (string * typ) list list -> thm list ->
     thm list -> int -> RuleCases.cases_tactic
   val nominal_induct_method: Method.src -> Proof.context -> Method.method
--- a/src/HOL/Nominal/nominal_inductive.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Nominal/nominal_inductive.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -562,17 +562,17 @@
             [ind_case_names, RuleCases.consumes 1]);
         val ([strong_induct'], thy') = thy |>
           Sign.add_path rec_name |>
-          PureThy.add_thms [(("strong_induct", #1 strong_induct), #2 strong_induct)];
+          PureThy.add_thms [((Binding.name "strong_induct", #1 strong_induct), #2 strong_induct)];
         val strong_inducts =
           ProjectRule.projects ctxt (1 upto length names) strong_induct'
       in
         thy' |>
-        PureThy.add_thmss [(("strong_inducts", strong_inducts),
+        PureThy.add_thmss [((Binding.name "strong_inducts", strong_inducts),
           [ind_case_names, RuleCases.consumes 1])] |> snd |>
         Sign.parent_path |>
         fold (fn ((name, elim), (_, cases)) =>
           Sign.add_path (Sign.base_name name) #>
-          PureThy.add_thms [(("strong_cases", elim),
+          PureThy.add_thms [((Binding.name "strong_cases", elim),
             [RuleCases.case_names (map snd cases),
              RuleCases.consumes 1])] #> snd #>
           Sign.parent_path) (strong_cases ~~ induct_cases')
@@ -653,7 +653,7 @@
   in
     fold (fn (name, ths) =>
       Sign.add_path (Sign.base_name name) #>
-      PureThy.add_thmss [(("eqvt", ths), [NominalThmDecls.eqvt_add])] #> snd #>
+      PureThy.add_thmss [((Binding.name "eqvt", ths), [NominalThmDecls.eqvt_add])] #> snd #>
       Sign.parent_path) (names ~~ transp thss) thy
   end;
 
--- a/src/HOL/Nominal/nominal_inductive2.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Nominal/nominal_inductive2.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -458,12 +458,12 @@
             [ind_case_names, RuleCases.consumes 1]);
         val ([strong_induct'], thy') = thy |>
           Sign.add_path rec_name |>
-          PureThy.add_thms [(("strong_induct", #1 strong_induct), #2 strong_induct)];
+          PureThy.add_thms [((Binding.name "strong_induct", #1 strong_induct), #2 strong_induct)];
         val strong_inducts =
           ProjectRule.projects ctxt (1 upto length names) strong_induct'
       in
         thy' |>
-        PureThy.add_thmss [(("strong_inducts", strong_inducts),
+        PureThy.add_thmss [((Binding.name "strong_inducts", strong_inducts),
           [ind_case_names, RuleCases.consumes 1])] |> snd |>
         Sign.parent_path
       end))
--- a/src/HOL/Nominal/nominal_package.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Nominal/nominal_package.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -490,13 +490,13 @@
             (full_new_type_names' ~~ tyvars) thy
         end) atoms |>
       PureThy.add_thmss
-        [((space_implode "_" new_type_names ^ "_unfolded_perm_eq",
+        [((Binding.name (space_implode "_" new_type_names ^ "_unfolded_perm_eq"),
           unfolded_perm_eq_thms), [Simplifier.simp_add]),
-         ((space_implode "_" new_type_names ^ "_perm_empty",
+         ((Binding.name (space_implode "_" new_type_names ^ "_perm_empty"),
           perm_empty_thms), [Simplifier.simp_add]),
-         ((space_implode "_" new_type_names ^ "_perm_append",
+         ((Binding.name (space_implode "_" new_type_names ^ "_perm_append"),
           perm_append_thms), [Simplifier.simp_add]),
-         ((space_implode "_" new_type_names ^ "_perm_eq",
+         ((Binding.name (space_implode "_" new_type_names ^ "_perm_eq"),
           perm_eq_thms), [Simplifier.simp_add])];
 
     (**** Define representing sets ****)
@@ -627,7 +627,7 @@
           val pi = Free ("pi", permT);
           val T = Type (Sign.intern_type thy name, map TFree tvs);
         in apfst (pair r o hd)
-          (PureThy.add_defs_unchecked true [(("prm_" ^ name ^ "_def", Logic.mk_equals
+          (PureThy.add_defs_unchecked true [((Binding.name ("prm_" ^ name ^ "_def"), Logic.mk_equals
             (Const ("Nominal.perm", permT --> T --> T) $ pi $ Free ("x", T),
              Const (Sign.intern_const thy ("Abs_" ^ name), U --> T) $
                (Const ("Nominal.perm", permT --> U --> U) $ pi $
@@ -801,7 +801,7 @@
         val def_name = (Sign.base_name cname) ^ "_def";
         val ([def_thm], thy') = thy |>
           Sign.add_consts_i [(cname', constrT, mx)] |>
-          (PureThy.add_defs false o map Thm.no_attributes) [(def_name, def)]
+          (PureThy.add_defs false o map Thm.no_attributes) [(Binding.name def_name, def)]
       in (thy', defs @ [def_thm], eqns @ [eqn]) end;
 
     fun dt_constr_defs ((thy, defs, eqns, dist_lemmas), ((((((_, (_, _, constrs)),
@@ -1114,8 +1114,8 @@
  
     val (_, thy9) = thy8 |>
       Sign.add_path big_name |>
-      PureThy.add_thms [(("induct", dt_induct), [case_names_induct])] ||>>
-      PureThy.add_thmss [(("inducts", projections dt_induct), [case_names_induct])] ||>
+      PureThy.add_thms [((Binding.name "induct", dt_induct), [case_names_induct])] ||>>
+      PureThy.add_thmss [((Binding.name "inducts", projections dt_induct), [case_names_induct])] ||>
       Sign.parent_path ||>>
       DatatypeAux.store_thmss_atts "distinct" new_type_names simp_atts distinct_thms ||>>
       DatatypeAux.store_thmss "constr_rep" new_type_names constr_rep_thmss ||>>
@@ -1405,9 +1405,9 @@
 
     val (_, thy10) = thy9 |>
       Sign.add_path big_name |>
-      PureThy.add_thms [(("strong_induct'", induct_aux), [])] ||>>
-      PureThy.add_thms [(("strong_induct", induct), [case_names_induct])] ||>>
-      PureThy.add_thmss [(("strong_inducts", projections induct), [case_names_induct])];
+      PureThy.add_thms [((Binding.name "strong_induct'", induct_aux), [])] ||>>
+      PureThy.add_thms [((Binding.name "strong_induct", induct), [case_names_induct])] ||>>
+      PureThy.add_thmss [((Binding.name "strong_inducts", projections induct), [case_names_induct])];
 
     (**** recursion combinator ****)
 
@@ -2015,7 +2015,7 @@
           (Sign.base_name name, rec_fn_Ts @ [T] ---> T', NoSyn))
           (reccomb_names ~~ recTs ~~ rec_result_Ts))
       |> (PureThy.add_defs false o map Thm.no_attributes) (map (fn ((((name, comb), set), T), T') =>
-          ((Sign.base_name name) ^ "_def", Logic.mk_equals (comb, absfree ("x", T,
+          (Binding.name (Sign.base_name name ^ "_def"), Logic.mk_equals (comb, absfree ("x", T,
            Const ("The", (T' --> HOLogic.boolT) --> T') $ absfree ("y", T',
              set $ Free ("x", T) $ Free ("y", T'))))))
                (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts));
@@ -2052,12 +2052,12 @@
     (* FIXME: theorems are stored in database for testing only *)
     val (_, thy13) = thy12 |>
       PureThy.add_thmss
-        [(("rec_equiv", List.concat rec_equiv_thms), []),
-         (("rec_equiv'", List.concat rec_equiv_thms'), []),
-         (("rec_fin_supp", List.concat rec_fin_supp_thms), []),
-         (("rec_fresh", List.concat rec_fresh_thms), []),
-         (("rec_unique", map standard rec_unique_thms), []),
-         (("recs", rec_thms), [])] ||>
+        [((Binding.name "rec_equiv", List.concat rec_equiv_thms), []),
+         ((Binding.name "rec_equiv'", List.concat rec_equiv_thms'), []),
+         ((Binding.name "rec_fin_supp", List.concat rec_fin_supp_thms), []),
+         ((Binding.name "rec_fresh", List.concat rec_fresh_thms), []),
+         ((Binding.name "rec_unique", map standard rec_unique_thms), []),
+         ((Binding.name "recs", rec_thms), [])] ||>
       Sign.parent_path ||>
       map_nominal_datatypes (fold Symtab.update dt_infos);
 
--- a/src/HOL/Nominal/nominal_primrec.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Nominal/nominal_primrec.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -9,8 +9,8 @@
 signature NOMINAL_PRIMREC =
 sig
   val add_primrec: term list option -> term option ->
-    (Binding.T * typ option * mixfix) list ->
-    (Binding.T * typ option * mixfix) list ->
+    (binding * typ option * mixfix) list ->
+    (binding * typ option * mixfix) list ->
     (Attrib.binding * term) list -> local_theory -> Proof.state
 end;
 
--- a/src/HOL/Nominal/nominal_thmdecls.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Nominal/nominal_thmdecls.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -187,8 +187,8 @@
         "equivariance theorem declaration (without checking the form of the lemma)"),
       ("fresh", Attrib.add_del_args fresh_add fresh_del, "freshness theorem declaration"),
       ("bij",   Attrib.add_del_args bij_add bij_del,     "bijection theorem declaration")] #>
-  PureThy.add_thms_dynamic ("eqvts", #eqvts o Data.get) #>
-  PureThy.add_thms_dynamic ("freshs", #freshs o Data.get) #>
-  PureThy.add_thms_dynamic ("bijs", #bijs o Data.get);
+  PureThy.add_thms_dynamic (Binding.name "eqvts", #eqvts o Data.get) #>
+  PureThy.add_thms_dynamic (Binding.name "freshs", #freshs o Data.get) #>
+  PureThy.add_thms_dynamic (Binding.name "bijs", #bijs o Data.get);
 
 end;
--- a/src/HOL/NumberTheory/IntPrimes.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/NumberTheory/IntPrimes.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -146,7 +146,7 @@
 lemma zcong_trans:
     "[a = b] (mod m) ==> [b = c] (mod m) ==> [a = c] (mod m)"
   unfolding zcong_def
-  apply (auto elim!: dvdE simp add: ring_simps)
+  apply (auto elim!: dvdE simp add: algebra_simps)
   unfolding left_distrib [symmetric]
   apply (rule dvd_mult dvd_refl)+
   done
@@ -255,7 +255,7 @@
 
 lemma zcong_iff_lin: "([a = b] (mod m)) = (\<exists>k. b = a + m * k)"
   unfolding zcong_def
-  apply (auto elim!: dvdE simp add: ring_simps)
+  apply (auto elim!: dvdE simp add: algebra_simps)
   apply (rule_tac x = "-k" in exI) apply simp
   done
 
--- a/src/HOL/OrderedGroup.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/OrderedGroup.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -22,17 +22,34 @@
   \end{itemize}
 *}
 
+ML{*
+structure AlgebraSimps =
+  NamedThmsFun(val name = "algebra_simps"
+               val description = "algebra simplification rules");
+*}
+
+setup AlgebraSimps.setup
+
+text{* The rewrites accumulated in @{text algebra_simps} deal with the
+classical algebraic structures of groups, rings and family. They simplify
+terms by multiplying everything out (in case of a ring) and bringing sums and
+products into a canonical form (by ordered rewriting). As a result it decides
+group and ring equalities but also helps with inequalities.
+
+Of course it also works for fields, but it knows nothing about multiplicative
+inverses or division. This is catered for by @{text field_simps}. *}
+
 subsection {* Semigroups and Monoids *}
 
 class semigroup_add = plus +
-  assumes add_assoc: "(a + b) + c = a + (b + c)"
+  assumes add_assoc[algebra_simps]: "(a + b) + c = a + (b + c)"
 
 class ab_semigroup_add = semigroup_add +
-  assumes add_commute: "a + b = b + a"
+  assumes add_commute[algebra_simps]: "a + b = b + a"
 begin
 
-lemma add_left_commute: "a + (b + c) = b + (a + c)"
-  by (rule mk_left_commute [of "plus", OF add_assoc add_commute])
+lemma add_left_commute[algebra_simps]: "a + (b + c) = b + (a + c)"
+by (rule mk_left_commute [of "plus", OF add_assoc add_commute])
 
 theorems add_ac = add_assoc add_commute add_left_commute
 
@@ -41,14 +58,14 @@
 theorems add_ac = add_assoc add_commute add_left_commute
 
 class semigroup_mult = times +
-  assumes mult_assoc: "(a * b) * c = a * (b * c)"
+  assumes mult_assoc[algebra_simps]: "(a * b) * c = a * (b * c)"
 
 class ab_semigroup_mult = semigroup_mult +
-  assumes mult_commute: "a * b = b * a"
+  assumes mult_commute[algebra_simps]: "a * b = b * a"
 begin
 
-lemma mult_left_commute: "a * (b * c) = b * (a * c)"
-  by (rule mk_left_commute [of "times", OF mult_assoc mult_commute])
+lemma mult_left_commute[algebra_simps]: "a * (b * c) = b * (a * c)"
+by (rule mk_left_commute [of "times", OF mult_assoc mult_commute])
 
 theorems mult_ac = mult_assoc mult_commute mult_left_commute
 
@@ -57,24 +74,20 @@
 theorems mult_ac = mult_assoc mult_commute mult_left_commute
 
 class ab_semigroup_idem_mult = ab_semigroup_mult +
-  assumes mult_idem: "x * x = x"
+  assumes mult_idem[simp]: "x * x = x"
 begin
 
-lemma mult_left_idem: "x * (x * y) = x * y"
+lemma mult_left_idem[simp]: "x * (x * y) = x * y"
   unfolding mult_assoc [symmetric, of x] mult_idem ..
 
-lemmas mult_ac_idem = mult_ac mult_idem mult_left_idem
-
 end
 
-lemmas mult_ac_idem = mult_ac mult_idem mult_left_idem
-
 class monoid_add = zero + semigroup_add +
   assumes add_0_left [simp]: "0 + a = a"
     and add_0_right [simp]: "a + 0 = a"
 
 lemma zero_reorient: "0 = x \<longleftrightarrow> x = 0"
-  by (rule eq_commute)
+by (rule eq_commute)
 
 class comm_monoid_add = zero + ab_semigroup_add +
   assumes add_0: "0 + a = a"
@@ -90,7 +103,7 @@
   assumes mult_1_right [simp]: "a * 1 = a"
 
 lemma one_reorient: "1 = x \<longleftrightarrow> x = 1"
-  by (rule eq_commute)
+by (rule eq_commute)
 
 class comm_monoid_mult = one + ab_semigroup_mult +
   assumes mult_1: "1 * a = a"
@@ -108,11 +121,11 @@
 
 lemma add_left_cancel [simp]:
   "a + b = a + c \<longleftrightarrow> b = c"
-  by (blast dest: add_left_imp_eq)
+by (blast dest: add_left_imp_eq)
 
 lemma add_right_cancel [simp]:
   "b + a = c + a \<longleftrightarrow> b = c"
-  by (blast dest: add_right_imp_eq)
+by (blast dest: add_right_imp_eq)
 
 end
 
@@ -142,7 +155,7 @@
 begin
 
 lemma minus_add_cancel: "- a + (a + b) = b"
-  by (simp add: add_assoc[symmetric])
+by (simp add: add_assoc[symmetric])
 
 lemma minus_zero [simp]: "- 0 = 0"
 proof -
@@ -176,8 +189,7 @@
 qed
 
 lemma equals_zero_I:
-  assumes "a + b = 0"
-  shows "- a = b"
+  assumes "a + b = 0" shows "- a = b"
 proof -
   have "- a = - a + (a + b)" using assms by simp
   also have "\<dots> = b" by (simp add: add_assoc[symmetric])
@@ -185,23 +197,22 @@
 qed
 
 lemma diff_self [simp]: "a - a = 0"
-  by (simp add: diff_minus)
+by (simp add: diff_minus)
 
 lemma diff_0 [simp]: "0 - a = - a"
-  by (simp add: diff_minus)
+by (simp add: diff_minus)
 
 lemma diff_0_right [simp]: "a - 0 = a" 
-  by (simp add: diff_minus)
+by (simp add: diff_minus)
 
 lemma diff_minus_eq_add [simp]: "a - - b = a + b"
-  by (simp add: diff_minus)
+by (simp add: diff_minus)
 
 lemma neg_equal_iff_equal [simp]:
   "- a = - b \<longleftrightarrow> a = b" 
 proof 
   assume "- a = - b"
-  hence "- (- a) = - (- b)"
-    by simp
+  hence "- (- a) = - (- b)" by simp
   thus "a = b" by simp
 next
   assume "a = b"
@@ -210,11 +221,11 @@
 
 lemma neg_equal_0_iff_equal [simp]:
   "- a = 0 \<longleftrightarrow> a = 0"
-  by (subst neg_equal_iff_equal [symmetric], simp)
+by (subst neg_equal_iff_equal [symmetric], simp)
 
 lemma neg_0_equal_iff_equal [simp]:
   "0 = - a \<longleftrightarrow> 0 = a"
-  by (subst neg_equal_iff_equal [symmetric], simp)
+by (subst neg_equal_iff_equal [symmetric], simp)
 
 text{*The next two equations can make the simplifier loop!*}
 
@@ -233,10 +244,12 @@
 qed
 
 lemma diff_add_cancel: "a - b + b = a"
-  by (simp add: diff_minus add_assoc)
+by (simp add: diff_minus add_assoc)
 
 lemma add_diff_cancel: "a + b - b = a"
-  by (simp add: diff_minus add_assoc)
+by (simp add: diff_minus add_assoc)
+
+declare diff_minus[symmetric, algebra_simps]
 
 end
 
@@ -257,43 +270,38 @@
   then show "b = c" by simp
 qed
 
-lemma uminus_add_conv_diff:
+lemma uminus_add_conv_diff[algebra_simps]:
   "- a + b = b - a"
-  by (simp add:diff_minus add_commute)
+by (simp add:diff_minus add_commute)
 
 lemma minus_add_distrib [simp]:
   "- (a + b) = - a + - b"
-  by (rule equals_zero_I) (simp add: add_ac)
+by (rule equals_zero_I) (simp add: add_ac)
 
 lemma minus_diff_eq [simp]:
   "- (a - b) = b - a"
-  by (simp add: diff_minus add_commute)
-
-lemma add_diff_eq: "a + (b - c) = (a + b) - c"
-  by (simp add: diff_minus add_ac)
+by (simp add: diff_minus add_commute)
 
-lemma diff_add_eq: "(a - b) + c = (a + c) - b"
-  by (simp add: diff_minus add_ac)
+lemma add_diff_eq[algebra_simps]: "a + (b - c) = (a + b) - c"
+by (simp add: diff_minus add_ac)
 
-lemma diff_eq_eq: "a - b = c \<longleftrightarrow> a = c + b"
-  by (auto simp add: diff_minus add_assoc)
+lemma diff_add_eq[algebra_simps]: "(a - b) + c = (a + c) - b"
+by (simp add: diff_minus add_ac)
 
-lemma eq_diff_eq: "a = c - b \<longleftrightarrow> a + b = c"
-  by (auto simp add: diff_minus add_assoc)
+lemma diff_eq_eq[algebra_simps]: "a - b = c \<longleftrightarrow> a = c + b"
+by (auto simp add: diff_minus add_assoc)
 
-lemma diff_diff_eq: "(a - b) - c = a - (b + c)"
-  by (simp add: diff_minus add_ac)
+lemma eq_diff_eq[algebra_simps]: "a = c - b \<longleftrightarrow> a + b = c"
+by (auto simp add: diff_minus add_assoc)
 
-lemma diff_diff_eq2: "a - (b - c) = (a + c) - b"
-  by (simp add: diff_minus add_ac)
+lemma diff_diff_eq[algebra_simps]: "(a - b) - c = a - (b + c)"
+by (simp add: diff_minus add_ac)
 
-lemmas compare_rls =
-       diff_minus [symmetric]
-       add_diff_eq diff_add_eq diff_diff_eq diff_diff_eq2
-       diff_eq_eq eq_diff_eq
+lemma diff_diff_eq2[algebra_simps]: "a - (b - c) = (a + c) - b"
+by (simp add: diff_minus add_ac)
 
 lemma eq_iff_diff_eq_0: "a = b \<longleftrightarrow> a - b = 0"
-  by (simp add: compare_rls)
+by (simp add: algebra_simps)
 
 end
 
@@ -305,7 +313,7 @@
 
 lemma add_right_mono:
   "a \<le> b \<Longrightarrow> a + c \<le> b + c"
-  by (simp add: add_commute [of _ c] add_left_mono)
+by (simp add: add_commute [of _ c] add_left_mono)
 
 text {* non-strict, in both arguments *}
 lemma add_mono:
@@ -322,11 +330,11 @@
 
 lemma add_strict_left_mono:
   "a < b \<Longrightarrow> c + a < c + b"
-  by (auto simp add: less_le add_left_mono)
+by (auto simp add: less_le add_left_mono)
 
 lemma add_strict_right_mono:
   "a < b \<Longrightarrow> a + c < b + c"
-  by (simp add: add_commute [of _ c] add_strict_left_mono)
+by (simp add: add_commute [of _ c] add_strict_left_mono)
 
 text{*Strict monotonicity in both arguments*}
 lemma add_strict_mono:
@@ -355,8 +363,7 @@
 begin
 
 lemma add_less_imp_less_left:
-   assumes less: "c + a < c + b"
-   shows "a < b"
+  assumes less: "c + a < c + b" shows "a < b"
 proof -
   from less have le: "c + a <= c + b" by (simp add: order_le_less)
   have "a <= b" 
@@ -381,23 +388,23 @@
 
 lemma add_less_cancel_left [simp]:
   "c + a < c + b \<longleftrightarrow> a < b"
-  by (blast intro: add_less_imp_less_left add_strict_left_mono) 
+by (blast intro: add_less_imp_less_left add_strict_left_mono) 
 
 lemma add_less_cancel_right [simp]:
   "a + c < b + c \<longleftrightarrow> a < b"
-  by (blast intro: add_less_imp_less_right add_strict_right_mono)
+by (blast intro: add_less_imp_less_right add_strict_right_mono)
 
 lemma add_le_cancel_left [simp]:
   "c + a \<le> c + b \<longleftrightarrow> a \<le> b"
-  by (auto, drule add_le_imp_le_left, simp_all add: add_left_mono) 
+by (auto, drule add_le_imp_le_left, simp_all add: add_left_mono) 
 
 lemma add_le_cancel_right [simp]:
   "a + c \<le> b + c \<longleftrightarrow> a \<le> b"
-  by (simp add: add_commute [of a c] add_commute [of b c])
+by (simp add: add_commute [of a c] add_commute [of b c])
 
 lemma add_le_imp_le_right:
   "a + c \<le> b + c \<Longrightarrow> a \<le> b"
-  by simp
+by simp
 
 lemma max_add_distrib_left:
   "max x y + z = max (x + z) (y + z)"
@@ -416,8 +423,7 @@
 begin
 
 lemma add_pos_nonneg:
-  assumes "0 < a" and "0 \<le> b"
-    shows "0 < a + b"
+  assumes "0 < a" and "0 \<le> b" shows "0 < a + b"
 proof -
   have "0 + 0 < a + b" 
     using assms by (rule add_less_le_mono)
@@ -425,13 +431,11 @@
 qed
 
 lemma add_pos_pos:
-  assumes "0 < a" and "0 < b"
-    shows "0 < a + b"
-  by (rule add_pos_nonneg) (insert assms, auto)
+  assumes "0 < a" and "0 < b" shows "0 < a + b"
+by (rule add_pos_nonneg) (insert assms, auto)
 
 lemma add_nonneg_pos:
-  assumes "0 \<le> a" and "0 < b"
-    shows "0 < a + b"
+  assumes "0 \<le> a" and "0 < b" shows "0 < a + b"
 proof -
   have "0 + 0 < a + b" 
     using assms by (rule add_le_less_mono)
@@ -439,8 +443,7 @@
 qed
 
 lemma add_nonneg_nonneg:
-  assumes "0 \<le> a" and "0 \<le> b"
-    shows "0 \<le> a + b"
+  assumes "0 \<le> a" and "0 \<le> b" shows "0 \<le> a + b"
 proof -
   have "0 + 0 \<le> a + b" 
     using assms by (rule add_mono)
@@ -448,8 +451,7 @@
 qed
 
 lemma add_neg_nonpos: 
-  assumes "a < 0" and "b \<le> 0"
-  shows "a + b < 0"
+  assumes "a < 0" and "b \<le> 0" shows "a + b < 0"
 proof -
   have "a + b < 0 + 0"
     using assms by (rule add_less_le_mono)
@@ -457,13 +459,11 @@
 qed
 
 lemma add_neg_neg: 
-  assumes "a < 0" and "b < 0"
-  shows "a + b < 0"
-  by (rule add_neg_nonpos) (insert assms, auto)
+  assumes "a < 0" and "b < 0" shows "a + b < 0"
+by (rule add_neg_nonpos) (insert assms, auto)
 
 lemma add_nonpos_neg:
-  assumes "a \<le> 0" and "b < 0"
-  shows "a + b < 0"
+  assumes "a \<le> 0" and "b < 0" shows "a + b < 0"
 proof -
   have "a + b < 0 + 0"
     using assms by (rule add_le_less_mono)
@@ -471,8 +471,7 @@
 qed
 
 lemma add_nonpos_nonpos:
-  assumes "a \<le> 0" and "b \<le> 0"
-  shows "a + b \<le> 0"
+  assumes "a \<le> 0" and "b \<le> 0" shows "a + b \<le> 0"
 proof -
   have "a + b \<le> 0 + 0"
     using assms by (rule add_mono)
@@ -500,31 +499,25 @@
 
 lemma max_diff_distrib_left:
   shows "max x y - z = max (x - z) (y - z)"
-  by (simp add: diff_minus, rule max_add_distrib_left) 
+by (simp add: diff_minus, rule max_add_distrib_left) 
 
 lemma min_diff_distrib_left:
   shows "min x y - z = min (x - z) (y - z)"
-  by (simp add: diff_minus, rule min_add_distrib_left) 
+by (simp add: diff_minus, rule min_add_distrib_left) 
 
 lemma le_imp_neg_le:
-  assumes "a \<le> b"
-  shows "-b \<le> -a"
+  assumes "a \<le> b" shows "-b \<le> -a"
 proof -
-  have "-a+a \<le> -a+b"
-    using `a \<le> b` by (rule add_left_mono) 
-  hence "0 \<le> -a+b"
-    by simp
-  hence "0 + (-b) \<le> (-a + b) + (-b)"
-    by (rule add_right_mono) 
-  thus ?thesis
-    by (simp add: add_assoc)
+  have "-a+a \<le> -a+b" using `a \<le> b` by (rule add_left_mono) 
+  hence "0 \<le> -a+b" by simp
+  hence "0 + (-b) \<le> (-a + b) + (-b)" by (rule add_right_mono) 
+  thus ?thesis by (simp add: add_assoc)
 qed
 
 lemma neg_le_iff_le [simp]: "- b \<le> - a \<longleftrightarrow> a \<le> b"
 proof 
   assume "- b \<le> - a"
-  hence "- (- a) \<le> - (- b)"
-    by (rule le_imp_neg_le)
+  hence "- (- a) \<le> - (- b)" by (rule le_imp_neg_le)
   thus "a\<le>b" by simp
 next
   assume "a\<le>b"
@@ -532,19 +525,19 @@
 qed
 
 lemma neg_le_0_iff_le [simp]: "- a \<le> 0 \<longleftrightarrow> 0 \<le> a"
-  by (subst neg_le_iff_le [symmetric], simp)
+by (subst neg_le_iff_le [symmetric], simp)
 
 lemma neg_0_le_iff_le [simp]: "0 \<le> - a \<longleftrightarrow> a \<le> 0"
-  by (subst neg_le_iff_le [symmetric], simp)
+by (subst neg_le_iff_le [symmetric], simp)
 
 lemma neg_less_iff_less [simp]: "- b < - a \<longleftrightarrow> a < b"
-  by (force simp add: less_le) 
+by (force simp add: less_le) 
 
 lemma neg_less_0_iff_less [simp]: "- a < 0 \<longleftrightarrow> 0 < a"
-  by (subst neg_less_iff_less [symmetric], simp)
+by (subst neg_less_iff_less [symmetric], simp)
 
 lemma neg_0_less_iff_less [simp]: "0 < - a \<longleftrightarrow> a < 0"
-  by (subst neg_less_iff_less [symmetric], simp)
+by (subst neg_less_iff_less [symmetric], simp)
 
 text{*The next several equations can make the simplifier loop!*}
 
@@ -573,7 +566,7 @@
 qed
 
 lemma minus_le_iff: "- a \<le> b \<longleftrightarrow> - b \<le> a"
-  by (auto simp add: le_less minus_less_iff)
+by (auto simp add: le_less minus_less_iff)
 
 lemma less_iff_diff_less_0: "a < b \<longleftrightarrow> a - b < 0"
 proof -
@@ -583,56 +576,34 @@
   finally show ?thesis .
 qed
 
-lemma diff_less_eq: "a - b < c \<longleftrightarrow> a < c + b"
+lemma diff_less_eq[algebra_simps]: "a - b < c \<longleftrightarrow> a < c + b"
 apply (subst less_iff_diff_less_0 [of a])
 apply (rule less_iff_diff_less_0 [of _ c, THEN ssubst])
 apply (simp add: diff_minus add_ac)
 done
 
-lemma less_diff_eq: "a < c - b \<longleftrightarrow> a + b < c"
+lemma less_diff_eq[algebra_simps]: "a < c - b \<longleftrightarrow> a + b < c"
 apply (subst less_iff_diff_less_0 [of "plus a b"])
 apply (subst less_iff_diff_less_0 [of a])
 apply (simp add: diff_minus add_ac)
 done
 
-lemma diff_le_eq: "a - b \<le> c \<longleftrightarrow> a \<le> c + b"
-  by (auto simp add: le_less diff_less_eq diff_add_cancel add_diff_cancel)
-
-lemma le_diff_eq: "a \<le> c - b \<longleftrightarrow> a + b \<le> c"
-  by (auto simp add: le_less less_diff_eq diff_add_cancel add_diff_cancel)
+lemma diff_le_eq[algebra_simps]: "a - b \<le> c \<longleftrightarrow> a \<le> c + b"
+by (auto simp add: le_less diff_less_eq diff_add_cancel add_diff_cancel)
 
-lemmas compare_rls =
-       diff_minus [symmetric]
-       add_diff_eq diff_add_eq diff_diff_eq diff_diff_eq2
-       diff_less_eq less_diff_eq diff_le_eq le_diff_eq
-       diff_eq_eq eq_diff_eq
-
-text{*This list of rewrites simplifies (in)equalities by bringing subtractions
-  to the top and then moving negative terms to the other side.
-  Use with @{text add_ac}*}
-lemmas (in -) compare_rls =
-       diff_minus [symmetric]
-       add_diff_eq diff_add_eq diff_diff_eq diff_diff_eq2
-       diff_less_eq less_diff_eq diff_le_eq le_diff_eq
-       diff_eq_eq eq_diff_eq
+lemma le_diff_eq[algebra_simps]: "a \<le> c - b \<longleftrightarrow> a + b \<le> c"
+by (auto simp add: le_less less_diff_eq diff_add_cancel add_diff_cancel)
 
 lemma le_iff_diff_le_0: "a \<le> b \<longleftrightarrow> a - b \<le> 0"
-  by (simp add: compare_rls)
+by (simp add: algebra_simps)
 
-lemmas group_simps =
-  add_ac
-  add_diff_eq diff_add_eq diff_diff_eq diff_diff_eq2
-  diff_eq_eq eq_diff_eq diff_minus [symmetric] uminus_add_conv_diff
-  diff_less_eq less_diff_eq diff_le_eq le_diff_eq
+text{*Legacy - use @{text algebra_simps} *}
+lemmas group_simps = algebra_simps
 
 end
 
-lemmas group_simps =
-  mult_ac
-  add_ac
-  add_diff_eq diff_add_eq diff_diff_eq diff_diff_eq2
-  diff_eq_eq eq_diff_eq diff_minus [symmetric] uminus_add_conv_diff
-  diff_less_eq less_diff_eq diff_le_eq le_diff_eq
+text{*Legacy - use @{text algebra_simps} *}
+lemmas group_simps = algebra_simps
 
 class ordered_ab_semigroup_add =
   linorder + pordered_ab_semigroup_add
@@ -766,8 +737,7 @@
   unfolding neg_le_0_iff_le by simp
 
 lemma abs_of_nonneg [simp]:
-  assumes nonneg: "0 \<le> a"
-  shows "\<bar>a\<bar> = a"
+  assumes nonneg: "0 \<le> a" shows "\<bar>a\<bar> = a"
 proof (rule antisym)
   from nonneg le_imp_neg_le have "- a \<le> 0" by simp
   from this nonneg have "- a \<le> a" by (rule order_trans)
@@ -775,8 +745,8 @@
 qed (rule abs_ge_self)
 
 lemma abs_idempotent [simp]: "\<bar>\<bar>a\<bar>\<bar> = \<bar>a\<bar>"
-  by (rule antisym)
-    (auto intro!: abs_ge_self abs_leI order_trans [of "uminus (abs a)" zero "abs a"])
+by (rule antisym)
+   (auto intro!: abs_ge_self abs_leI order_trans [of "uminus (abs a)" zero "abs a"])
 
 lemma abs_eq_0 [simp]: "\<bar>a\<bar> = 0 \<longleftrightarrow> a = 0"
 proof -
@@ -792,7 +762,7 @@
 qed
 
 lemma abs_zero [simp]: "\<bar>0\<bar> = 0"
-  by simp
+by simp
 
 lemma abs_0_eq [simp, noatp]: "0 = \<bar>a\<bar> \<longleftrightarrow> a = 0"
 proof -
@@ -811,7 +781,7 @@
 qed
 
 lemma zero_less_abs_iff [simp]: "0 < \<bar>a\<bar> \<longleftrightarrow> a \<noteq> 0"
-  by (simp add: less_le)
+by (simp add: less_le)
 
 lemma abs_not_less_zero [simp]: "\<not> \<bar>a\<bar> < 0"
 proof -
@@ -834,11 +804,10 @@
 qed
 
 lemma abs_of_pos: "0 < a \<Longrightarrow> \<bar>a\<bar> = a"
-  by (rule abs_of_nonneg, rule less_imp_le)
+by (rule abs_of_nonneg, rule less_imp_le)
 
 lemma abs_of_nonpos [simp]:
-  assumes "a \<le> 0"
-  shows "\<bar>a\<bar> = - a"
+  assumes "a \<le> 0" shows "\<bar>a\<bar> = - a"
 proof -
   let ?b = "- a"
   have "- ?b \<le> 0 \<Longrightarrow> \<bar>- ?b\<bar> = - (- ?b)"
@@ -849,24 +818,24 @@
 qed
   
 lemma abs_of_neg: "a < 0 \<Longrightarrow> \<bar>a\<bar> = - a"
-  by (rule abs_of_nonpos, rule less_imp_le)
+by (rule abs_of_nonpos, rule less_imp_le)
 
 lemma abs_le_D1: "\<bar>a\<bar> \<le> b \<Longrightarrow> a \<le> b"
-  by (insert abs_ge_self, blast intro: order_trans)
+by (insert abs_ge_self, blast intro: order_trans)
 
 lemma abs_le_D2: "\<bar>a\<bar> \<le> b \<Longrightarrow> - a \<le> b"
-  by (insert abs_le_D1 [of "uminus a"], simp)
+by (insert abs_le_D1 [of "uminus a"], simp)
 
 lemma abs_le_iff: "\<bar>a\<bar> \<le> b \<longleftrightarrow> a \<le> b \<and> - a \<le> b"
-  by (blast intro: abs_leI dest: abs_le_D1 abs_le_D2)
+by (blast intro: abs_leI dest: abs_le_D1 abs_le_D2)
 
 lemma abs_triangle_ineq2: "\<bar>a\<bar> - \<bar>b\<bar> \<le> \<bar>a - b\<bar>"
-  apply (simp add: compare_rls)
-  apply (subgoal_tac "abs a = abs (plus (minus a b) b)")
+  apply (simp add: algebra_simps)
+  apply (subgoal_tac "abs a = abs (plus b (minus a b))")
   apply (erule ssubst)
   apply (rule abs_triangle_ineq)
-  apply (rule arg_cong) back
-  apply (simp add: compare_rls)
+  apply (rule arg_cong[of _ _ abs])
+  apply (simp add: algebra_simps)
 done
 
 lemma abs_triangle_ineq3: "\<bar>\<bar>a\<bar> - \<bar>b\<bar>\<bar> \<le> \<bar>a - b\<bar>"
@@ -879,12 +848,9 @@
 
 lemma abs_triangle_ineq4: "\<bar>a - b\<bar> \<le> \<bar>a\<bar> + \<bar>b\<bar>"
 proof -
-  have "abs(a - b) = abs(a + - b)"
-    by (subst diff_minus, rule refl)
-  also have "... <= abs a + abs (- b)"
-    by (rule abs_triangle_ineq)
-  finally show ?thesis
-    by simp
+  have "abs(a - b) = abs(a + - b)" by (subst diff_minus, rule refl)
+  also have "... <= abs a + abs (- b)" by (rule abs_triangle_ineq)
+  finally show ?thesis by simp
 qed
 
 lemma abs_diff_triangle_ineq: "\<bar>a + b - (c + d)\<bar> \<le> \<bar>a - c\<bar> + \<bar>b - d\<bar>"
@@ -999,23 +965,19 @@
 qed
 
 lemma neg_inf_eq_sup: "- inf a b = sup (-a) (-b)"
-  by (simp add: inf_eq_neg_sup)
+by (simp add: inf_eq_neg_sup)
 
 lemma neg_sup_eq_inf: "- sup a b = inf (-a) (-b)"
-  by (simp add: sup_eq_neg_inf)
+by (simp add: sup_eq_neg_inf)
 
 lemma add_eq_inf_sup: "a + b = sup a b + inf a b"
 proof -
   have "0 = - inf 0 (a-b) + inf (a-b) 0" by (simp add: inf_commute)
   hence "0 = sup 0 (b-a) + inf (a-b) 0" by (simp add: inf_eq_neg_sup)
   hence "0 = (-a + sup a b) + (inf a b + (-b))"
-    apply (simp add: add_sup_distrib_left add_inf_distrib_right)
-    by (simp add: diff_minus add_commute)
-  thus ?thesis
-    apply (simp add: compare_rls)
-    apply (subst add_left_cancel [symmetric, of "plus a b" "plus (sup a b) (inf a b)" "uminus a"])
-    apply (simp only: add_assoc, simp add: add_assoc[symmetric])
-    done
+    by (simp add: add_sup_distrib_left add_inf_distrib_right)
+       (simp add: algebra_simps)
+  thus ?thesis by (simp add: algebra_simps)
 qed
 
 subsection {* Positive Part, Negative Part, Absolute Value *}
@@ -1044,13 +1006,13 @@
 qed
 
 lemma prts: "a = pprt a + nprt a"
-  by (simp add: pprt_def nprt_def add_eq_inf_sup[symmetric])
+by (simp add: pprt_def nprt_def add_eq_inf_sup[symmetric])
 
 lemma zero_le_pprt[simp]: "0 \<le> pprt a"
-  by (simp add: pprt_def)
+by (simp add: pprt_def)
 
 lemma nprt_le_zero[simp]: "nprt a \<le> 0"
-  by (simp add: nprt_def)
+by (simp add: nprt_def)
 
 lemma le_eq_neg: "a \<le> - b \<longleftrightarrow> a + b \<le> 0" (is "?l = ?r")
 proof -
@@ -1071,16 +1033,16 @@
 lemma nprt_0[simp]: "nprt 0 = 0" by (simp add: nprt_def)
 
 lemma pprt_eq_id [simp, noatp]: "0 \<le> x \<Longrightarrow> pprt x = x"
-  by (simp add: pprt_def le_iff_sup sup_ACI)
+by (simp add: pprt_def le_iff_sup sup_ACI)
 
 lemma nprt_eq_id [simp, noatp]: "x \<le> 0 \<Longrightarrow> nprt x = x"
-  by (simp add: nprt_def le_iff_inf inf_ACI)
+by (simp add: nprt_def le_iff_inf inf_ACI)
 
 lemma pprt_eq_0 [simp, noatp]: "x \<le> 0 \<Longrightarrow> pprt x = 0"
-  by (simp add: pprt_def le_iff_sup sup_ACI)
+by (simp add: pprt_def le_iff_sup sup_ACI)
 
 lemma nprt_eq_0 [simp, noatp]: "0 \<le> x \<Longrightarrow> nprt x = 0"
-  by (simp add: nprt_def le_iff_inf inf_ACI)
+by (simp add: nprt_def le_iff_inf inf_ACI)
 
 lemma sup_0_imp_0: "sup a (- a) = 0 \<Longrightarrow> a = 0"
 proof -
@@ -1105,10 +1067,10 @@
 done
 
 lemma inf_0_eq_0 [simp, noatp]: "inf a (- a) = 0 \<longleftrightarrow> a = 0"
-  by (rule, erule inf_0_imp_0) simp
+by (rule, erule inf_0_imp_0) simp
 
 lemma sup_0_eq_0 [simp, noatp]: "sup a (- a) = 0 \<longleftrightarrow> a = 0"
-  by (rule, erule sup_0_imp_0) simp
+by (rule, erule sup_0_imp_0) simp
 
 lemma zero_le_double_add_iff_zero_le_single_add [simp]:
   "0 \<le> a + a \<longleftrightarrow> 0 \<le> a"
@@ -1190,22 +1152,22 @@
 qed
 
 lemma zero_le_iff_zero_nprt: "0 \<le> a \<longleftrightarrow> nprt a = 0"
-  by (simp add: le_iff_inf nprt_def inf_commute)
+by (simp add: le_iff_inf nprt_def inf_commute)
 
 lemma le_zero_iff_zero_pprt: "a \<le> 0 \<longleftrightarrow> pprt a = 0"
-  by (simp add: le_iff_sup pprt_def sup_commute)
+by (simp add: le_iff_sup pprt_def sup_commute)
 
 lemma le_zero_iff_pprt_id: "0 \<le> a \<longleftrightarrow> pprt a = a"
-  by (simp add: le_iff_sup pprt_def sup_commute)
+by (simp add: le_iff_sup pprt_def sup_commute)
 
 lemma zero_le_iff_nprt_id: "a \<le> 0 \<longleftrightarrow> nprt a = a"
-  by (simp add: le_iff_inf nprt_def inf_commute)
+by (simp add: le_iff_inf nprt_def inf_commute)
 
 lemma pprt_mono [simp, noatp]: "a \<le> b \<Longrightarrow> pprt a \<le> pprt b"
-  by (simp add: le_iff_sup pprt_def sup_ACI sup_assoc [symmetric, of a])
+by (simp add: le_iff_sup pprt_def sup_ACI sup_assoc [symmetric, of a])
 
 lemma nprt_mono [simp, noatp]: "a \<le> b \<Longrightarrow> nprt a \<le> nprt b"
-  by (simp add: le_iff_inf nprt_def inf_ACI inf_assoc [symmetric, of a])
+by (simp add: le_iff_inf nprt_def inf_ACI inf_assoc [symmetric, of a])
 
 end
 
@@ -1231,7 +1193,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
@@ -1240,37 +1202,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
@@ -1287,7 +1238,7 @@
 lemma abs_if_lattice:
   fixes a :: "'a\<Colon>{lordered_ab_group_add_abs, linorder}"
   shows "\<bar>a\<bar> = (if a < 0 then - a else a)"
-  by auto
+by auto
 
 
 text {* Needed for abelian cancellation simprocs: *}
@@ -1333,7 +1284,7 @@
   "a + b <= (c::'a::lordered_ab_group_add_abs) \<Longrightarrow> a <= c + abs b" 
 proof -
   assume "a+b <= c"
-  hence 2: "a <= c+(-b)" by (simp add: group_simps)
+  hence 2: "a <= c+(-b)" by (simp add: algebra_simps)
   have 3: "(-b) <= abs b" by (rule abs_ge_minus_self)
   show ?thesis by (rule le_add_right_mono[OF 2 3])
 qed
--- a/src/HOL/Orderings.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Orderings.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Orderings.thy
-    ID:         $Id$
     Author:     Tobias Nipkow, Markus Wenzel, and Larry Paulson
 *)
 
--- a/src/HOL/PReal.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/PReal.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -718,7 +718,7 @@
     case (Suc k)
       from this obtain b where "b \<in> A" "b + of_nat k * u \<in> A" ..
       hence "b + of_int (int k)*u + u \<in> A" by (simp add: prems)
-      thus ?case by (force simp add: left_distrib add_ac prems) 
+      thus ?case by (force simp add: algebra_simps prems) 
   qed
 next
   case (neg n)
@@ -815,7 +815,7 @@
     proof -
       have "r + ?d < r + (r * ?d)/y" by (simp add: dless)
       also with ypos have "... = (r/y) * (y + ?d)"
-	by (simp only: right_distrib divide_inverse mult_ac, simp)
+	by (simp only: algebra_simps divide_inverse, simp)
       also have "... = r*x" using ypos
 	by (simp add: times_divide_eq_left) 
       finally show "r + ?d < r*x" .
--- a/src/HOL/Parity.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Parity.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -5,10 +5,10 @@
 header {* Even and Odd for int and nat *}
 
 theory Parity
-imports Plain Presburger
+imports Plain Presburger Main
 begin
 
-class even_odd = type + 
+class even_odd = 
   fixes even :: "'a \<Rightarrow> bool"
 
 abbreviation
--- a/src/HOL/Plain.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Plain.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,7 +1,7 @@
 header {* Plain HOL *}
 
 theory Plain
-imports Datatype FunDef Record SAT Extraction
+imports Datatype FunDef Record SAT Extraction Divides
 begin
 
 text {*
@@ -9,6 +9,9 @@
   include @{text Hilbert_Choice}.
 *}
 
+instance option :: (finite) finite
+  by default (simp add: insert_None_conv_UNIV [symmetric])
+
 ML {* path_add "~~/src/HOL/Library" *}
 
 end
--- a/src/HOL/Polynomial.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Polynomial.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -6,7 +6,7 @@
 header {* Univariate Polynomials *}
 
 theory Polynomial
-imports Plain SetInterval
+imports Plain SetInterval Main
 begin
 
 subsection {* Definition of type @{text poly} *}
@@ -208,7 +208,7 @@
 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]:
+  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)
 
@@ -293,6 +293,14 @@
 
 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
 
@@ -345,19 +353,22 @@
   "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: "degree (p + q) \<le> max (degree p) (degree q)"
+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)
+  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 (rule ord_le_eq_trans [OF degree_add_le])
-  apply simp
+  apply (simp add: degree_add_le)
   apply (rule le_degree)
   apply (simp add: coeff_eq_0)
   done
@@ -370,13 +381,17 @@
 lemma degree_minus [simp]: "degree (- p) = degree p"
   unfolding degree_def by simp
 
-lemma degree_diff_le: "degree (p - q) \<le> max (degree p) (degree q)"
+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 (auto intro: le_less_trans degree_diff_le)
+  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
@@ -427,11 +442,11 @@
 
 lemma smult_add_right:
   "smult a (p + q) = smult a p + smult a q"
-  by (rule poly_ext, simp add: ring_simps)
+  by (rule poly_ext, simp add: algebra_simps)
 
 lemma smult_add_left:
   "smult (a + b) p = smult a p + smult b p"
-  by (rule poly_ext, simp add: ring_simps)
+  by (rule poly_ext, simp add: algebra_simps)
 
 lemma smult_minus_right [simp]:
   "smult (a::'a::comm_ring) (- p) = - smult a p"
@@ -443,11 +458,11 @@
 
 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)
+  by (rule poly_ext, simp add: algebra_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)
+  by (rule poly_ext, simp add: algebra_simps)
 
 lemmas smult_distribs =
   smult_add_left smult_add_right
@@ -460,60 +475,20 @@
 lemma smult_monom: "smult a (monom b n) = monom (a * b) n"
   by (induct n, simp add: monom_0, simp add: monom_Suc)
 
+lemma degree_smult_eq [simp]:
+  fixes a :: "'a::idom"
+  shows "degree (smult a p) = (if a = 0 then 0 else degree p)"
+  by (cases "a = 0", simp, simp add: degree_def)
+
+lemma smult_eq_0_iff [simp]:
+  fixes a :: "'a::idom"
+  shows "smult a p = 0 \<longleftrightarrow> a = 0 \<or> p = 0"
+  by (simp add: expand_poly_eq)
+
 
 subsection {* Multiplication of polynomials *}
 
-lemma Poly_mult_lemma:
-  fixes f g :: "nat \<Rightarrow> 'a::comm_semiring_0" and m n :: nat
-  assumes "\<forall>i>m. f i = 0"
-  assumes "\<forall>j>n. g j = 0"
-  shows "\<forall>k>m+n. (\<Sum>i\<le>k. f i * g (k-i)) = 0"
-proof (clarify)
-  fix k :: nat
-  assume "m + n < k"
-  show "(\<Sum>i\<le>k. f i * g (k - i)) = 0"
-  proof (rule setsum_0' [rule_format])
-    fix i :: nat
-    assume "i \<in> {..k}" hence "i \<le> k" by simp
-    with `m + n < k` have "m < i \<or> n < k - i" by arith
-    thus "f i * g (k - i) = 0"
-      using prems by auto
-  qed
-qed
-
-lemma Poly_mult:
-  fixes f g :: "nat \<Rightarrow> 'a::comm_semiring_0"
-  shows "\<lbrakk>f \<in> Poly; g \<in> Poly\<rbrakk> \<Longrightarrow> (\<lambda>n. \<Sum>i\<le>n. f i * g (n-i)) \<in> Poly"
-  unfolding Poly_def
-  by (safe, rule exI, rule Poly_mult_lemma)
-
-lemma poly_mult_assoc_lemma:
-  fixes k :: nat and f :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a::comm_monoid_add"
-  shows "(\<Sum>j\<le>k. \<Sum>i\<le>j. f i (j - i) (n - j)) =
-         (\<Sum>j\<le>k. \<Sum>i\<le>k - j. f j i (n - j - i))"
-proof (induct k)
-  case 0 show ?case by simp
-next
-  case (Suc k) thus ?case
-    by (simp add: Suc_diff_le setsum_addf add_assoc
-             cong: strong_setsum_cong)
-qed
-
-lemma poly_mult_commute_lemma:
-  fixes n :: nat and f :: "nat \<Rightarrow> nat \<Rightarrow> 'a::comm_monoid_add"
-  shows "(\<Sum>i\<le>n. f i (n - i)) = (\<Sum>i\<le>n. f (n - i) i)"
-proof (rule setsum_reindex_cong)
-  show "inj_on (\<lambda>i. n - i) {..n}"
-    by (rule inj_onI) simp
-  show "{..n} = (\<lambda>i. n - i) ` {..n}"
-    by (auto, rule_tac x="n - x" in image_eqI, simp_all)
-next
-  fix i assume "i \<in> {..n}"
-  hence "n - (n - i) = i" by simp
-  thus "f (n - i) i = f (n - i) (n - (n - i))" by simp
-qed
-
-text {* TODO: move to appropriate theory *}
+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))"
@@ -537,69 +512,71 @@
 begin
 
 definition
-  times_poly_def:
-    "p * q = Abs_poly (\<lambda>n. \<Sum>i\<le>n. coeff p i * coeff q (n-i))"
+  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 coeff_mult:
-  "coeff (p * q) n = (\<Sum>i\<le>n. coeff p i * coeff q (n-i))"
-  unfolding times_poly_def
-  by (simp add: Abs_poly_inverse coeff Poly_mult)
+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: algebra_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 algebra_simps)
 
 instance proof
   fix p q r :: "'a poly"
   show 0: "0 * p = 0"
-    by (simp add: expand_poly_eq coeff_mult)
+    by (rule mult_poly_0_left)
   show "p * 0 = 0"
-    by (simp add: expand_poly_eq coeff_mult)
+    by (rule mult_poly_0_right)
   show "(p + q) * r = p * r + q * r"
-    by (simp add: expand_poly_eq coeff_mult left_distrib setsum_addf)
+    by (rule mult_poly_add_left)
   show "(p * q) * r = p * (q * r)"
-  proof (rule poly_ext)
-    fix n :: nat
-    have "(\<Sum>j\<le>n. \<Sum>i\<le>j. coeff p i * coeff q (j - i) * coeff r (n - j)) =
-          (\<Sum>j\<le>n. \<Sum>i\<le>n - j. coeff p j * coeff q i * coeff r (n - j - i))"
-      by (rule poly_mult_assoc_lemma)
-    thus "coeff ((p * q) * r) n = coeff (p * (q * r)) n"
-      by (simp add: coeff_mult setsum_right_distrib
-                    setsum_left_distrib mult_assoc)
-  qed
+    by (induct p, simp add: mult_poly_0, simp add: mult_poly_add_left)
   show "p * q = q * p"
-  proof (rule poly_ext)
-    fix n :: nat
-    have "(\<Sum>i\<le>n. coeff p i * coeff q (n - i)) =
-          (\<Sum>i\<le>n. coeff p (n - i) * coeff q i)"
-      by (rule poly_mult_commute_lemma)
-    thus "coeff (p * q) n = coeff (q * p) n"
-      by (simp add: coeff_mult mult_commute)
-  qed
+    by (induct p, simp add: mult_poly_0, simp)
 qed
 
 end
 
-lemma degree_mult_le: "degree (p * q) \<le> degree p + degree q"
-apply (rule degree_le, simp add: coeff_mult)
-apply (rule Poly_mult_lemma)
-apply (simp_all add: coeff_eq_0)
-done
+instance poly :: (comm_semiring_0_cancel) comm_semiring_0_cancel ..
 
-lemma mult_pCons_left [simp]:
-  "pCons a p * q = smult a q + pCons 0 (p * q)"
-apply (rule poly_ext)
-apply (case_tac n)
-apply (simp add: coeff_mult)
-apply (simp add: coeff_mult setsum_atMost_Suc_shift
-            del: setsum_atMost_Suc)
-done
+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 mult_pCons_right [simp]:
-  "p * pCons a q = smult a p + pCons 0 (p * q)"
-  using mult_pCons_left [of a q p] by (simp add: mult_commute)
-
-lemma mult_smult_left [simp]: "smult a p * q = smult a (p * q)"
-  by (induct p, simp, simp add: smult_add_right)
-
-lemma mult_smult_right [simp]: "p * smult a q = smult a (p * q)"
-  by (induct q, simp, simp add: smult_add_right)
+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)
@@ -625,6 +602,8 @@
 
 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)
@@ -697,18 +676,19 @@
 subsection {* Long division of polynomials *}
 
 definition
-  divmod_poly_rel :: "'a::field poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<Rightarrow> bool"
+  pdivmod_rel :: "'a::field poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<Rightarrow> bool"
 where
-  "divmod_poly_rel x y q r \<longleftrightarrow>
+  [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 divmod_poly_rel_0:
-  "divmod_poly_rel 0 y 0 0"
-  unfolding divmod_poly_rel_def by simp
+lemma pdivmod_rel_0:
+  "pdivmod_rel 0 y 0 0"
+  unfolding pdivmod_rel_def by simp
 
-lemma divmod_poly_rel_by_0:
-  "divmod_poly_rel x 0 0 x"
-  unfolding divmod_poly_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"
@@ -734,64 +714,61 @@
   then show ?thesis ..
 qed
 
-lemma divmod_poly_rel_pCons:
-  assumes rel: "divmod_poly_rel x y q r"
+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 "divmod_poly_rel (pCons a x) y (pCons b q) (pCons a r - smult b y)"
-    (is "divmod_poly_rel ?x y ?q ?r")
+  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 divmod_poly_rel_def by simp_all
+    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)
-    have "degree ?r \<le> max (degree (pCons a r)) (degree (smult b y))"
-      by (rule degree_diff_le)
-    also have "\<dots> \<le> degree y"
-    proof (rule min_max.le_supI)
+    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
-    finally show "degree ?r \<le> degree y" .
   next
     show "coeff ?r (degree y) = 0"
       using `y \<noteq> 0` unfolding b by simp
   qed
 
   from 1 2 show ?thesis
-    unfolding divmod_poly_rel_def
+    unfolding pdivmod_rel_def
     using `y \<noteq> 0` by simp
 qed
 
-lemma divmod_poly_rel_exists: "\<exists>q r. divmod_poly_rel x y q r"
+lemma pdivmod_rel_exists: "\<exists>q r. pdivmod_rel x y q r"
 apply (cases "y = 0")
-apply (fast intro!: divmod_poly_rel_by_0)
+apply (fast intro!: pdivmod_rel_by_0)
 apply (induct x)
-apply (fast intro!: divmod_poly_rel_0)
-apply (fast intro!: divmod_poly_rel_pCons)
+apply (fast intro!: pdivmod_rel_0)
+apply (fast intro!: pdivmod_rel_pCons)
 done
 
-lemma divmod_poly_rel_unique:
-  assumes 1: "divmod_poly_rel x y q1 r1"
-  assumes 2: "divmod_poly_rel x y q2 r2"
+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: divmod_poly_rel_def)
+    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 divmod_poly_rel_def by simp_all
+    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 divmod_poly_rel_def by simp_all
+    unfolding pdivmod_rel_def by simp_all
   from q1 q2 have q3: "(q1 - q2) * y = r2 - r1"
-    by (simp add: ring_simps)
+    by (simp add: algebra_simps)
   from r1 r2 have r3: "(r2 - r1) = 0 \<or> degree (r2 - r1) < degree y"
     by (auto intro: degree_diff_less)
 
@@ -810,36 +787,42 @@
   qed
 qed
 
-lemmas divmod_poly_rel_unique_div =
-  divmod_poly_rel_unique [THEN conjunct1, standard]
+lemma pdivmod_rel_0_iff: "pdivmod_rel 0 y q r \<longleftrightarrow> q = 0 \<and> r = 0"
+by (auto dest: pdivmod_rel_unique intro: pdivmod_rel_0)
+
+lemma pdivmod_rel_by_0_iff: "pdivmod_rel x 0 q r \<longleftrightarrow> q = 0 \<and> r = x"
+by (auto dest: pdivmod_rel_unique intro: pdivmod_rel_by_0)
 
-lemmas divmod_poly_rel_unique_mod =
-  divmod_poly_rel_unique [THEN conjunct2, standard]
+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. divmod_poly_rel x y q r)"
+  [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. divmod_poly_rel x y q r)"
+  [code del]: "x mod y = (THE r. \<exists>q. pdivmod_rel x y q r)"
 
 lemma div_poly_eq:
-  "divmod_poly_rel x y q r \<Longrightarrow> x div y = q"
+  "pdivmod_rel x y q r \<Longrightarrow> x div y = q"
 unfolding div_poly_def
-by (fast elim: divmod_poly_rel_unique_div)
+by (fast elim: pdivmod_rel_unique_div)
 
 lemma mod_poly_eq:
-  "divmod_poly_rel x y q r \<Longrightarrow> x mod y = r"
+  "pdivmod_rel x y q r \<Longrightarrow> x mod y = r"
 unfolding mod_poly_def
-by (fast elim: divmod_poly_rel_unique_mod)
+by (fast elim: pdivmod_rel_unique_mod)
 
-lemma divmod_poly_rel:
-  "divmod_poly_rel x y (x div y) (x mod y)"
+lemma pdivmod_rel:
+  "pdivmod_rel x y (x div y) (x mod y)"
 proof -
-  from divmod_poly_rel_exists
-    obtain q r where "divmod_poly_rel x y q r" by fast
+  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
@@ -847,26 +830,26 @@
 instance proof
   fix x y :: "'a poly"
   show "x div y * y + x mod y = x"
-    using divmod_poly_rel [of x y]
-    by (simp add: divmod_poly_rel_def)
+    using pdivmod_rel [of x y]
+    by (simp add: pdivmod_rel_def)
 next
   fix x :: "'a poly"
-  have "divmod_poly_rel x 0 0 x"
-    by (rule divmod_poly_rel_by_0)
+  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 "divmod_poly_rel 0 y 0 0"
-    by (rule divmod_poly_rel_0)
+  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 "divmod_poly_rel (x + z * y) y (z + x div y) (x mod y)"
-    using divmod_poly_rel [of x y]
-    by (simp add: divmod_poly_rel_def left_distrib)
+  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
@@ -875,25 +858,73 @@
 
 lemma degree_mod_less:
   "y \<noteq> 0 \<Longrightarrow> x mod y = 0 \<or> degree (x mod y) < degree y"
-  using divmod_poly_rel [of x y]
-  unfolding divmod_poly_rel_def by simp
+  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 "divmod_poly_rel x y 0 x"
-    by (simp add: divmod_poly_rel_def)
+  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 "divmod_poly_rel x y 0 x"
-    by (simp add: divmod_poly_rel_def)
+  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 pdivmod_rel_smult_left:
+  "pdivmod_rel x y q r
+    \<Longrightarrow> pdivmod_rel (smult a x) y (smult a q) (smult a r)"
+  unfolding pdivmod_rel_def by (simp add: smult_add_right)
+
+lemma div_smult_left: "(smult a x) div y = smult a (x div y)"
+  by (rule div_poly_eq, rule pdivmod_rel_smult_left, rule pdivmod_rel)
+
+lemma mod_smult_left: "(smult a x) mod y = smult a (x mod y)"
+  by (rule mod_poly_eq, rule pdivmod_rel_smult_left, rule pdivmod_rel)
+
+lemma pdivmod_rel_smult_right:
+  "\<lbrakk>a \<noteq> 0; pdivmod_rel x y q r\<rbrakk>
+    \<Longrightarrow> pdivmod_rel x (smult a y) (smult (inverse a) q) r"
+  unfolding pdivmod_rel_def by simp
+
+lemma div_smult_right:
+  "a \<noteq> 0 \<Longrightarrow> x div (smult a y) = smult (inverse a) (x div y)"
+  by (rule div_poly_eq, erule pdivmod_rel_smult_right, rule pdivmod_rel)
+
+lemma mod_smult_right: "a \<noteq> 0 \<Longrightarrow> x mod (smult a y) = x mod y"
+  by (rule mod_poly_eq, erule pdivmod_rel_smult_right, rule pdivmod_rel)
+
+lemma pdivmod_rel_mult:
+  "\<lbrakk>pdivmod_rel x y q r; pdivmod_rel q z q' r'\<rbrakk>
+    \<Longrightarrow> pdivmod_rel x (y * z) q' (y * r' + r)"
+apply (cases "z = 0", simp add: pdivmod_rel_def)
+apply (cases "y = 0", simp add: pdivmod_rel_by_0_iff pdivmod_rel_0_iff)
+apply (cases "r = 0")
+apply (cases "r' = 0")
+apply (simp add: pdivmod_rel_def)
+apply (simp add: pdivmod_rel_def ring_simps degree_mult_eq)
+apply (cases "r' = 0")
+apply (simp add: pdivmod_rel_def degree_mult_eq)
+apply (simp add: pdivmod_rel_def ring_simps)
+apply (simp add: degree_mult_eq degree_add_less)
+done
+
+lemma poly_div_mult_right:
+  fixes x y z :: "'a::field poly"
+  shows "x div (y * z) = (x div y) div z"
+  by (rule div_poly_eq, rule pdivmod_rel_mult, (rule pdivmod_rel)+)
+
+lemma poly_mod_mult_right:
+  fixes x y z :: "'a::field poly"
+  shows "x mod (y * z) = y * (x div y mod z) + x mod y"
+  by (rule mod_poly_eq, rule pdivmod_rel_mult, (rule pdivmod_rel)+)
+
 lemma mod_pCons:
   fixes a and x
   assumes y: "y \<noteq> 0"
@@ -901,7 +932,7 @@
   shows "(pCons a x) mod y = (pCons a (x mod y) - smult b y)"
 unfolding b
 apply (rule mod_poly_eq)
-apply (rule divmod_poly_rel_pCons [OF divmod_poly_rel y refl])
+apply (rule pdivmod_rel_pCons [OF pdivmod_rel y refl])
 done
 
 
@@ -927,7 +958,7 @@
 
 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)
+  apply (case_tac q, simp, simp add: algebra_simps)
   done
 
 lemma poly_minus [simp]:
@@ -944,10 +975,10 @@
   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)
+  by (induct p, simp, simp add: algebra_simps)
 
 lemma poly_mult [simp]: "poly (p * q) x = poly p x * poly q x"
-  by (induct p, simp_all, simp add: ring_simps)
+  by (induct p, simp_all, simp add: algebra_simps)
 
 lemma poly_power [simp]:
   fixes p :: "'a::{comm_semiring_1,recpower} poly"
@@ -959,7 +990,7 @@
 
 definition
   synthetic_divmod :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a poly \<times> 'a"
-where
+where [code del]:
   "synthetic_divmod p c =
     poly_rec (0, 0) (\<lambda>a p (q, r). (pCons r q, a + c * r)) p"
 
@@ -1016,7 +1047,7 @@
   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)
+  by (simp add: algebra_simps)
 
 lemma poly_eq_0_iff_dvd:
   fixes c :: "'a::idom"
@@ -1068,4 +1099,98 @@
   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/Power.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Power.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -11,7 +11,7 @@
 imports Nat
 begin
 
-class power = type +
+class power =
   fixes power :: "'a \<Rightarrow> nat \<Rightarrow> 'a"            (infixr "^" 80)
 
 subsection{*Powers for Arbitrary Monoids*}
--- a/src/HOL/Presburger.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Presburger.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -59,7 +59,7 @@
   "(d::'a::{comm_ring,Ring_and_Field.dvd}) dvd D \<Longrightarrow> \<forall>x k. (d dvd x + t) = (d dvd (x - k*D) + t)"
   "(d::'a::{comm_ring,Ring_and_Field.dvd}) dvd D \<Longrightarrow> \<forall>x k. (\<not>d dvd x + t) = (\<not>d dvd (x - k*D) + t)"
   "\<forall>x k. F = F"
-apply (auto elim!: dvdE simp add: ring_simps)
+apply (auto elim!: dvdE simp add: algebra_simps)
 unfolding mult_assoc [symmetric] left_distrib [symmetric] left_diff_distrib [symmetric]
 unfolding dvd_def mult_commute [of d] 
 by auto
@@ -101,7 +101,7 @@
   {fix x assume nob: "\<forall>j\<in>{1 .. D}. \<forall>b\<in>B. x \<noteq> b + j" and g: "x > t" and ng: "\<not> (x - D) > t"
     hence "x -t \<le> D" and "1 \<le> x - t" by simp+
       hence "\<exists>j \<in> {1 .. D}. x - t = j" by auto
-      hence "\<exists>j \<in> {1 .. D}. x = t + j" by (simp add: ring_simps)
+      hence "\<exists>j \<in> {1 .. D}. x = t + j" by (simp add: algebra_simps)
       with nob tB have "False" by simp}
   thus "\<forall>x.(\<forall>j\<in>{1 .. D}. \<forall>b\<in>B. x \<noteq> b + j)\<longrightarrow> (x > t) \<longrightarrow> (x - D > t)" by blast
 next
@@ -109,7 +109,7 @@
   {fix x assume nob: "\<forall>j\<in>{1 .. D}. \<forall>b\<in>B. x \<noteq> b + j" and g: "x \<ge> t" and ng: "\<not> (x - D) \<ge> t"
     hence "x - (t - 1) \<le> D" and "1 \<le> x - (t - 1)" by simp+
       hence "\<exists>j \<in> {1 .. D}. x - (t - 1) = j" by auto
-      hence "\<exists>j \<in> {1 .. D}. x = (t - 1) + j" by (simp add: ring_simps)
+      hence "\<exists>j \<in> {1 .. D}. x = (t - 1) + j" by (simp add: algebra_simps)
       with nob tB have "False" by simp}
   thus "\<forall>x.(\<forall>j\<in>{1 .. D}. \<forall>b\<in>B. x \<noteq> b + j)\<longrightarrow> (x \<ge> t) \<longrightarrow> (x - D \<ge> t)" by blast
 next
@@ -119,7 +119,7 @@
 next
   assume d: "d dvd D"
   {fix x assume H: "\<not>(d dvd x + t)" with d have "\<not> d dvd (x - D) + t"
-      by (clarsimp simp add: dvd_def,erule_tac x= "ka + k" in allE,simp add: ring_simps)}
+      by (clarsimp simp add: dvd_def,erule_tac x= "ka + k" in allE,simp add: algebra_simps)}
   thus "\<forall>(x::int).(\<forall>j\<in>{1 .. D}. \<forall>b\<in>B. x \<noteq> b + j)\<longrightarrow> (\<not>d dvd x+t) \<longrightarrow> (\<not>d dvd (x - D) + t)" by auto
 qed blast
 
@@ -158,26 +158,26 @@
   {fix x assume nob: "\<forall>j\<in>{1 .. D}. \<forall>b\<in>A. x \<noteq> b - j" and g: "x < t" and ng: "\<not> (x + D) < t"
     hence "t - x \<le> D" and "1 \<le> t - x" by simp+
       hence "\<exists>j \<in> {1 .. D}. t - x = j" by auto
-      hence "\<exists>j \<in> {1 .. D}. x = t - j" by (auto simp add: ring_simps) 
+      hence "\<exists>j \<in> {1 .. D}. x = t - j" by (auto simp add: algebra_simps) 
       with nob tA have "False" by simp}
   thus "\<forall>x.(\<forall>j\<in>{1 .. D}. \<forall>b\<in>A. x \<noteq> b - j)\<longrightarrow> (x < t) \<longrightarrow> (x + D < t)" by blast
 next
   assume dp: "D > 0" and tA:"t + 1\<in> A"
   {fix x assume nob: "\<forall>j\<in>{1 .. D}. \<forall>b\<in>A. x \<noteq> b - j" and g: "x \<le> t" and ng: "\<not> (x + D) \<le> t"
-    hence "(t + 1) - x \<le> D" and "1 \<le> (t + 1) - x" by (simp_all add: ring_simps)
+    hence "(t + 1) - x \<le> D" and "1 \<le> (t + 1) - x" by (simp_all add: algebra_simps)
       hence "\<exists>j \<in> {1 .. D}. (t + 1) - x = j" by auto
-      hence "\<exists>j \<in> {1 .. D}. x = (t + 1) - j" by (auto simp add: ring_simps)
+      hence "\<exists>j \<in> {1 .. D}. x = (t + 1) - j" by (auto simp add: algebra_simps)
       with nob tA have "False" by simp}
   thus "\<forall>x.(\<forall>j\<in>{1 .. D}. \<forall>b\<in>A. x \<noteq> b - j)\<longrightarrow> (x \<le> t) \<longrightarrow> (x + D \<le> t)" by blast
 next
   assume d: "d dvd D"
   {fix x assume H: "d dvd x + t" with d have "d dvd (x + D) + t"
-      by (clarsimp simp add: dvd_def,rule_tac x= "ka + k" in exI,simp add: ring_simps)}
+      by (clarsimp simp add: dvd_def,rule_tac x= "ka + k" in exI,simp add: algebra_simps)}
   thus "\<forall>(x::int).(\<forall>j\<in>{1 .. D}. \<forall>b\<in>A. x \<noteq> b - j)\<longrightarrow> (d dvd x+t) \<longrightarrow> (d dvd (x + D) + t)" by simp
 next
   assume d: "d dvd D"
   {fix x assume H: "\<not>(d dvd x + t)" with d have "\<not>d dvd (x + D) + t"
-      by (clarsimp simp add: dvd_def,erule_tac x= "ka - k" in allE,simp add: ring_simps)}
+      by (clarsimp simp add: dvd_def,erule_tac x= "ka - k" in allE,simp add: algebra_simps)}
   thus "\<forall>(x::int).(\<forall>j\<in>{1 .. D}. \<forall>b\<in>A. x \<noteq> b - j)\<longrightarrow> (\<not>d dvd x+t) \<longrightarrow> (\<not>d dvd (x + D) + t)" by auto
 qed blast
 
@@ -304,7 +304,7 @@
   from ePeqP1 obtain z where P1eqP: "\<forall>x>z. P x = P' x" ..
   let ?w' = "x + (abs(x-z)+1) * d"
   let ?w = "x - (-(abs(x-z) + 1))*d"
-  have ww'[simp]: "?w = ?w'" by (simp add: ring_simps)
+  have ww'[simp]: "?w = ?w'" by (simp add: algebra_simps)
   from dpos have w: "?w > z" by(simp only: ww' incr_lemma)
   hence "P' x = P' ?w" using P1eqP1 by blast
   also have "\<dots> = P(?w)" using w P1eqP by blast
--- a/src/HOL/RComplete.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/RComplete.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -376,7 +376,7 @@
   hence "real (Suc n) * (inverse (real (Suc n)) * x) < real (Suc n) * 1"
     by (rule mult_strict_left_mono) simp
   hence "x < real (Suc n)"
-    by (simp add: ring_simps)
+    by (simp add: algebra_simps)
   thus "\<exists>(n::nat). x < real n" ..
 qed
 
@@ -391,9 +391,9 @@
   hence "y * inverse x * x < real n * x"
     using x_greater_zero by (simp add: mult_strict_right_mono)
   hence "x * inverse x * y < x * real n"
-    by (simp add: ring_simps)
+    by (simp add: algebra_simps)
   hence "y < real (n::nat) * x"
-    using x_not_zero by (simp add: ring_simps)
+    using x_not_zero by (simp add: algebra_simps)
   thus "\<exists>(n::nat). y < real n * x" ..
 qed
 
@@ -1084,9 +1084,7 @@
 done
 
 lemma real_natfloor_gt_diff_one: "x - 1 < real(natfloor x)"
-  apply (simp add: compare_rls)
-  apply (rule real_natfloor_add_one_gt)
-done
+using real_natfloor_add_one_gt by (simp add: algebra_simps)
 
 lemma ge_natfloor_plus_one_imp_gt: "natfloor z + 1 <= n ==> z < real n"
   apply (subgoal_tac "z < real(natfloor z) + 1")
@@ -1279,7 +1277,7 @@
     by simp
   also have "... = real((natfloor x) div y) + real((natfloor x) mod y) /
     real y + (x - real(natfloor x)) / real y"
-    by (auto simp add: ring_simps add_divide_distrib
+    by (auto simp add: algebra_simps add_divide_distrib
       diff_divide_distrib prems)
   finally have "natfloor (x / real y) = natfloor(...)" by simp
   also have "... = natfloor(real((natfloor x) mod y) /
@@ -1293,7 +1291,7 @@
     apply simp
     apply (simp add: prems)
     apply (rule divide_nonneg_pos)
-    apply (simp add: compare_rls)
+    apply (simp add: algebra_simps)
     apply (rule real_natfloor_le)
     apply (insert prems, auto)
     done
@@ -1306,7 +1304,7 @@
     apply force
     apply (force simp add: prems)
     apply (rule divide_nonneg_pos)
-    apply (simp add: compare_rls)
+    apply (simp add: algebra_simps)
     apply (rule real_natfloor_le)
     apply (auto simp add: prems)
     apply (insert prems, arith)
@@ -1314,8 +1312,8 @@
     apply (subgoal_tac "real y = real y - 1 + 1")
     apply (erule ssubst)
     apply (rule add_le_less_mono)
-    apply (simp add: compare_rls)
-    apply (subgoal_tac "real(natfloor x mod y) + 1 =
+    apply (simp add: algebra_simps)
+    apply (subgoal_tac "1 + real(natfloor x mod y) =
       real(natfloor x mod y + 1)")
     apply (erule ssubst)
     apply (subst real_of_nat_le_iff)
@@ -1323,9 +1321,8 @@
     apply arith
     apply (rule mod_less_divisor)
     apply auto
-    apply (simp add: compare_rls)
-    apply (subst add_commute)
-    apply (rule real_natfloor_add_one_gt)
+    using real_natfloor_add_one_gt
+    apply (simp add: algebra_simps)
     done
   finally show ?thesis by simp
 qed
--- a/src/HOL/ROOT.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/ROOT.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,6 +1,7 @@
 (* Classical Higher-order Logic -- batteries included *)
 
 use_thy "Main";
+share_common_data ();
 use_thy "Complex_Main";
 
 val HOL_proofs = ! Proofterm.proofs;
--- a/src/HOL/Rational.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Rational.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -172,7 +172,7 @@
     by (cases q) (simp add: One_rat_def eq_rat)
 next
   fix q r s :: rat show "(q + r) + s = q + (r + s)"
-    by (cases q, cases r, cases s) (simp add: eq_rat ring_simps)
+    by (cases q, cases r, cases s) (simp add: eq_rat algebra_simps)
 next
   fix q r :: rat show "q + r = r + q"
     by (cases q, cases r) (simp add: eq_rat)
@@ -187,7 +187,7 @@
     by (cases q, cases r) (simp add: eq_rat)
 next
   fix q r s :: rat show "(q + r) * s = q * s + r * s"
-    by (cases q, cases r, cases s) (simp add: eq_rat ring_simps)
+    by (cases q, cases r, cases s) (simp add: eq_rat algebra_simps)
 next
   show "(0::rat) \<noteq> 1" by (simp add: Zero_rat_def One_rat_def eq_rat)
 next
--- a/src/HOL/RealDef.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/RealDef.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -202,18 +202,17 @@
 
 lemma real_mult_assoc: "((z1::real) * z2) * z3 = z1 * (z2 * z3)"
 apply (cases z1, cases z2, cases z3)
-apply (simp add: real_mult right_distrib add_ac mult_ac)
+apply (simp add: real_mult algebra_simps)
 done
 
 lemma real_mult_1: "(1::real) * z = z"
 apply (cases z)
-apply (simp add: real_mult real_one_def right_distrib
-                  mult_1_right mult_ac add_ac)
+apply (simp add: real_mult real_one_def algebra_simps)
 done
 
 lemma real_add_mult_distrib: "((z1::real) + z2) * w = (z1 * w) + (z2 * w)"
 apply (cases z1, cases z2, cases w)
-apply (simp add: real_add real_mult right_distrib add_ac mult_ac)
+apply (simp add: real_add real_mult algebra_simps)
 done
 
 text{*one and zero are distinct*}
@@ -253,7 +252,7 @@
 apply (rule_tac [2]
         x = "Abs_Real (realrel``{(inverse (D) + 1, 1)})" 
        in exI)
-apply (auto simp add: real_mult preal_mult_inverse_right ring_simps)
+apply (auto simp add: real_mult preal_mult_inverse_right algebra_simps)
 done
 
 lemma real_mult_inverse_left: "x \<noteq> 0 ==> inverse(x)*x = (1::real)"
@@ -373,7 +372,7 @@
   assumes le: "x \<le> y" shows "z + x \<le> z + (y::real)"
 proof -
   have "z + x - (z + y) = (z + -z) + (x - y)" 
-    by (simp add: diff_minus add_ac) 
+    by (simp add: algebra_simps) 
   with le show ?thesis 
     by (simp add: real_le_eq_diff[of x] real_le_eq_diff[of "z+x"] diff_minus)
 qed
@@ -391,8 +390,7 @@
                   real_zero_def real_le real_mult)
   --{*Reduce to the (simpler) @{text "\<le>"} relation *}
 apply (auto dest!: less_add_left_Ex
-     simp add: add_ac mult_ac
-          right_distrib preal_self_less_add_left)
+     simp add: algebra_simps preal_self_less_add_left)
 done
 
 lemma real_mult_less_mono2: "[| (0::real) < z; x < y |] ==> z * x < z * y"
@@ -437,11 +435,11 @@
 
 lemma real_of_preal_add:
      "real_of_preal ((x::preal) + y) = real_of_preal x + real_of_preal y"
-by (simp add: real_of_preal_def real_add left_distrib add_ac)
+by (simp add: real_of_preal_def real_add algebra_simps)
 
 lemma real_of_preal_mult:
      "real_of_preal ((x::preal) * y) = real_of_preal x* real_of_preal y"
-by (simp add: real_of_preal_def real_mult right_distrib add_ac mult_ac)
+by (simp add: real_of_preal_def real_mult algebra_simps)
 
 
 text{*Gleason prop 9-4.4 p 127*}
@@ -650,7 +648,7 @@
   then have "real x / real d = ... / real d"
     by simp
   then show ?thesis
-    by (auto simp add: add_divide_distrib ring_simps prems)
+    by (auto simp add: add_divide_distrib algebra_simps prems)
 qed
 
 lemma real_of_int_div: "(d::int) ~= 0 ==> d dvd n ==>
@@ -665,13 +663,13 @@
   apply (case_tac "x = 0")
   apply simp
   apply (case_tac "0 < x")
-  apply (simp add: compare_rls)
+  apply (simp add: algebra_simps)
   apply (subst real_of_int_div_aux)
   apply simp
   apply simp
   apply (subst zero_le_divide_iff)
   apply auto
-  apply (simp add: compare_rls)
+  apply (simp add: algebra_simps)
   apply (subst real_of_int_div_aux)
   apply simp
   apply simp
@@ -683,7 +681,7 @@
   "real (n::int) / real (x) - real (n div x) <= 1"
   apply(case_tac "x = 0")
   apply simp
-  apply (simp add: compare_rls)
+  apply (simp add: algebra_simps)
   apply (subst real_of_int_div_aux)
   apply assumption
   apply simp
@@ -795,7 +793,7 @@
   then have "real x / real d = \<dots> / real d"
     by simp
   then show ?thesis
-    by (auto simp add: add_divide_distrib ring_simps prems)
+    by (auto simp add: add_divide_distrib algebra_simps prems)
 qed
 
 lemma real_of_nat_div: "0 < (d::nat) ==> d dvd n ==>
@@ -810,7 +808,7 @@
   "0 <= real (n::nat) / real (x) - real (n div x)"
 apply(case_tac "x = 0")
  apply (simp)
-apply (simp add: compare_rls)
+apply (simp add: algebra_simps)
 apply (subst real_of_nat_div_aux)
  apply simp
 apply simp
@@ -822,14 +820,14 @@
   "real (n::nat) / real (x) - real (n div x) <= 1"
 apply(case_tac "x = 0")
 apply (simp)
-apply (simp add: compare_rls)
+apply (simp add: algebra_simps)
 apply (subst real_of_nat_div_aux)
  apply simp
 apply simp
 done
 
 lemma real_of_nat_div4: "real (n div x) <= real (n::nat) / real x" 
-  by (insert real_of_nat_div2 [of n x], simp)
+by (insert real_of_nat_div2 [of n x], simp)
 
 lemma real_of_int_real_of_nat: "real (int n) = real n"
 by (simp add: real_of_nat_def real_of_int_def int_eq_of_nat)
--- a/src/HOL/RealPow.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/RealPow.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -58,7 +58,7 @@
 lemma realpow_two_diff:
      "(x::real)^Suc (Suc 0) - y^Suc (Suc 0) = (x - y) * (x + y)"
 apply (unfold real_diff_def)
-apply (simp add: ring_simps)
+apply (simp add: algebra_simps)
 done
 
 lemma realpow_two_disj:
--- a/src/HOL/RealVector.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/RealVector.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -124,7 +124,7 @@
 
 subsection {* Real vector spaces *}
 
-class scaleR = type +
+class scaleR =
   fixes scaleR :: "real \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "*\<^sub>R" 75)
 begin
 
@@ -418,7 +418,7 @@
 
 subsection {* Real normed vector spaces *}
 
-class norm = type +
+class norm =
   fixes norm :: "'a \<Rightarrow> real"
 
 instantiation real :: norm
--- a/src/HOL/Recdef.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Recdef.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,12 +1,11 @@
 (*  Title:      HOL/Recdef.thy
-    ID:         $Id$
     Author:     Konrad Slind and Markus Wenzel, TU Muenchen
 *)
 
 header {* TFL: recursive function definitions *}
 
 theory Recdef
-imports FunDef
+imports FunDef Plain
 uses
   ("Tools/TFL/casesplit.ML")
   ("Tools/TFL/utils.ML")
--- a/src/HOL/Relation.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Relation.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Relation.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1996  University of Cambridge
 *)
@@ -7,7 +6,7 @@
 header {* Relations *}
 
 theory Relation
-imports Product_Type
+imports Datatype Finite_Set
 begin
 
 subsection {* Definitions *}
@@ -379,6 +378,12 @@
 lemma fst_eq_Domain: "fst ` R = Domain R";
 by (auto intro!:image_eqI)
 
+lemma Domain_dprod [simp]: "Domain (dprod r s) = uprod (Domain r) (Domain s)"
+by auto
+
+lemma Domain_dsum [simp]: "Domain (dsum r s) = usum (Domain r) (Domain s)"
+by auto
+
 
 subsection {* Range *}
 
@@ -566,6 +571,31 @@
   done
 
 
+subsection {* Finiteness *}
+
+lemma finite_converse [iff]: "finite (r^-1) = finite r"
+  apply (subgoal_tac "r^-1 = (%(x,y). (y,x))`r")
+   apply simp
+   apply (rule iffI)
+    apply (erule finite_imageD [unfolded inj_on_def])
+    apply (simp split add: split_split)
+   apply (erule finite_imageI)
+  apply (simp add: converse_def image_def, auto)
+  apply (rule bexI)
+   prefer 2 apply assumption
+  apply simp
+  done
+
+text {* \paragraph{Finiteness of transitive closure} (Thanks to Sidi
+Ehmety) *}
+
+lemma finite_Field: "finite r ==> finite (Field r)"
+  -- {* A finite relation has a finite field (@{text "= domain \<union> range"}. *}
+  apply (induct set: finite)
+   apply (auto simp add: Field_def Domain_insert Range_insert)
+  done
+
+
 subsection {* Version of @{text lfp_induct} for binary relations *}
 
 lemmas lfp_induct2 = 
--- a/src/HOL/Relation_Power.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Relation_Power.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Relation_Power.thy
-    ID:         $Id$
     Author:     Tobias Nipkow
     Copyright   1996  TU Muenchen
 *)
@@ -7,7 +6,7 @@
 header{*Powers of Relations and Functions*}
 
 theory Relation_Power
-imports Power Transitive_Closure
+imports Power Transitive_Closure Plain
 begin
 
 instance
--- a/src/HOL/Ring_and_Field.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Ring_and_Field.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -24,14 +24,14 @@
 *}
 
 class semiring = ab_semigroup_add + semigroup_mult +
-  assumes left_distrib: "(a + b) * c = a * c + b * c"
-  assumes right_distrib: "a * (b + c) = a * b + a * c"
+  assumes left_distrib[algebra_simps]: "(a + b) * c = a * c + b * c"
+  assumes right_distrib[algebra_simps]: "a * (b + c) = a * b + a * c"
 begin
 
 text{*For the @{text combine_numerals} simproc*}
 lemma combine_common_factor:
   "a * e + (b * e + c) = (a + b) * e + c"
-  by (simp add: left_distrib add_ac)
+by (simp add: left_distrib add_ac)
 
 end
 
@@ -47,16 +47,12 @@
 subclass semiring_0
 proof
   fix a :: 'a
-  have "0 * a + 0 * a = 0 * a + 0"
-    by (simp add: left_distrib [symmetric])
-  thus "0 * a = 0"
-    by (simp only: add_left_cancel)
+  have "0 * a + 0 * a = 0 * a + 0" by (simp add: left_distrib [symmetric])
+  thus "0 * a = 0" by (simp only: add_left_cancel)
 next
   fix a :: 'a
-  have "a * 0 + a * 0 = a * 0 + 0"
-    by (simp add: right_distrib [symmetric])
-  thus "a * 0 = 0"
-    by (simp only: add_left_cancel)
+  have "a * 0 + a * 0 = a * 0 + 0" by (simp add: right_distrib [symmetric])
+  thus "a * 0 = 0" by (simp only: add_left_cancel)
 qed
 
 end
@@ -98,7 +94,7 @@
 begin
 
 lemma one_neq_zero [simp]: "1 \<noteq> 0"
-  by (rule not_sym) (rule zero_neq_one)
+by (rule not_sym) (rule zero_neq_one)
 
 end
 
@@ -142,7 +138,7 @@
 qed
 
 lemma dvd_0_left_iff [noatp, simp]: "0 dvd a \<longleftrightarrow> a = 0"
-  by (auto intro: dvd_refl elim!: dvdE)
+by (auto intro: dvd_refl elim!: dvdE)
 
 lemma dvd_0_right [iff]: "a dvd 0"
 proof
@@ -150,10 +146,10 @@
 qed
 
 lemma one_dvd [simp]: "1 dvd a"
-  by (auto intro!: dvdI)
+by (auto intro!: dvdI)
 
 lemma dvd_mult: "a dvd c \<Longrightarrow> a dvd (b * c)"
-  by (auto intro!: mult_left_commute dvdI elim!: dvdE)
+by (auto intro!: mult_left_commute dvdI elim!: dvdE)
 
 lemma dvd_mult2: "a dvd b \<Longrightarrow> a dvd (b * c)"
   apply (subst mult_commute)
@@ -161,10 +157,10 @@
   done
 
 lemma dvd_triv_right [simp]: "a dvd b * a"
-  by (rule dvd_mult) (rule dvd_refl)
+by (rule dvd_mult) (rule dvd_refl)
 
 lemma dvd_triv_left [simp]: "a dvd a * b"
-  by (rule dvd_mult2) (rule dvd_refl)
+by (rule dvd_mult2) (rule dvd_refl)
 
 lemma mult_dvd_mono:
   assumes ab: "a dvd b"
@@ -178,13 +174,13 @@
 qed
 
 lemma dvd_mult_left: "a * b dvd c \<Longrightarrow> a dvd c"
-  by (simp add: dvd_def mult_assoc, blast)
+by (simp add: dvd_def mult_assoc, blast)
 
 lemma dvd_mult_right: "a * b dvd c \<Longrightarrow> b dvd c"
   unfolding mult_ac [of a] by (rule dvd_mult_left)
 
 lemma dvd_0_left: "0 dvd a \<Longrightarrow> a = 0"
-  by simp
+by simp
 
 lemma dvd_add:
   assumes ab: "a dvd b"
@@ -230,43 +226,40 @@
 text {* Distribution rules *}
 
 lemma minus_mult_left: "- (a * b) = - a * b"
-  by (rule equals_zero_I) (simp add: left_distrib [symmetric]) 
+by (rule equals_zero_I) (simp add: left_distrib [symmetric]) 
 
 lemma minus_mult_right: "- (a * b) = a * - b"
-  by (rule equals_zero_I) (simp add: right_distrib [symmetric]) 
+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
+by simp
 
 lemma minus_mult_commute: "- a * b = a * - b"
-  by simp
-
-lemma right_diff_distrib: "a * (b - c) = a * b - a * c"
-  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)
+by simp
+
+lemma right_diff_distrib[algebra_simps]: "a * (b - c) = a * b - a * c"
+by (simp add: right_distrib diff_minus)
+
+lemma left_diff_distrib[algebra_simps]: "(a - b) * c = a * c - b * c"
+by (simp add: left_distrib diff_minus)
 
 lemmas ring_distribs =
   right_distrib left_distrib left_diff_distrib right_diff_distrib
 
-lemmas ring_simps =
-  add_ac
-  add_diff_eq diff_add_eq diff_diff_eq diff_diff_eq2
-  diff_eq_eq eq_diff_eq diff_minus [symmetric] uminus_add_conv_diff
-  ring_distribs
+text{*Legacy - use @{text algebra_simps} *}
+lemmas ring_simps = algebra_simps
 
 lemma eq_add_iff1:
   "a * e + c = b * e + d \<longleftrightarrow> (a - b) * e + c = d"
-  by (simp add: ring_simps)
+by (simp add: algebra_simps)
 
 lemma eq_add_iff2:
   "a * e + c = b * e + d \<longleftrightarrow> c = (b - a) * e + d"
-  by (simp add: ring_simps)
+by (simp add: algebra_simps)
 
 end
 
@@ -320,7 +313,7 @@
 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)
+by (simp add: diff_minus dvd_add dvd_minus_iff)
 
 end
 
@@ -341,18 +334,16 @@
   "a * c = b * c \<longleftrightarrow> c = 0 \<or> a = b"
 proof -
   have "(a * c = b * c) = ((a - b) * c = 0)"
-    by (simp add: ring_distribs right_minus_eq)
-  thus ?thesis
-    by (simp add: disj_commute right_minus_eq)
+    by (simp add: algebra_simps right_minus_eq)
+  thus ?thesis by (simp add: disj_commute right_minus_eq)
 qed
 
 lemma mult_cancel_left [simp, noatp]:
   "c * a = c * b \<longleftrightarrow> c = 0 \<or> a = b"
 proof -
   have "(c * a = c * b) = (c * (a - b) = 0)"
-    by (simp add: ring_distribs right_minus_eq)
-  thus ?thesis
-    by (simp add: right_minus_eq)
+    by (simp add: algebra_simps right_minus_eq)
+  thus ?thesis by (simp add: right_minus_eq)
 qed
 
 end
@@ -362,19 +353,19 @@
 
 lemma mult_cancel_right1 [simp]:
   "c = b * c \<longleftrightarrow> c = 0 \<or> b = 1"
-  by (insert mult_cancel_right [of 1 c b], force)
+by (insert mult_cancel_right [of 1 c b], force)
 
 lemma mult_cancel_right2 [simp]:
   "a * c = c \<longleftrightarrow> c = 0 \<or> a = 1"
-  by (insert mult_cancel_right [of a c 1], simp)
+by (insert mult_cancel_right [of a c 1], simp)
  
 lemma mult_cancel_left1 [simp]:
   "c = c * b \<longleftrightarrow> c = 0 \<or> b = 1"
-  by (insert mult_cancel_left [of c 1 b], force)
+by (insert mult_cancel_left [of c 1 b], force)
 
 lemma mult_cancel_left2 [simp]:
   "c * a = c \<longleftrightarrow> c = 0 \<or> a = 1"
-  by (insert mult_cancel_left [of c a 1], simp)
+by (insert mult_cancel_left [of c a 1], simp)
 
 end
 
@@ -397,14 +388,11 @@
   show "a * b \<noteq> 0"
   proof
     assume ab: "a * b = 0"
-    hence "0 = inverse a * (a * b) * inverse b"
-      by simp
+    hence "0 = inverse a * (a * b) * inverse b" by simp
     also have "\<dots> = (inverse a * a) * (b * inverse b)"
       by (simp only: mult_assoc)
-    also have "\<dots> = 1"
-      using a b by simp
-    finally show False
-      by simp
+    also have "\<dots> = 1" using a b by simp
+    finally show False by simp
   qed
 qed
 
@@ -437,45 +425,41 @@
 
 lemma nonzero_inverse_minus_eq:
   "a \<noteq> 0 \<Longrightarrow> inverse (- a) = - inverse a"
-  by (rule inverse_unique) simp
+by (rule inverse_unique) simp
 
 lemma nonzero_inverse_inverse_eq:
   "a \<noteq> 0 \<Longrightarrow> inverse (inverse a) = a"
-  by (rule inverse_unique) simp
+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)
+  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
+by (rule inverse_unique) simp
 
 lemma nonzero_inverse_mult_distrib: 
   assumes "a \<noteq> 0" and "b \<noteq> 0"
   shows "inverse (a * b) = inverse b * inverse a"
 proof -
-  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 (rule inverse_unique)
+  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 (rule inverse_unique)
 qed
 
 lemma division_ring_inverse_add:
   "a \<noteq> 0 \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> inverse a + inverse b = inverse a * (a + b) * inverse b"
-  by (simp add: ring_simps mult_assoc)
+by (simp add: algebra_simps)
 
 lemma division_ring_inverse_diff:
   "a \<noteq> 0 \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> inverse a - inverse b = inverse a * (b - a) * inverse b"
-  by (simp add: ring_simps mult_assoc)
+by (simp add: algebra_simps)
 
 end
 
@@ -508,19 +492,19 @@
 qed
 
 lemma nonzero_inverse_eq_divide: "a \<noteq> 0 \<Longrightarrow> inverse a = 1 / a"
-  by (simp add: divide_inverse)
+by (simp add: divide_inverse)
 
 lemma divide_self [simp]: "a \<noteq> 0 \<Longrightarrow> a / a = 1"
-  by (simp add: divide_inverse)
+by (simp add: divide_inverse)
 
 lemma divide_zero_left [simp]: "0 / a = 0"
-  by (simp add: divide_inverse)
+by (simp add: divide_inverse)
 
 lemma inverse_eq_divide: "inverse a = 1 / a"
-  by (simp add: divide_inverse)
+by (simp add: divide_inverse)
 
 lemma add_divide_distrib: "(a+b) / c = a/c + b/c"
-  by (simp add: divide_inverse ring_distribs) 
+by (simp add: divide_inverse algebra_simps) 
 
 end
 
@@ -529,11 +513,11 @@
 
 lemma divide_zero [simp]:
   "a / 0 = (0::'a::{field,division_by_zero})"
-  by (simp add: divide_inverse)
+by (simp add: divide_inverse)
 
 lemma divide_self_if [simp]:
   "a / (a::'a::{field,division_by_zero}) = (if a=0 then 0 else 1)"
-  by simp
+by simp
 
 class mult_mono = times + zero + ord +
   assumes mult_left_mono: "a \<le> b \<Longrightarrow> 0 \<le> c \<Longrightarrow> c * a \<le> c * b"
@@ -566,16 +550,16 @@
 subclass pordered_semiring ..
 
 lemma mult_nonneg_nonneg: "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> a * b"
-  by (drule mult_left_mono [of zero b], auto)
+by (drule mult_left_mono [of zero b], auto)
 
 lemma mult_nonneg_nonpos: "0 \<le> a \<Longrightarrow> b \<le> 0 \<Longrightarrow> a * b \<le> 0"
-  by (drule mult_left_mono [of b zero], auto)
+by (drule mult_left_mono [of b zero], auto)
 
 lemma mult_nonneg_nonpos2: "0 \<le> a \<Longrightarrow> b \<le> 0 \<Longrightarrow> b * a \<le> 0" 
-  by (drule mult_right_mono [of b zero], auto)
+by (drule mult_right_mono [of b zero], auto)
 
 lemma split_mult_neg_le: "(0 \<le> a & b \<le> 0) | (a \<le> 0 & 0 \<le> b) \<Longrightarrow> a * b \<le> 0" 
-  by (auto simp add: mult_nonneg_nonpos mult_nonneg_nonpos2)
+by (auto simp add: mult_nonneg_nonpos mult_nonneg_nonpos2)
 
 end
 
@@ -588,11 +572,11 @@
 
 lemma mult_left_less_imp_less:
   "c * a < c * b \<Longrightarrow> 0 \<le> c \<Longrightarrow> a < b"
-  by (force simp add: mult_left_mono not_le [symmetric])
+by (force simp add: mult_left_mono not_le [symmetric])
  
 lemma mult_right_less_imp_less:
   "a * c < b * c \<Longrightarrow> 0 \<le> c \<Longrightarrow> a < b"
-  by (force simp add: mult_right_mono not_le [symmetric])
+by (force simp add: mult_right_mono not_le [symmetric])
 
 end
 
@@ -617,23 +601,23 @@
 
 lemma mult_left_le_imp_le:
   "c * a \<le> c * b \<Longrightarrow> 0 < c \<Longrightarrow> a \<le> b"
-  by (force simp add: mult_strict_left_mono _not_less [symmetric])
+by (force simp add: mult_strict_left_mono _not_less [symmetric])
  
 lemma mult_right_le_imp_le:
   "a * c \<le> b * c \<Longrightarrow> 0 < c \<Longrightarrow> a \<le> b"
-  by (force simp add: mult_strict_right_mono not_less [symmetric])
+by (force simp add: mult_strict_right_mono not_less [symmetric])
 
 lemma mult_pos_pos:
   "0 < a \<Longrightarrow> 0 < b \<Longrightarrow> 0 < a * b"
-  by (drule mult_strict_left_mono [of zero b], auto)
+by (drule mult_strict_left_mono [of zero b], auto)
 
 lemma mult_pos_neg:
   "0 < a \<Longrightarrow> b < 0 \<Longrightarrow> a * b < 0"
-  by (drule mult_strict_left_mono [of b zero], auto)
+by (drule mult_strict_left_mono [of b zero], auto)
 
 lemma mult_pos_neg2:
   "0 < a \<Longrightarrow> b < 0 \<Longrightarrow> b * a < 0" 
-  by (drule mult_strict_right_mono [of b zero], auto)
+by (drule mult_strict_right_mono [of b zero], auto)
 
 lemma zero_less_mult_pos:
   "0 < a * b \<Longrightarrow> 0 < a \<Longrightarrow> 0 < b"
@@ -666,7 +650,7 @@
 lemma mult_strict_mono':
   assumes "a < b" and "c < d" and "0 \<le> a" and "0 \<le> c"
   shows "a * c < b * d"
-  by (rule mult_strict_mono) (insert assms, auto)
+by (rule mult_strict_mono) (insert assms, auto)
 
 lemma mult_less_le_imp_less:
   assumes "a < b" and "c \<le> d" and "0 \<le> a" and "0 < c"
@@ -697,8 +681,7 @@
   assume "\<not>  a < b"
   hence "b \<le> a" by (simp add: linorder_not_less)
   hence "c * b \<le> c * a" using nonneg by (rule mult_left_mono)
-  with this and less show False 
-    by (simp add: not_less [symmetric])
+  with this and less show False by (simp add: not_less [symmetric])
 qed
 
 lemma mult_less_imp_less_right:
@@ -708,8 +691,7 @@
   assume "\<not> a < b"
   hence "b \<le> a" by (simp add: linorder_not_less)
   hence "b * c \<le> a * c" using nonneg by (rule mult_right_mono)
-  with this and less show False 
-    by (simp add: not_less [symmetric])
+  with this and less show False by (simp add: not_less [symmetric])
 qed  
 
 end
@@ -768,23 +750,24 @@
 
 subclass pordered_ab_group_add ..
 
-lemmas ring_simps = ring_simps group_simps
+text{*Legacy - use @{text algebra_simps} *}
+lemmas ring_simps = algebra_simps
 
 lemma less_add_iff1:
   "a * e + c < b * e + d \<longleftrightarrow> (a - b) * e + c < d"
-  by (simp add: ring_simps)
+by (simp add: algebra_simps)
 
 lemma less_add_iff2:
   "a * e + c < b * e + d \<longleftrightarrow> c < (b - a) * e + d"
-  by (simp add: ring_simps)
+by (simp add: algebra_simps)
 
 lemma le_add_iff1:
   "a * e + c \<le> b * e + d \<longleftrightarrow> (a - b) * e + c \<le> d"
-  by (simp add: ring_simps)
+by (simp add: algebra_simps)
 
 lemma le_add_iff2:
   "a * e + c \<le> b * e + d \<longleftrightarrow> c \<le> (b - a) * e + d"
-  by (simp add: ring_simps)
+by (simp add: algebra_simps)
 
 lemma mult_left_mono_neg:
   "b \<le> a \<Longrightarrow> c \<le> 0 \<Longrightarrow> c * a \<le> c * b"
@@ -800,11 +783,11 @@
 
 lemma mult_nonpos_nonpos:
   "a \<le> 0 \<Longrightarrow> b \<le> 0 \<Longrightarrow> 0 \<le> a * b"
-  by (drule mult_right_mono_neg [of a zero b]) auto
+by (drule mult_right_mono_neg [of a zero b]) auto
 
 lemma split_mult_pos_le:
   "(0 \<le> a \<and> 0 \<le> b) \<or> (a \<le> 0 \<and> b \<le> 0) \<Longrightarrow> 0 \<le> a * b"
-  by (auto simp add: mult_nonneg_nonneg mult_nonpos_nonpos)
+by (auto simp add: mult_nonneg_nonneg mult_nonpos_nonpos)
 
 end
 
@@ -827,7 +810,7 @@
 proof
   fix a b
   show "\<bar>a + b\<bar> \<le> \<bar>a\<bar> + \<bar>b\<bar>"
-  by (auto simp add: abs_if not_less neg_less_eq_nonneg less_eq_neg_nonpos)
+by (auto simp add: abs_if not_less neg_less_eq_nonneg less_eq_neg_nonpos)
    (auto simp del: minus_add_distrib simp add: minus_add_distrib [symmetric]
      neg_less_eq_nonneg less_eq_neg_nonpos, auto intro: add_nonneg_nonneg,
       auto intro!: less_imp_le add_neg_neg)
@@ -858,7 +841,7 @@
 
 lemma mult_neg_neg:
   "a < 0 \<Longrightarrow> b < 0 \<Longrightarrow> 0 < a * b"
-  by (drule mult_strict_right_mono_neg, auto)
+by (drule mult_strict_right_mono_neg, auto)
 
 subclass ring_no_zero_divisors
 proof
@@ -903,7 +886,7 @@
 
 lemma zero_le_mult_iff:
   "0 \<le> a * b \<longleftrightarrow> 0 \<le> a \<and> 0 \<le> b \<or> a \<le> 0 \<and> b \<le> 0"
-  by (auto simp add: eq_commute [of 0] le_less not_less zero_less_mult_iff)
+by (auto simp add: eq_commute [of 0] le_less not_less zero_less_mult_iff)
 
 lemma mult_less_0_iff:
   "a * b < 0 \<longleftrightarrow> 0 < a \<and> b < 0 \<or> a < 0 \<and> 0 < b"
@@ -918,10 +901,10 @@
   done
 
 lemma zero_le_square [simp]: "0 \<le> a * a"
-  by (simp add: zero_le_mult_iff linear)
+by (simp add: zero_le_mult_iff linear)
 
 lemma not_square_less_zero [simp]: "\<not> (a * a < 0)"
-  by (simp add: not_less)
+by (simp add: not_less)
 
 text{*Cancellation laws for @{term "c*a < c*b"} and @{term "a*c < b*c"},
    also with the relations @{text "\<le>"} and equality.*}
@@ -968,19 +951,16 @@
 
 lemma mult_le_cancel_right:
    "a * c \<le> b * c \<longleftrightarrow> (0 < c \<longrightarrow> a \<le> b) \<and> (c < 0 \<longrightarrow> b \<le> a)"
-  by (simp add: not_less [symmetric] mult_less_cancel_right_disj)
+by (simp add: not_less [symmetric] mult_less_cancel_right_disj)
 
 lemma mult_le_cancel_left:
   "c * a \<le> c * b \<longleftrightarrow> (0 < c \<longrightarrow> a \<le> b) \<and> (c < 0 \<longrightarrow> b \<le> a)"
-  by (simp add: not_less [symmetric] mult_less_cancel_left_disj)
+by (simp add: not_less [symmetric] mult_less_cancel_left_disj)
 
 end
 
-text{*This list of rewrites simplifies ring terms by multiplying
-everything out and bringing sums and products into a canonical form
-(by ordered rewriting). As a result it decides ring equalities but
-also helps with inequalities. *}
-lemmas ring_simps = group_simps ring_distribs
+text{*Legacy - use @{text algebra_simps} *}
+lemmas ring_simps = algebra_simps
 
 
 class pordered_comm_ring = comm_ring + pordered_comm_semiring
@@ -1001,13 +981,13 @@
   using add_strict_mono [of zero a b c] by simp
 
 lemma zero_le_one [simp]: "0 \<le> 1"
-  by (rule zero_less_one [THEN less_imp_le]) 
+by (rule zero_less_one [THEN less_imp_le]) 
 
 lemma not_one_le_zero [simp]: "\<not> 1 \<le> 0"
-  by (simp add: not_le) 
+by (simp add: not_le) 
 
 lemma not_one_less_zero [simp]: "\<not> 1 < 0"
-  by (simp add: not_less) 
+by (simp add: not_less) 
 
 lemma less_1_mult:
   assumes "1 < m" and "1 < n"
@@ -1041,35 +1021,35 @@
 
 lemma mult_le_cancel_right1:
   "c \<le> b * c \<longleftrightarrow> (0 < c \<longrightarrow> 1 \<le> b) \<and> (c < 0 \<longrightarrow> b \<le> 1)"
-  by (insert mult_le_cancel_right [of 1 c b], simp)
+by (insert mult_le_cancel_right [of 1 c b], simp)
 
 lemma mult_le_cancel_right2:
   "a * c \<le> c \<longleftrightarrow> (0 < c \<longrightarrow> a \<le> 1) \<and> (c < 0 \<longrightarrow> 1 \<le> a)"
-  by (insert mult_le_cancel_right [of a c 1], simp)
+by (insert mult_le_cancel_right [of a c 1], simp)
 
 lemma mult_le_cancel_left1:
   "c \<le> c * b \<longleftrightarrow> (0 < c \<longrightarrow> 1 \<le> b) \<and> (c < 0 \<longrightarrow> b \<le> 1)"
-  by (insert mult_le_cancel_left [of c 1 b], simp)
+by (insert mult_le_cancel_left [of c 1 b], simp)
 
 lemma mult_le_cancel_left2:
   "c * a \<le> c \<longleftrightarrow> (0 < c \<longrightarrow> a \<le> 1) \<and> (c < 0 \<longrightarrow> 1 \<le> a)"
-  by (insert mult_le_cancel_left [of c a 1], simp)
+by (insert mult_le_cancel_left [of c a 1], simp)
 
 lemma mult_less_cancel_right1:
   "c < b * c \<longleftrightarrow> (0 \<le> c \<longrightarrow> 1 < b) \<and> (c \<le> 0 \<longrightarrow> b < 1)"
-  by (insert mult_less_cancel_right [of 1 c b], simp)
+by (insert mult_less_cancel_right [of 1 c b], simp)
 
 lemma mult_less_cancel_right2:
   "a * c < c \<longleftrightarrow> (0 \<le> c \<longrightarrow> a < 1) \<and> (c \<le> 0 \<longrightarrow> 1 < a)"
-  by (insert mult_less_cancel_right [of a c 1], simp)
+by (insert mult_less_cancel_right [of a c 1], simp)
 
 lemma mult_less_cancel_left1:
   "c < c * b \<longleftrightarrow> (0 \<le> c \<longrightarrow> 1 < b) \<and> (c \<le> 0 \<longrightarrow> b < 1)"
-  by (insert mult_less_cancel_left [of c 1 b], simp)
+by (insert mult_less_cancel_left [of c 1 b], simp)
 
 lemma mult_less_cancel_left2:
   "c * a < c \<longleftrightarrow> (0 \<le> c \<longrightarrow> a < 1) \<and> (c \<le> 0 \<longrightarrow> 1 < a)"
-  by (insert mult_less_cancel_left [of c a 1], simp)
+by (insert mult_less_cancel_left [of c a 1], simp)
 
 lemma sgn_sgn [simp]:
   "sgn (sgn a) = sgn a"
@@ -1089,7 +1069,10 @@
 
 lemma sgn_times:
   "sgn (a * b) = sgn a * sgn b"
-  by (auto simp add: sgn_if zero_less_mult_iff)
+by (auto simp add: sgn_if zero_less_mult_iff)
+
+lemma abs_sgn: "abs k = k * sgn k"
+  unfolding sgn_if abs_if by auto
 
 end
 
@@ -1150,12 +1133,10 @@
      "inverse(a*b) = inverse(a) * inverse(b::'a::{field,division_by_zero})"
   proof cases
     assume "a \<noteq> 0 & b \<noteq> 0" 
-    thus ?thesis
-      by (simp add: nonzero_inverse_mult_distrib mult_commute)
+    thus ?thesis by (simp add: nonzero_inverse_mult_distrib mult_commute)
   next
     assume "~ (a \<noteq> 0 & b \<noteq> 0)" 
-    thus ?thesis
-      by force
+    thus ?thesis by force
   qed
 
 text{*There is no slick version using division by zero.*}
@@ -1182,10 +1163,8 @@
     by (simp add: divide_inverse nonzero_inverse_mult_distrib)
   also have "... =  a * inverse b * (inverse c * c)"
     by (simp only: mult_ac)
-  also have "... =  a * inverse b"
-    by simp
-    finally show ?thesis 
-    by (simp add: divide_inverse)
+  also have "... =  a * inverse b" by simp
+    finally show ?thesis by (simp add: divide_inverse)
 qed
 
 lemma mult_divide_mult_cancel_left:
@@ -1346,14 +1325,14 @@
 
 lemma divide_eq_imp: "(c::'a::{division_by_zero,field}) ~= 0 ==>
     b = a * c ==> b / c = a"
-  by (subst divide_eq_eq, simp)
+by (subst divide_eq_eq, simp)
 
 lemma eq_divide_imp: "(c::'a::{division_by_zero,field}) ~= 0 ==>
     a * c = b ==> a = b / c"
-  by (subst eq_divide_eq, simp)
-
-
-lemmas field_eq_simps = ring_simps
+by (subst eq_divide_eq, simp)
+
+
+lemmas field_eq_simps = algebra_simps
   (* pull / out*)
   add_divide_eq_iff divide_add_eq_iff
   diff_divide_eq_iff divide_diff_eq_iff
@@ -1475,12 +1454,10 @@
 shows "inverse b < inverse (a::'a::ordered_field)"
 proof (rule ccontr)
   assume "~ inverse b < inverse a"
-  hence "inverse a \<le> inverse b"
-    by (simp add: linorder_not_less)
+  hence "inverse a \<le> inverse b" by (simp add: linorder_not_less)
   hence "~ (a < b)"
     by (simp add: linorder_not_less inverse_le_imp_le [OF _ apos])
-  thus False
-    by (rule notE [OF _ less])
+  thus False by (rule notE [OF _ less])
 qed
 
 lemma inverse_less_imp_less:
@@ -1711,9 +1688,10 @@
 
 subsection{*Field simplification*}
 
-text{* Lemmas @{text field_simps} multiply with denominators in
-in(equations) if they can be proved to be non-zero (for equations) or
-positive/negative (for inequations). *}
+text{* Lemmas @{text field_simps} multiply with denominators in in(equations)
+if they can be proved to be non-zero (for equations) or positive/negative
+(for inequations). Can be too aggressive and is therefore separate from the
+more benign @{text algebra_simps}. *}
 
 lemmas field_simps = field_eq_simps
   (* multiply ineqn *)
@@ -1981,15 +1959,15 @@
 
 lemma mult_right_le_one_le: "0 <= (x::'a::ordered_idom) ==> 0 <= y ==> y <= 1
     ==> x * y <= x"
-  by (auto simp add: mult_compare_simps);
+by (auto simp add: mult_compare_simps);
 
 lemma mult_left_le_one_le: "0 <= (x::'a::ordered_idom) ==> 0 <= y ==> y <= 1
     ==> y * x <= x"
-  by (auto simp add: mult_compare_simps);
+by (auto simp add: mult_compare_simps);
 
 lemma mult_imp_div_pos_le: "0 < (y::'a::ordered_field) ==> x <= z * y ==>
     x / y <= z";
-  by (subst pos_divide_le_eq, assumption+);
+by (subst pos_divide_le_eq, assumption+);
 
 lemma mult_imp_le_div_pos: "0 < (y::'a::ordered_field) ==> z * y <= x ==>
     z <= x / y"
@@ -2054,7 +2032,7 @@
 qed
 
 lemma zero_less_two: "0 < 1 + 1"
-  by (blast intro: less_trans zero_less_one less_add_one)
+by (blast intro: less_trans zero_less_one less_add_one)
 
 end
 
@@ -2086,7 +2064,7 @@
 end
 
 lemma abs_one [simp]: "abs 1 = (1::'a::ordered_idom)"
-  by (simp add: abs_if zero_less_one [THEN order_less_not_sym])
+by (simp add: abs_if zero_less_one [THEN order_less_not_sym])
 
 class pordered_ring_abs = pordered_ring + pordered_ab_group_add_abs +
   assumes abs_eq_mult:
@@ -2106,14 +2084,14 @@
   let ?x = "pprt a * pprt b - pprt a * nprt b - nprt a * pprt b + nprt a * nprt b"
   let ?y = "pprt a * pprt b + pprt a * nprt b + nprt a * pprt b + nprt a * nprt b"
   have a: "(abs a) * (abs b) = ?x"
-    by (simp only: abs_prts[of a] abs_prts[of b] ring_simps)
+    by (simp only: abs_prts[of a] abs_prts[of b] algebra_simps)
   {
     fix u v :: 'a
     have bh: "\<lbrakk>u = a; v = b\<rbrakk> \<Longrightarrow> 
               u * v = pprt a * pprt b + pprt a * nprt b + 
                       nprt a * pprt b + nprt a * nprt b"
       apply (subst prts[of u], subst prts[of v])
-      apply (simp add: ring_simps) 
+      apply (simp add: algebra_simps) 
       done
   }
   note b = this[OF refl[of a] refl[of b]]
@@ -2166,7 +2144,7 @@
       apply (simp_all add: mulprts abs_prts)
       apply (insert prems)
       apply (auto simp add: 
-	ring_simps 
+	algebra_simps 
 	iffD1[OF zero_le_iff_zero_nprt] iffD1[OF le_zero_iff_zero_pprt]
 	iffD1[OF le_zero_iff_pprt_id] iffD1[OF zero_le_iff_nprt_id])
 	apply(drule (1) mult_nonneg_nonpos[of a b], simp)
@@ -2178,7 +2156,7 @@
     then show ?thesis
       apply (simp_all add: mulprts abs_prts)
       apply (insert prems)
-      apply (auto simp add: ring_simps)
+      apply (auto simp add: algebra_simps)
       apply(drule (1) mult_nonneg_nonneg[of a b],simp)
       apply(drule (1) mult_nonpos_nonpos[of a b],simp)
       done
@@ -2191,10 +2169,10 @@
   equal_neg_zero neg_equal_zero mult_less_0_iff)
 
 lemma abs_mult: "abs (a * b) = abs a * abs (b::'a::ordered_idom)" 
-  by (simp add: abs_eq_mult linorder_linear)
+by (simp add: abs_eq_mult linorder_linear)
 
 lemma abs_mult_self: "abs a * abs a = a * (a::'a::ordered_idom)"
-  by (simp add: abs_if) 
+by (simp add: abs_if) 
 
 lemma nonzero_abs_inverse:
      "a \<noteq> 0 ==> abs (inverse (a::'a::ordered_field)) = inverse (abs a)"
@@ -2268,7 +2246,7 @@
     apply simp
     done
   then have "a * b = pprt a * pprt b + pprt a * nprt b + nprt a * pprt b + nprt a * nprt b"
-    by (simp add: ring_simps)
+    by (simp add: algebra_simps)
   moreover have "pprt a * pprt b <= pprt a2 * pprt b2"
     by (simp_all add: prems mult_mono)
   moreover have "pprt a * nprt b <= pprt a1 * nprt b2"
--- a/src/HOL/SEQ.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/SEQ.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -388,7 +388,7 @@
 lemma inverse_diff_inverse:
   "\<lbrakk>(a::'a::division_ring) \<noteq> 0; b \<noteq> 0\<rbrakk>
    \<Longrightarrow> inverse a - inverse b = - (inverse a * (a - b) * inverse b)"
-by (simp add: ring_simps)
+by (simp add: algebra_simps)
 
 lemma Bseq_inverse_lemma:
   fixes x :: "'a::real_normed_div_algebra"
--- a/src/HOL/SetInterval.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/SetInterval.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -532,15 +532,15 @@
   done
 
 lemma card_atLeastAtMost_int [simp]: "card {l..u} = nat (u - l + 1)"
-  apply (subst atLeastLessThanPlusOne_atLeastAtMost_int [THEN sym])
-  apply (auto simp add: compare_rls)
-  done
+apply (subst atLeastLessThanPlusOne_atLeastAtMost_int [THEN sym])
+apply (auto simp add: algebra_simps)
+done
 
 lemma card_greaterThanAtMost_int [simp]: "card {l<..u} = nat (u - l)"
-  by (subst atLeastPlusOneAtMost_greaterThanAtMost_int [THEN sym], simp)
+by (subst atLeastPlusOneAtMost_greaterThanAtMost_int [THEN sym], simp)
 
 lemma card_greaterThanLessThan_int [simp]: "card {l<..<u} = nat (u - (l + 1))"
-  by (subst atLeastPlusOneLessThan_greaterThanLessThan_int [THEN sym], simp)
+by (subst atLeastPlusOneLessThan_greaterThanLessThan_int [THEN sym], simp)
 
 lemma finite_M_bounded_by_nat: "finite {k. P k \<and> k < (i::nat)}"
 proof -
@@ -806,7 +806,7 @@
 lemma setsum_head_upt_Suc:
   "m < n \<Longrightarrow> setsum f {m..<n} = f m + setsum f {Suc m..<n}"
 apply(insert setsum_head_Suc[of m "n - 1" f])
-apply (simp add: atLeastLessThanSuc_atLeastAtMost[symmetric] ring_simps)
+apply (simp add: atLeastLessThanSuc_atLeastAtMost[symmetric] algebra_simps)
 done
 
 
@@ -870,7 +870,7 @@
   show ?case by simp
 next
   case (Suc n)
-  then show ?case by (simp add: ring_simps)
+  then show ?case by (simp add: algebra_simps)
 qed
 
 theorem arith_series_general:
@@ -894,11 +894,11 @@
   have "(1+1)*?n*a + d*(1+1)*(\<Sum>i\<in>{1..n - 1}. ?I i) = ((1+1)*?n*a + d*?I (n - 1)*?I n)"
     by (simp only: mult_ac gauss_sum [of "n - 1"])
        (simp add:  mult_ac trans [OF add_commute of_nat_Suc [symmetric]])
-  finally show ?thesis by (simp add: mult_ac add_ac right_distrib)
+  finally show ?thesis by (simp add: algebra_simps)
 next
   assume "\<not>(n > 1)"
   hence "n = 1 \<or> n = 0" by auto
-  thus ?thesis by (auto simp: mult_ac right_distrib)
+  thus ?thesis by (auto simp: algebra_simps)
 qed
 
 lemma arith_series_nat:
--- a/src/HOL/SizeChange/Kleene_Algebras.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/SizeChange/Kleene_Algebras.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -11,7 +11,7 @@
 
 text {* A type class of kleene algebras *}
 
-class star = type +
+class star =
   fixes star :: "'a \<Rightarrow> 'a"
 
 class idem_add = ab_semigroup_add +
--- a/src/HOL/Statespace/StateSpaceEx.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Statespace/StateSpaceEx.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -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/Tools/ComputeNumeral.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Tools/ComputeNumeral.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,5 @@
 theory ComputeNumeral
-imports ComputeHOL "~~/src/HOL/Real/Float"
+imports ComputeHOL Float
 begin
 
 (* normalization of bit strings *)
@@ -55,7 +55,7 @@
 lemma multb0x: "(Int.Bit0 x) * y = Int.Bit0 (x * y)" by (rule mult_Bit0)
 lemma multxb0: "x * (Int.Bit0 y) = Int.Bit0 (x * y)" unfolding Bit0_def by simp
 lemma multb1: "(Int.Bit1 x) * (Int.Bit1 y) = Int.Bit1 (Int.Bit0 (x * y) + x + y)"
-  unfolding Bit0_def Bit1_def by (simp add: ring_simps)
+  unfolding Bit0_def Bit1_def by (simp add: algebra_simps)
 lemmas bitmul = mult_Pls mult_Min mult_Pls_right mult_Min_right multb0x multxb0 multb1
 
 lemmas bitarith = bitnorm bitiszero bitneg bitlezero biteq bitless bitle bitsucc bitpred bituminus bitadd bitmul 
@@ -151,18 +151,18 @@
   by (auto simp only: adjust_def)
 
 lemma negateSnd: "negateSnd (q, r) = (q, -r)" 
-  by (auto simp only: negateSnd_def)
+  by (simp add: negateSnd_def)
 
-lemma divAlg: "divAlg (a, b) = (if 0\<le>a then
+lemma divmod: "IntDiv.divmod a b = (if 0\<le>a then
                   if 0\<le>b then posDivAlg a b
                   else if a=0 then (0, 0)
                        else negateSnd (negDivAlg (-a) (-b))
                else 
                   if 0<b then negDivAlg a b
                   else negateSnd (posDivAlg (-a) (-b)))"
-  by (auto simp only: divAlg_def)
+  by (auto simp only: IntDiv.divmod_def)
 
-lemmas compute_div_mod = div_def mod_def divAlg adjust negateSnd posDivAlg.simps negDivAlg.simps
+lemmas compute_div_mod = div_def mod_def divmod adjust negateSnd posDivAlg.simps negDivAlg.simps
 
 
 
@@ -193,6 +193,3 @@
                          compute_natarith compute_numberarith max_def min_def compute_num_conversions zpowerarith compute_div_mod compute_even
 
 end
-
-
-
--- a/src/HOL/Tools/TFL/tfl.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Tools/TFL/tfl.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -390,7 +390,7 @@
                           (wfrec $ map_types poly_tvars R)
                       $ functional
      val def_term = mk_const_def thy (x, Ty, wfrec_R_M)
-     val ([def], thy') = PureThy.add_defs false [Thm.no_attributes (def_name, def_term)] thy
+     val ([def], thy') = PureThy.add_defs false [Thm.no_attributes (Binding.name def_name, def_term)] thy
  in (thy', def) end;
 end;
 
@@ -549,7 +549,7 @@
      val ([def0], theory) =
        thy
        |> PureThy.add_defs false
-            [Thm.no_attributes (fid ^ "_def", defn)]
+            [Thm.no_attributes (Binding.name (fid ^ "_def"), defn)]
      val def = Thm.freezeT def0;
      val dummy = if !trace then writeln ("DEF = " ^ Display.string_of_thm def)
                            else ()
--- a/src/HOL/Tools/atp_manager.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Tools/atp_manager.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -19,7 +19,7 @@
   val kill: unit -> unit
   val info: unit -> unit
   val messages: int option -> unit
-  type prover = int -> Proof.state -> bool * string
+  type prover = int -> int -> Proof.state -> bool * string
   val add_prover: string -> prover -> theory -> theory
   val print_provers: theory -> unit
   val sledgehammer: string list -> Proof.state -> unit
@@ -35,9 +35,9 @@
 
 local
 
-val atps = ref "e";
+val atps = ref "e remote_vampire";
 val max_atps = ref 5;   (* ~1 means infinite number of atps *)
-val timeout = ref 60;
+val timeout = ref 100;
 
 in
 
@@ -89,13 +89,14 @@
   oldest_heap: ThreadHeap.T,
   active: (Thread.thread * (Time.time * Time.time * string)) list,
   cancelling: (Thread.thread * (Time.time * Time.time * string)) list,
-  messages: string list};
+  messages: string list,
+  store: string list};
 
-fun make_state timeout_heap oldest_heap active cancelling messages =
+fun make_state timeout_heap oldest_heap active cancelling messages store =
   State {timeout_heap = timeout_heap, oldest_heap = oldest_heap,
-    active = active, cancelling = cancelling, messages = messages};
+    active = active, cancelling = cancelling, messages = messages, store = store};
 
-val state = Synchronized.var "atp_manager" (make_state ThreadHeap.empty ThreadHeap.empty [] [] []);
+val state = Synchronized.var "atp_manager" (make_state ThreadHeap.empty ThreadHeap.empty [] [] [] []);
 
 
 (* the managing thread *)
@@ -106,29 +107,27 @@
 
 (* unregister thread *)
 
-fun unregister (success, message) thread = Synchronized.change_result state
-  (fn state as State {timeout_heap, oldest_heap, active, cancelling, messages} =>
+fun unregister (success, message) thread = Synchronized.change state
+  (fn state as State {timeout_heap, oldest_heap, active, cancelling, messages, store} =>
     (case lookup_thread active thread of
       SOME (birthtime, _, description) =>
         let
           val (group, active') =
             if success then List.partition (fn (_, (tb, _, _)) => tb = birthtime) active
             else List.partition (fn (th, _) => Thread.equal (th, thread)) active
-          val others = delete_thread thread group
 
           val now = Time.now ()
           val cancelling' =
-            fold (fn (th, (tb, _, desc)) => update_thread (th, (tb, now, desc))) others cancelling
+            fold (fn (th, (tb, _, desc)) => update_thread (th, (tb, now, desc))) group cancelling
 
-          val msg = description ^ "\n" ^ message
-          val message' = "Sledgehammer: " ^ msg ^
-            (if null others then ""
-             else "\nInterrupted " ^ string_of_int (length others) ^ " other group members")
-          val messages' = msg ::
-            (if length messages <= message_store_limit then messages
-             else #1 (chop message_store_limit messages))
-        in (message', make_state timeout_heap oldest_heap active' cancelling' messages') end
-    | NONE => ("", state)));
+          val message' = description ^ "\n" ^ message ^
+            (if length group <= 1 then ""
+             else "\nInterrupted " ^ string_of_int (length group - 1) ^ " other group members")
+          val store' = message' ::
+            (if length store <= message_store_limit then store
+             else #1 (chop message_store_limit store))
+        in make_state timeout_heap oldest_heap active' cancelling' (message' :: messages) store' end
+    | NONE =>state));
 
 
 (* kill excessive atp threads *)
@@ -142,13 +141,13 @@
 fun kill_oldest () =
   let exception Unchanged in
     Synchronized.change_result state
-      (fn State {timeout_heap, oldest_heap, active, cancelling, messages} =>
+      (fn State {timeout_heap, oldest_heap, active, cancelling, messages, store} =>
         if ThreadHeap.is_empty oldest_heap orelse not (excessive_atps active)
         then raise Unchanged
         else
           let val ((_, oldest_thread), oldest_heap') = ThreadHeap.min_elem oldest_heap
-          in (oldest_thread, make_state timeout_heap oldest_heap' active cancelling messages) end)
-      |> (priority o unregister (false, "Interrupted (maximum number of ATPs exceeded)"))
+          in (oldest_thread, make_state timeout_heap oldest_heap' active cancelling messages store) end)
+      |> unregister (false, "Interrupted (maximum number of ATPs exceeded)")
     handle Unchanged => ()
   end;
 
@@ -160,6 +159,13 @@
 
 end;
 
+fun print_new_messages () =
+  let val to_print = Synchronized.change_result state
+    (fn State {timeout_heap, oldest_heap, active, cancelling, messages, store} =>
+    (messages, make_state timeout_heap oldest_heap active cancelling [] store))  
+  in if null to_print then ()
+  else priority ("Sledgehammer: " ^ (space_implode "\n\n" to_print)) end;
+
 
 (* start a watching thread which runs forever -- only one may exist *)
 
@@ -176,8 +182,8 @@
           NONE => SOME (Time.+ (Time.now (), max_wait_time))
         | SOME (time, _) => SOME time)
 
-      (* action: cancel find threads whose timeout is reached, and interrupt cancelling threads *)
-      fun action (State {timeout_heap, oldest_heap, active, cancelling, messages}) =
+      (* action: find threads whose timeout is reached, and interrupt cancelling threads *)
+      fun action (State {timeout_heap, oldest_heap, active, cancelling, messages, store}) =
         let val (timeout_threads, timeout_heap') =
           ThreadHeap.upto (Time.now (), Thread.self ()) timeout_heap
         in
@@ -187,15 +193,16 @@
             let
               val _ = List.app (SimpleThread.interrupt o #1) cancelling
               val cancelling' = filter (Thread.isActive o #1) cancelling
-              val state' = make_state timeout_heap' oldest_heap active cancelling' messages
+              val state' = make_state timeout_heap' oldest_heap active cancelling' messages store
             in SOME (map #2 timeout_threads, state') end
         end
     in
       while true do
        (Synchronized.timed_access state time_limit action
         |> these
-        |> List.app (priority o unregister (false, "Interrupted (reached timeout)"));
+        |> List.app (unregister (false, "Interrupted (reached timeout)"));
         kill_excessive ();
+        print_new_messages ();
         (*give threads time to respond to interrupt*)
         OS.Process.sleep min_wait_time)
     end)));
@@ -206,12 +213,12 @@
 fun register birthtime deadtime (thread, desc) =
  (check_thread_manager ();
   Synchronized.change state
-    (fn State {timeout_heap, oldest_heap, active, cancelling, messages} =>
+    (fn State {timeout_heap, oldest_heap, active, cancelling, messages, store} =>
       let
         val timeout_heap' = ThreadHeap.insert (deadtime, thread) timeout_heap
         val oldest_heap' = ThreadHeap.insert (birthtime, thread) oldest_heap
         val active' = update_thread (thread, (birthtime, deadtime, desc)) active
-      in make_state timeout_heap' oldest_heap' active' cancelling messages end));
+      in make_state timeout_heap' oldest_heap' active' cancelling messages store end));
 
 
 
@@ -220,9 +227,9 @@
 (* kill: move all threads to cancelling *)
 
 fun kill () = Synchronized.change state
-  (fn State {timeout_heap, oldest_heap, active, cancelling, messages} =>
+  (fn State {timeout_heap, oldest_heap, active, cancelling, messages, store} =>
     let val formerly_active = map (fn (th, (tb, _, desc)) => (th, (tb, Time.now (), desc))) active
-    in make_state timeout_heap oldest_heap [] (formerly_active @ cancelling) messages end);
+    in make_state timeout_heap oldest_heap [] (formerly_active @ cancelling) messages store end);
 
 
 (* ATP info *)
@@ -253,7 +260,7 @@
 fun messages opt_limit =
   let
     val limit = the_default message_display_limit opt_limit;
-    val State {messages = msgs, ...} = Synchronized.value state
+    val State {store = msgs, ...} = Synchronized.value state
     val header = "Recent ATP messages" ^
       (if length msgs <= limit then ":" else " (" ^ string_of_int limit ^ " displayed):");
   in writeln (space_implode "\n\n" (header :: #1 (chop limit msgs))) end;
@@ -264,7 +271,7 @@
 
 (* named provers *)
 
-type prover = int -> Proof.state -> bool * string;
+type prover = int -> int -> Proof.state -> bool * string;
 
 fun err_dup_prover name = error ("Duplicate prover: " ^ quote name);
 
@@ -300,12 +307,12 @@
         val _ = SimpleThread.fork true (fn () =>
           let
             val _ = register birthtime deadtime (Thread.self (), desc)
-            val result = prover i proof_state
+            val result = prover (get_timeout ()) i proof_state
               handle ResHolClause.TOO_TRIVIAL
                 => (true, "Empty clause: Try this command: " ^ Markup.markup Markup.sendback "apply metis")
               | ERROR msg
                 => (false, "Error: " ^ msg)
-            val _ = priority (unregister result (Thread.self ()))
+            val _ = unregister result (Thread.self ())
           in () end handle Interrupt => ())
       in () end);
 
--- a/src/HOL/Tools/atp_wrapper.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Tools/atp_wrapper.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -12,7 +12,7 @@
   val external_prover:
     ((int -> Path.T) -> Proof.context * thm list * thm -> string list * ResHolClause.axiom_name Vector.vector list) ->
     Path.T * string ->
-    (string * int -> bool) ->
+    (string -> string option) ->
     (string * string vector * Proof.context * thm * int -> string) ->
     AtpManager.prover
   val tptp_prover_opts_full: int -> bool -> bool -> Path.T * string -> AtpManager.prover
@@ -30,8 +30,8 @@
   val eprover_full: AtpManager.prover
   val spass_opts: int -> bool  -> AtpManager.prover
   val spass: AtpManager.prover
-  val remote_prover_opts: int -> bool -> string -> string -> AtpManager.prover
-  val remote_prover: string -> string -> AtpManager.prover
+  val remote_prover_opts: int -> bool -> string -> AtpManager.prover
+  val remote_prover: string -> AtpManager.prover
 end;
 
 structure AtpWrapper: ATP_WRAPPER =
@@ -47,7 +47,7 @@
 
 (* basic template *)
 
-fun external_prover write_problem_files (cmd, args) check_success produce_answer subgoalno state =
+fun external_prover write_problem_files (cmd, args) find_failure produce_answer timeout subgoalno state =
   let
     (* path to unique problem file *)
     val destdir' = ! destdir
@@ -70,19 +70,18 @@
       if File.exists cmd then File.shell_path cmd ^ " " ^ args
       else error ("Bad executable: " ^ Path.implode cmd)
     val (proof, rc) = system_out (cmdline ^ " " ^ nth filenames (subgoalno - 1))
-    val _ =
-      if rc <> 0 then
-        warning ("Sledgehammer prover exited with return code " ^ string_of_int rc ^ "\n" ^ cmdline)
-      else ()
 
     (* remove *temporary* files *)
     val _ = if destdir' = "" then List.app OS.FileSys.remove filenames else ()
-
-    val success = check_success (proof, rc)
+    
+    (* check for success and print out some information on failure *)
+    val failure = find_failure proof
+    val success = rc = 0 andalso is_none failure
     val message =
-      if success
-      then "Try this command: " ^ produce_answer (proof, thm_names, ctxt, goal, subgoalno)
-      else "Failed."
+      if isSome failure then "Could not prove: " ^ the failure
+      else if rc <> 0
+      then "Exited with return code " ^ string_of_int rc ^ ": " ^ proof 
+      else "Try this command: " ^ produce_answer (proof, thm_names, ctxt, goal, subgoalno)
   in (success, message) end;
 
 
@@ -95,7 +94,7 @@
   external_prover
     (ResAtp.write_problem_files ResHolClause.tptp_write_file max_new theory_const)
     command
-    ResReconstruct.check_success_e_vamp_spass
+    ResReconstruct.find_failure_e_vamp_spass
     (if full then ResReconstruct.structured_proof else ResReconstruct.lemma_list_tstp);
 
 (*arbitrary ATP with TPTP input/output and problemfile as last argument*)
@@ -115,15 +114,19 @@
 
 (*NB: Vampire does not work without explicit timelimit*)
 
-fun vampire_opts max_new theory_const = tptp_prover_opts
+fun vampire_opts max_new theory_const timeout = tptp_prover_opts
   max_new theory_const
-  (Path.explode "$VAMPIRE_HOME/vampire", "--output_syntax tptp --mode casc -t 3600");
+  (Path.explode "$VAMPIRE_HOME/vampire",
+               ("--output_syntax tptp --mode casc -t " ^ string_of_int timeout))
+  timeout;
 
 val vampire = vampire_opts 60 false;
 
-fun vampire_opts_full max_new theory_const = full_prover_opts
+fun vampire_opts_full max_new theory_const timeout = full_prover_opts
   max_new theory_const
-  (Path.explode "$VAMPIRE_HOME/vampire", "--output_syntax tptp --mode casc -t 3600");
+  (Path.explode "$VAMPIRE_HOME/vampire",
+               ("--output_syntax tptp --mode casc -t " ^ string_of_int timeout))
+  timeout;
 
 val vampire_full = vampire_opts 60 false;
 
@@ -148,7 +151,7 @@
 fun spass_opts max_new theory_const = external_prover
   (ResAtp.write_problem_files ResHolClause.dfg_write_file max_new theory_const)
   (Path.explode "$SPASS_HOME/SPASS", "-Auto -SOS=1 -PGiven=0 -PProblem=0 -Splits=0 -FullRed=0 -DocProof")
-  ResReconstruct.check_success_e_vamp_spass
+  ResReconstruct.find_failure_e_vamp_spass
   ResReconstruct.lemma_list_dfg;
 
 val spass = spass_opts 40 true;
@@ -156,9 +159,10 @@
 
 (* remote prover invocation via SystemOnTPTP *)
 
-fun remote_prover_opts max_new theory_const name command =
+fun remote_prover_opts max_new theory_const args timeout =
   tptp_prover_opts max_new theory_const
-  (Path.explode "$ISABELLE_HOME/contrib/SystemOnTPTP/remote", name ^ " " ^ command);
+  (Path.explode "$ISABELLE_HOME/contrib/SystemOnTPTP/remote", args ^ " -t " ^ string_of_int timeout)
+  timeout;
 
 val remote_prover = remote_prover_opts 60 false;
 
--- a/src/HOL/Tools/datatype_abs_proofs.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Tools/datatype_abs_proofs.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -238,7 +238,7 @@
           (Sign.base_name name, reccomb_fn_Ts @ [T] ---> T', NoSyn))
           (reccomb_names ~~ recTs ~~ rec_result_Ts))
       |> (PureThy.add_defs false o map Thm.no_attributes) (map (fn ((((name, comb), set), T), T') =>
-          ((Sign.base_name name) ^ "_def", Logic.mk_equals (comb, absfree ("x", T,
+          (Binding.name (Sign.base_name name ^ "_def"), Logic.mk_equals (comb, absfree ("x", T,
            Const ("The", (T' --> HOLogic.boolT) --> T') $ absfree ("y", T',
              set $ Free ("x", T) $ Free ("y", T'))))))
                (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts))
@@ -262,7 +262,7 @@
   in
     thy2
     |> Sign.add_path (space_implode "_" new_type_names)
-    |> PureThy.add_thmss [(("recs", rec_thms), [])]
+    |> PureThy.add_thmss [((Binding.name "recs", rec_thms), [])]
     ||> Sign.parent_path
     ||> Theory.checkpoint
     |-> (fn thms => pair (reccomb_names, Library.flat thms))
@@ -316,7 +316,7 @@
             fns2 @ (List.concat (Library.drop (i + 1, case_dummy_fns)));
           val reccomb = Const (recname, (map fastype_of fns) @ [T] ---> T');
           val decl = ((Binding.name (Sign.base_name name), caseT), NoSyn);
-          val def = ((Sign.base_name name) ^ "_def",
+          val def = (Binding.name (Sign.base_name name ^ "_def"),
             Logic.mk_equals (list_comb (Const (name, caseT), fns1),
               list_comb (reccomb, (List.concat (Library.take (i, case_dummy_fns))) @
                 fns2 @ (List.concat (Library.drop (i + 1, case_dummy_fns))) )));
--- a/src/HOL/Tools/datatype_aux.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Tools/datatype_aux.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -76,7 +76,7 @@
 fun store_thmss_atts label tnames attss thmss =
   fold_map (fn ((tname, atts), thms) =>
     Sign.add_path tname
-    #> PureThy.add_thmss [((label, thms), atts)]
+    #> PureThy.add_thmss [((Binding.name label, thms), atts)]
     #-> (fn thm::_ => Sign.parent_path #> pair thm)) (tnames ~~ attss ~~ thmss)
   ##> Theory.checkpoint;
 
@@ -85,7 +85,7 @@
 fun store_thms_atts label tnames attss thmss =
   fold_map (fn ((tname, atts), thms) =>
     Sign.add_path tname
-    #> PureThy.add_thms [((label, thms), atts)]
+    #> PureThy.add_thms [((Binding.name label, thms), atts)]
     #-> (fn thm::_ => Sign.parent_path #> pair thm)) (tnames ~~ attss ~~ thmss)
   ##> Theory.checkpoint;
 
--- a/src/HOL/Tools/datatype_case.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Tools/datatype_case.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -419,21 +419,21 @@
 
 (* destruct nested patterns *)
 
-fun strip_case' dest (pat, rhs) =
+fun strip_case'' dest (pat, rhs) =
   case dest (Term.add_free_names pat []) rhs of
     SOME (exp as Free _, clauses) =>
       if member op aconv (OldTerm.term_frees pat) exp andalso
         not (exists (fn (_, rhs') =>
           member op aconv (OldTerm.term_frees rhs') exp) clauses)
       then
-        maps (strip_case' dest) (map (fn (pat', rhs') =>
+        maps (strip_case'' dest) (map (fn (pat', rhs') =>
           (subst_free [(exp, pat')] pat, rhs')) clauses)
       else [(pat, rhs)]
   | _ => [(pat, rhs)];
 
 fun gen_strip_case dest t = case dest [] t of
     SOME (x, clauses) =>
-      SOME (x, maps (strip_case' dest) clauses)
+      SOME (x, maps (strip_case'' dest) clauses)
   | NONE => NONE;
 
 val strip_case = gen_strip_case oo dest_case;
--- a/src/HOL/Tools/datatype_package.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Tools/datatype_package.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -196,13 +196,13 @@
 
 fun add_rules simps case_thms rec_thms inject distinct
                   weak_case_congs cong_att =
-  PureThy.add_thmss [(("simps", simps), []),
-    (("", flat case_thms @
+  PureThy.add_thmss [((Binding.name "simps", simps), []),
+    ((Binding.empty, flat case_thms @
           flat distinct @ rec_thms), [Simplifier.simp_add]),
-    (("", rec_thms), [Code.add_default_eqn_attribute]),
-    (("", flat inject), [iff_add]),
-    (("", map (fn th => th RS notE) (flat distinct)), [Classical.safe_elim NONE]),
-    (("", weak_case_congs), [cong_att])]
+    ((Binding.empty, rec_thms), [Code.add_default_eqn_attribute]),
+    ((Binding.empty, flat inject), [iff_add]),
+    ((Binding.empty, map (fn th => th RS notE) (flat distinct)), [Classical.safe_elim NONE]),
+    ((Binding.empty, weak_case_congs), [cong_att])]
   #> snd;
 
 
@@ -213,15 +213,15 @@
     val inducts = ProjectRule.projections (ProofContext.init thy) induction;
 
     fun named_rules (name, {index, exhaustion, ...}: datatype_info) =
-      [(("", nth inducts index), [Induct.induct_type name]),
-       (("", exhaustion), [Induct.cases_type name])];
+      [((Binding.empty, nth inducts index), [Induct.induct_type name]),
+       ((Binding.empty, exhaustion), [Induct.cases_type name])];
     fun unnamed_rule i =
-      (("", nth inducts i), [Thm.kind_internal, Induct.induct_type ""]);
+      ((Binding.empty, nth inducts i), [Thm.kind_internal, Induct.induct_type ""]);
   in
     thy |> PureThy.add_thms
       (maps named_rules infos @
         map unnamed_rule (length infos upto length inducts - 1)) |> snd
-    |> PureThy.add_thmss [(("inducts", inducts), [])] |> snd
+    |> PureThy.add_thmss [((Binding.name "inducts", inducts), [])] |> snd
   end;
 
 
@@ -451,7 +451,7 @@
       |> store_thmss "inject" new_type_names inject
       ||>> store_thmss "distinct" new_type_names distinct
       ||> Sign.add_path (space_implode "_" new_type_names)
-      ||>> PureThy.add_thms [(("induct", induct), [case_names_induct])];
+      ||>> PureThy.add_thms [((Binding.name "induct", induct), [case_names_induct])];
 
     val dt_infos = map (make_dt_info alt_names descr sorts induct' reccomb_names rec_thms)
       ((0 upto length descr - 1) ~~ descr ~~ case_names ~~ case_thms ~~ casedist_thms ~~
--- a/src/HOL/Tools/datatype_realizer.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Tools/datatype_realizer.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -130,7 +130,7 @@
     val vs = map (fn i => List.nth (pnames, i)) is;
     val (thm', thy') = thy
       |> Sign.absolute_path
-      |> PureThy.store_thm (space_implode "_" (ind_name :: vs @ ["correctness"]), thm)
+      |> PureThy.store_thm (Binding.name (space_implode "_" (ind_name :: vs @ ["correctness"])), thm)
       ||> Sign.restore_naming thy;
 
     val ivs = rev (Term.add_vars (Logic.varify (DatatypeProp.make_ind [descr] sorts)) []);
@@ -196,7 +196,7 @@
     val exh_name = Thm.get_name exhaustion;
     val (thm', thy') = thy
       |> Sign.absolute_path
-      |> PureThy.store_thm (exh_name ^ "_P_correctness", thm)
+      |> PureThy.store_thm (Binding.name (exh_name ^ "_P_correctness"), thm)
       ||> Sign.restore_naming thy;
 
     val P = Var (("P", 0), rT' --> HOLogic.boolT);
--- a/src/HOL/Tools/datatype_rep_proofs.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Tools/datatype_rep_proofs.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -242,7 +242,7 @@
         val ([def_thm], thy') =
           thy
           |> Sign.add_consts_i [(cname', constrT, mx)]
-          |> (PureThy.add_defs false o map Thm.no_attributes) [(def_name, def)];
+          |> (PureThy.add_defs false o map Thm.no_attributes) [(Binding.name def_name, def)];
 
       in (thy', defs @ [def_thm], eqns @ [eqn], i + 1) end;
 
@@ -343,7 +343,7 @@
         
         val (fs, eqns, isos) = Library.foldl process_dt (([], [], []), ds);
         val fTs = map fastype_of fs;
-        val defs = map (fn (rec_name, (T, iso_name)) => ((Sign.base_name iso_name) ^ "_def",
+        val defs = map (fn (rec_name, (T, iso_name)) => (Binding.name (Sign.base_name iso_name ^ "_def"),
           Logic.mk_equals (Const (iso_name, T --> Univ_elT),
             list_comb (Const (rec_name, fTs @ [T] ---> Univ_elT), fs)))) (rec_names ~~ isos);
         val (def_thms, thy') =
@@ -631,7 +631,7 @@
     val ([dt_induct'], thy7) =
       thy6
       |> Sign.add_path big_name
-      |> PureThy.add_thms [(("induct", dt_induct), [case_names_induct])]
+      |> PureThy.add_thms [((Binding.name "induct", dt_induct), [case_names_induct])]
       ||> Sign.parent_path
       ||> Theory.checkpoint;
 
--- a/src/HOL/Tools/function_package/fundef_package.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Tools/function_package/fundef_package.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -9,14 +9,14 @@
 
 signature FUNDEF_PACKAGE =
 sig
-    val add_fundef :  (Binding.T * string option * mixfix) list
+    val add_fundef :  (binding * string option * mixfix) list
                       -> (Attrib.binding * string) list 
                       -> FundefCommon.fundef_config
                       -> bool list
                       -> local_theory
                       -> Proof.state
 
-    val add_fundef_i:  (Binding.T * typ option * mixfix) list
+    val add_fundef_i:  (binding * typ option * mixfix) list
                        -> (Attrib.binding * term) list
                        -> FundefCommon.fundef_config
                        -> bool list
--- a/src/HOL/Tools/function_package/size.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Tools/function_package/size.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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') =
@@ -145,7 +144,7 @@
            (size_names ~~ recTs1))
       |> PureThy.add_defs false
         (map (Thm.no_attributes o apsnd (Logic.mk_equals o apsnd (app fs)))
-           (def_names ~~ (size_fns ~~ rec_combs1)))
+           (map Binding.name def_names ~~ (size_fns ~~ rec_combs1)))
       ||> TheoryTarget.instantiation
            (map (#1 o snd) descr', map dest_TFree paramTs, [HOLogic.class_size])
       ||>> fold_map define_overloaded
@@ -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)
@@ -209,7 +208,7 @@
       prove_size_eqs is_rec_type overloaded_size_fns (K NONE) simpset3;
 
     val ([size_thms], thy'') =  PureThy.add_thmss
-      [(("size", size_eqns),
+      [((Binding.name "size", size_eqns),
         [Simplifier.simp_add, Thm.declaration_attribute
               (fn thm => Context.mapping (Code.add_default_eqn thm) I)])] thy'
 
--- a/src/HOL/Tools/inductive_package.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Tools/inductive_package.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -38,17 +38,17 @@
     thm list list * local_theory
   type inductive_flags
   val add_inductive_i:
-    inductive_flags -> ((Binding.T * typ) * mixfix) list ->
+    inductive_flags -> ((binding * typ) * mixfix) list ->
     (string * typ) list -> (Attrib.binding * term) list -> thm list -> local_theory ->
     inductive_result * local_theory
   val add_inductive: bool -> bool ->
-    (Binding.T * string option * mixfix) list ->
-    (Binding.T * string option * mixfix) list ->
+    (binding * string option * mixfix) list ->
+    (binding * string option * mixfix) list ->
     (Attrib.binding * string) list ->
     (Facts.ref * Attrib.src list) list ->
     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 ->
+    ((binding * typ) * mixfix) list -> (string * typ) list -> (Attrib.binding * term) list ->
     thm list -> theory -> inductive_result * theory
   val arities_of: thm -> (string * int) list
   val params_of: thm -> term list
@@ -63,16 +63,16 @@
 sig
   include BASIC_INDUCTIVE_PACKAGE
   type add_ind_def
-  val declare_rules: string -> Binding.T -> bool -> bool -> string list ->
-    thm list -> Binding.T list -> Attrib.src list list -> (thm * string list) list ->
+  val declare_rules: string -> binding -> bool -> bool -> string list ->
+    thm list -> binding list -> Attrib.src list list -> (thm * string list) list ->
     thm -> local_theory -> thm list * thm list * thm * local_theory
   val add_ind_def: add_ind_def
   val gen_add_inductive_i: add_ind_def -> inductive_flags ->
-    ((Binding.T * typ) * mixfix) list -> (string * typ) list -> (Attrib.binding * term) list ->
+    ((binding * typ) * mixfix) list -> (string * typ) list -> (Attrib.binding * term) list ->
     thm list -> local_theory -> inductive_result * local_theory
   val gen_add_inductive: add_ind_def -> bool -> bool ->
-    (Binding.T * string option * mixfix) list ->
-    (Binding.T * string option * mixfix) list ->
+    (binding * string option * mixfix) list ->
+    (binding * string option * mixfix) list ->
     (Attrib.binding * string) list -> (Facts.ref * Attrib.src list) list ->
     bool -> local_theory -> inductive_result * local_theory
   val gen_ind_decl: add_ind_def -> bool ->
@@ -720,13 +720,13 @@
   in (intrs', elims', induct', ctxt3) end;
 
 type inductive_flags =
-  {quiet_mode: bool, verbose: bool, kind: string, alt_name: Binding.T,
+  {quiet_mode: bool, verbose: bool, kind: string, alt_name: binding,
    coind: bool, no_elim: bool, no_ind: bool, skip_mono: bool, fork_mono: bool}
 
 type add_ind_def =
   inductive_flags ->
   term list -> (Attrib.binding * term) list -> thm list ->
-  term list -> (Binding.T * mixfix) list ->
+  term list -> (binding * 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, fork_mono}
--- a/src/HOL/Tools/inductive_realizer.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Tools/inductive_realizer.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -391,14 +391,14 @@
            REPEAT ((resolve_tac prems THEN_ALL_NEW EVERY'
              [K (rewrite_goals_tac rews), ObjectLogic.atomize_prems_tac,
               DEPTH_SOLVE_1 o FIRST' [atac, etac allE, etac impE]]) 1)]);
-        val (thm', thy') = PureThy.store_thm (space_implode "_"
-          (NameSpace.qualified qualifier "induct" :: vs' @ Ps @ ["correctness"]), thm) thy;
+        val (thm', thy') = PureThy.store_thm (Binding.name (space_implode "_"
+          (NameSpace.qualified qualifier "induct" :: vs' @ Ps @ ["correctness"])), thm) thy;
         val thms = map (fn th => zero_var_indexes (rotate_prems ~1 (th RS mp)))
           (DatatypeAux.split_conj_thm thm');
         val ([thms'], thy'') = PureThy.add_thmss
-          [((space_implode "_"
+          [((Binding.name (space_implode "_"
              (NameSpace.qualified qualifier "inducts" :: vs' @ Ps @
-               ["correctness"]), thms), [])] thy';
+               ["correctness"])), thms), [])] thy';
         val realizers = inducts ~~ thms' ~~ rlzs ~~ rs;
       in
         Extraction.add_realizers_i
@@ -451,8 +451,8 @@
            rewrite_goals_tac rews,
            REPEAT ((resolve_tac prems THEN_ALL_NEW (ObjectLogic.atomize_prems_tac THEN'
              DEPTH_SOLVE_1 o FIRST' [atac, etac allE, etac impE])) 1)]);
-        val (thm', thy') = PureThy.store_thm (space_implode "_"
-          (name_of_thm elim :: vs @ Ps @ ["correctness"]), thm) thy
+        val (thm', thy') = PureThy.store_thm (Binding.name (space_implode "_"
+          (name_of_thm elim :: vs @ Ps @ ["correctness"])), thm) thy
       in
         Extraction.add_realizers_i
           [mk_realizer thy' (vs @ Ps) (name_of_thm elim, elim, thm', rlz, r)] thy'
--- a/src/HOL/Tools/inductive_set_package.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Tools/inductive_set_package.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -12,13 +12,13 @@
   val pred_set_conv_att: attribute
   val add_inductive_i:
     InductivePackage.inductive_flags ->
-    ((Binding.T * typ) * mixfix) list ->
+    ((binding * typ) * mixfix) list ->
     (string * typ) list ->
     (Attrib.binding * term) list -> thm list ->
     local_theory -> InductivePackage.inductive_result * local_theory
   val add_inductive: bool -> bool ->
-    (Binding.T * string option * mixfix) list ->
-    (Binding.T * string option * mixfix) list ->
+    (binding * string option * mixfix) list ->
+    (binding * string option * mixfix) list ->
     (Attrib.binding * string) list -> (Facts.ref * Attrib.src list) list ->
     bool -> local_theory -> InductivePackage.inductive_result * local_theory
   val codegen_preproc: theory -> thm list -> thm list
--- a/src/HOL/Tools/lin_arith.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Tools/lin_arith.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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/old_primrec_package.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Tools/old_primrec_package.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -305,11 +305,11 @@
   end;
 
 fun thy_note ((name, atts), thms) =
-  PureThy.add_thmss [((name, thms), atts)] #-> (fn [thms] => pair (name, thms));
+  PureThy.add_thmss [((Binding.name name, thms), atts)] #-> (fn [thms] => pair (name, thms));
 fun thy_def false ((name, atts), t) =
-      PureThy.add_defs false [((name, t), atts)] #-> (fn [thm] => pair (name, thm))
+      PureThy.add_defs false [((Binding.name name, t), atts)] #-> (fn [thm] => pair (name, thm))
   | thy_def true ((name, atts), t) =
-      PureThy.add_defs_unchecked false [((name, t), atts)] #-> (fn [thm] => pair (name, thm));
+      PureThy.add_defs_unchecked false [((Binding.name name, t), atts)] #-> (fn [thm] => pair (name, thm));
 
 in
 
--- a/src/HOL/Tools/primrec_package.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Tools/primrec_package.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -7,12 +7,12 @@
 
 signature PRIMREC_PACKAGE =
 sig
-  val add_primrec: (Binding.T * typ option * mixfix) list ->
+  val add_primrec: (binding * typ option * mixfix) list ->
     (Attrib.binding * term) list -> local_theory -> thm list * local_theory
-  val add_primrec_global: (Binding.T * typ option * mixfix) list ->
+  val add_primrec_global: (binding * typ option * mixfix) list ->
     (Attrib.binding * term) list -> theory -> thm list * theory
   val add_primrec_overloaded: (string * (string * typ) * bool) list ->
-    (Binding.T * typ option * mixfix) list ->
+    (binding * typ option * mixfix) list ->
     (Attrib.binding * term) list -> theory -> thm list * theory
 end;
 
--- a/src/HOL/Tools/recdef_package.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Tools/recdef_package.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/recdef_package.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Wrapper module for Konrad Slind's TFL package.
@@ -16,10 +15,10 @@
   val cong_del: attribute
   val wf_add: attribute
   val wf_del: attribute
-  val add_recdef: bool -> xstring -> string -> ((bstring * string) * Attrib.src list) list ->
+  val add_recdef: bool -> xstring -> string -> ((binding * string) * Attrib.src list) list ->
     Attrib.src option -> theory -> theory
       * {simps: thm list, rules: thm list list, induct: thm, tcs: term list}
-  val add_recdef_i: bool -> xstring -> term -> ((bstring * term) * attribute list) list ->
+  val add_recdef_i: bool -> xstring -> term -> ((binding * term) * attribute list) list ->
     theory -> theory * {simps: thm list, rules: thm list list, induct: thm, tcs: term list}
   val defer_recdef: xstring -> string list -> (Facts.ref * Attrib.src list) list
     -> theory -> theory * {induct_rules: thm}
@@ -214,8 +213,8 @@
       thy
       |> Sign.add_path bname
       |> PureThy.add_thmss
-        ((("simps", List.concat rules), simp_att) :: ((eq_names ~~ rules) ~~ eq_atts))
-      ||>> PureThy.add_thms [(("induct", induct), [])];
+        (((Binding.name "simps", List.concat rules), simp_att) :: ((eq_names ~~ rules) ~~ eq_atts))
+      ||>> PureThy.add_thms [((Binding.name "induct", induct), [])];
     val result = {simps = simps', rules = rules', induct = induct', tcs = tcs};
     val thy =
       thy
@@ -243,7 +242,7 @@
     val ([induct_rules'], thy3) =
       thy2
       |> Sign.add_path bname
-      |> PureThy.add_thms [(("induct_rules", induct_rules), [])]
+      |> PureThy.add_thms [((Binding.name "induct_rules", induct_rules), [])]
       ||> Sign.parent_path;
   in (thy3, {induct_rules = induct_rules'}) end;
 
@@ -299,7 +298,7 @@
 
 val recdef_decl =
   Scan.optional (P.$$$ "(" -- P.!!! (P.$$$ "permissive" -- P.$$$ ")") >> K false) true --
-  P.name -- P.term -- Scan.repeat1 ((SpecParse.opt_thm_name ":" >> apfst Binding.base_name) -- P.prop)
+  P.name -- P.term -- Scan.repeat1 (SpecParse.opt_thm_name ":" -- P.prop)
     -- Scan.option hints
   >> (fn ((((p, f), R), eqs), src) => #1 o add_recdef p f R (map P.triple_swap eqs) src);
 
--- a/src/HOL/Tools/record_package.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Tools/record_package.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1534,8 +1534,10 @@
       |> extension_typedef name repT (alphas@[zeta])
       ||> Sign.add_consts_i
             (map Syntax.no_syn ((apfst base ext_decl)::dest_decls@upd_decls))
-      ||>> PureThy.add_defs false (map Thm.no_attributes (ext_spec::dest_specs))
-      ||>> PureThy.add_defs false (map Thm.no_attributes upd_specs)
+      ||>> PureThy.add_defs false
+            (map (Thm.no_attributes o apfst Binding.name) (ext_spec :: dest_specs))
+      ||>> PureThy.add_defs false
+            (map (Thm.no_attributes o apfst Binding.name) upd_specs)
       |-> (fn args as ((_, dest_defs), upd_defs) =>
           fold Code.add_default_eqn dest_defs
           #> fold Code.add_default_eqn upd_defs
@@ -1693,14 +1695,14 @@
           [dest_convs',upd_convs']),
       thm_thy) =
       defs_thy
-      |> (PureThy.add_thms o map Thm.no_attributes)
+      |> (PureThy.add_thms o map (Thm.no_attributes o apfst Binding.name))
            [("ext_inject", inject),
             ("ext_induct", induct),
             ("ext_cases", cases),
             ("ext_surjective", surjective),
             ("ext_split", split_meta)]
-      ||>> (PureThy.add_thmss o map Thm.no_attributes)
-              [("dest_convs",dest_convs_standard),("upd_convs",upd_convs)]
+      ||>> (PureThy.add_thmss o map (Thm.no_attributes o apfst Binding.name))
+            [("dest_convs", dest_convs_standard), ("upd_convs", upd_convs)]
 
   in (thm_thy,extT,induct',inject',dest_convs',split_meta',upd_convs')
   end;
@@ -1938,9 +1940,9 @@
           (map2 (fn (x, T) => fn mx => (x, T, mx)) sel_decls (field_syntax @ [Syntax.NoSyn]))
       |> (Sign.add_consts_i o map Syntax.no_syn)
           (upd_decls @ [make_decl, fields_decl, extend_decl, truncate_decl])
-      |> ((PureThy.add_defs false o map Thm.no_attributes) sel_specs)
-      ||>> ((PureThy.add_defs false o map Thm.no_attributes) upd_specs)
-      ||>> ((PureThy.add_defs false o map Thm.no_attributes)
+      |> ((PureThy.add_defs false o map (Thm.no_attributes o apfst Binding.name)) sel_specs)
+      ||>> ((PureThy.add_defs false o map (Thm.no_attributes o apfst Binding.name)) upd_specs)
+      ||>> ((PureThy.add_defs false o map (Thm.no_attributes o apfst Binding.name))
              [make_spec, fields_spec, extend_spec, truncate_spec])
       |-> (fn defs as ((sel_defs, upd_defs), derived_defs) =>
           fold Code.add_default_eqn sel_defs
@@ -2164,17 +2166,17 @@
     val ((([sel_convs',upd_convs',sel_defs',upd_defs',[split_meta',split_object',split_ex'],derived_defs'],
             [surjective',equality']),[induct_scheme',induct',cases_scheme',cases']), thms_thy) =
       defs_thy
-      |> (PureThy.add_thmss o map Thm.no_attributes)
+      |> (PureThy.add_thmss o map (Thm.no_attributes o apfst Binding.name))
          [("select_convs", sel_convs_standard),
           ("update_convs", upd_convs),
           ("select_defs", sel_defs),
           ("update_defs", upd_defs),
           ("splits", [split_meta_standard,split_object,split_ex]),
           ("defs", derived_defs)]
-      ||>> (PureThy.add_thms o map Thm.no_attributes)
+      ||>> (PureThy.add_thms o map (Thm.no_attributes o apfst Binding.name))
           [("surjective", surjective),
            ("equality", equality)]
-      ||>> PureThy.add_thms
+      ||>> (PureThy.add_thms o (map o apfst o apfst) Binding.name)
         [(("induct_scheme", induct_scheme), induct_type_global (suffix schemeN name)),
          (("induct", induct), induct_type_global name),
          (("cases_scheme", cases_scheme), cases_type_global (suffix schemeN name)),
@@ -2186,8 +2188,8 @@
     val final_thy =
       thms_thy
       |> (snd oo PureThy.add_thmss)
-          [(("simps", sel_upd_simps), [Simplifier.simp_add]),
-           (("iffs",iffs), [iff_add])]
+          [((Binding.name "simps", sel_upd_simps), [Simplifier.simp_add]),
+           ((Binding.name "iffs", iffs), [iff_add])]
       |> put_record name (make_record_info args parent fields extension induct_scheme')
       |> put_sel_upd (names @ [full_moreN]) sel_upd_simps
       |> add_record_equalities extension_id equality'
--- a/src/HOL/Tools/res_axioms.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Tools/res_axioms.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -84,7 +84,7 @@
             val (c, thy') =
               Sign.declare_const [Markup.property_internal] ((Binding.name cname, cT), NoSyn) thy
             val cdef = cname ^ "_def"
-            val thy'' = Theory.add_defs_i true false [(cdef, Logic.mk_equals (c, rhs))] thy'
+            val thy'' = Theory.add_defs_i true false [(Binding.name cdef, Logic.mk_equals (c, rhs))] thy'
             val ax = Thm.axiom thy'' (Sign.full_bname thy'' cdef)
           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 =
--- a/src/HOL/Tools/res_reconstruct.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Tools/res_reconstruct.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -15,7 +15,7 @@
   val strip_prefix: string -> string -> string option
   val setup: Context.theory -> Context.theory
   (* extracting lemma list*)
-  val check_success_e_vamp_spass: string * int -> bool
+  val find_failure_e_vamp_spass: string -> string option
   val lemma_list_dfg: string * string vector * Proof.context * Thm.thm * int -> string
   val lemma_list_tstp: string * string vector * Proof.context * Thm.thm * int -> string
   (* structured proofs *)
@@ -463,11 +463,12 @@
   val failure_strings_vampire = ["Satisfiability detected", "Refutation not found", "CANNOT PROVE"];
   val failure_strings_SPASS = ["SPASS beiseite: Completion found.","SPASS beiseite: Ran out of time.",
   "SPASS beiseite: Maximal number of loops exceeded."];
-  fun check_success_e_vamp_spass (proof, rc) =
-    not (exists (fn s => String.isSubstring s proof)
-        (failure_strings_E @ failure_strings_vampire @ failure_strings_SPASS))
-    andalso (rc = 0);
-
+  fun find_failure_e_vamp_spass proof =
+    let val failures =
+      map_filter (fn s => if String.isSubstring s proof then SOME s else NONE) 
+         (failure_strings_E @ failure_strings_vampire @ failure_strings_SPASS)
+    in if null failures then NONE else SOME (hd failures) end;
+    
   (*=== EXTRACTING PROOF-TEXT === *)
   
   val begin_proof_strings = ["# SZS output start CNFRefutation.",
--- a/src/HOL/Tools/specification_package.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Tools/specification_package.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -28,7 +28,7 @@
                                else thname
                 val def_eq = Logic.mk_equals (Const(cname_full,ctype),
                                               HOLogic.choice_const ctype $  P)
-                val (thms, thy') = PureThy.add_defs covld [((cdefname,def_eq),[])] thy
+                val (thms, thy') = PureThy.add_defs covld [((Binding.name cdefname, def_eq),[])] thy
                 val thm' = [thm,hd thms] MRS @{thm exE_some}
             in
                 mk_definitional cos (thy',thm')
@@ -39,7 +39,7 @@
         let
             fun process [] (thy,tm) =
                 let
-                    val (thms, thy') = PureThy.add_axioms [((axname,HOLogic.mk_Trueprop tm),[])] thy
+                    val (thms, thy') = PureThy.add_axioms [((Binding.name axname, HOLogic.mk_Trueprop tm),[])] thy
                 in
                     (thy',hd thms)
                 end
@@ -184,7 +184,7 @@
                             if name = ""
                             then arg |> Library.swap
                             else (writeln ("  " ^ name ^ ": " ^ (Display.string_of_thm thm));
-                                  PureThy.store_thm (name, thm) thy)
+                                  PureThy.store_thm (Binding.name name, thm) thy)
                     in
                         args |> apsnd (remove_alls frees)
                              |> apsnd undo_imps
--- a/src/HOL/Tools/typedef_package.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Tools/typedef_package.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -112,7 +112,8 @@
       if def then
         theory
         |> Sign.add_consts_i [(name, setT', NoSyn)]
-        |> PureThy.add_defs false [Thm.no_attributes ((PrimitiveDefs.mk_defpair (setC, set)))]
+        |> PureThy.add_defs false [Thm.no_attributes (apfst (Binding.name)
+            (PrimitiveDefs.mk_defpair (setC, set)))]
         |-> (fn [th] => pair (SOME th))
       else (NONE, theory);
     fun contract_def NONE th = th
@@ -130,7 +131,7 @@
          (Abs_name, oldT --> newT, NoSyn)]
       #> add_def
       #-> (fn set_def =>
-        PureThy.add_axioms [((typedef_name, typedef_prop),
+        PureThy.add_axioms [((Binding.name typedef_name, typedef_prop),
           [Thm.rule_attribute (K (fn cond_axm => contract_def set_def inhabited RS cond_axm))])]
         ##>> pair set_def)
       ##> Theory.add_deps "" (dest_Const RepC) typedef_deps
@@ -143,7 +144,7 @@
             thy1
             |> Sign.add_path name
             |> PureThy.add_thms
-              ([((Rep_name, make @{thm type_definition.Rep}), []),
+              ((map o apfst o apfst) Binding.name [((Rep_name, make @{thm type_definition.Rep}), []),
                 ((Rep_name ^ "_inverse", make @{thm type_definition.Rep_inverse}), []),
                 ((Abs_name ^ "_inverse", make @{thm type_definition.Abs_inverse}), []),
                 ((Rep_name ^ "_inject", make @{thm type_definition.Rep_inject}), []),
--- a/src/HOL/Transcendental.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Transcendental.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -40,7 +40,7 @@
 apply (simp add: right_distrib del: setsum_op_ivl_Suc)
 apply (subst mult_left_commute [where a="x - y"])
 apply (erule subst)
-apply (simp add: power_Suc ring_simps)
+apply (simp add: power_Suc algebra_simps)
 done
 
 lemma lemma_realpow_rev_sumr:
@@ -146,8 +146,7 @@
   fixes z :: "'a :: {recpower,comm_ring}" shows
   "(\<Sum>p=0..<m. (((z + h) ^ (m - p)) * (z ^ p)) - (z ^ m)) =  
    (\<Sum>p=0..<m. (z ^ p) * (((z + h) ^ (m - p)) - (z ^ (m - p))))"
-by (auto simp add: right_distrib diff_minus power_add [symmetric] mult_ac
-  cong: strong_setsum_cong)
+by(auto simp add: algebra_simps power_add [symmetric] cong: strong_setsum_cong)
 
 lemma sumr_diff_mult_const2:
   "setsum f {0..<n} - of_nat n * (r::'a::ring_1) = (\<Sum>i = 0..<n. f i - r)"
@@ -406,7 +405,7 @@
       apply (rule summable_diff [OF B A])
       apply (rule sums_summable [OF diffs_equiv [OF C]])
       apply (rule arg_cong [where f="suminf"], rule ext)
-      apply (simp add: ring_simps)
+      apply (simp add: algebra_simps)
       done
   next
     show "(\<lambda>h. \<Sum>n. c n * (((x + h) ^ n - x ^ n) / h -
@@ -1122,7 +1121,7 @@
 apply (safe, rule lemma_DERIV_subst)
 apply (best intro!: DERIV_intros intro: DERIV_chain2) 
   --{*replaces the old @{text DERIV_tac}*}
-apply (auto simp add: diff_minus left_distrib right_distrib mult_ac add_ac)
+apply (auto simp add: algebra_simps)
 done
 
 lemma sin_cos_add [simp]:
@@ -1146,8 +1145,8 @@
 lemma lemma_DERIV_sin_cos_minus:
     "\<forall>x. DERIV (%x. (sin(-x) + (sin x)) ^ 2 + (cos(-x) - (cos x)) ^ 2) x :> 0"
 apply (safe, rule lemma_DERIV_subst)
-apply (best intro!: DERIV_intros intro: DERIV_chain2) 
-apply (auto simp add: diff_minus left_distrib right_distrib mult_ac add_ac)
+apply (best intro!: DERIV_intros intro: DERIV_chain2)
+apply (simp add: algebra_simps)
 done
 
 lemma sin_cos_minus: 
@@ -1520,9 +1519,8 @@
 apply (drule pi_gt_zero [THEN reals_Archimedean4], safe)
 apply (subgoal_tac "0 \<le> x - real n * pi & 
                     (x - real n * pi) \<le> pi & (cos (x - real n * pi) = 0) ")
-apply (auto simp add: compare_rls) 
-  prefer 3 apply (simp add: cos_diff) 
- prefer 2 apply (simp add: real_of_nat_Suc left_distrib) 
+apply (auto simp add: algebra_simps real_of_nat_Suc)
+ prefer 2 apply (simp add: cos_diff)
 apply (simp add: cos_diff)
 apply (subgoal_tac "EX! x. 0 \<le> x & x \<le> pi & cos x = 0")
 apply (rule_tac [2] cos_total, safe)
@@ -1530,7 +1528,7 @@
 apply (drule_tac x = "pi/2" in spec)
 apply (simp add: cos_diff)
 apply (rule_tac x = "Suc (2 * n)" in exI)
-apply (simp add: real_of_nat_Suc left_distrib, auto)
+apply (simp add: real_of_nat_Suc algebra_simps, auto)
 done
 
 lemma sin_zero_lemma:
@@ -1601,7 +1599,7 @@
 apply (rule_tac c1 = "cos x * cos y" in real_mult_right_cancel [THEN subst])
 apply (auto simp del: inverse_mult_distrib 
             simp add: mult_assoc left_diff_distrib cos_add)
-done  
+done
 
 lemma add_tan_eq: 
       "[| cos x \<noteq> 0; cos y \<noteq> 0 |]  
@@ -1982,7 +1980,7 @@
   also have "\<dots> = (?c * ?c - ?s * ?s) * ?c - (?s * ?c + ?c * ?s) * ?s"
     by (simp only: cos_add sin_add)
   also have "\<dots> = ?c * (?c\<twosuperior> - 3 * ?s\<twosuperior>)"
-    by (simp add: ring_simps power2_eq_square)
+    by (simp add: algebra_simps power2_eq_square)
   finally have "?c\<twosuperior> = (sqrt 3 / 2)\<twosuperior>"
     using pos_c by (simp add: sin_squared_eq power_divide)
   thus ?thesis
@@ -2051,7 +2049,7 @@
 lemma sin_cos_npi [simp]: "sin (real (Suc (2 * n)) * pi / 2) = (-1) ^ n"
 proof -
   have "sin ((real n + 1/2) * pi) = cos (real n * pi)"
-    by (auto simp add: right_distrib sin_add left_distrib mult_ac)
+    by (auto simp add: algebra_simps sin_add)
   thus ?thesis
     by (simp add: real_of_nat_Suc left_distrib add_divide_distrib 
                   mult_commute [of pi])
--- a/src/HOL/Transitive_Closure.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Transitive_Closure.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Transitive_Closure.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1992  University of Cambridge
 *)
@@ -568,6 +567,22 @@
    apply auto
   done
 
+lemma trancl_subset_Field2: "r^+ <= Field r \<times> Field r"
+  apply clarify
+  apply (erule trancl_induct)
+   apply (auto simp add: Field_def)
+  done
+
+lemma finite_trancl: "finite (r^+) = finite r"
+  apply auto
+   prefer 2
+   apply (rule trancl_subset_Field2 [THEN finite_subset])
+   apply (rule finite_SigmaI)
+    prefer 3
+    apply (blast intro: r_into_trancl' finite_subset)
+   apply (auto simp add: finite_Field)
+  done
+
 text {* More about converse @{text rtrancl} and @{text trancl}, should
   be merged with main body. *}
 
--- a/src/HOL/Typedef.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Typedef.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -123,7 +123,7 @@
 text {* This class is just a workaround for classes without parameters;
   it shall disappear as soon as possible. *}
 
-class itself = type + 
+class itself = 
   fixes itself :: "'a itself"
 
 setup {*
--- a/src/HOL/Typerep.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Typerep.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
-(*  Title:      HOL/Library/RType.thy
-    ID:         $Id$
+(*  Title:      HOL/Typerep.thy
     Author:     Florian Haftmann, TU Muenchen
 *)
 
@@ -15,9 +14,7 @@
   fixes typerep :: "'a\<Colon>{} itself \<Rightarrow> typerep"
 begin
 
-definition
-  typerep_of :: "'a \<Rightarrow> typerep"
-where
+definition typerep_of :: "'a \<Rightarrow> typerep" where
   [simp]: "typerep_of x = typerep TYPE('a)"
 
 end
--- a/src/HOL/Univ_Poly.thy	Thu Jan 29 12:05:19 2009 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1050 +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 "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/Wellfounded.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Wellfounded.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
-(*  ID:         $Id$
-    Author:     Tobias Nipkow
+(*  Author:     Tobias Nipkow
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Author:     Konrad Slind, Alexander Krauss
     Copyright   1992-2008  University of Cambridge and TU Muenchen
@@ -8,7 +7,7 @@
 header {*Well-founded Recursion*}
 
 theory Wellfounded
-imports Finite_Set Nat
+imports Finite_Set Transitive_Closure Nat
 uses ("Tools/function_package/size.ML")
 begin
 
--- a/src/HOL/Word/BinBoolList.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Word/BinBoolList.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (* 
-  ID:     $Id$
   Author: Jeremy Dawson, NICTA
 
   contains theorems to do with integers, expressed using Pls, Min, BIT,
@@ -281,7 +280,7 @@
   apply clarsimp
   apply safe
   apply (erule allE, erule xtr8 [rotated],
-         simp add: numeral_simps ring_simps cong add : number_of_False_cong)+
+         simp add: numeral_simps algebra_simps cong add : number_of_False_cong)+
   done
 
 lemma bl_to_bin_lt2p: "bl_to_bin bs < (2 ^ length bs)"
@@ -299,7 +298,7 @@
   apply clarsimp
   apply safe
    apply (erule allE, erule preorder_class.order_trans [rotated],
-          simp add: numeral_simps ring_simps cong add : number_of_False_cong)+
+          simp add: numeral_simps algebra_simps cong add : number_of_False_cong)+
   done
 
 lemma bl_to_bin_ge0: "bl_to_bin bs >= 0"
@@ -1138,4 +1137,3 @@
   done
 
 end
-
--- a/src/HOL/Word/BinGeneral.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Word/BinGeneral.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (* 
-  ID:     $Id$
   Author: Jeremy Dawson, NICTA
 
   contains basic definition to do with integers
--- a/src/HOL/Word/BinOperations.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Word/BinOperations.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (* 
-  ID:     $Id$
   Author: Jeremy Dawson and Gerwin Klein, NICTA
 
   definition and basic theorems for bit-wise logical operations 
--- a/src/HOL/Word/BitSyntax.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Word/BitSyntax.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (* 
-  ID:     $Id$
   Author: Brian Huffman, PSU and Gerwin Klein, NICTA
 
   Syntactic class for bitwise operations.
@@ -12,7 +11,7 @@
 imports BinGeneral
 begin
 
-class bit = type +
+class bit =
   fixes bitNOT :: "'a \<Rightarrow> 'a"       ("NOT _" [70] 71)
     and bitAND :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "AND" 64)
     and bitOR  :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "OR"  59)
--- a/src/HOL/Word/Examples/WordExamples.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Word/Examples/WordExamples.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (* 
-  ID:     $Id$
   Author: Gerwin Klein, NICTA
 
   Examples demonstrating and testing various word operations.
@@ -7,9 +6,14 @@
 
 header "Examples of word operations"
 
-theory WordExamples imports WordMain
+theory WordExamples
+imports Word
 begin
 
+types word32 = "32 word"
+types word8 = "8 word"
+types byte = word8
+
 -- "modulus"
 
 lemma "(27 :: 4 word) = -5" by simp
--- a/src/HOL/Word/ROOT.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Word/ROOT.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,2 +1,1 @@
-no_document use_thys ["Infinite_Set"];
-use_thy "WordMain";
+use_thy "Word";
--- a/src/HOL/Word/Size.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Word/Size.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (* 
-    ID:         $Id$
     Author:     John Matthews, Galois Connections, Inc., copyright 2006
 
     A typeclass for parameterizing types by size.
@@ -18,7 +17,7 @@
   some duplication with the definitions in @{text "Numeral_Type"}.
 *}
 
-class len0 = type +
+class len0 =
   fixes len_of :: "'a itself \<Rightarrow> nat"
 
 text {* 
--- a/src/HOL/Word/TdThs.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Word/TdThs.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (* 
-    ID:         $Id$
     Author:     Jeremy Dawson and Gerwin Klein, NICTA
 
   consequences of type definition theorems, 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Word/Word.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -0,0 +1,13 @@
+(*  Title:      HOL/Word/Word.thy
+    Author:     Gerwin Klein, NICTA
+*)
+
+header {* Word Library interafce *}
+
+theory Word
+imports WordGenLib
+begin
+
+text {* see @{text "Examples/WordExamples.thy"} for examples *}
+
+end
--- a/src/HOL/Word/WordArith.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Word/WordArith.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (* 
-    ID:         $Id$
     Author:     Jeremy Dawson and Gerwin Klein, NICTA
 
   contains arithmetic theorems for word, instantiations to
@@ -22,7 +21,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 = 
@@ -1002,7 +1001,7 @@
   apply (auto simp add: unat_def uint_sub_if')
    apply (rule nat_diff_distrib)
     prefer 3
-    apply (simp add: group_simps)
+    apply (simp add: algebra_simps)
     apply (rule nat_diff_distrib [THEN trans])
       prefer 3
       apply (subst nat_add_distrib)
--- a/src/HOL/Word/WordBitwise.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Word/WordBitwise.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (* 
-    ID:         $Id$
     Author:     Jeremy Dawson and Gerwin Klein, NICTA
 
   contains theorems to do with bit-wise (logical) operations on words
--- a/src/HOL/Word/WordDefinition.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Word/WordDefinition.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (* 
-  ID:     $Id$
   Author: Jeremy Dawson and Gerwin Klein, NICTA
   
   Basic definition of word type and basic theorems following from 
@@ -12,8 +11,50 @@
 imports Size BinBoolList TdThs
 begin
 
-typedef (open word) 'a word
-  = "{(0::int) ..< 2^len_of TYPE('a::len0)}" by auto
+subsection {* Type definition *}
+
+typedef (open word) 'a word = "{(0::int) ..< 2^len_of TYPE('a::len0)}"
+  morphisms uint Abs_word by auto
+
+definition word_of_int :: "int \<Rightarrow> 'a\<Colon>len0 word" where
+  -- {* representation of words using unsigned or signed bins, 
+        only difference in these is the type class *}
+  [code del]: "word_of_int w = Abs_word (bintrunc (len_of TYPE ('a)) w)" 
+
+lemma uint_word_of_int [code]: "uint (word_of_int w \<Colon> 'a\<Colon>len0 word) = w mod 2 ^ len_of TYPE('a)"
+  by (auto simp add: word_of_int_def bintrunc_mod2p intro: Abs_word_inverse)
+
+code_datatype word_of_int
+
+
+subsection {* Type conversions and casting *}
+
+definition sint :: "'a :: len word => int" where
+  -- {* treats the most-significant-bit as a sign bit *}
+  sint_uint: "sint w = sbintrunc (len_of TYPE ('a) - 1) (uint w)"
+
+definition unat :: "'a :: len0 word => nat" where
+  "unat w = nat (uint w)"
+
+definition uints :: "nat => int set" where
+  -- "the sets of integers representing the words"
+  "uints n = range (bintrunc n)"
+
+definition sints :: "nat => int set" where
+  "sints n = range (sbintrunc (n - 1))"
+
+definition unats :: "nat => nat set" where
+  "unats n = {i. i < 2 ^ n}"
+
+definition norm_sint :: "nat => int => int" where
+  "norm_sint n w = (w + 2 ^ (n - 1)) mod 2 ^ n - 2 ^ (n - 1)"
+
+definition scast :: "'a :: len word => 'b :: len word" where
+  -- "cast a word to a different length"
+  "scast w = word_of_int (sint w)"
+
+definition ucast :: "'a :: len0 word => 'b :: len0 word" where
+  "ucast w = word_of_int (uint w)"
 
 instantiation word :: (len0) size
 begin
@@ -25,83 +66,39 @@
 
 end
 
-definition
-  -- {* representation of words using unsigned or signed bins, 
-        only difference in these is the type class *}
-  word_of_int :: "int \<Rightarrow> 'a\<Colon>len0 word"
-where
-  [code del]: "word_of_int w = Abs_word (bintrunc (len_of TYPE ('a)) w)" 
-
-code_datatype word_of_int
-
+definition source_size :: "('a :: len0 word => 'b) => nat" where
+  -- "whether a cast (or other) function is to a longer or shorter length"
+  "source_size c = (let arb = undefined ; x = c arb in size arb)"  
 
-subsection "Type conversions and casting"
+definition target_size :: "('a => 'b :: len0 word) => nat" where
+  "target_size c = size (c undefined)"
 
-constdefs
-  -- {* uint and sint cast a word to an integer,
-        uint treats the word as unsigned,
-        sint treats the most-significant-bit as a sign bit *}
-  uint :: "'a :: len0 word => int"
-  "uint w == Rep_word w"
-  sint :: "'a :: len word => int"
-  sint_uint: "sint w == sbintrunc (len_of TYPE ('a) - 1) (uint w)"
-  unat :: "'a :: len0 word => nat"
-  "unat w == nat (uint w)"
+definition is_up :: "('a :: len0 word => 'b :: len0 word) => bool" where
+  "is_up c \<longleftrightarrow> source_size c <= target_size c"
 
-  -- "the sets of integers representing the words"
-  uints :: "nat => int set"
-  "uints n == range (bintrunc n)"
-  sints :: "nat => int set"
-  "sints n == range (sbintrunc (n - 1))"
-  unats :: "nat => nat set"
-  "unats n == {i. i < 2 ^ n}"
-  norm_sint :: "nat => int => int"
-  "norm_sint n w == (w + 2 ^ (n - 1)) mod 2 ^ n - 2 ^ (n - 1)"
+definition is_down :: "('a :: len0 word => 'b :: len0 word) => bool" where
+  "is_down c \<longleftrightarrow> target_size c <= source_size c"
 
-  -- "cast a word to a different length"
-  scast :: "'a :: len word => 'b :: len word"
-  "scast w == word_of_int (sint w)"
-  ucast :: "'a :: len0 word => 'b :: len0 word"
-  "ucast w == word_of_int (uint w)"
+definition of_bl :: "bool list => 'a :: len0 word" where
+  "of_bl bl = word_of_int (bl_to_bin bl)"
 
-  -- "whether a cast (or other) function is to a longer or shorter length"
-  source_size :: "('a :: len0 word => 'b) => nat"
-  "source_size c == let arb = undefined ; x = c arb in size arb"  
-  target_size :: "('a => 'b :: len0 word) => nat"
-  "target_size c == size (c undefined)"
-  is_up :: "('a :: len0 word => 'b :: len0 word) => bool"
-  "is_up c == source_size c <= target_size c"
-  is_down :: "('a :: len0 word => 'b :: len0 word) => bool"
-  "is_down c == target_size c <= source_size c"
+definition to_bl :: "'a :: len0 word => bool list" where
+  "to_bl w = bin_to_bl (len_of TYPE ('a)) (uint w)"
 
-constdefs
-  of_bl :: "bool list => 'a :: len0 word" 
-  "of_bl bl == word_of_int (bl_to_bin bl)"
-  to_bl :: "'a :: len0 word => bool list"
-  "to_bl w == 
-  bin_to_bl (len_of TYPE ('a)) (uint w)"
+definition word_reverse :: "'a :: len0 word => 'a word" where
+  "word_reverse w = of_bl (rev (to_bl w))"
 
-  word_reverse :: "'a :: len0 word => 'a word"
-  "word_reverse w == of_bl (rev (to_bl w))"
-
-constdefs
-  word_int_case :: "(int => 'b) => ('a :: len0 word) => 'b"
-  "word_int_case f w == f (uint w)"
+definition word_int_case :: "(int => 'b) => ('a :: len0 word) => 'b" where
+  "word_int_case f w = f (uint w)"
 
 syntax
   of_int :: "int => 'a"
 translations
-  "case x of of_int y => b" == "word_int_case (%y. b) x"
+  "case x of of_int y => b" == "CONST word_int_case (%y. b) x"
 
 
 subsection  "Arithmetic operations"
 
-declare uint_def [code del]
-
-lemma [code]: "uint (word_of_int w \<Colon> 'a\<Colon>len0 word) = bintrunc (len_of TYPE('a)) w"
-  by (auto simp add: uint_def word_of_int_def intro!: Abs_word_inverse)
-    (insert range_bintrunc, auto)
-
 instantiation word :: (len0) "{number, uminus, minus, plus, one, zero, times, Divides.div, power, ord, bit}"
 begin
 
@@ -186,8 +183,6 @@
 
 subsection "Bit-wise operations"
 
-
-
 instantiation word :: (len0) bits
 begin
 
@@ -337,21 +332,21 @@
   unfolding atLeastLessThan_alt by auto
 
 lemma 
-  Rep_word_0:"0 <= Rep_word x" and 
-  Rep_word_lt: "Rep_word (x::'a::len0 word) < 2 ^ len_of TYPE('a)"
-  by (auto simp: Rep_word [simplified])
+  uint_0:"0 <= uint x" and 
+  uint_lt: "uint (x::'a::len0 word) < 2 ^ len_of TYPE('a)"
+  by (auto simp: uint [simplified])
 
-lemma Rep_word_mod_same:
-  "Rep_word x mod 2 ^ len_of TYPE('a) = Rep_word (x::'a::len0 word)"
-  by (simp add: int_mod_eq Rep_word_lt Rep_word_0)
+lemma uint_mod_same:
+  "uint x mod 2 ^ len_of TYPE('a) = uint (x::'a::len0 word)"
+  by (simp add: int_mod_eq uint_lt uint_0)
 
 lemma td_ext_uint: 
   "td_ext (uint :: 'a word => int) word_of_int (uints (len_of TYPE('a::len0))) 
     (%w::int. w mod 2 ^ len_of TYPE('a))"
   apply (unfold td_ext_def')
-  apply (simp add: uints_num uint_def word_of_int_def bintrunc_mod2p)
-  apply (simp add: Rep_word_mod_same Rep_word_0 Rep_word_lt
-                   word.Rep_word_inverse word.Abs_word_inverse int_mod_lem)
+  apply (simp add: uints_num word_of_int_def bintrunc_mod2p)
+  apply (simp add: uint_mod_same uint_0 uint_lt
+                   word.uint_inverse word.Abs_word_inverse int_mod_lem)
   done
 
 lemmas int_word_uint = td_ext_uint [THEN td_ext.eq_norm, standard]
@@ -793,10 +788,7 @@
 lemmas is_down = is_down_def [unfolded source_size target_size]
 lemmas is_up = is_up_def [unfolded source_size target_size]
 
-lemmas is_up_down = 
-  trans [OF is_up [THEN meta_eq_to_obj_eq] 
-            is_down [THEN meta_eq_to_obj_eq, symmetric], 
-         standard]
+lemmas is_up_down =  trans [OF is_up is_down [symmetric], standard]
 
 lemma down_cast_same': "uc = ucast ==> is_down uc ==> uc = scast"
   apply (unfold is_down)
@@ -950,4 +942,17 @@
 lemmas word_log_defs = word_and_def word_or_def word_xor_def word_not_def
 lemmas word_log_bin_defs = word_log_defs
 
+text {* Executable equality *}
+
+instantiation word :: ("{len0}") eq
+begin
+
+definition eq_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> bool" where
+  "eq_word k l \<longleftrightarrow> HOL.eq (uint k) (uint l)"
+
+instance proof
+qed (simp add: eq eq_word_def)
+
 end
+
+end
--- a/src/HOL/Word/WordGenLib.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Word/WordGenLib.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (* Author: Gerwin Klein, Jeremy Dawson
-   $Id$
 
    Miscellaneous additional library definitions and lemmas for 
    the word type. Instantiation to boolean algebras, definition
@@ -452,4 +451,13 @@
   "1 + n \<noteq> (0::'a::len word) \<Longrightarrow> unat (1 + n) = Suc (unat n)"
   by unat_arith
 
+
+lemmas word_no_1 [simp] = word_1_no [symmetric, unfolded BIT_simps]
+lemmas word_no_0 [simp] = word_0_no [symmetric]
+
+declare word_0_bl [simp]
+declare bin_to_bl_def [simp]
+declare to_bl_0 [simp]
+declare of_bl_True [simp]
+
 end
--- a/src/HOL/Word/WordMain.thy	Thu Jan 29 12:05:19 2009 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,30 +0,0 @@
-(* 
-  ID:     $Id$
-  Author: Gerwin Klein, NICTA
-
-  The main interface of the word library to other theories.
-*)
-
-header {* Main Word Library *}
-
-theory WordMain
-imports WordGenLib
-begin
-
-lemmas word_no_1 [simp] = word_1_no [symmetric, unfolded BIT_simps]
-lemmas word_no_0 [simp] = word_0_no [symmetric]
-
-declare word_0_bl [simp]
-declare bin_to_bl_def [simp]
-declare to_bl_0 [simp]
-declare of_bl_True [simp]
-
-text "Examples"
-
-types word32 = "32 word"
-types word8 = "8 word"
-types byte = word8
-
-text {* for more see WordExamples.thy *}
-
-end
--- a/src/HOL/Word/WordShift.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/Word/WordShift.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (* 
-    ID:         $Id$
     Author:     Jeremy Dawson and Gerwin Klein, NICTA
 
   contains theorems to do with shifting, rotating, splitting words
--- a/src/HOL/ZF/LProd.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/ZF/LProd.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -45,9 +45,9 @@
   case (lprod_list ah at bh bt a b)
   from prems have transR: "trans R" by auto
   have as: "multiset_of (ah @ a # at) = multiset_of (ah @ at) + {#a#}" (is "_ = ?ma + _")
-    by (simp add: ring_simps)
+    by (simp add: algebra_simps)
   have bs: "multiset_of (bh @ b # bt) = multiset_of (bh @ bt) + {#b#}" (is "_ = ?mb + _")
-    by (simp add: ring_simps)
+    by (simp add: algebra_simps)
   from prems have "(?ma, ?mb) \<in> mult R"
     by auto
   with mult_implies_one_step[OF transR] have 
@@ -66,7 +66,7 @@
     then show ?thesis
       apply (simp only: as bs)
       apply (simp only: decomposed True)
-      apply (simp add: ring_simps)
+      apply (simp add: algebra_simps)
       done
   next
     case False
@@ -78,7 +78,7 @@
     then show ?thesis
       apply (simp only: as bs)
       apply (simp only: decomposed)
-      apply (simp add: ring_simps)
+      apply (simp add: algebra_simps)
       done
   qed
 qed
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/base.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -0,0 +1,2 @@
+(*side-entry for HOL-Base*)
+use_thy "Code_Setup";
--- a/src/HOL/ex/Lagrange.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/ex/Lagrange.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -35,7 +35,7 @@
    sq (x1*y2 + x2*y1 + x3*y4 - x4*y3)  +
    sq (x1*y3 - x2*y4 + x3*y1 + x4*y2)  +
    sq (x1*y4 + x2*y3 - x3*y2 + x4*y1)"
-by (simp add: sq_def ring_simps)
+by (simp add: sq_def algebra_simps)
 
 
 text {* A challenge by John Harrison. Takes about 12s on a 1.6GHz machine. *}
@@ -51,6 +51,6 @@
       sq (p1*u2 + q1*t2 - r1*w2 + s1*v2 - t1*q2 + u1*p2 - v1*s2 + w1*r2) +
       sq (p1*v2 + q1*w2 + r1*t2 - s1*u2 - t1*r2 + u1*s2 + v1*p2 - w1*q2) +
       sq (p1*w2 - q1*v2 + r1*u2 + s1*t2 - t1*s2 - u1*r2 + v1*q2 + w1*p2)"
-by (simp add: sq_def ring_simps)
+by (simp add: sq_def algebra_simps)
 
 end
--- a/src/HOL/ex/LocaleTest2.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/ex/LocaleTest2.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -625,9 +625,6 @@
 lemma "gcd x y dvd x"
   apply (rule nat_dvd.meet_left) done
 
-print_interps dpo
-print_interps dlat
-
 
 subsection {* Group example with defined operations @{text inv} and @{text unit} *}
 
--- a/src/HOL/ex/MIR.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/ex/MIR.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -709,11 +709,11 @@
 next
   case (2 n c t)  hence gd: "g dvd c" by simp
   from gp have gnz: "g \<noteq> 0" by simp
-  from prems show ?case by (simp add: real_of_int_div[OF gnz gd] ring_simps)
+  from prems show ?case by (simp add: real_of_int_div[OF gnz gd] algebra_simps)
 next
   case (3 c s t)  hence gd: "g dvd c" by simp
   from gp have gnz: "g \<noteq> 0" by simp
-  from prems show ?case by (simp add: real_of_int_div[OF gnz gd] ring_simps) 
+  from prems show ?case by (simp add: real_of_int_div[OF gnz gd] algebra_simps) 
 qed (auto simp add: numgcd_def gp)
 consts ismaxcoeff:: "num \<Rightarrow> int \<Rightarrow> bool"
 recdef ismaxcoeff "measure size"
@@ -850,12 +850,12 @@
 lemma numadd[simp]: "Inum bs (numadd (t,s)) = Inum bs (Add t s)"
 apply (induct t s rule: numadd.induct, simp_all add: Let_def)
  apply (case_tac "c1+c2 = 0",case_tac "n1 \<le> n2", simp_all)
-  apply (case_tac "n1 = n2", simp_all add: ring_simps)
+  apply (case_tac "n1 = n2", simp_all add: algebra_simps)
   apply (simp only: left_distrib[symmetric])
  apply simp
 apply (case_tac "lex_bnd t1 t2", simp_all)
  apply (case_tac "c1+c2 = 0")
-  by (case_tac "t1 = t2", simp_all add: ring_simps left_distrib[symmetric] real_of_int_mult[symmetric] real_of_int_add[symmetric]del: real_of_int_mult real_of_int_add left_distrib)
+  by (case_tac "t1 = t2", simp_all add: algebra_simps left_distrib[symmetric] real_of_int_mult[symmetric] real_of_int_add[symmetric]del: real_of_int_mult real_of_int_add left_distrib)
 
 lemma numadd_nb[simp]: "\<lbrakk> numbound0 t ; numbound0 s\<rbrakk> \<Longrightarrow> numbound0 (numadd (t,s))"
 by (induct t s rule: numadd.induct, auto simp add: Let_def)
@@ -868,7 +868,7 @@
   "nummul t = (\<lambda> i. Mul i t)"
 
 lemma nummul[simp]: "\<And> i. Inum bs (nummul t i) = Inum bs (Mul i t)"
-by (induct t rule: nummul.induct, auto simp add: ring_simps)
+by (induct t rule: nummul.induct, auto simp add: algebra_simps)
 
 lemma nummul_nb[simp]: "\<And> i. numbound0 t \<Longrightarrow> numbound0 (nummul t i)"
 by (induct t rule: nummul.induct, auto)
@@ -928,7 +928,7 @@
   with prems(1) have b:"Inum bs (Add ?bv ?bi) = Inum bs b" and bii: "isint ?bi bs" by blast+
   from prems(2) have tibi: "ti = CF c a ?bi" by (simp add: Let_def split_def)
   from prems(2) b[symmetric] bii show ?case by (auto simp add: Let_def split_def isint_Floor isint_add isint_Mul isint_CF)
-qed (auto simp add: Let_def isint_iff isint_Floor isint_add isint_Mul split_def ring_simps)
+qed (auto simp add: Let_def isint_iff isint_Floor isint_add isint_Mul split_def algebra_simps)
 
 lemma split_int_nb: "numbound0 t \<Longrightarrow> numbound0 (fst (split_int t)) \<and> numbound0 (snd (split_int t)) "
 by (induct t rule: split_int.induct, auto simp add: Let_def split_def)
@@ -1768,7 +1768,7 @@
   have th: "(real a + b >0) = (real (-a) + (-b)< 0)" by arith
   show ?thesis using myless[rule_format, where b="real (floor b)"] 
     by (simp only:th split_int_less_real'[where a="-a" and b="-b"]) 
-    (simp add: ring_simps diff_def[symmetric],arith)
+    (simp add: algebra_simps diff_def[symmetric],arith)
 qed
 
 lemma split_int_le_real: 
@@ -1795,7 +1795,7 @@
 proof- 
   have th: "(real a + b \<ge>0) = (real (-a) + (-b) \<le> 0)" by arith
   show ?thesis by (simp only: th split_int_le_real'[where a="-a" and b="-b"])
-    (simp add: ring_simps diff_def[symmetric],arith)
+    (simp add: algebra_simps diff_def[symmetric],arith)
 qed
 
 lemma split_int_eq_real: "(real (a::int) = b) = ( a = floor b \<and> b = real (floor b))" (is "?l = ?r")
@@ -2281,9 +2281,9 @@
       (is "?ri rdvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri rdvd ?rt")
       hence "\<exists> (l::int). ?rt = ?ri * (real l)" by (simp add: rdvd_def)
       hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real l)+?rc*(?rk * (real i) * (real di))" 
-	by (simp add: ring_simps di_def)
+	by (simp add: algebra_simps di_def)
       hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real (l + c*k*di))"
-	by (simp add: ring_simps)
+	by (simp add: algebra_simps)
       hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri* (real l)" by blast
       thus "real i rdvd real c * real x + Inum (real x # bs) e" using rdvd_def by simp
     next
@@ -2292,7 +2292,7 @@
       hence "\<exists> (l::int). ?rc*?rx+?e = ?ri * (real l)" by (simp add: rdvd_def)
       hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real d)" by simp
       hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real i * real di)" by (simp add: di_def)
-      hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real (l - c*k*di))" by (simp add: ring_simps)
+      hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real (l - c*k*di))" by (simp add: algebra_simps)
       hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l)"
 	by blast
       thus "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e" using rdvd_def by simp
@@ -2308,9 +2308,9 @@
       (is "?ri rdvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri rdvd ?rt")
       hence "\<exists> (l::int). ?rt = ?ri * (real l)" by (simp add: rdvd_def)
       hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real l)+?rc*(?rk * (real i) * (real di))" 
-	by (simp add: ring_simps di_def)
+	by (simp add: algebra_simps di_def)
       hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real (l + c*k*di))"
-	by (simp add: ring_simps)
+	by (simp add: algebra_simps)
       hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri* (real l)" by blast
       thus "real i rdvd real c * real x + Inum (real x # bs) e" using rdvd_def by simp
     next
@@ -2319,7 +2319,7 @@
       hence "\<exists> (l::int). ?rc*?rx+?e = ?ri * (real l)" by (simp add: rdvd_def)
       hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real d)" by simp
       hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real i * real di)" by (simp add: di_def)
-      hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real (l - c*k*di))" by (simp add: ring_simps)
+      hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real (l - c*k*di))" by (simp add: algebra_simps)
       hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l)"
 	by blast
       thus "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e" using rdvd_def by simp
@@ -2450,16 +2450,16 @@
   have th: "(real j rdvd real c * real x - Inum (real x # bs) e) =
        (real j rdvd - (real c * real x - Inum (real x # bs) e))"
     by (simp only: rdvd_minus[symmetric])
-  from prems show  ?case
-    by (simp add: ring_simps th[simplified ring_simps]
+  from prems th show  ?case
+    by (simp add: algebra_simps
       numbound0_I[where bs="bs" and b'="real x" and b="- real x"])
 next
     case (10 j c e)
   have th: "(real j rdvd real c * real x - Inum (real x # bs) e) =
        (real j rdvd - (real c * real x - Inum (real x # bs) e))"
     by (simp only: rdvd_minus[symmetric])
-  from prems show  ?case
-    by (simp add: ring_simps th[simplified ring_simps]
+  from prems th show  ?case
+    by (simp add: algebra_simps
       numbound0_I[where bs="bs" and b'="real x" and b="- real x"])
 qed (auto simp add: numbound0_I[where bs="bs" and b="real x" and b'="- real x"] nth_pos2)
 
@@ -2541,7 +2541,7 @@
     hence "(real l * real x + real (l div c) * Inum (real x # bs) e < (0\<Colon>real)) =
           (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e < 0)"
       by simp
-    also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) < (real (l div c)) * 0)" by (simp add: ring_simps)
+    also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) < (real (l div c)) * 0)" by (simp add: algebra_simps)
     also have "\<dots> = (real c * real x + Inum (real x # bs) e < 0)"
     using mult_less_0_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp
   finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] be  isint_Mul[OF ei] by simp
@@ -2559,7 +2559,7 @@
     hence "(real l * real x + real (l div c) * Inum (real x # bs) e \<le> (0\<Colon>real)) =
           (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e \<le> 0)"
       by simp
-    also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) \<le> (real (l div c)) * 0)" by (simp add: ring_simps)
+    also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) \<le> (real (l div c)) * 0)" by (simp add: algebra_simps)
     also have "\<dots> = (real c * real x + Inum (real x # bs) e \<le> 0)"
     using mult_le_0_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp
   finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"]  be  isint_Mul[OF ei] by simp
@@ -2577,7 +2577,7 @@
     hence "(real l * real x + real (l div c) * Inum (real x # bs) e > (0\<Colon>real)) =
           (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e > 0)"
       by simp
-    also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) > (real (l div c)) * 0)" by (simp add: ring_simps)
+    also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) > (real (l div c)) * 0)" by (simp add: algebra_simps)
     also have "\<dots> = (real c * real x + Inum (real x # bs) e > 0)"
     using zero_less_mult_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp
   finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"]  be  isint_Mul[OF ei] by simp
@@ -2595,7 +2595,7 @@
     hence "(real l * real x + real (l div c) * Inum (real x # bs) e \<ge> (0\<Colon>real)) =
           (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e \<ge> 0)"
       by simp
-    also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) \<ge> (real (l div c)) * 0)" by (simp add: ring_simps)
+    also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) \<ge> (real (l div c)) * 0)" by (simp add: algebra_simps)
     also have "\<dots> = (real c * real x + Inum (real x # bs) e \<ge> 0)"
     using zero_le_mult_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp
   finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"]  be  isint_Mul[OF ei] by simp
@@ -2613,7 +2613,7 @@
     hence "(real l * real x + real (l div c) * Inum (real x # bs) e = (0\<Colon>real)) =
           (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e = 0)"
       by simp
-    also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) = (real (l div c)) * 0)" by (simp add: ring_simps)
+    also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) = (real (l div c)) * 0)" by (simp add: algebra_simps)
     also have "\<dots> = (real c * real x + Inum (real x # bs) e = 0)"
     using mult_eq_0_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp
   finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"]  be  isint_Mul[OF ei] by simp
@@ -2631,7 +2631,7 @@
     hence "(real l * real x + real (l div c) * Inum (real x # bs) e \<noteq> (0\<Colon>real)) =
           (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e \<noteq> 0)"
       by simp
-    also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) \<noteq> (real (l div c)) * 0)" by (simp add: ring_simps)
+    also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) \<noteq> (real (l div c)) * 0)" by (simp add: algebra_simps)
     also have "\<dots> = (real c * real x + Inum (real x # bs) e \<noteq> 0)"
     using zero_le_mult_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp
   finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"]  be  isint_Mul[OF ei] by simp
@@ -2647,7 +2647,7 @@
     hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] 
       by simp
     hence "(\<exists> (k::int). real l * real x + real (l div c) * Inum (real x # bs) e = (real (l div c) * real j) * real k) = (\<exists> (k::int). real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e = (real (l div c) * real j) * real k)"  by simp
-    also have "\<dots> = (\<exists> (k::int). real (l div c) * (real c * real x + Inum (real x # bs) e - real j * real k) = real (l div c)*0)" by (simp add: ring_simps)
+    also have "\<dots> = (\<exists> (k::int). real (l div c) * (real c * real x + Inum (real x # bs) e - real j * real k) = real (l div c)*0)" by (simp add: algebra_simps)
     also fix k have "\<dots> = (\<exists> (k::int). real c * real x + Inum (real x # bs) e - real j * real k = 0)"
     using zero_le_mult_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e - real j * real k"] ldcp by simp
   also have "\<dots> = (\<exists> (k::int). real c * real x + Inum (real x # bs) e = real j * real k)" by simp
@@ -2664,7 +2664,7 @@
     hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] 
       by simp
     hence "(\<exists> (k::int). real l * real x + real (l div c) * Inum (real x # bs) e = (real (l div c) * real j) * real k) = (\<exists> (k::int). real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e = (real (l div c) * real j) * real k)"  by simp
-    also have "\<dots> = (\<exists> (k::int). real (l div c) * (real c * real x + Inum (real x # bs) e - real j * real k) = real (l div c)*0)" by (simp add: ring_simps)
+    also have "\<dots> = (\<exists> (k::int). real (l div c) * (real c * real x + Inum (real x # bs) e - real j * real k) = real (l div c)*0)" by (simp add: algebra_simps)
     also fix k have "\<dots> = (\<exists> (k::int). real c * real x + Inum (real x # bs) e - real j * real k = 0)"
     using zero_le_mult_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e - real j * real k"] ldcp by simp
   also have "\<dots> = (\<exists> (k::int). real c * real x + Inum (real x # bs) e = real j * real k)" by simp
@@ -2719,7 +2719,7 @@
       hence "x + floor ?e \<ge> 1 \<and> x + floor ?e \<le> d"  by simp
       hence "\<exists> (j::int) \<in> {1 .. d}. j = x + floor ?e" by simp
       hence "\<exists> (j::int) \<in> {1 .. d}. real x = real (- floor ?e + j)" 
-	by (simp only: real_of_int_inject) (simp add: ring_simps)
+	by (simp only: real_of_int_inject) (simp add: algebra_simps)
       hence "\<exists> (j::int) \<in> {1 .. d}. real x = - ?e + real j" 
 	by (simp add: ie[simplified isint_iff])
       with nob have ?case by auto}
@@ -2744,7 +2744,7 @@
 	using ie by simp
       hence "x + floor ?e +1 \<ge> 1 \<and> x + floor ?e + 1 \<le> d"  by simp
       hence "\<exists> (j::int) \<in> {1 .. d}. j = x + floor ?e + 1" by simp
-      hence "\<exists> (j::int) \<in> {1 .. d}. x= - floor ?e - 1 + j" by (simp add: ring_simps)
+      hence "\<exists> (j::int) \<in> {1 .. d}. x= - floor ?e - 1 + j" by (simp add: algebra_simps)
       hence "\<exists> (j::int) \<in> {1 .. d}. real x= real (- floor ?e - 1 + j)" 
 	by (simp only: real_of_int_inject)
       hence "\<exists> (j::int) \<in> {1 .. d}. real x= - ?e - 1 + real j" 
@@ -2759,7 +2759,7 @@
     have vb: "?v \<in> set (\<beta> (Eq (CN 0 c e)))" by simp
     from p have "real x= - ?e" by (simp add: c1) with prems(11) show ?case using dp
       by simp (erule ballE[where x="1"],
-	simp_all add:ring_simps numbound0_I[OF bn,where b="real x"and b'="a"and bs="bs"])
+	simp_all add:algebra_simps numbound0_I[OF bn,where b="real x"and b'="a"and bs="bs"])
 next
   case (4 c e)hence p: "Ifm (real x #bs) (NEq (CN 0 c e))" (is "?p x") and c1: "c=1" 
     and bn:"numbound0 e" and ie1: "isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+
@@ -2983,7 +2983,7 @@
       from prems have "?I (real x) (?s (Eq (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k = 0)"
 	using real_of_int_div[OF knz kdt]
 	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
-	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps)
+	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps)
       also have "\<dots> = (?I ?tk (Eq (CN 0 c e)))" using nonzero_eq_divide_eq[OF knz', where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
 	by (simp add: ti)
@@ -3005,7 +3005,7 @@
       from prems have "?I (real x) (?s (NEq (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k \<noteq> 0)"
 	using real_of_int_div[OF knz kdt]
 	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
-	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps)
+	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps)
       also have "\<dots> = (?I ?tk (NEq (CN 0 c e)))" using nonzero_eq_divide_eq[OF knz', where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
 	by (simp add: ti)
@@ -3027,7 +3027,7 @@
       from prems have "?I (real x) (?s (Lt (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k < 0)"
 	using real_of_int_div[OF knz kdt]
 	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
-	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps)
+	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps)
       also have "\<dots> = (?I ?tk (Lt (CN 0 c e)))" using pos_less_divide_eq[OF kpos, where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
 	by (simp add: ti)
@@ -3049,7 +3049,7 @@
       from prems have "?I (real x) (?s (Le (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k \<le> 0)"
 	using real_of_int_div[OF knz kdt]
 	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
-	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps)
+	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps)
       also have "\<dots> = (?I ?tk (Le (CN 0 c e)))" using pos_le_divide_eq[OF kpos, where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
 	by (simp add: ti)
@@ -3071,7 +3071,7 @@
       from prems have "?I (real x) (?s (Gt (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k > 0)"
 	using real_of_int_div[OF knz kdt]
 	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
-	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps)
+	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps)
       also have "\<dots> = (?I ?tk (Gt (CN 0 c e)))" using pos_divide_less_eq[OF kpos, where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
 	by (simp add: ti)
@@ -3093,7 +3093,7 @@
       from prems have "?I (real x) (?s (Ge (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k \<ge> 0)"
 	using real_of_int_div[OF knz kdt]
 	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
-	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps)
+	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps)
       also have "\<dots> = (?I ?tk (Ge (CN 0 c e)))" using pos_divide_le_eq[OF kpos, where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
 	by (simp add: ti)
@@ -3114,7 +3114,7 @@
       from prems have "?I (real x) (?s (Dvd i (CN 0 c e))) = (real i * real k rdvd (real c * (?N (real x) t / real k) + ?N (real x) e)* real k)"
 	using real_of_int_div[OF knz kdt]
 	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
-	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps)
+	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps)
       also have "\<dots> = (?I ?tk (Dvd i (CN 0 c e)))" using rdvd_mult[OF knz, where n="i"] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
 	by (simp add: ti)
@@ -3135,7 +3135,7 @@
       from prems have "?I (real x) (?s (NDvd i (CN 0 c e))) = (\<not> (real i * real k rdvd (real c * (?N (real x) t / real k) + ?N (real x) e)* real k))"
 	using real_of_int_div[OF knz kdt]
 	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
-	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps)
+	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps)
       also have "\<dots> = (?I ?tk (NDvd i (CN 0 c e)))" using rdvd_mult[OF knz, where n="i"] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
 	by (simp add: ti)
@@ -3154,7 +3154,7 @@
   from kp have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
     {assume kdc: "k dvd c" from prems have  ?case using real_of_int_div[OF knz kdc] by simp } 
     moreover 
-    {assume nkdc: "\<not> k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] nonzero_eq_divide_eq[OF knz', where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: ring_simps)}
+    {assume nkdc: "\<not> k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] nonzero_eq_divide_eq[OF knz', where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: algebra_simps)}
     ultimately show ?case by blast 
 next
   case (4 c e)   
@@ -3162,7 +3162,7 @@
   from kp have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
     {assume kdc: "k dvd c" from prems have  ?case using real_of_int_div[OF knz kdc] by simp } 
     moreover 
-    {assume nkdc: "\<not> k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] nonzero_eq_divide_eq[OF knz', where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: ring_simps)}
+    {assume nkdc: "\<not> k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] nonzero_eq_divide_eq[OF knz', where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: algebra_simps)}
     ultimately show ?case by blast 
 next
   case (5 c e)   
@@ -3170,7 +3170,7 @@
   from kp have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
     {assume kdc: "k dvd c" from prems have  ?case using real_of_int_div[OF knz kdc] by simp } 
     moreover 
-    {assume nkdc: "\<not> k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] pos_less_divide_eq[OF kp, where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: ring_simps)}
+    {assume nkdc: "\<not> k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] pos_less_divide_eq[OF kp, where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: algebra_simps)}
     ultimately show ?case by blast 
 next
   case (6 c e)    
@@ -3178,7 +3178,7 @@
   from kp have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
     {assume kdc: "k dvd c" from prems have  ?case using real_of_int_div[OF knz kdc] by simp } 
     moreover 
-    {assume nkdc: "\<not> k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] pos_le_divide_eq[OF kp, where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: ring_simps)}
+    {assume nkdc: "\<not> k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] pos_le_divide_eq[OF kp, where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: algebra_simps)}
     ultimately show ?case by blast 
 next
   case (7 c e)    
@@ -3186,7 +3186,7 @@
   from kp have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
     {assume kdc: "k dvd c" from prems have  ?case using real_of_int_div[OF knz kdc] by simp } 
     moreover 
-    {assume nkdc: "\<not> k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] pos_divide_less_eq[OF kp, where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: ring_simps)}
+    {assume nkdc: "\<not> k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] pos_divide_less_eq[OF kp, where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: algebra_simps)}
     ultimately show ?case by blast 
 next
   case (8 c e)    
@@ -3194,7 +3194,7 @@
   from kp have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
     {assume kdc: "k dvd c" from prems have  ?case using real_of_int_div[OF knz kdc] by simp } 
     moreover 
-    {assume nkdc: "\<not> k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] pos_divide_le_eq[OF kp, where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: ring_simps)}
+    {assume nkdc: "\<not> k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] pos_divide_le_eq[OF kp, where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: algebra_simps)}
     ultimately show ?case by blast 
 next
   case (9 i c e)
@@ -3206,7 +3206,7 @@
     hence "Ifm (real (x*k)#bs) (a\<rho> (Dvd i (CN 0 c e)) k) = 
       (real i * real k rdvd (real c * real x + Inum (real x#bs) e) * real k)" 
       using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] 
-      by (simp add: ring_simps)
+      by (simp add: algebra_simps)
     also have "\<dots> = (Ifm (real x#bs) (Dvd i (CN 0 c e)))" by (simp add: rdvd_mult[OF knz, where n="i"])
     finally have ?case . }
   ultimately show ?case by blast 
@@ -3220,7 +3220,7 @@
     hence "Ifm (real (x*k)#bs) (a\<rho> (NDvd i (CN 0 c e)) k) = 
       (\<not> (real i * real k rdvd (real c * real x + Inum (real x#bs) e) * real k))" 
       using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] 
-      by (simp add: ring_simps)
+      by (simp add: algebra_simps)
     also have "\<dots> = (Ifm (real x#bs) (NDvd i (CN 0 c e)))" by (simp add: rdvd_mult[OF knz, where n="i"])
     finally have ?case . }
   ultimately show ?case by blast 
@@ -3233,7 +3233,7 @@
 proof-
   have "(\<exists> x. ?D x \<and> ?P' x) = (\<exists> x. k dvd x \<and> ?P' x)" using int_rdvd_iff by simp
   also have "\<dots> = (\<exists>x. ?P' (x*k))" using unity_coeff_ex[where P="?P'" and l="k", simplified]
-    by (simp add: ring_simps)
+    by (simp add: algebra_simps)
   also have "\<dots> = (\<exists> x. ?P x)" using a\<rho> iszlfm_gen[OF lp] kp by auto
   finally show ?thesis .
 qed
@@ -3297,7 +3297,7 @@
     by simp+
   {assume "real (c*i) \<noteq> - ?N i e + real (c*d)"
     with numbound0_I[OF nb, where bs="bs" and b="real i - real d" and b'="real i"]
-    have ?case by (simp add: ring_simps)}
+    have ?case by (simp add: algebra_simps)}
   moreover
   {assume pi: "real (c*i) = - ?N i e + real (c*d)"
     from mult_strict_left_mono[OF dp cp] have d: "(c*d) \<in> {1 .. c*d}" by simp
@@ -3309,27 +3309,27 @@
     real_of_int_mult]
   show ?case using prems dp 
     by (simp add: add: numbound0_I[where bs="bs" and b="real i - real d" and b'="real i"] 
-      ring_simps)
+      algebra_simps)
 next
   case (6 c e)  hence cp: "c > 0" by simp
   from prems mult_strict_left_mono[OF dp cp, simplified real_of_int_less_iff[symmetric] 
     real_of_int_mult]
   show ?case using prems dp 
     by (simp add: add: numbound0_I[where bs="bs" and b="real i - real d" and b'="real i"] 
-      ring_simps)
+      algebra_simps)
 next
   case (7 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)"
     and nob: "\<forall> j\<in> {1 .. c*d}. real (c*i) \<noteq> - ?N i e + real j"
     and pi: "real (c*i) + ?N i e > 0" and cp': "real c >0"
     by simp+
   let ?fe = "floor (?N i e)"
-  from pi cp have th:"(real i +?N i e / real c)*real c > 0" by (simp add: ring_simps)
+  from pi cp have th:"(real i +?N i e / real c)*real c > 0" by (simp add: algebra_simps)
   from pi ei[simplified isint_iff] have "real (c*i + ?fe) > real (0::int)" by simp
   hence pi': "c*i + ?fe > 0" by (simp only: real_of_int_less_iff[symmetric])
   have "real (c*i) + ?N i e > real (c*d) \<or> real (c*i) + ?N i e \<le> real (c*d)" by auto
   moreover
   {assume "real (c*i) + ?N i e > real (c*d)" hence ?case
-      by (simp add: ring_simps 
+      by (simp add: algebra_simps 
 	numbound0_I[OF nb,where bs="bs" and b="real i - real d" and b'="real i"])} 
   moreover 
   {assume H:"real (c*i) + ?N i e \<le> real (c*d)"
@@ -3337,7 +3337,7 @@
     hence pid: "c*i + ?fe \<le> c*d" by (simp only: real_of_int_le_iff)
     with pi' have "\<exists> j1\<in> {1 .. c*d}. c*i + ?fe = j1" by auto
     hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = - ?N i e + real j1" 
-      by (simp only: diff_def[symmetric] real_of_int_mult real_of_int_add real_of_int_inject[symmetric] ei[simplified isint_iff] ring_simps)
+      by (simp only: diff_def[symmetric] real_of_int_mult real_of_int_add real_of_int_inject[symmetric] ei[simplified isint_iff] algebra_simps)
     with nob  have ?case by blast }
   ultimately show ?case by blast
 next
@@ -3346,13 +3346,13 @@
     and pi: "real (c*i) + ?N i e \<ge> 0" and cp': "real c >0"
     by simp+
   let ?fe = "floor (?N i e)"
-  from pi cp have th:"(real i +?N i e / real c)*real c \<ge> 0" by (simp add: ring_simps)
+  from pi cp have th:"(real i +?N i e / real c)*real c \<ge> 0" by (simp add: algebra_simps)
   from pi ei[simplified isint_iff] have "real (c*i + ?fe) \<ge> real (0::int)" by simp
   hence pi': "c*i + 1 + ?fe \<ge> 1" by (simp only: real_of_int_le_iff[symmetric])
   have "real (c*i) + ?N i e \<ge> real (c*d) \<or> real (c*i) + ?N i e < real (c*d)" by auto
   moreover
   {assume "real (c*i) + ?N i e \<ge> real (c*d)" hence ?case
-      by (simp add: ring_simps 
+      by (simp add: algebra_simps 
 	numbound0_I[OF nb,where bs="bs" and b="real i - real d" and b'="real i"])} 
   moreover 
   {assume H:"real (c*i) + ?N i e < real (c*d)"
@@ -3360,9 +3360,9 @@
     hence pid: "c*i + 1 + ?fe \<le> c*d" by (simp only: real_of_int_le_iff)
     with pi' have "\<exists> j1\<in> {1 .. c*d}. c*i + 1+ ?fe = j1" by auto
     hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) + 1= - ?N i e + real j1"
-      by (simp only: diff_def[symmetric] real_of_int_mult real_of_int_add real_of_int_inject[symmetric] ei[simplified isint_iff] ring_simps real_of_one) 
+      by (simp only: diff_def[symmetric] real_of_int_mult real_of_int_add real_of_int_inject[symmetric] ei[simplified isint_iff] algebra_simps real_of_one) 
     hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = (- ?N i e + real j1) - 1"
-      by (simp only: ring_simps diff_def[symmetric])
+      by (simp only: algebra_simps diff_def[symmetric])
         hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = - 1 - ?N i e + real j1"
 	  by (simp only: add_ac diff_def)
     with nob  have ?case by blast }
@@ -3383,10 +3383,10 @@
       using int_rdvd_iff[where i="j" and t="(c*i - c*d + floor ?e)",symmetric, simplified]
       ie by simp
     also have "\<dots> = (real j rdvd real (c*(i - d)) + ?e)" 
-      using ie by (simp add:ring_simps)
+      using ie by (simp add:algebra_simps)
     finally show ?case 
       using numbound0_I[OF bn,where b="real i - real d" and b'="real i" and bs="bs"] p 
-      by (simp add: ring_simps)
+      by (simp add: algebra_simps)
 next
   case (10 j c e)   hence p: "\<not> (real j rdvd real (c*i) + ?N i e)" (is "?p x") and cp: "c > 0" and bn:"numbound0 e"  by simp+
     let ?e = "Inum (real i # bs) e"
@@ -3403,10 +3403,10 @@
       using int_rdvd_iff[where i="j" and t="(c*i - c*d + floor ?e)",symmetric, simplified]
       ie by simp
     also have "\<dots> = Not (real j rdvd real (c*(i - d)) + ?e)" 
-      using ie by (simp add:ring_simps)
+      using ie by (simp add:algebra_simps)
     finally show ?case 
       using numbound0_I[OF bn,where b="real i - real d" and b'="real i" and bs="bs"] p 
-      by (simp add: ring_simps)
+      by (simp add: algebra_simps)
 qed(auto simp add: numbound0_I[where bs="bs" and b="real i - real d" and b'="real i"] nth_pos2)
 
 lemma \<sigma>_nb: assumes lp: "iszlfm p (a#bs)" and nb: "numbound0 t"
@@ -3649,7 +3649,7 @@
 	from H 
 	have "?I (?p (p',n',s') j) \<longrightarrow> 
 	  (((?N ?nxs \<ge> real ?l) \<and> (?N ?nxs < real (?l + 1))) \<and> (?N a = ?N ?nxs ))" 
-	  by (simp add: fp_def np ring_simps numsub numadd numfloor)
+	  by (simp add: fp_def np algebra_simps numsub numadd numfloor)
 	also have "\<dots> \<longrightarrow> ((floor (?N ?nxs) = ?l) \<and> (?N a = ?N ?nxs ))"
 	  using floor_int_eq[where x="?N ?nxs" and n="?l"] by simp
 	moreover
@@ -3675,7 +3675,7 @@
 	from H 
 	have "?I (?p (p',n',s') j) \<longrightarrow> 
 	  (((?N ?nxs \<ge> real ?l) \<and> (?N ?nxs < real (?l + 1))) \<and> (?N a = ?N ?nxs ))" 
-	  by (simp add: np fp_def ring_simps numneg numfloor numadd numsub)
+	  by (simp add: np fp_def algebra_simps numneg numfloor numadd numsub)
 	also have "\<dots> \<longrightarrow> ((floor (?N ?nxs) = ?l) \<and> (?N a = ?N ?nxs ))"
 	  using floor_int_eq[where x="?N ?nxs" and n="?l"] by simp
 	moreover
@@ -3691,7 +3691,7 @@
   apply (erule_tac x = "(aa, aaa, ba)" in ballE) apply simp_all
   apply (erule_tac x = "(ab, ac, baa)" in ballE) apply simp_all
   done
-qed (auto simp add: Let_def split_def ring_simps conj_rl)
+qed (auto simp add: Let_def split_def algebra_simps conj_rl)
 
 lemma real_in_int_intervals: 
   assumes xb: "real m \<le> x \<and> x < real ((n::int) + 1)"
@@ -3795,7 +3795,7 @@
     hence "\<exists> j\<in> {0 .. n}. 0 \<le> real n *x + ?N s - ?N (Floor s) - real j \<and> real n *x + ?N s - ?N (Floor s) - real (j+1) < 0"
       by(simp only: myl[rule_format, where b="real n * x + Inum (x # bs) s - Inum (x # bs) (Floor s)"] less_iff_diff_less_0[where a="real n *x + ?N s - ?N (Floor s)"]) 
     hence "\<exists> j\<in> {0.. n}. ?I (?p (p,n,s) j)"
-      using pns by (simp add: fp_def np ring_simps numsub numadd)
+      using pns by (simp add: fp_def np algebra_simps numsub numadd)
     then obtain "j" where j_def: "j\<in> {0 .. n} \<and> ?I (?p (p,n,s) j)" by blast
     hence "\<exists>x \<in> {?p (p,n,s) j |j. 0\<le> j \<and> j \<le> n }. ?I x" by auto
     hence ?case using pns 
@@ -3813,7 +3813,7 @@
       have "real n *x + ?N s \<ge> real n + ?N s" by simp 
       moreover from real_of_int_floor_le[where r="?N s"]  have "real n + ?N s \<ge> real n + ?N (Floor s)" by simp
       ultimately have "real n *x + ?N s \<ge> ?N (Floor s) + real n" 
-	by (simp only: ring_simps)}
+	by (simp only: algebra_simps)}
     ultimately have "?N (Floor s) + real n \<le> real n *x + ?N s\<and> real n *x + ?N s < ?N (Floor s) + real (1::int)" by simp
     hence th: "real n \<le> real n *x + ?N s - ?N (Floor s) \<and> real n *x + ?N s - ?N (Floor s) < real (1::int)" by simp
     have th1: "\<forall> (a::real). (- a > 0) = (a < 0)" by auto
@@ -3919,7 +3919,7 @@
   fix a n s
   assume H: "?N a = ?N (CN 0 n s)"
   show "?I (lt n s) = ?I (Lt a)" using H by (cases "n=0", (simp add: lt_def))
-  (cases "n > 0", simp_all add: lt_def ring_simps myless[rule_format, where b="0"])
+  (cases "n > 0", simp_all add: lt_def algebra_simps myless[rule_format, where b="0"])
 qed
 
 lemma lt_l: "isrlfm (rsplit lt a)"
@@ -3931,7 +3931,7 @@
   fix a n s
   assume H: "?N a = ?N (CN 0 n s)"
   show "?I (le n s) = ?I (Le a)" using H by (cases "n=0", (simp add: le_def))
-  (cases "n > 0", simp_all add: le_def ring_simps myl[rule_format, where b="0"])
+  (cases "n > 0", simp_all add: le_def algebra_simps myl[rule_format, where b="0"])
 qed
 
 lemma le_l: "isrlfm (rsplit le a)"
@@ -3943,7 +3943,7 @@
   fix a n s
   assume H: "?N a = ?N (CN 0 n s)"
   show "?I (gt n s) = ?I (Gt a)" using H by (cases "n=0", (simp add: gt_def))
-  (cases "n > 0", simp_all add: gt_def ring_simps myless[rule_format, where b="0"])
+  (cases "n > 0", simp_all add: gt_def algebra_simps myless[rule_format, where b="0"])
 qed
 lemma gt_l: "isrlfm (rsplit gt a)"
   by (rule rsplit_l[where f="gt" and a="a"], auto simp add: gt_def) 
@@ -3954,7 +3954,7 @@
   fix a n s 
   assume H: "?N a = ?N (CN 0 n s)"
   show "?I (ge n s) = ?I (Ge a)" using H by (cases "n=0", (simp add: ge_def))
-  (cases "n > 0", simp_all add: ge_def ring_simps myl[rule_format, where b="0"])
+  (cases "n > 0", simp_all add: ge_def algebra_simps myl[rule_format, where b="0"])
 qed
 lemma ge_l: "isrlfm (rsplit ge a)"
   by (rule rsplit_l[where f="ge" and a="a"], auto simp add: ge_def) 
@@ -3964,7 +3964,7 @@
 proof(clarify)
   fix a n s 
   assume H: "?N a = ?N (CN 0 n s)"
-  show "?I (eq n s) = ?I (Eq a)" using H by (auto simp add: eq_def ring_simps)
+  show "?I (eq n s) = ?I (Eq a)" using H by (auto simp add: eq_def algebra_simps)
 qed
 lemma eq_l: "isrlfm (rsplit eq a)"
   by (rule rsplit_l[where f="eq" and a="a"], auto simp add: eq_def) 
@@ -3974,7 +3974,7 @@
 proof(clarify)
   fix a n s bs
   assume H: "?N a = ?N (CN 0 n s)"
-  show "?I (neq n s) = ?I (NEq a)" using H by (auto simp add: neq_def ring_simps)
+  show "?I (neq n s) = ?I (NEq a)" using H by (auto simp add: neq_def algebra_simps)
 qed
 
 lemma neq_l: "isrlfm (rsplit neq a)"
@@ -4012,10 +4012,10 @@
   also have "\<dots> = (?DE \<and> ?a \<ge> 0 \<and> ?a < n)" by(simp only: small_le[OF ss0 ss1] small_lt[OF ss0 ss1])
   also have "\<dots> = (?DE \<and> (\<exists> j\<in> {0 .. (n - 1)}. ?a = j))" by simp
   also have "\<dots> = (?DE \<and> (\<exists> j\<in> {0 .. (n - 1)}. real (\<lfloor>real n * u - s\<rfloor>) = real j - real \<lfloor>s\<rfloor> ))"
-    by (simp only: ring_simps real_of_int_diff[symmetric] real_of_int_inject del: real_of_int_diff)
+    by (simp only: algebra_simps real_of_int_diff[symmetric] real_of_int_inject del: real_of_int_diff)
   also have "\<dots> = ((\<exists> j\<in> {0 .. (n - 1)}. real n * u - s = real j - real \<lfloor>s\<rfloor> \<and> real i rdvd real n * u - s))" using int_rdvd_iff[where i="i" and t="\<lfloor>real n * u - s\<rfloor>"]
     by (auto cong: conj_cong)
-  also have "\<dots> = ?rhs" by(simp cong: conj_cong) (simp add: ring_simps )
+  also have "\<dots> = ?rhs" by(simp cong: conj_cong) (simp add: algebra_simps )
   finally show ?thesis .
 qed
 
@@ -4038,7 +4038,7 @@
   from foldr_disj_map[where xs="iupt(0,n - 1)" and bs="x#bs" and f="?f"]
   have "Ifm (x#bs) (DVDJ i n s) = (\<exists> j\<in> {0 .. (n - 1)}. Ifm (x#bs) (?f j))" 
     by (simp add: iupt_set np DVDJ_def del: iupt.simps)
-  also have "\<dots> = (\<exists> j\<in> {0 .. (n - 1)}. real n * x = (- ?s) - real (floor (- ?s)) + real j \<and> real i rdvd real (j - floor (- ?s)))" by (simp add: ring_simps diff_def[symmetric])
+  also have "\<dots> = (\<exists> j\<in> {0 .. (n - 1)}. real n * x = (- ?s) - real (floor (- ?s)) + real j \<and> real i rdvd real (j - floor (- ?s)))" by (simp add: algebra_simps diff_def[symmetric])
   also from rdvd01_cs[OF xp x1 np, where i="i" and s="-?s"] 
   have "\<dots> = (real i rdvd real n * x - (-?s))" by simp
   finally show ?thesis by simp
@@ -4053,7 +4053,7 @@
   from foldr_conj_map[where xs="iupt(0,n - 1)" and bs="x#bs" and f="?f"]
   have "Ifm (x#bs) (NDVDJ i n s) = (\<forall> j\<in> {0 .. (n - 1)}. Ifm (x#bs) (?f j))" 
     by (simp add: iupt_set np NDVDJ_def del: iupt.simps)
-  also have "\<dots> = (\<not> (\<exists> j\<in> {0 .. (n - 1)}. real n * x = (- ?s) - real (floor (- ?s)) + real j \<and> real i rdvd real (j - floor (- ?s))))" by (simp add: ring_simps diff_def[symmetric])
+  also have "\<dots> = (\<not> (\<exists> j\<in> {0 .. (n - 1)}. real n * x = (- ?s) - real (floor (- ?s)) + real j \<and> real i rdvd real (j - floor (- ?s))))" by (simp add: algebra_simps diff_def[symmetric])
   also from rdvd01_cs[OF xp x1 np, where i="i" and s="-?s"] 
   have "\<dots> = (\<not> (real i rdvd real n * x - (-?s)))" by simp
   finally show ?thesis by simp
@@ -4652,40 +4652,40 @@
     using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
   also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) < 0)"
     by (simp only: pos_less_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" 
-      and b="0", simplified divide_zero_left]) (simp only: ring_simps)
+      and b="0", simplified divide_zero_left]) (simp only: algebra_simps)
   also have "\<dots> = (real c *?t + ?n* (?N x e) < 0)"
     using np by simp 
-  finally show ?case using nbt nb by (simp add: ring_simps)
+  finally show ?case using nbt nb by (simp add: algebra_simps)
 next
   case (6 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+
   have "?I ?u (Le (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \<le> 0)"
     using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
   also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) \<le> 0)"
     by (simp only: pos_le_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" 
-      and b="0", simplified divide_zero_left]) (simp only: ring_simps)
+      and b="0", simplified divide_zero_left]) (simp only: algebra_simps)
   also have "\<dots> = (real c *?t + ?n* (?N x e) \<le> 0)"
     using np by simp 
-  finally show ?case using nbt nb by (simp add: ring_simps)
+  finally show ?case using nbt nb by (simp add: algebra_simps)
 next
   case (7 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+
   have "?I ?u (Gt (CN 0 c e)) = (real c *(?t/?n) + (?N x e) > 0)"
     using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
   also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) > 0)"
     by (simp only: pos_divide_less_eq[OF np, where a="real c *(?t/?n) + (?N x e)" 
-      and b="0", simplified divide_zero_left]) (simp only: ring_simps)
+      and b="0", simplified divide_zero_left]) (simp only: algebra_simps)
   also have "\<dots> = (real c *?t + ?n* (?N x e) > 0)"
     using np by simp 
-  finally show ?case using nbt nb by (simp add: ring_simps)
+  finally show ?case using nbt nb by (simp add: algebra_simps)
 next
   case (8 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+
   have "?I ?u (Ge (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \<ge> 0)"
     using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
   also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) \<ge> 0)"
     by (simp only: pos_divide_le_eq[OF np, where a="real c *(?t/?n) + (?N x e)" 
-      and b="0", simplified divide_zero_left]) (simp only: ring_simps)
+      and b="0", simplified divide_zero_left]) (simp only: algebra_simps)
   also have "\<dots> = (real c *?t + ?n* (?N x e) \<ge> 0)"
     using np by simp 
-  finally show ?case using nbt nb by (simp add: ring_simps)
+  finally show ?case using nbt nb by (simp add: algebra_simps)
 next
   case (3 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+
   from np have np: "real n \<noteq> 0" by simp
@@ -4693,10 +4693,10 @@
     using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
   also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) = 0)"
     by (simp only: nonzero_eq_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" 
-      and b="0", simplified divide_zero_left]) (simp only: ring_simps)
+      and b="0", simplified divide_zero_left]) (simp only: algebra_simps)
   also have "\<dots> = (real c *?t + ?n* (?N x e) = 0)"
     using np by simp 
-  finally show ?case using nbt nb by (simp add: ring_simps)
+  finally show ?case using nbt nb by (simp add: algebra_simps)
 next
   case (4 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+
   from np have np: "real n \<noteq> 0" by simp
@@ -4704,10 +4704,10 @@
     using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
   also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) \<noteq> 0)"
     by (simp only: nonzero_eq_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" 
-      and b="0", simplified divide_zero_left]) (simp only: ring_simps)
+      and b="0", simplified divide_zero_left]) (simp only: algebra_simps)
   also have "\<dots> = (real c *?t + ?n* (?N x e) \<noteq> 0)"
     using np by simp 
-  finally show ?case using nbt nb by (simp add: ring_simps)
+  finally show ?case using nbt nb by (simp add: algebra_simps)
 qed(simp_all add: nbt numbound0_I[where bs ="bs" and b="(Inum (x#bs) t)/ real n" and b'="x"] nth_pos2)
 
 lemma \<Upsilon>_l:
@@ -4758,7 +4758,7 @@
 using lp px noS
 proof (induct p rule: isrlfm.induct)
   case (5 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+
-    from prems have "x * real c + ?N x e < 0" by (simp add: ring_simps)
+    from prems have "x * real c + ?N x e < 0" by (simp add: algebra_simps)
     hence pxc: "x < (- ?N x e) / real c" 
       by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="-?N x e"])
     from prems have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
@@ -4767,7 +4767,7 @@
     moreover {assume y: "y < (-?N x e)/ real c"
       hence "y * real c < - ?N x e"
 	by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric])
-      hence "real c * y + ?N x e < 0" by (simp add: ring_simps)
+      hence "real c * y + ?N x e < 0" by (simp add: algebra_simps)
       hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
     moreover {assume y: "y > (- ?N x e) / real c" 
       with yu have eu: "u > (- ?N x e) / real c" by auto
@@ -4777,7 +4777,7 @@
     ultimately show ?case by blast
 next
   case (6 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp +
-    from prems have "x * real c + ?N x e \<le> 0" by (simp add: ring_simps)
+    from prems have "x * real c + ?N x e \<le> 0" by (simp add: algebra_simps)
     hence pxc: "x \<le> (- ?N x e) / real c" 
       by (simp only: pos_le_divide_eq[OF cp, where a="x" and b="-?N x e"])
     from prems have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
@@ -4786,7 +4786,7 @@
     moreover {assume y: "y < (-?N x e)/ real c"
       hence "y * real c < - ?N x e"
 	by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric])
-      hence "real c * y + ?N x e < 0" by (simp add: ring_simps)
+      hence "real c * y + ?N x e < 0" by (simp add: algebra_simps)
       hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
     moreover {assume y: "y > (- ?N x e) / real c" 
       with yu have eu: "u > (- ?N x e) / real c" by auto
@@ -4796,7 +4796,7 @@
     ultimately show ?case by blast
 next
   case (7 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+
-    from prems have "x * real c + ?N x e > 0" by (simp add: ring_simps)
+    from prems have "x * real c + ?N x e > 0" by (simp add: algebra_simps)
     hence pxc: "x > (- ?N x e) / real c" 
       by (simp only: pos_divide_less_eq[OF cp, where a="x" and b="-?N x e"])
     from prems have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
@@ -4805,7 +4805,7 @@
     moreover {assume y: "y > (-?N x e)/ real c"
       hence "y * real c > - ?N x e"
 	by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric])
-      hence "real c * y + ?N x e > 0" by (simp add: ring_simps)
+      hence "real c * y + ?N x e > 0" by (simp add: algebra_simps)
       hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
     moreover {assume y: "y < (- ?N x e) / real c" 
       with ly have eu: "l < (- ?N x e) / real c" by auto
@@ -4815,7 +4815,7 @@
     ultimately show ?case by blast
 next
   case (8 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+
-    from prems have "x * real c + ?N x e \<ge> 0" by (simp add: ring_simps)
+    from prems have "x * real c + ?N x e \<ge> 0" by (simp add: algebra_simps)
     hence pxc: "x \<ge> (- ?N x e) / real c" 
       by (simp only: pos_divide_le_eq[OF cp, where a="x" and b="-?N x e"])
     from prems have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
@@ -4824,7 +4824,7 @@
     moreover {assume y: "y > (-?N x e)/ real c"
       hence "y * real c > - ?N x e"
 	by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric])
-      hence "real c * y + ?N x e > 0" by (simp add: ring_simps)
+      hence "real c * y + ?N x e > 0" by (simp add: algebra_simps)
       hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
     moreover {assume y: "y < (- ?N x e) / real c" 
       with ly have eu: "l < (- ?N x e) / real c" by auto
@@ -4835,7 +4835,7 @@
 next
   case (3 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+
     from cp have cnz: "real c \<noteq> 0" by simp
-    from prems have "x * real c + ?N x e = 0" by (simp add: ring_simps)
+    from prems have "x * real c + ?N x e = 0" by (simp add: algebra_simps)
     hence pxc: "x = (- ?N x e) / real c" 
       by (simp only: nonzero_eq_divide_eq[OF cnz, where a="x" and b="-?N x e"])
     from prems have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
@@ -4848,9 +4848,9 @@
     with ly yu have yne: "y \<noteq> - ?N x e / real c" by auto
     hence "y* real c \<noteq> -?N x e"      
       by (simp only: nonzero_eq_divide_eq[OF cnz, where a="y" and b="-?N x e"]) simp
-    hence "y* real c + ?N x e \<noteq> 0" by (simp add: ring_simps)
+    hence "y* real c + ?N x e \<noteq> 0" by (simp add: algebra_simps)
     thus ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] 
-      by (simp add: ring_simps)
+      by (simp add: algebra_simps)
 qed (auto simp add: nth_pos2 numbound0_I[where bs="bs" and b="y" and b'="x"])
 
 lemma finite_set_intervals:
@@ -5013,7 +5013,7 @@
 	by (simp add: mult_commute)
       from tnb snb have st_nb: "numbound0 ?st" by simp
       have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
-	using mnp mp np by (simp add: ring_simps add_divide_distrib)
+	using mnp mp np by (simp add: algebra_simps add_divide_distrib)
       from \<upsilon>_I[OF lp mnp st_nb, where x="x" and bs="bs"] 
       have "?I x (\<upsilon> p (?st,2*n*m)) = ?I ((?N t / real n + ?N s / real m) /2) p" by (simp only: st[symmetric])}
     with rinf_\<Upsilon>[OF lp nmi npi px] have "?F" by blast hence "?D" by blast}
@@ -5082,7 +5082,7 @@
 
 lemma exsplitnum: 
   "Inum (x#y#bs) (exsplitnum t) = Inum ((x+y) #bs) t"
-  by(induct t rule: exsplitnum.induct) (simp_all add: ring_simps)
+  by(induct t rule: exsplitnum.induct) (simp_all add: algebra_simps)
 
 lemma exsplit: 
   assumes qfp: "qfree p"
@@ -5173,14 +5173,14 @@
   from Ul th have mnz: "m \<noteq> 0" by auto
   from Ul th have  nnz: "n \<noteq> 0" by auto  
   have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
-   using mnz nnz by (simp add: ring_simps add_divide_distrib)
+   using mnz nnz by (simp add: algebra_simps add_divide_distrib)
  
   thus "(real m *  Inum (x # bs) t + real n * Inum (x # bs) s) /
        (2 * real n * real m)
        \<in> (\<lambda>((t, n), s, m).
              (Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2) `
          (set U \<times> set U)"using mnz nnz th  
-    apply (auto simp add: th add_divide_distrib ring_simps split_def image_def)
+    apply (auto simp add: th add_divide_distrib algebra_simps split_def image_def)
     by (rule_tac x="(s,m)" in bexI,simp_all) 
   (rule_tac x="(t,n)" in bexI,simp_all)
 next
@@ -5191,7 +5191,7 @@
   from Ul smU have mnz: "m \<noteq> 0" by auto
   from Ul tnU have  nnz: "n \<noteq> 0" by auto  
   have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
-   using mnz nnz by (simp add: ring_simps add_divide_distrib)
+   using mnz nnz by (simp add: algebra_simps add_divide_distrib)
  let ?P = "\<lambda> (t',n') (s',m'). (Inum (x # bs) t / real n + Inum (x # bs) s / real m)/2 = (Inum (x # bs) t' / real n' + Inum (x # bs) s' / real m')/2"
  have Pc:"\<forall> a b. ?P a b = ?P b a"
    by auto
@@ -5204,7 +5204,7 @@
  from ts'_U Up have mnz': "m' \<noteq> 0" and nnz': "n'\<noteq> 0" by auto
  let ?st' = "Add (Mul m' t') (Mul n' s')"
    have st': "(?N t' / real n' + ?N s' / real m')/2 = ?N ?st' / real (2*n'*m')"
-   using mnz' nnz' by (simp add: ring_simps add_divide_distrib)
+   using mnz' nnz' by (simp add: algebra_simps add_divide_distrib)
  from Pts' have 
    "(Inum (x # bs) t / real n + Inum (x # bs) s / real m)/2 = (Inum (x # bs) t' / real n' + Inum (x # bs) s' / real m')/2" by simp
  also have "\<dots> = ((\<lambda>(t, n). Inum (x # bs) t / real n) ((\<lambda>((t, n), s, m). (Add (Mul m t) (Mul n s), 2 * n * m)) ((t',n'),(s',m'))))" by (simp add: st')
@@ -5234,7 +5234,7 @@
       by (simp add: mult_commute real_of_int_mult[symmetric] del: real_of_int_mult)
     from tnb snb have stnb: "numbound0 ?st" by simp
   have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
-   using mp np by (simp add: ring_simps add_divide_distrib)
+   using mp np by (simp add: algebra_simps add_divide_distrib)
   from tnU smU UU' have "?g ((t,n),(s,m)) \<in> ?f ` U'" by blast
   hence "\<exists> (t',n') \<in> U'. ?g ((t,n),(s,m)) = ?f (t',n')"
     by auto (rule_tac x="(a,b)" in bexI, auto)
@@ -5262,7 +5262,7 @@
       by (simp add: mult_commute real_of_int_mult[symmetric] del: real_of_int_mult)
     from tnb snb have stnb: "numbound0 ?st" by simp
   have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
-   using mp np by (simp add: ring_simps add_divide_distrib)
+   using mp np by (simp add: algebra_simps add_divide_distrib)
   from U' tnU' have tnb': "numbound0 t'" and np': "real n' > 0" by auto
   from \<upsilon>_I[OF lp np' tnb', where bs="bs" and x="x",simplified th[simplified split_def fst_conv snd_conv] st] Pt'
   have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real (2 * n * m) # bs) p" by simp
--- a/src/HOL/ex/Numeral.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/ex/Numeral.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -699,7 +699,7 @@
 begin
 
 subclass semiring_1_minus
-  proof qed (simp_all add: ring_simps)
+  proof qed (simp_all add: algebra_simps)
 
 lemma Dig_zero_minus_of_num [numeral]:
   "0 - of_num n = - of_num n"
@@ -783,7 +783,7 @@
   "sub (Dig1 m) (Dig1 n) = dup (sub m n)"
   "sub (Dig1 m) (Dig0 n) = dup (sub m n) + 1"
   "sub (Dig0 m) (Dig1 n) = dup (sub m n) - 1"
-  apply (simp_all add: dup_def ring_simps)
+  apply (simp_all add: dup_def algebra_simps)
   apply (simp_all add: of_num_plus Dig_one_plus_dec)[4]
   apply (simp_all add: of_num.simps)
   done
--- a/src/HOL/ex/Quickcheck.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/ex/Quickcheck.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -200,7 +200,7 @@
             in
               lthy
               |> LocalTheory.theory (Code.del_eqns c
-                   #> PureThy.add_thm ((fst (dest_Free random') ^ "_code", thm), [Thm.kind_internal])
+                   #> PureThy.add_thm ((Binding.name (fst (dest_Free random') ^ "_code"), thm), [Thm.kind_internal])
                    #-> Code.add_eqn)
             end;
         in
--- a/src/HOL/ex/ROOT.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/ex/ROOT.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -63,7 +63,6 @@
   "Dense_Linear_Order_Ex",
   "PresburgerEx",
   "Reflected_Presburger",
-  "Reflection",
   "ReflectionEx",
   "BinEx",
   "Sqrt",
--- a/src/HOL/ex/ReflectedFerrack.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/ex/ReflectedFerrack.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -532,7 +532,7 @@
 next
   case (2 n c t)  hence gd: "g dvd c" by simp
   from gp have gnz: "g \<noteq> 0" by simp
-  from prems show ?case by (simp add: real_of_int_div[OF gnz gd] ring_simps)
+  from prems show ?case by (simp add: real_of_int_div[OF gnz gd] algebra_simps)
 qed (auto simp add: numgcd_def gp)
 consts ismaxcoeff:: "num \<Rightarrow> int \<Rightarrow> bool"
 recdef ismaxcoeff "measure size"
@@ -636,7 +636,7 @@
 lemma numadd[simp]: "Inum bs (numadd (t,s)) = Inum bs (Add t s)"
 apply (induct t s rule: numadd.induct, simp_all add: Let_def)
 apply (case_tac "c1+c2 = 0",case_tac "n1 \<le> n2", simp_all)
-apply (case_tac "n1 = n2", simp_all add: ring_simps)
+apply (case_tac "n1 = n2", simp_all add: algebra_simps)
 by (simp only: left_distrib[symmetric],simp)
 
 lemma numadd_nb[simp]: "\<lbrakk> numbound0 t ; numbound0 s\<rbrakk> \<Longrightarrow> numbound0 (numadd (t,s))"
@@ -648,7 +648,7 @@
   "nummul t = (\<lambda> i. Mul i t)"
 
 lemma nummul[simp]: "\<And> i. Inum bs (nummul t i) = Inum bs (Mul i t)"
-by (induct t rule: nummul.induct, auto simp add: ring_simps)
+by (induct t rule: nummul.induct, auto simp add: algebra_simps)
 
 lemma nummul_nb[simp]: "\<And> i. numbound0 t \<Longrightarrow> numbound0 (nummul t i)"
 by (induct t rule: nummul.induct, auto )
@@ -1001,10 +1001,10 @@
     by(cases "rsplit0 a",auto simp add: Let_def split_def)
   have "Inum bs ((split (CN 0)) (rsplit0 (Add a b))) = 
     Inum bs ((split (CN 0)) ?sa)+Inum bs ((split (CN 0)) ?sb)"
-    by (simp add: Let_def split_def ring_simps)
+    by (simp add: Let_def split_def algebra_simps)
   also have "\<dots> = Inum bs a + Inum bs b" using prems by (cases "rsplit0 a", simp_all)
   finally show ?case using nb by simp 
-qed(auto simp add: Let_def split_def ring_simps , simp add: right_distrib[symmetric])
+qed(auto simp add: Let_def split_def algebra_simps , simp add: right_distrib[symmetric])
 
     (* Linearize a formula*)
 definition
@@ -1396,40 +1396,40 @@
     using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
   also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) < 0)"
     by (simp only: pos_less_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" 
-      and b="0", simplified divide_zero_left]) (simp only: ring_simps)
+      and b="0", simplified divide_zero_left]) (simp only: algebra_simps)
   also have "\<dots> = (real c *?t + ?n* (?N x e) < 0)"
     using np by simp 
-  finally show ?case using nbt nb by (simp add: ring_simps)
+  finally show ?case using nbt nb by (simp add: algebra_simps)
 next
   case (6 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+
   have "?I ?u (Le (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \<le> 0)"
     using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
   also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) \<le> 0)"
     by (simp only: pos_le_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" 
-      and b="0", simplified divide_zero_left]) (simp only: ring_simps)
+      and b="0", simplified divide_zero_left]) (simp only: algebra_simps)
   also have "\<dots> = (real c *?t + ?n* (?N x e) \<le> 0)"
     using np by simp 
-  finally show ?case using nbt nb by (simp add: ring_simps)
+  finally show ?case using nbt nb by (simp add: algebra_simps)
 next
   case (7 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+
   have "?I ?u (Gt (CN 0 c e)) = (real c *(?t/?n) + (?N x e) > 0)"
     using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
   also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) > 0)"
     by (simp only: pos_divide_less_eq[OF np, where a="real c *(?t/?n) + (?N x e)" 
-      and b="0", simplified divide_zero_left]) (simp only: ring_simps)
+      and b="0", simplified divide_zero_left]) (simp only: algebra_simps)
   also have "\<dots> = (real c *?t + ?n* (?N x e) > 0)"
     using np by simp 
-  finally show ?case using nbt nb by (simp add: ring_simps)
+  finally show ?case using nbt nb by (simp add: algebra_simps)
 next
   case (8 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+
   have "?I ?u (Ge (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \<ge> 0)"
     using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
   also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) \<ge> 0)"
     by (simp only: pos_divide_le_eq[OF np, where a="real c *(?t/?n) + (?N x e)" 
-      and b="0", simplified divide_zero_left]) (simp only: ring_simps)
+      and b="0", simplified divide_zero_left]) (simp only: algebra_simps)
   also have "\<dots> = (real c *?t + ?n* (?N x e) \<ge> 0)"
     using np by simp 
-  finally show ?case using nbt nb by (simp add: ring_simps)
+  finally show ?case using nbt nb by (simp add: algebra_simps)
 next
   case (3 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+
   from np have np: "real n \<noteq> 0" by simp
@@ -1437,10 +1437,10 @@
     using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
   also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) = 0)"
     by (simp only: nonzero_eq_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" 
-      and b="0", simplified divide_zero_left]) (simp only: ring_simps)
+      and b="0", simplified divide_zero_left]) (simp only: algebra_simps)
   also have "\<dots> = (real c *?t + ?n* (?N x e) = 0)"
     using np by simp 
-  finally show ?case using nbt nb by (simp add: ring_simps)
+  finally show ?case using nbt nb by (simp add: algebra_simps)
 next
   case (4 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+
   from np have np: "real n \<noteq> 0" by simp
@@ -1448,10 +1448,10 @@
     using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
   also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) \<noteq> 0)"
     by (simp only: nonzero_eq_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" 
-      and b="0", simplified divide_zero_left]) (simp only: ring_simps)
+      and b="0", simplified divide_zero_left]) (simp only: algebra_simps)
   also have "\<dots> = (real c *?t + ?n* (?N x e) \<noteq> 0)"
     using np by simp 
-  finally show ?case using nbt nb by (simp add: ring_simps)
+  finally show ?case using nbt nb by (simp add: algebra_simps)
 qed(simp_all add: nbt numbound0_I[where bs ="bs" and b="(Inum (x#bs) t)/ real n" and b'="x"] nth_pos2)
 
 lemma uset_l:
@@ -1502,7 +1502,7 @@
 using lp px noS
 proof (induct p rule: isrlfm.induct)
   case (5 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+
-    from prems have "x * real c + ?N x e < 0" by (simp add: ring_simps)
+    from prems have "x * real c + ?N x e < 0" by (simp add: algebra_simps)
     hence pxc: "x < (- ?N x e) / real c" 
       by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="-?N x e"])
     from prems have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
@@ -1511,7 +1511,7 @@
     moreover {assume y: "y < (-?N x e)/ real c"
       hence "y * real c < - ?N x e"
 	by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric])
-      hence "real c * y + ?N x e < 0" by (simp add: ring_simps)
+      hence "real c * y + ?N x e < 0" by (simp add: algebra_simps)
       hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
     moreover {assume y: "y > (- ?N x e) / real c" 
       with yu have eu: "u > (- ?N x e) / real c" by auto
@@ -1521,7 +1521,7 @@
     ultimately show ?case by blast
 next
   case (6 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp +
-    from prems have "x * real c + ?N x e \<le> 0" by (simp add: ring_simps)
+    from prems have "x * real c + ?N x e \<le> 0" by (simp add: algebra_simps)
     hence pxc: "x \<le> (- ?N x e) / real c" 
       by (simp only: pos_le_divide_eq[OF cp, where a="x" and b="-?N x e"])
     from prems have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
@@ -1530,7 +1530,7 @@
     moreover {assume y: "y < (-?N x e)/ real c"
       hence "y * real c < - ?N x e"
 	by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric])
-      hence "real c * y + ?N x e < 0" by (simp add: ring_simps)
+      hence "real c * y + ?N x e < 0" by (simp add: algebra_simps)
       hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
     moreover {assume y: "y > (- ?N x e) / real c" 
       with yu have eu: "u > (- ?N x e) / real c" by auto
@@ -1540,7 +1540,7 @@
     ultimately show ?case by blast
 next
   case (7 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+
-    from prems have "x * real c + ?N x e > 0" by (simp add: ring_simps)
+    from prems have "x * real c + ?N x e > 0" by (simp add: algebra_simps)
     hence pxc: "x > (- ?N x e) / real c" 
       by (simp only: pos_divide_less_eq[OF cp, where a="x" and b="-?N x e"])
     from prems have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
@@ -1549,7 +1549,7 @@
     moreover {assume y: "y > (-?N x e)/ real c"
       hence "y * real c > - ?N x e"
 	by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric])
-      hence "real c * y + ?N x e > 0" by (simp add: ring_simps)
+      hence "real c * y + ?N x e > 0" by (simp add: algebra_simps)
       hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
     moreover {assume y: "y < (- ?N x e) / real c" 
       with ly have eu: "l < (- ?N x e) / real c" by auto
@@ -1559,7 +1559,7 @@
     ultimately show ?case by blast
 next
   case (8 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+
-    from prems have "x * real c + ?N x e \<ge> 0" by (simp add: ring_simps)
+    from prems have "x * real c + ?N x e \<ge> 0" by (simp add: algebra_simps)
     hence pxc: "x \<ge> (- ?N x e) / real c" 
       by (simp only: pos_divide_le_eq[OF cp, where a="x" and b="-?N x e"])
     from prems have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
@@ -1568,7 +1568,7 @@
     moreover {assume y: "y > (-?N x e)/ real c"
       hence "y * real c > - ?N x e"
 	by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric])
-      hence "real c * y + ?N x e > 0" by (simp add: ring_simps)
+      hence "real c * y + ?N x e > 0" by (simp add: algebra_simps)
       hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
     moreover {assume y: "y < (- ?N x e) / real c" 
       with ly have eu: "l < (- ?N x e) / real c" by auto
@@ -1579,7 +1579,7 @@
 next
   case (3 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+
     from cp have cnz: "real c \<noteq> 0" by simp
-    from prems have "x * real c + ?N x e = 0" by (simp add: ring_simps)
+    from prems have "x * real c + ?N x e = 0" by (simp add: algebra_simps)
     hence pxc: "x = (- ?N x e) / real c" 
       by (simp only: nonzero_eq_divide_eq[OF cnz, where a="x" and b="-?N x e"])
     from prems have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
@@ -1592,9 +1592,9 @@
     with ly yu have yne: "y \<noteq> - ?N x e / real c" by auto
     hence "y* real c \<noteq> -?N x e"      
       by (simp only: nonzero_eq_divide_eq[OF cnz, where a="y" and b="-?N x e"]) simp
-    hence "y* real c + ?N x e \<noteq> 0" by (simp add: ring_simps)
+    hence "y* real c + ?N x e \<noteq> 0" by (simp add: algebra_simps)
     thus ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] 
-      by (simp add: ring_simps)
+      by (simp add: algebra_simps)
 qed (auto simp add: nth_pos2 numbound0_I[where bs="bs" and b="y" and b'="x"])
 
 lemma finite_set_intervals:
@@ -1757,7 +1757,7 @@
 	by (simp add: mult_commute)
       from tnb snb have st_nb: "numbound0 ?st" by simp
       have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
-	using mnp mp np by (simp add: ring_simps add_divide_distrib)
+	using mnp mp np by (simp add: algebra_simps add_divide_distrib)
       from usubst_I[OF lp mnp st_nb, where x="x" and bs="bs"] 
       have "?I x (usubst p (?st,2*n*m)) = ?I ((?N t / real n + ?N s / real m) /2) p" by (simp only: st[symmetric])}
     with rinf_uset[OF lp nmi npi px] have "?F" by blast hence "?D" by blast}
@@ -1801,14 +1801,14 @@
   from Ul th have mnz: "m \<noteq> 0" by auto
   from Ul th have  nnz: "n \<noteq> 0" by auto  
   have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
-   using mnz nnz by (simp add: ring_simps add_divide_distrib)
+   using mnz nnz by (simp add: algebra_simps add_divide_distrib)
  
   thus "(real m *  Inum (x # bs) t + real n * Inum (x # bs) s) /
        (2 * real n * real m)
        \<in> (\<lambda>((t, n), s, m).
              (Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2) `
          (set U \<times> set U)"using mnz nnz th  
-    apply (auto simp add: th add_divide_distrib ring_simps split_def image_def)
+    apply (auto simp add: th add_divide_distrib algebra_simps split_def image_def)
     by (rule_tac x="(s,m)" in bexI,simp_all) 
   (rule_tac x="(t,n)" in bexI,simp_all)
 next
@@ -1819,7 +1819,7 @@
   from Ul smU have mnz: "m \<noteq> 0" by auto
   from Ul tnU have  nnz: "n \<noteq> 0" by auto  
   have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
-   using mnz nnz by (simp add: ring_simps add_divide_distrib)
+   using mnz nnz by (simp add: algebra_simps add_divide_distrib)
  let ?P = "\<lambda> (t',n') (s',m'). (Inum (x # bs) t / real n + Inum (x # bs) s / real m)/2 = (Inum (x # bs) t' / real n' + Inum (x # bs) s' / real m')/2"
  have Pc:"\<forall> a b. ?P a b = ?P b a"
    by auto
@@ -1832,7 +1832,7 @@
  from ts'_U Up have mnz': "m' \<noteq> 0" and nnz': "n'\<noteq> 0" by auto
  let ?st' = "Add (Mul m' t') (Mul n' s')"
    have st': "(?N t' / real n' + ?N s' / real m')/2 = ?N ?st' / real (2*n'*m')"
-   using mnz' nnz' by (simp add: ring_simps add_divide_distrib)
+   using mnz' nnz' by (simp add: algebra_simps add_divide_distrib)
  from Pts' have 
    "(Inum (x # bs) t / real n + Inum (x # bs) s / real m)/2 = (Inum (x # bs) t' / real n' + Inum (x # bs) s' / real m')/2" by simp
  also have "\<dots> = ((\<lambda>(t, n). Inum (x # bs) t / real n) ((\<lambda>((t, n), s, m). (Add (Mul m t) (Mul n s), 2 * n * m)) ((t',n'),(s',m'))))" by (simp add: st')
@@ -1862,7 +1862,7 @@
       by (simp add: mult_commute real_of_int_mult[symmetric] del: real_of_int_mult)
     from tnb snb have stnb: "numbound0 ?st" by simp
   have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
-   using mp np by (simp add: ring_simps add_divide_distrib)
+   using mp np by (simp add: algebra_simps add_divide_distrib)
   from tnU smU UU' have "?g ((t,n),(s,m)) \<in> ?f ` U'" by blast
   hence "\<exists> (t',n') \<in> U'. ?g ((t,n),(s,m)) = ?f (t',n')"
     by auto (rule_tac x="(a,b)" in bexI, auto)
@@ -1890,7 +1890,7 @@
       by (simp add: mult_commute real_of_int_mult[symmetric] del: real_of_int_mult)
     from tnb snb have stnb: "numbound0 ?st" by simp
   have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
-   using mp np by (simp add: ring_simps add_divide_distrib)
+   using mp np by (simp add: algebra_simps add_divide_distrib)
   from U' tnU' have tnb': "numbound0 t'" and np': "real n' > 0" by auto
   from usubst_I[OF lp np' tnb', where bs="bs" and x="x",simplified th[simplified split_def fst_conv snd_conv] st] Pt'
   have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real (2 * n * m) # bs) p" by simp
--- a/src/HOL/ex/Reflected_Presburger.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/ex/Reflected_Presburger.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -437,7 +437,7 @@
 apply (induct t s rule: numadd.induct, simp_all add: Let_def)
 apply (case_tac "c1+c2 = 0",case_tac "n1 \<le> n2", simp_all)
  apply (case_tac "n1 = n2")
-  apply(simp_all add: ring_simps)
+  apply(simp_all add: algebra_simps)
 apply(simp add: left_distrib[symmetric])
 done
 
@@ -452,7 +452,7 @@
   | "nummul i t = Mul i t"
 
 lemma nummul: "\<And> i. Inum bs (nummul i t) = Inum bs (Mul i t)"
-by (induct t rule: nummul.induct, auto simp add: ring_simps numadd)
+by (induct t rule: nummul.induct, auto simp add: algebra_simps numadd)
 
 lemma nummul_nb: "\<And> i. numbound0 t \<Longrightarrow> numbound0 (nummul i t)"
 by (induct t rule: nummul.induct, auto simp add: numadd_nb)
@@ -912,7 +912,7 @@
   have Ia:"Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto 
   let ?N = "\<lambda> t. Inum (i#bs) t"
   from prems Ia nb  show ?case 
-    apply (auto simp add: Let_def split_def ring_simps) 
+    apply (auto simp add: Let_def split_def algebra_simps) 
     apply (cases "?r",auto)
     apply (case_tac nat, auto)
     done
@@ -925,7 +925,7 @@
   have Ia:"Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto 
   let ?N = "\<lambda> t. Inum (i#bs) t"
   from prems Ia nb  show ?case 
-    apply (auto simp add: Let_def split_def ring_simps) 
+    apply (auto simp add: Let_def split_def algebra_simps) 
     apply (cases "?r",auto)
     apply (case_tac nat, auto)
     done
@@ -938,7 +938,7 @@
   have Ia:"Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto 
   let ?N = "\<lambda> t. Inum (i#bs) t"
   from prems Ia nb  show ?case 
-    apply (auto simp add: Let_def split_def ring_simps) 
+    apply (auto simp add: Let_def split_def algebra_simps) 
     apply (cases "?r",auto)
     apply (case_tac nat, auto)
     done
@@ -951,7 +951,7 @@
   have Ia:"Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto 
   let ?N = "\<lambda> t. Inum (i#bs) t"
   from prems Ia nb  show ?case 
-    apply (auto simp add: Let_def split_def ring_simps) 
+    apply (auto simp add: Let_def split_def algebra_simps) 
     apply (cases "?r",auto)
     apply (case_tac nat, auto)
     done
@@ -964,7 +964,7 @@
   have Ia:"Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto 
   let ?N = "\<lambda> t. Inum (i#bs) t"
   from prems Ia nb  show ?case 
-    apply (auto simp add: Let_def split_def ring_simps) 
+    apply (auto simp add: Let_def split_def algebra_simps) 
     apply (cases "?r",auto)
     apply (case_tac nat, auto)
     done
@@ -977,7 +977,7 @@
   have Ia:"Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto 
   let ?N = "\<lambda> t. Inum (i#bs) t"
   from prems Ia nb  show ?case 
-    apply (auto simp add: Let_def split_def ring_simps) 
+    apply (auto simp add: Let_def split_def algebra_simps) 
     apply (cases "?r",auto)
     apply (case_tac nat, auto)
     done
@@ -996,7 +996,7 @@
   moreover
   {assume "?c=0" and "j\<noteq>0" hence ?case 
       using zsplit0_I[OF spl, where x="i" and bs="bs"] zdvd_abs1[where i="j"]
-    apply (auto simp add: Let_def split_def ring_simps) 
+    apply (auto simp add: Let_def split_def algebra_simps) 
     apply (cases "?r",auto)
     apply (case_tac nat, auto)
     done}
@@ -1027,7 +1027,7 @@
   moreover
   {assume "?c=0" and "j\<noteq>0" hence ?case 
       using zsplit0_I[OF spl, where x="i" and bs="bs"] zdvd_abs1[where i="j"]
-    apply (auto simp add: Let_def split_def ring_simps) 
+    apply (auto simp add: Let_def split_def algebra_simps) 
     apply (cases "?r",auto)
     apply (case_tac nat, auto)
     done}
@@ -1301,9 +1301,9 @@
       (is "?ri dvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri dvd ?rt")
       hence "\<exists> (l::int). ?rt = i * l" by (simp add: dvd_def)
       hence "\<exists> (l::int). c*x+ ?I x e = i*l+c*(k * i*di)" 
-	by (simp add: ring_simps di_def)
+	by (simp add: algebra_simps di_def)
       hence "\<exists> (l::int). c*x+ ?I x e = i*(l + c*k*di)"
-	by (simp add: ring_simps)
+	by (simp add: algebra_simps)
       hence "\<exists> (l::int). c*x+ ?I x e = i*l" by blast
       thus "i dvd c*x + Inum (x # bs) e" by (simp add: dvd_def) 
     next
@@ -1312,7 +1312,7 @@
       hence "\<exists> (l::int). c*x+?e = i*l" by (simp add: dvd_def)
       hence "\<exists> (l::int). c*x - c*(k*d) +?e = i*l - c*(k*d)" by simp
       hence "\<exists> (l::int). c*x - c*(k*d) +?e = i*l - c*(k*i*di)" by (simp add: di_def)
-      hence "\<exists> (l::int). c*x - c*(k*d) +?e = i*((l - c*k*di))" by (simp add: ring_simps)
+      hence "\<exists> (l::int). c*x - c*(k*d) +?e = i*((l - c*k*di))" by (simp add: algebra_simps)
       hence "\<exists> (l::int). c*x - c * (k*d) +?e = i*l"
 	by blast
       thus "i dvd c*x - c*(k*d) + Inum (x # bs) e" by (simp add: dvd_def)
@@ -1328,9 +1328,9 @@
       (is "?ri dvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri dvd ?rt")
       hence "\<exists> (l::int). ?rt = i * l" by (simp add: dvd_def)
       hence "\<exists> (l::int). c*x+ ?I x e = i*l+c*(k * i*di)" 
-	by (simp add: ring_simps di_def)
+	by (simp add: algebra_simps di_def)
       hence "\<exists> (l::int). c*x+ ?I x e = i*(l + c*k*di)"
-	by (simp add: ring_simps)
+	by (simp add: algebra_simps)
       hence "\<exists> (l::int). c*x+ ?I x e = i*l" by blast
       thus "i dvd c*x + Inum (x # bs) e" by (simp add: dvd_def) 
     next
@@ -1339,7 +1339,7 @@
       hence "\<exists> (l::int). c*x+?e = i*l" by (simp add: dvd_def)
       hence "\<exists> (l::int). c*x - c*(k*d) +?e = i*l - c*(k*d)" by simp
       hence "\<exists> (l::int). c*x - c*(k*d) +?e = i*l - c*(k*i*di)" by (simp add: di_def)
-      hence "\<exists> (l::int). c*x - c*(k*d) +?e = i*((l - c*k*di))" by (simp add: ring_simps)
+      hence "\<exists> (l::int). c*x - c*(k*d) +?e = i*((l - c*k*di))" by (simp add: algebra_simps)
       hence "\<exists> (l::int). c*x - c * (k*d) +?e = i*l"
 	by blast
       thus "i dvd c*x - c*(k*d) + Inum (x # bs) e" by (simp add: dvd_def)
@@ -1363,7 +1363,7 @@
     by (simp only: zdvd_zminus_iff)
   also have "\<dots> = (j dvd (c* (- x)) + ?e)"
     apply (simp only: minus_mult_right[symmetric] minus_mult_left[symmetric] diff_def zadd_ac zminus_zadd_distrib)
-    by (simp add: ring_simps)
+    by (simp add: algebra_simps)
   also have "\<dots> = Ifm bbs ((- x)#bs) (Dvd j (CN 0 c e))"
     using numbound0_I[OF nb, where bs="bs" and b="x" and b'="- x"]
     by simp
@@ -1375,7 +1375,7 @@
     by (simp only: zdvd_zminus_iff)
   also have "\<dots> = (j dvd (c* (- x)) + ?e)"
     apply (simp only: minus_mult_right[symmetric] minus_mult_left[symmetric] diff_def zadd_ac zminus_zadd_distrib)
-    by (simp add: ring_simps)
+    by (simp add: algebra_simps)
   also have "\<dots> = Ifm bbs ((- x)#bs) (Dvd j (CN 0 c e))"
     using numbound0_I[OF nb, where bs="bs" and b="x" and b'="- x"]
     by simp
@@ -1443,7 +1443,7 @@
     hence "(l*x + (l div c) * Inum (x # bs) e < 0) =
           ((c * (l div c)) * x + (l div c) * Inum (x # bs) e < 0)"
       by simp
-    also have "\<dots> = ((l div c) * (c*x + Inum (x # bs) e) < (l div c) * 0)" by (simp add: ring_simps)
+    also have "\<dots> = ((l div c) * (c*x + Inum (x # bs) e) < (l div c) * 0)" by (simp add: algebra_simps)
     also have "\<dots> = (c*x + Inum (x # bs) e < 0)"
     using mult_less_0_iff [where a="(l div c)" and b="c*x + Inum (x # bs) e"] ldcp by simp
   finally show ?case using numbound0_I[OF be,where b="l*x" and b'="x" and bs="bs"] be  by simp
@@ -1461,7 +1461,7 @@
     hence "(l*x + (l div c) * Inum (x# bs) e \<le> 0) =
           ((c * (l div c)) * x + (l div c) * Inum (x # bs) e \<le> 0)"
       by simp
-    also have "\<dots> = ((l div c) * (c * x + Inum (x # bs) e) \<le> ((l div c)) * 0)" by (simp add: ring_simps)
+    also have "\<dots> = ((l div c) * (c * x + Inum (x # bs) e) \<le> ((l div c)) * 0)" by (simp add: algebra_simps)
     also have "\<dots> = (c*x + Inum (x # bs) e \<le> 0)"
     using mult_le_0_iff [where a="(l div c)" and b="c*x + Inum (x # bs) e"] ldcp by simp
   finally show ?case using numbound0_I[OF be,where b="l*x" and b'="x" and bs="bs"]  be by simp
@@ -1479,7 +1479,7 @@
     hence "(l*x + (l div c)* Inum (x # bs) e > 0) =
           ((c * (l div c)) * x + (l div c) * Inum (x # bs) e > 0)"
       by simp
-    also have "\<dots> = ((l div c) * (c * x + Inum (x # bs) e) > ((l div c)) * 0)" by (simp add: ring_simps)
+    also have "\<dots> = ((l div c) * (c * x + Inum (x # bs) e) > ((l div c)) * 0)" by (simp add: algebra_simps)
     also have "\<dots> = (c * x + Inum (x # bs) e > 0)"
     using zero_less_mult_iff [where a="(l div c)" and b="c * x + Inum (x # bs) e"] ldcp by simp
   finally show ?case using numbound0_I[OF be,where b="(l * x)" and b'="x" and bs="bs"]  be  by simp
@@ -1498,7 +1498,7 @@
           ((c*(l div c))*x + (l div c)* Inum (x # bs) e \<ge> 0)"
       by simp
     also have "\<dots> = ((l div c)*(c*x + Inum (x # bs) e) \<ge> ((l div c)) * 0)" 
-      by (simp add: ring_simps)
+      by (simp add: algebra_simps)
     also have "\<dots> = (c*x + Inum (x # bs) e \<ge> 0)" using ldcp 
       zero_le_mult_iff [where a="l div c" and b="c*x + Inum (x # bs) e"] by simp
   finally show ?case using be numbound0_I[OF be,where b="l*x" and b'="x" and bs="bs"]  
@@ -1517,7 +1517,7 @@
     hence "(l * x + (l div c) * Inum (x # bs) e = 0) =
           ((c * (l div c)) * x + (l div c) * Inum (x # bs) e = 0)"
       by simp
-    also have "\<dots> = ((l div c) * (c * x + Inum (x # bs) e) = ((l div c)) * 0)" by (simp add: ring_simps)
+    also have "\<dots> = ((l div c) * (c * x + Inum (x # bs) e) = ((l div c)) * 0)" by (simp add: algebra_simps)
     also have "\<dots> = (c * x + Inum (x # bs) e = 0)"
     using mult_eq_0_iff [where a="(l div c)" and b="c * x + Inum (x # bs) e"] ldcp by simp
   finally show ?case using numbound0_I[OF be,where b="(l * x)" and b'="x" and bs="bs"]  be  by simp
@@ -1535,7 +1535,7 @@
     hence "(l * x + (l div c) * Inum (x # bs) e \<noteq> 0) =
           ((c * (l div c)) * x + (l div c) * Inum (x # bs) e \<noteq> 0)"
       by simp
-    also have "\<dots> = ((l div c) * (c * x + Inum (x # bs) e) \<noteq> ((l div c)) * 0)" by (simp add: ring_simps)
+    also have "\<dots> = ((l div c) * (c * x + Inum (x # bs) e) \<noteq> ((l div c)) * 0)" by (simp add: algebra_simps)
     also have "\<dots> = (c * x + Inum (x # bs) e \<noteq> 0)"
     using zero_le_mult_iff [where a="(l div c)" and b="c * x + Inum (x # bs) e"] ldcp by simp
   finally show ?case using numbound0_I[OF be,where b="(l * x)" and b'="x" and bs="bs"]  be  by simp
@@ -1551,7 +1551,7 @@
     hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] 
       by simp
     hence "(\<exists> (k::int). l * x + (l div c) * Inum (x # bs) e = ((l div c) * j) * k) = (\<exists> (k::int). (c * (l div c)) * x + (l div c) * Inum (x # bs) e = ((l div c) * j) * k)"  by simp
-    also have "\<dots> = (\<exists> (k::int). (l div c) * (c * x + Inum (x # bs) e - j * k) = (l div c)*0)" by (simp add: ring_simps)
+    also have "\<dots> = (\<exists> (k::int). (l div c) * (c * x + Inum (x # bs) e - j * k) = (l div c)*0)" by (simp add: algebra_simps)
     also fix k have "\<dots> = (\<exists> (k::int). c * x + Inum (x # bs) e - j * k = 0)"
     using zero_le_mult_iff [where a="(l div c)" and b="c * x + Inum (x # bs) e - j * k"] ldcp by simp
   also have "\<dots> = (\<exists> (k::int). c * x + Inum (x # bs) e = j * k)" by simp
@@ -1568,7 +1568,7 @@
     hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] 
       by simp
     hence "(\<exists> (k::int). l * x + (l div c) * Inum (x # bs) e = ((l div c) * j) * k) = (\<exists> (k::int). (c * (l div c)) * x + (l div c) * Inum (x # bs) e = ((l div c) * j) * k)"  by simp
-    also have "\<dots> = (\<exists> (k::int). (l div c) * (c * x + Inum (x # bs) e - j * k) = (l div c)*0)" by (simp add: ring_simps)
+    also have "\<dots> = (\<exists> (k::int). (l div c) * (c * x + Inum (x # bs) e - j * k) = (l div c)*0)" by (simp add: algebra_simps)
     also fix k have "\<dots> = (\<exists> (k::int). c * x + Inum (x # bs) e - j * k = 0)"
     using zero_le_mult_iff [where a="(l div c)" and b="c * x + Inum (x # bs) e - j * k"] ldcp by simp
   also have "\<dots> = (\<exists> (k::int). c * x + Inum (x # bs) e = j * k)" by simp
@@ -1617,7 +1617,7 @@
       hence "x + ?e \<ge> 1 \<and> x + ?e \<le> d"  by simp
       hence "\<exists> (j::int) \<in> {1 .. d}. j = x + ?e" by simp
       hence "\<exists> (j::int) \<in> {1 .. d}. x = (- ?e + j)" 
-	by (simp add: ring_simps)
+	by (simp add: algebra_simps)
       with nob have ?case by auto}
     ultimately show ?case by blast
 next
@@ -1636,7 +1636,7 @@
       from H p have "x + ?e \<ge> 0 \<and> x + ?e < d" by (simp add: c1)
       hence "x + ?e +1 \<ge> 1 \<and> x + ?e + 1 \<le> d"  by simp
       hence "\<exists> (j::int) \<in> {1 .. d}. j = x + ?e + 1" by simp
-      hence "\<exists> (j::int) \<in> {1 .. d}. x= - ?e - 1 + j" by (simp add: ring_simps)
+      hence "\<exists> (j::int) \<in> {1 .. d}. x= - ?e - 1 + j" by (simp add: algebra_simps)
       with nob have ?case by simp }
     ultimately show ?case by blast
 next
@@ -1646,7 +1646,7 @@
     have vb: "?v \<in> set (\<beta> (Eq (CN 0 c e)))" by simp
     from p have "x= - ?e" by (simp add: c1) with prems(11) show ?case using dp
       by simp (erule ballE[where x="1"],
-	simp_all add:ring_simps numbound0_I[OF bn,where b="x"and b'="a"and bs="bs"])
+	simp_all add:algebra_simps numbound0_I[OF bn,where b="x"and b'="a"and bs="bs"])
 next
   case (4 c e)hence p: "Ifm bbs (x #bs) (NEq (CN 0 c e))" (is "?p x") and c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+
     let ?e = "Inum (x # bs) e"
--- a/src/HOL/ex/Reflection.thy	Thu Jan 29 12:05:19 2009 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,48 +0,0 @@
-(*
-    ID:         $Id$
-    Author:     Amine Chaieb, TU Muenchen
-*)
-
-header {* Generic reflection and reification *}
-
-theory Reflection
-imports Main
-  uses "reflection_data.ML" ("reflection.ML")
-begin
-
-setup {* Reify_Data.setup*}
-
-
-lemma ext2: "(\<forall>x. f x = g x) \<Longrightarrow> f = g"
-  by (blast intro: ext)
-
-use "reflection.ML"
-
-method_setup reify = {*
-  fn src =>
-    Method.syntax (Attrib.thms --
-      Scan.option (Scan.lift (Args.$$$ "(") |-- Args.term --| Scan.lift (Args.$$$ ")") )) src #>
-  (fn ((eqs, to), ctxt) => Method.SIMPLE_METHOD' (Reflection.genreify_tac ctxt (eqs @ (fst (Reify_Data.get ctxt))) to))
-*} "partial automatic reification"
-
-method_setup reflection = {* 
-let 
-fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ();
-val onlyN = "only";
-val rulesN = "rules";
-val any_keyword = keyword onlyN || keyword rulesN;
-val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat;
-val terms = thms >> map (term_of o Drule.dest_term);
-fun optional scan = Scan.optional scan [];
-in
-fn src =>
-    Method.syntax (thms -- optional (keyword rulesN |-- thms) -- Scan.option (keyword onlyN |-- Args.term)) src #> 
-  (fn (((eqs,ths),to), ctxt) => 
-      let 
-        val (ceqs,cths) = Reify_Data.get ctxt 
-        val corr_thms = ths@cths
-        val raw_eqs = eqs@ceqs
-      in Method.SIMPLE_METHOD' (Reflection.reflection_tac ctxt corr_thms raw_eqs to) 
-     end) end
-*} "reflection method"
-end
--- a/src/HOL/ex/ReflectionEx.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOL/ex/ReflectionEx.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,9 +1,9 @@
-(*
-    ID:         $Id$
+(*  Title:      HOL/ex/ReflectionEx.thy
     Author:     Amine Chaieb, TU Muenchen
 *)
 
 header {* Examples for generic reflection and reification *}
+
 theory ReflectionEx
 imports Reflection
 begin
@@ -183,7 +183,7 @@
 lemma lin_add: "Inum (lin_add (t,s)) bs = Inum (Add t s) bs"
 apply (induct t s rule: lin_add.induct, simp_all add: Let_def)
 apply (case_tac "c1+c2 = 0",case_tac "n1 \<le> n2", simp_all)
-by (case_tac "n1 = n2", simp_all add: ring_simps)
+by (case_tac "n1 = n2", simp_all add: algebra_simps)
 
 consts lin_mul :: "num \<Rightarrow> nat \<Rightarrow> num"
 recdef lin_mul "measure size "
@@ -192,7 +192,7 @@
   "lin_mul t = (\<lambda> i. Mul i t)"
 
 lemma lin_mul: "Inum (lin_mul t i) bs = Inum (Mul i t) bs"
-by (induct t arbitrary: i rule: lin_mul.induct, auto simp add: ring_simps)
+by (induct t arbitrary: i rule: lin_mul.induct, auto simp add: algebra_simps)
 
 consts linum:: "num \<Rightarrow> num"
 recdef linum "measure num_size"
--- a/src/HOL/ex/reflection.ML	Thu Jan 29 12:05:19 2009 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,325 +0,0 @@
-(*  Author:     Amine Chaieb, TU Muenchen
-
-A trial for automatical reification.
-*)
-
-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
-
-val ext2 = thm "ext2";
-val nth_Cons_0 = thm "nth_Cons_0";
-val nth_Cons_Suc = thm "nth_Cons_Suc";
-
-  (* Make a congruence rule out of a defining equation for the interpretation *)
-  (* th is one defining equation of f, i.e.
-     th is "f (Cp ?t1 ... ?tn) = P(f ?t1, .., f ?tn)" *)
-  (* Cp is a constructor pattern and P is a pattern *)
-
-  (* The result is:
-      [|?A1 = f ?t1 ; .. ; ?An= f ?tn |] ==> P (?A1, .., ?An) = f (Cp ?t1 .. ?tn) *)
-  (*  + the a list of names of the A1 .. An, Those are fresh in the ctxt*)
-
-
-fun mk_congeq ctxt fs th = 
-  let 
-   val (f as Const(fN,fT)) = th |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq 
-				|> fst |> strip_comb |> fst
-   val thy = ProofContext.theory_of ctxt
-   val cert = Thm.cterm_of thy
-   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 => 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 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'
-   val tys = map fastype_of fterms
-   val vs = map Free (xs ~~ tys)
-   val env = fterms ~~ vs
-		    (* FIXME!!!!*)	
-   fun replace_fterms (t as t1 $ t2) =
-       (case AList.lookup (op aconv) env t of
-	    SOME v => v
-	  | NONE => replace_fterms t1 $ replace_fterms t2)
-     | replace_fterms t = (case AList.lookup (op aconv) env t of
-			       SOME v => v
-			     | NONE => t)
-      
-   fun mk_def (Abs(x,xT,t),v) = HOLogic.mk_Trueprop ((HOLogic.all_const xT)$ Abs(x,xT,HOLogic.mk_eq(v$(Bound 0), t)))
-     | mk_def (t, v) = HOLogic.mk_Trueprop (HOLogic.mk_eq (v, t))
-   fun tryext x = (x RS ext2 handle THM _ =>  x)
-   val cong = (Goal.prove ctxt'' [] (map mk_def env)
-			  (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, replace_fterms rhs)))
-			  (fn x => LocalDefs.unfold_tac (#context x) (map tryext (#prems x)) 
-							THEN rtac th' 1)) RS sym
-	      
-   val (cong' :: vars') = 
-       Variable.export ctxt'' ctxt (cong :: map (Drule.mk_term o cert) vs)
-   val vs' = map (fst o fst o Term.dest_Var o Thm.term_of o Drule.dest_term) vars'
-					      
-  in  (vs', cong') end; 
- (* congs is a list of pairs (P,th) where th is a theorem for *)
-        (* [| f p1 = A1; ...; f pn = An|] ==> f (C p1 .. pn) = P *)
-val FWD = curry (op OF);
-
- (* da is the decomposition for atoms, ie. it returns ([],g) where g
- returns the right instance f (AtC n) = t , where AtC is the Atoms
- constructor and n is the number of the atom corresponding to t *)
-
-(* Generic decomp for reification : matches the actual term with the
-rhs of one cong rule. The result of the matching guides the
-proof synthesis: The matches of the introduced Variables A1 .. An are
-processed recursively
- The rest is instantiated in the cong rule,i.e. no reification is needed *)
-
-exception REIF of string;
-
-val bds = ref ([]: (typ * ((term list) * (term list))) list);
-
-fun index_of t = 
- let 
-  val tt = HOLogic.listT (fastype_of t)
- in 
-  (case AList.lookup Type.could_unify (!bds) tt of
-    NONE => error "index_of : type not found in environements!"
-  | SOME (tbs,tats) =>
-    let
-     val i = find_index_eq t tats
-     val j = find_index_eq t tbs 
-    in (if j= ~1 then 
-	    if i= ~1 
-	    then (bds := AList.update Type.could_unify (tt,(tbs,tats@[t])) (!bds) ; 
-		  length tbs + length tats) 
-	    else i else j)
-    end)
- end;
-
-fun dest_listT (Type ("List.list", [T])) = T;
-
-fun decomp_genreif da cgns (t,ctxt) =
- let 
-  val thy = ProofContext.theory_of ctxt 
-  val cert = cterm_of thy
-  fun tryabsdecomp (s,ctxt) = 
-   (case s of 
-     Abs(xn,xT,ta) => 
-     (let
-       val ([xn],ctxt') = Variable.variant_fixes ["x"] ctxt
-       val (xn,ta) = variant_abs (xn,xT,ta)
-       val x = Free(xn,xT)
-       val _ = (case AList.lookup Type.could_unify (!bds) (HOLogic.listT xT)
-		 of NONE => error "tryabsdecomp: Type not found in the Environement"
-		  | SOME (bsT,atsT) => 
-		    (bds := AList.update Type.could_unify (HOLogic.listT xT, ((x::bsT), atsT)) (!bds)))
-      in ([(ta, ctxt')] , 
-	  fn [th] => ((let val (bsT,asT) = the(AList.lookup Type.could_unify (!bds) (HOLogic.listT xT))
-		       in (bds := AList.update Type.could_unify (HOLogic.listT xT,(tl bsT,asT)) (!bds))
-		       end) ; 
-		      hd (Variable.export ctxt' ctxt [(forall_intr (cert x) th) COMP allI])))
-	end)
-    | _ => da (s,ctxt))
-  in 
-  (case cgns of 
-    [] => tryabsdecomp (t,ctxt)
-  | ((vns,cong)::congs) => ((let
-        val cert = cterm_of thy
-	val certy = ctyp_of thy
-        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), Vartab.empty)
-        val (fnvs,invs) = List.partition (fn ((vn,_),_) => vn mem vns) (Vartab.dest tmenv)
-        val (fts,its) = 
-	    (map (snd o snd) fnvs,
-             map (fn ((vn,vi),(tT,t)) => (cert(Var ((vn,vi),tT)), cert t)) invs)
-	val ctyenv = map (fn ((vn,vi),(s,ty)) => (certy (TVar((vn,vi),s)), certy ty)) (Vartab.dest tyenv)
-    in (fts ~~ (replicate (length fts) ctxt), FWD (instantiate (ctyenv, its) cong))
-    end)
-      handle MATCH => decomp_genreif da congs (t,ctxt)))
-  end;
-
- (* looks for the atoms equation and instantiates it with the right number *)
-
-
-fun mk_decompatom eqs (t,ctxt) =
-let 
- val tT = fastype_of t
- fun isat eq = 
-  let 
-   val rhs = eq |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd
-   in exists_Const 
-	  (fn (n,ty) => n="List.nth" 
-			andalso 
-			AList.defined Type.could_unify (!bds) (domain_type ty)) rhs 
-	  andalso Type.could_unify (fastype_of rhs, tT)
-   end
- fun get_nths t acc = 
-  case t of
-    Const("List.nth",_)$vs$n => insert (fn ((a,_),(b,_)) => a aconv b) (t,(vs,n)) acc
-  | t1$t2 => get_nths t1 (get_nths t2 acc)
-  | Abs(_,_,t') => get_nths t'  acc
-  | _ => acc
-
- fun 
-   tryeqs [] = error "Can not find the atoms equation"
- | tryeqs (eq::eqs) = ((
-  let 
-   val rhs = eq |> prop_of |> HOLogic.dest_Trueprop  |> HOLogic.dest_eq |> snd
-   val nths = get_nths rhs []
-   val (vss,ns) = fold_rev (fn (_,(vs,n)) => fn (vss,ns) => 
-                             (insert (op aconv) vs vss, insert (op aconv) n ns)) nths ([],[]) 
-   val (vsns, ctxt') = Variable.variant_fixes (replicate (length vss) "vs") ctxt
-   val (xns, ctxt'') = Variable.variant_fixes (replicate (length nths) "x") ctxt' 
-   val thy = ProofContext.theory_of ctxt''
-   val cert = cterm_of thy
-   val certT = ctyp_of thy
-   val vsns_map = vss ~~ vsns
-   val xns_map = (fst (split_list nths)) ~~ xns
-   val subst = map (fn (nt, xn) => (nt, Var ((xn,0), fastype_of nt))) xns_map
-   val rhs_P = subst_free subst rhs
-   val (tyenv, tmenv) = Pattern.match 
-	                    thy (rhs_P, t)
-	                    (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)) 
-                      (Vartab.dest tyenv)
-   val tml = Vartab.dest tmenv
-   val t's = map (fn xn => snd (valOf (AList.lookup (op =) tml (xn,0)))) xns (* FIXME : Express with sbst*)
-   val subst_ns = map (fn (Const _ $ vs $ n, Var (xn0,T)) => 
-                          (cert n, snd (valOf (AList.lookup (op =) tml xn0)) 
-                             |> (index_of #> HOLogic.mk_nat #> cert))) 
-                      subst
-   val subst_vs = 
-    let 
-     fun ty (Const _ $ (vs as Var (vsn,lT)) $ n, Var (xn0,T)) = (certT T, certT (sbsT T))
-     fun h (Const _ $ (vs as Var (vsn,lT)) $ n, Var (xn0,T)) = 
-      let 
-       val cns = sbst (Const("List.list.Cons", T --> lT --> lT))
-       val lT' = sbsT lT
-       val (bsT,asT) = the (AList.lookup Type.could_unify (!bds) lT)
-       val vsn = valOf (AList.lookup (op =) vsns_map vs)
-       val cvs = cert (fold_rev (fn x => fn xs => cns$x$xs) bsT (Free (vsn, lT')))
-      in (cert vs, cvs) end
-    in map h subst end
-   val cts = map (fn ((vn,vi),(tT,t)) => (cert(Var ((vn,vi),tT)), cert t)) 
-                 (fold (AList.delete (fn (((a: string),_),(b,_)) => a = b)) 
-                       (map (fn n => (n,0)) xns) tml)
-   val substt = 
-    let val ih = Drule.cterm_rule (Thm.instantiate (subst_ty,[]))
-    in map (fn (v,t) => (ih v, ih t)) (subst_ns@subst_vs@cts)  end
-   val th = (instantiate (subst_ty, substt)  eq) RS sym
-  in  hd (Variable.export ctxt'' ctxt [th]) end)
- handle MATCH => tryeqs eqs)
-in ([], fn _ => tryeqs (filter isat eqs))
-end;
-
-  (* Generic reification procedure: *)
-  (* creates all needed cong rules and then just uses the theorem synthesis *)
-
-  fun mk_congs ctxt raw_eqs = 
- let
-  val fs = fold_rev (fn eq =>
-		     insert (op =) (eq |> prop_of |> HOLogic.dest_Trueprop 
-			 |> HOLogic.dest_eq |> fst |> strip_comb 
-			 |> fst)) raw_eqs []
-  val tys = fold_rev (fn f => fn ts => (f |> fastype_of |> binder_types |> tl) 
-				    union ts) fs []
-  val _ = bds := AList.make (fn _ => ([],[])) tys
-  val (vs, ctxt') = Variable.variant_fixes (replicate (length tys) "vs") ctxt
-  val thy = ProofContext.theory_of ctxt'
-  val cert = cterm_of thy
-  val vstys = map (fn (t,v) => (t,SOME (cert (Free(v,t))))) 
-		  (tys ~~ vs)
-  val is_Var = can dest_Var
-  fun insteq eq vs = 
-   let
-     val subst = map (fn (v as Var(n,t)) => (cert v, (valOf o valOf) (AList.lookup (op =) vstys t)))  
-  (filter is_Var vs)
-   in Thm.instantiate ([],subst) eq
-   end
-  val eqs = map (fn eq => eq |> prop_of |> HOLogic.dest_Trueprop 
-			     |> HOLogic.dest_eq |> fst |> strip_comb |> snd |> tl
-			     |> (insteq eq)) raw_eqs
-  val (ps,congs) = split_list (map (mk_congeq ctxt' fs) eqs)
-in ps ~~ (Variable.export ctxt' ctxt congs)
-end
-
-fun partition P [] = ([],[])
-  | partition P (x::xs) = 
-     let val (yes,no) = partition P xs
-     in if P x then (x::yes,no) else (yes, x::no) end
-
-fun rearrange congs = 
-let 
- fun P (_, th) = 
-  let val @{term "Trueprop"}$(Const ("op =",_) $l$_) = concl_of th
-  in can dest_Var l end
- val (yes,no) = partition P congs 
- in no @ yes end
-
-
-
-fun genreif ctxt raw_eqs t =
- let 
-  val congs = rearrange (mk_congs ctxt raw_eqs)
-  val th = divide_and_conquer (decomp_genreif (mk_decompatom raw_eqs) congs) (t,ctxt)
-  fun is_listVar (Var (_,t)) = can dest_listT t
-       | is_listVar _ = false
-  val vars = th |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd
-	       |> strip_comb |> snd |> filter is_listVar
-  val cert = cterm_of (ProofContext.theory_of ctxt)
-  val cvs = map (fn (v as Var(n,t)) => (cert v, the (AList.lookup Type.could_unify (!bds) t) |> snd |> HOLogic.mk_list (dest_listT t) |> cert)) vars
-  val th' = instantiate ([], cvs) th
-  val t' = (fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) th'
-  val th'' = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (t, t')))
-			(fn _ => simp_tac (local_simpset_of ctxt) 1)
-  val _ = bds := []
-in FWD trans [th'',th']
-end
-
-
-fun genreflect ctxt conv corr_thms raw_eqs t =
-let 
-  val reifth = genreif ctxt raw_eqs t
-  fun trytrans [] = error "No suitable correctness theorem found"
-    | trytrans (th::ths) = 
-         (FWD trans [reifth, th RS sym] handle THM _ => trytrans ths)
-  val th = trytrans corr_thms
-  val ft = (Thm.dest_arg1 o Thm.dest_arg o Thm.dest_arg o cprop_of) th
-  val rth = conv ft
-in simplify (HOL_basic_ss addsimps raw_eqs addsimps [nth_Cons_0, nth_Cons_Suc])
-           (simplify (HOL_basic_ss addsimps [rth]) th)
-end
-
-fun genreify_tac ctxt eqs to i = (fn st =>
-  let
-    val P = HOLogic.dest_Trueprop (List.nth (prems_of st, i - 1))
-    val t = (case to of NONE => P | SOME x => x)
-    val th = (genreif ctxt eqs t) RS ssubst
-  in rtac th i st
-  end);
-
-    (* Reflection calls reification and uses the correctness *)
-        (* theorem assumed to be the dead of the list *)
-fun gen_reflection_tac ctxt conv corr_thms raw_eqs to i = (fn st =>
-  let
-    val P = HOLogic.dest_Trueprop (nth (prems_of st) (i - 1));
-    val t = the_default P to;
-    val th = genreflect ctxt conv corr_thms raw_eqs t
-      RS ssubst;
-  in (rtac th i THEN TRY(rtac TrueI i)) st end);
-
-fun reflection_tac ctxt = gen_reflection_tac ctxt Codegen.evaluation_conv;
-end
--- a/src/HOL/ex/reflection_data.ML	Thu Jan 29 12:05:19 2009 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,69 +0,0 @@
-(*  Title:      HOL/ex/reflection_data.ML
-    ID:         $Id$
-    Author:     Amine Chaieb, TU Muenchen
-
-Data for the reification and reflection methods.
-*)
-
-signature REIFY_DATA =
-sig
-  type entry = thm list * thm list
-  val get: Proof.context -> entry
-  val del: attribute
-  val add: attribute
-  val setup: theory -> theory
-end;
-
-structure Reify_Data : REIFY_DATA =
-struct
-
-type entry = thm list * thm list;
-
-structure Data = GenericDataFun
-(
-  type T = entry
-  val empty = ([], [])
-  val extend = I
-  fun merge _ = pairself Thm.merge_thms
-);
-
-val get = Data.get o Context.Proof;
-
-val add = Thm.declaration_attribute (fn th => fn context =>
-  context |> Data.map (apfst (Thm.add_thm th)))
-
-val del = Thm.declaration_attribute (fn th => fn context =>
-  context |> Data.map (apfst (Thm.del_thm th)))
-
-val radd = Thm.declaration_attribute (fn th => fn context =>
-  context |> Data.map (apsnd (Thm.add_thm th)))
-
-val rdel = Thm.declaration_attribute (fn th => fn context =>
-  context |> Data.map (apsnd (Thm.del_thm th)))
-
-(* concrete syntax *)
-(*
-local
-fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ();
-
-val constsN = "consts";
-val addN = "add";
-val delN = "del";
-val any_keyword = keyword constsN || keyword addN || keyword delN;
-val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat;
-val terms = thms >> map (term_of o Drule.dest_term);
-
-fun optional scan = Scan.optional scan [];
-
-in
-val att_syntax = Attrib.syntax
- ((Scan.lift (Args.$$$ "del") |-- thms) >> del ||
-  optional (keyword addN |-- thms) >> add)
-end;
-*)
-
-(* theory setup *)
- val setup =
-  Attrib.add_attributes [("reify", Attrib.add_del_args add del, "Reify data"),
-                         ("reflection", Attrib.add_del_args radd rdel, "Reflection data")];
-end;
--- a/src/HOLCF/Bifinite.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOLCF/Bifinite.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -10,7 +10,7 @@
 
 subsection {* Omega-profinite and bifinite domains *}
 
-class profinite = cpo +
+class profinite =
   fixes approx :: "nat \<Rightarrow> 'a \<rightarrow> 'a"
   assumes chain_approx [simp]: "chain approx"
   assumes lub_approx_app [simp]: "(\<Squnion>i. approx i\<cdot>x) = x"
--- a/src/HOLCF/Cfun.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOLCF/Cfun.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOLCF/CompactBasis.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOLCF/Cont.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOLCF/ConvexPD.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -296,18 +296,17 @@
 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)
 
 lemma convex_plus_left_absorb: "xs +\<natural> (xs +\<natural> ys) = xs +\<natural> ys"
 by (rule aci_convex_plus.mult_left_idem)
-
+(*
 lemmas convex_plus_aci = aci_convex_plus.mult_ac_idem
-
+*)
 lemma convex_unit_less_plus_iff [simp]:
   "{x}\<natural> \<sqsubseteq> ys +\<natural> zs \<longleftrightarrow> {x}\<natural> \<sqsubseteq> ys \<and> {x}\<natural> \<sqsubseteq> zs"
  apply (rule iffI)
--- a/src/HOLCF/Cprod.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOLCF/Cprod.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ /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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOLCF/Fixrec.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -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
--- a/src/HOLCF/HOLCF.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOLCF/HOLCF.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -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
--- a/src/HOLCF/IsaMakefile	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOLCF/IsaMakefile	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOLCF/LowerPD.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -250,18 +250,17 @@
 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)
 
 lemma lower_plus_left_absorb: "xs +\<flat> (xs +\<flat> ys) = xs +\<flat> ys"
 by (rule aci_lower_plus.mult_left_idem)
-
+(*
 lemmas lower_plus_aci = aci_lower_plus.mult_ac_idem
-
+*)
 lemma lower_plus_less1: "xs \<sqsubseteq> xs +\<flat> ys"
 apply (induct xs ys rule: lower_pd.principal_induct2, simp, simp)
 apply (simp add: PDPlus_lower_less)
--- a/src/HOLCF/Pcpo.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOLCF/Pcpo.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -12,9 +12,9 @@
 
 text {* The class cpo of chain complete partial orders *}
 
-axclass cpo < po
+class cpo = po +
         -- {* class axiom: *}
-  cpo:   "chain S \<Longrightarrow> \<exists>x. range S <<| x"
+  assumes cpo:   "chain S \<Longrightarrow> \<exists>x :: 'a::po. range S <<| x"
 
 text {* in cpo's everthing equal to THE lub has lub properties for every chain *}
 
@@ -170,8 +170,8 @@
 
 text {* The class pcpo of pointed cpos *}
 
-axclass pcpo < cpo
-  least: "\<exists>x. \<forall>y. x \<sqsubseteq> y"
+class pcpo = cpo +
+  assumes least: "\<exists>x. \<forall>y. x \<sqsubseteq> y"
 
 definition
   UU :: "'a::pcpo" where
@@ -254,13 +254,13 @@
 
 text {* further useful classes for HOLCF domains *}
 
-axclass finite_po < finite, po
+class finite_po = finite + po
 
-axclass chfin < po
-  chfin: "chain Y \<Longrightarrow> \<exists>n. max_in_chain n Y"
+class chfin = po +
+  assumes chfin: "chain Y \<Longrightarrow> \<exists>n. max_in_chain n (Y :: nat => 'a::po)"
 
-axclass flat < pcpo
-  ax_flat: "x \<sqsubseteq> y \<Longrightarrow> (x = \<bottom>) \<or> (x = y)"
+class flat = pcpo +
+  assumes ax_flat: "(x :: 'a::pcpo) \<sqsubseteq> y \<Longrightarrow> x = \<bottom> \<or> x = y"
 
 text {* finite partial orders are chain-finite *}
 
@@ -310,11 +310,11 @@
 
 text {* Discrete cpos *}
 
-axclass discrete_cpo < sq_ord
-  discrete_cpo [simp]: "x \<sqsubseteq> y \<longleftrightarrow> x = y"
+class discrete_cpo = sq_ord +
+  assumes discrete_cpo [simp]: "x \<sqsubseteq> y \<longleftrightarrow> x = y"
 
-instance discrete_cpo < po
-by (intro_classes, simp_all)
+subclass (in discrete_cpo) po
+proof qed simp_all
 
 text {* In a discrete cpo, every chain is constant *}
 
--- a/src/HOLCF/Porder.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOLCF/Porder.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -10,7 +10,7 @@
 
 subsection {* Type class for partial orders *}
 
-class sq_ord = type +
+class sq_ord =
   fixes sq_le :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
 
 notation
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOLCF/Product_Cpo.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOLCF/Ssum.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:24:00 2009 +0000
@@ -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/cont_proc.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOLCF/Tools/cont_proc.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -7,7 +7,7 @@
   val is_lcf_term: term -> bool
   val cont_thms: term -> thm list
   val all_cont_thms: term -> thm list
-  val cont_tac: thm list ref -> int -> tactic
+  val cont_tac: int -> tactic
   val cont_proc: theory -> simproc
   val setup: theory -> theory
 end;
@@ -15,15 +15,6 @@
 structure ContProc: CONT_PROC =
 struct
 
-structure ContProcData = TheoryDataFun
-(
-  type T = thm list ref;
-  val empty = ref [];
-  val copy = I;
-  val extend = I;
-  fun merge _ _ = ref [];
-)
-
 (** theory context references **)
 
 val cont_K = @{thm cont_const};
@@ -107,26 +98,21 @@
   conditional rewrite rule with the unsolved subgoals as premises.
 *)
 
-fun cont_tac prev_cont_thms =
+val cont_tac =
   let
     val rules = [cont_K, cont_I, cont_R, cont_A, cont_L];
   
-    fun old_cont_tac i thm =
-      case !prev_cont_thms of
-        [] => no_tac thm
-      | (c::cs) => (prev_cont_thms := cs; rtac c i thm);
-
-    fun new_cont_tac f' i thm =
+    fun new_cont_tac f' i =
       case all_cont_thms f' of
-        [] => no_tac thm
-      | (c::cs) => (prev_cont_thms := cs; rtac c i thm);
+        [] => no_tac
+      | (c::cs) => rtac c i;
 
     fun cont_tac_of_term (Const (@{const_name cont}, _) $ f) =
       let
         val f' = Const (@{const_name Abs_CFun}, dummyT) $ f;
       in
         if is_lcf_term f'
-        then old_cont_tac ORELSE' new_cont_tac f'
+        then new_cont_tac f'
         else REPEAT_ALL_NEW (resolve_tac rules)
       end
       | cont_tac_of_term _ = K no_tac;
@@ -139,8 +125,7 @@
   fun solve_cont thy _ t =
     let
       val tr = instantiate' [] [SOME (cterm_of thy t)] Eq_TrueI;
-      val prev_thms = ContProcData.get thy
-    in Option.map fst (Seq.pull (cont_tac prev_thms 1 tr)) end
+    in Option.map fst (Seq.pull (cont_tac 1 tr)) end
 in
   fun cont_proc thy =
     Simplifier.simproc thy "cont_proc" ["cont f"] solve_cont;
--- a/src/HOLCF/Tools/domain/domain_axioms.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOLCF/Tools/domain/domain_axioms.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -111,10 +111,10 @@
 
 fun infer_props thy = map (apsnd (FixrecPackage.legacy_infer_prop thy));
 
-fun add_axioms_i x = snd o PureThy.add_axioms (map Thm.no_attributes x);
+fun add_axioms_i x = snd o PureThy.add_axioms (map (Thm.no_attributes o apfst Binding.name) x);
 fun add_axioms_infer axms thy = add_axioms_i (infer_props thy axms) thy;
 
-fun add_defs_i x = snd o (PureThy.add_defs false) (map Thm.no_attributes x);
+fun add_defs_i x = snd o (PureThy.add_defs false) (map (Thm.no_attributes o apfst Binding.name) x);
 fun add_defs_infer defs thy = add_defs_i (infer_props thy defs) thy;
 
 in (* local *)
--- a/src/HOLCF/Tools/domain/domain_extender.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOLCF/Tools/domain/domain_extender.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -134,7 +134,7 @@
   in
     theorems_thy
     |> Sign.add_path (Sign.base_name comp_dnam)
-    |> (snd o (PureThy.add_thmss [(("rews", List.concat rewss @ take_rews), [])]))
+    |> (snd o (PureThy.add_thmss [((Binding.name "rews", List.concat rewss @ take_rews), [])]))
     |> Sign.parent_path
   end;
 
--- a/src/HOLCF/Tools/domain/domain_theorems.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOLCF/Tools/domain/domain_theorems.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -607,7 +607,7 @@
 in
   thy
     |> Sign.add_path (Sign.base_name dname)
-    |> (snd o (PureThy.add_thmss (map Thm.no_attributes [
+    |> (snd o (PureThy.add_thmss (map (Thm.no_attributes o apfst Binding.name) [
         ("iso_rews" , iso_rews  ),
         ("exhaust"  , [exhaust] ),
         ("casedist" , [casedist]),
@@ -623,7 +623,7 @@
         ("injects" , injects ),
         ("copy_rews", copy_rews)])))
     |> (snd o PureThy.add_thmss
-        [(("match_rews", mat_rews), [Simplifier.simp_add])])
+        [((Binding.name "match_rews", mat_rews), [Simplifier.simp_add])])
     |> Sign.parent_path
     |> pair (iso_rews @ when_rews @ con_rews @ sel_rews @ dis_rews @
         pat_rews @ dist_les @ dist_eqs @ copy_rews)
@@ -1000,7 +1000,7 @@
 end; (* local *)
 
 in thy |> Sign.add_path comp_dnam
-       |> (snd o (PureThy.add_thmss (map Thm.no_attributes [
+       |> (snd o (PureThy.add_thmss (map (Thm.no_attributes o apfst Binding.name) [
 		("take_rews"  , take_rews  ),
 		("take_lemmas", take_lemmas),
 		("finites"    , finites    ),
--- a/src/HOLCF/Tools/fixrec_package.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOLCF/Tools/fixrec_package.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -9,9 +9,9 @@
   val legacy_infer_term: theory -> term -> term
   val legacy_infer_prop: theory -> term -> term
   val add_fixrec: bool -> (Attrib.binding * string) list list -> theory -> theory
-  val add_fixrec_i: bool -> ((Binding.T * attribute list) * term) list list -> theory -> theory
+  val add_fixrec_i: bool -> ((binding * attribute list) * term) list list -> theory -> theory
   val add_fixpat: Attrib.binding * string list -> theory -> theory
-  val add_fixpat_i: (Binding.T * attribute list) * term list -> theory -> theory
+  val add_fixpat_i: (binding * attribute list) * term list -> theory -> theory
 end;
 
 structure FixrecPackage: FIXREC_PACKAGE =
@@ -96,7 +96,7 @@
     
     val fixdefs = map (apsnd (legacy_infer_prop thy)) pre_fixdefs;
     val (fixdef_thms, thy') =
-      PureThy.add_defs false (map Thm.no_attributes fixdefs) thy;
+      PureThy.add_defs false (map (Thm.no_attributes o apfst Binding.name) fixdefs) thy;
     val ctuple_fixdef_thm = foldr1 (fn (x,y) => @{thm cpair_equalI} OF [x,y]) fixdef_thms;
     
     val ctuple_unfold = legacy_infer_term thy' (mk_trp (mk_ctuple lhss === mk_ctuple rhss));
@@ -114,7 +114,7 @@
         in (n^"_unfold", thmL) :: unfolds ns thmR end;
     val unfold_thms = unfolds names ctuple_unfold_thm;
     val thms = ctuple_induct_thm :: unfold_thms;
-    val (_, thy'') = PureThy.add_thms (map Thm.no_attributes thms) thy';
+    val (_, thy'') = PureThy.add_thms (map (Thm.no_attributes o apfst Binding.name) thms) thy';
   in
     (thy'', names, fixdef_thms, map snd unfold_thms)
   end;
@@ -241,14 +241,14 @@
   in
     if strict then let (* only prove simp rules if strict = true *)
       val eqn_blocks = unconcat lengths ((names ~~ eqn_ts') ~~ atts);
-      val simps = List.concat (map (make_simps thy') (unfold_thms ~~ eqn_blocks));
-      val (simp_thms, thy'') = PureThy.add_thms simps thy';
+      val simps = maps (make_simps thy') (unfold_thms ~~ eqn_blocks);
+      val (simp_thms, thy'') = PureThy.add_thms ((map o apfst o apfst) Binding.name simps) thy';
       
       val simp_names = map (fn name => name^"_simps") cnames;
       val simp_attribute = rpair [Simplifier.simp_add];
       val simps' = map simp_attribute (simp_names ~~ unconcat lengths simp_thms);
     in
-      (snd o PureThy.add_thmss simps') thy''
+      (snd o PureThy.add_thmss ((map o apfst o apfst) Binding.name simps')) thy''
     end
     else thy'
   end;
@@ -278,7 +278,7 @@
     val ts = map (prep_term thy) strings;
     val simps = map (fix_pat thy) ts;
   in
-    (snd o PureThy.add_thmss [((Binding.base_name name, simps), atts)]) thy
+    (snd o PureThy.add_thmss [((name, simps), atts)]) thy
   end;
 
 val add_fixpat = gen_add_fixpat Syntax.read_term_global Attrib.attribute;
--- a/src/HOLCF/Tools/pcpodef_package.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOLCF/Tools/pcpodef_package.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -97,12 +97,12 @@
         theory'
         |> Sign.add_path name
         |> PureThy.add_thms
-          ([(("adm_" ^ name, admissible'), []),
-            (("cont_" ^ Rep_name, @{thm typedef_cont_Rep} OF cpo_thms'), []),
-            (("cont_" ^ Abs_name, @{thm typedef_cont_Abs} OF cpo_thms'), []),
-            (("lub_" ^ name, @{thm typedef_lub} OF cpo_thms'), []),
-            (("thelub_" ^ name, @{thm typedef_thelub} OF cpo_thms'), []),
-            (("compact_" ^ name, @{thm typedef_compact} OF cpo_thms'), [])])
+          ([((Binding.name ("adm_" ^ name), admissible'), []),
+            ((Binding.name ("cont_" ^ Rep_name), @{thm typedef_cont_Rep} OF cpo_thms'), []),
+            ((Binding.name ("cont_" ^ Abs_name), @{thm typedef_cont_Abs} OF cpo_thms'), []),
+            ((Binding.name ("lub_" ^ name), @{thm typedef_lub} OF cpo_thms'), []),
+            ((Binding.name ("thelub_" ^ name), @{thm typedef_thelub} OF cpo_thms'), []),
+            ((Binding.name ("compact_" ^ name), @{thm typedef_compact} OF cpo_thms'), [])])
         |> snd
         |> Sign.parent_path
       end;
@@ -119,12 +119,12 @@
         theory'
         |> Sign.add_path name
         |> PureThy.add_thms
-            ([((Rep_name ^ "_strict", @{thm typedef_Rep_strict} OF pcpo_thms'), []),
-              ((Abs_name ^ "_strict", @{thm typedef_Abs_strict} OF pcpo_thms'), []),
-              ((Rep_name ^ "_strict_iff", @{thm typedef_Rep_strict_iff} OF pcpo_thms'), []),
-              ((Abs_name ^ "_strict_iff", @{thm typedef_Abs_strict_iff} OF pcpo_thms'), []),
-              ((Rep_name ^ "_defined", @{thm typedef_Rep_defined} OF pcpo_thms'), []),
-              ((Abs_name ^ "_defined", @{thm typedef_Abs_defined} OF pcpo_thms'), [])
+            ([((Binding.name (Rep_name ^ "_strict"), @{thm typedef_Rep_strict} OF pcpo_thms'), []),
+              ((Binding.name (Abs_name ^ "_strict"), @{thm typedef_Abs_strict} OF pcpo_thms'), []),
+              ((Binding.name (Rep_name ^ "_strict_iff"), @{thm typedef_Rep_strict_iff} OF pcpo_thms'), []),
+              ((Binding.name (Abs_name ^ "_strict_iff"), @{thm typedef_Abs_strict_iff} OF pcpo_thms'), []),
+              ((Binding.name (Rep_name ^ "_defined"), @{thm typedef_Rep_defined} OF pcpo_thms'), []),
+              ((Binding.name (Abs_name ^ "_defined"), @{thm typedef_Abs_defined} OF pcpo_thms'), [])
               ])
         |> snd
         |> Sign.parent_path
--- a/src/HOLCF/Up.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOLCF/Up.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOLCF/UpperPD.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -248,18 +248,17 @@
 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)
 
 lemma upper_plus_left_absorb: "xs +\<sharp> (xs +\<sharp> ys) = xs +\<sharp> ys"
 by (rule aci_upper_plus.mult_left_idem)
-
+(*
 lemmas upper_plus_aci = aci_upper_plus.mult_ac_idem
-
+*)
 lemma upper_plus_less1: "xs +\<sharp> ys \<sqsubseteq> xs"
 apply (induct xs ys rule: upper_pd.principal_induct2, simp, simp)
 apply (simp add: PDPlus_upper_less)
--- a/src/HOLCF/ex/Stream.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/HOLCF/ex/Stream.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -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/splitter.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Provers/splitter.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Concurrent/future.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -270,8 +270,24 @@
 
 (* 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_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 () =>
@@ -280,12 +296,13 @@
       val _ = Multithreading.self_critical () andalso
         error "Cannot join future values within critical section";
 
-      fun join_loop _ [] = ()
-        | join_loop name deps =
+      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_loop name deps'));
+            | SOME (work, deps') => (execute name work; join_deps name deps'));
+
       val _ =
         (case thread_data () of
           NONE =>
@@ -299,14 +316,14 @@
               val deps = map task_of pending;
               val _ = SYNCHRONIZED "join" (fn () =>
                 (change queue (Task_Queue.depend deps task); notify_all ()));
-              val _ = join_loop ("join_loop: " ^ name) deps;
-              val _ =
-                while not (forall is_finished pending)
-                do SYNCHRONIZED "join_task" (fn () => worker_wait ());
+              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);
 
--- a/src/Pure/Concurrent/mailbox.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Concurrent/mailbox.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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/simple_thread.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Concurrent/simple_thread.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Concurrent/simple_thread.ML
-    ID:         $Id$
     Author:     Makarius
 
 Simplified thread operations.
--- a/src/Pure/Concurrent/synchronized.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Concurrent/synchronized.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Concurrent/synchronized.ML
-    ID:         $Id$
     Author:     Fabian Immler and Makarius
 
 State variables with synchronized access.
--- a/src/Pure/General/ROOT.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/ROOT.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/General/ROOT.ML
-    ID:         $Id$
 
 Library of general tools.
 *)
--- a/src/Pure/General/alist.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/alist.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/General/alist.ML
-    ID:         $Id$
     Author:     Florian Haftmann, TU Muenchen
 
 Association lists -- lists of (key, value) pairs.
--- a/src/Pure/General/balanced_tree.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/balanced_tree.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/General/balanced_tree.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson and Makarius
 
 Balanced binary trees.
--- a/src/Pure/General/basics.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/basics.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/General/basics.ML
-    ID:         $Id$
     Author:     Florian Haftmann and Makarius, TU Muenchen
 
 Fundamental concepts.
--- a/src/Pure/General/binding.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/binding.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -6,6 +6,7 @@
 
 signature BASIC_BINDING =
 sig
+  type binding
   val long_names: bool ref
   val short_names: bool ref
   val unique_names: bool ref
@@ -14,22 +15,21 @@
 signature BINDING =
 sig
   include BASIC_BINDING
-  type T
-  val name_pos: string * Position.T -> T
-  val name: string -> T
-  val empty: T
-  val map_base: (string -> string) -> T -> T
-  val qualify: string -> T -> T
-  val add_prefix: bool -> string -> T -> T
-  val map_prefix: ((string * bool) list -> T -> T) -> T -> T
-  val is_empty: T -> bool
-  val base_name: T -> string
-  val pos_of: T -> Position.T
-  val dest: T -> (string * bool) list * string
+  val name_pos: string * Position.T -> binding
+  val name: string -> binding
+  val empty: binding
+  val map_base: (string -> string) -> binding -> binding
+  val qualify: string -> binding -> binding
+  val add_prefix: bool -> string -> binding -> binding
+  val map_prefix: ((string * bool) list -> binding -> binding) -> binding -> binding
+  val is_empty: binding -> bool
+  val base_name: binding -> string
+  val pos_of: binding -> Position.T
+  val dest: binding -> (string * bool) list * string
   val separator: string
   val is_qualified: string -> bool
-  val display: T -> string
-end
+  val display: binding -> string
+end;
 
 structure Binding : BINDING =
 struct
@@ -54,7 +54,7 @@
 
 (** binding representation **)
 
-datatype T = Binding of ((string * bool) list * string) * Position.T;
+datatype binding = Binding of ((string * bool) list * string) * Position.T;
   (* (prefix components (with mandatory flag), base name, position) *)
 
 fun name_pos (name, pos) = Binding (([], name), pos);
--- a/src/Pure/General/file.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/file.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/General/file.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 File system operations.
--- a/src/Pure/General/graph.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/graph.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/General/graph.ML
-    ID:         $Id$
     Author:     Markus Wenzel and Stefan Berghofer, TU Muenchen
 
 Directed graphs.
--- a/src/Pure/General/heap.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/heap.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,6 +1,5 @@
 (*  Title:      Pure/General/heap.ML
-    ID:         $Id$
-    Author:     Markus Wenzel, TU Muenchen
+    Author:     Lawrence C Paulson and Markus Wenzel
 
 Heaps over linearly ordered types.  See also Chris Okasaki: "Purely
 Functional Data Structures" (Chapter 3), Cambridge University Press,
--- a/src/Pure/General/integer.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/integer.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/General/integer.ML
-    ID:         $Id$
     Author:     Florian Haftmann, TU Muenchen
 
 Unbounded integers.
--- a/src/Pure/General/markup.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/markup.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -91,19 +91,11 @@
   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
@@ -269,24 +261,21 @@
 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 *)
--- a/src/Pure/General/markup.scala	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/markup.scala	Thu Jan 29 12:24:00 2009 +0000
@@ -119,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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/name_space.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -3,9 +3,10 @@
 
 Generic name spaces with declared and hidden entries.  Unknown names
 are considered global; no support for absolute addressing.
+Cf. Pure/General/binding.ML
 *)
 
-type bstring = string;    (*names to be bound*)
+type bstring = string;    (*simple names to be bound -- legacy*)
 type xstring = string;    (*external names*)
 
 signature NAME_SPACE =
@@ -31,8 +32,8 @@
   val merge: T * T -> T
   type naming
   val default_naming: naming
-  val declare: naming -> Binding.T -> T -> string * T
-  val full_name: naming -> Binding.T -> string
+  val declare: naming -> binding -> T -> string * T
+  val full_name: naming -> binding -> string
   val external_names: naming -> string -> string list
   val path_of: naming -> string
   val add_path: string -> naming -> naming
@@ -41,7 +42,7 @@
   val sticky_prefix: string -> naming -> naming
   type 'a table = T * 'a Symtab.table
   val empty_table: 'a table
-  val bind: naming -> Binding.T * 'a
+  val bind: naming -> binding * 'a
     -> 'a table -> string * 'a table (*exception Symtab.DUP*)
   val merge_tables: ('a * 'a -> bool) -> 'a table * 'a table -> 'a table
   val join_tables: (string -> 'a * 'a -> 'a)
--- a/src/Pure/General/ord_list.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/ord_list.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/General/ord_list.ML
-    ID:         $Id$
     Author:     Makarius
 
 Ordered lists without duplicates -- a light-weight representation of
--- a/src/Pure/General/output.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/output.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/General/output.ML
-    ID:         $Id$
     Author:     Makarius, Hagia Maria Sion Abbey (Jerusalem)
 
 Output channels and timing messages.
--- a/src/Pure/General/path.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/path.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/General/path.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Abstract algebra of file paths (external encoding in Unix style).
--- a/src/Pure/General/pretty.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/pretty.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/General/pretty.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Author:     Markus Wenzel, TU Munich
 
--- a/src/Pure/General/print_mode.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/print_mode.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/General/print_mode.ML
-    ID:         $Id$
     Author:     Makarius
 
 Generic print mode as thread-local value derived from global template;
--- a/src/Pure/General/properties.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/properties.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/General/properties.ML
-    ID:         $Id$
     Author:     Makarius
 
 Property lists.
--- a/src/Pure/General/queue.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/queue.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/General/queue.ML
-    ID:         $Id$
     Author:     Makarius
 
 Efficient queues.
--- a/src/Pure/General/scan.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/scan.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/General/scan.ML
-    ID:         $Id$
     Author:     Markus Wenzel and Tobias Nipkow, TU Muenchen
 
 Generic scanners (for potentially infinite input).
--- a/src/Pure/General/secure.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/secure.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/General/secure.ML
-    ID:         $Id$
     Author:     Makarius
 
 Secure critical operations.
--- a/src/Pure/General/seq.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/seq.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/General/seq.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Author:     Markus Wenzel, TU Munich
 
--- a/src/Pure/General/source.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/source.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/General/source.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Coalgebraic data sources -- efficient purely functional input streams.
--- a/src/Pure/General/stack.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/stack.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/General/stack.ML
-    ID:         $Id$
     Author:     Makarius
 
 Non-empty stacks.
--- a/src/Pure/General/swing.scala	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/swing.scala	Thu Jan 29 12:24:00 2009 +0000
@@ -10,9 +10,13 @@
 
 object Swing
 {
-  def now(body: => Unit) =
-    SwingUtilities.invokeAndWait(new Runnable { def run = body })
+  def now(body: => Unit) {
+    if (SwingUtilities.isEventDispatchThread) body
+    else SwingUtilities.invokeAndWait(new Runnable { def run = body })
+  }
 
-  def later(body: => Unit) =
-    SwingUtilities.invokeLater(new Runnable { def run = body })
+  def later(body: => Unit) {
+    if (SwingUtilities.isEventDispatchThread) body
+    else SwingUtilities.invokeLater(new Runnable { def run = body })
+  }
 }
--- a/src/Pure/General/symbol.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/symbol.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/General/symbol.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Generalized characters with infinitely many named symbols.
--- a/src/Pure/General/symbol.scala	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/symbol.scala	Thu Jan 29 12:24:00 2009 +0000
@@ -78,9 +78,9 @@
 
   /** Symbol interpretation **/
 
-  class Interpretation(isabelle_system: IsabelleSystem) {
-
-    private var symbols = new HashMap[String, HashMap[String, String]]
+  class Interpretation(symbol_decls: Iterator[String])
+  {
+    private val symbols = new HashMap[String, HashMap[String, String]]
     private var decoder: Recoder = null
     private var encoder: Recoder = null
 
@@ -94,10 +94,11 @@
     private val blank_pattern = compile(""" \s+ """)
     private val key_pattern = compile(""" (.+): """)
 
-    private def read_line(line: String) = {
-      def err() = error("Bad symbol specification (line " + line + ")")
+    private def read_decl(decl: String) = {
+      def err() = error("Bad symbol declaration: " + decl)
 
-      def read_props(props: List[String], tab: HashMap[String, String]): Unit = {
+      def read_props(props: List[String], tab: HashMap[String, String])
+      {
         props match {
           case Nil => ()
           case _ :: Nil => err()
@@ -112,8 +113,8 @@
         }
       }
 
-      if (!empty_pattern.matcher(line).matches) {
-        blank_pattern.split(line).toList match {
+      if (!empty_pattern.matcher(decl).matches) {
+        blank_pattern.split(decl).toList match {
           case Nil => err()
           case symbol :: props => {
             val tab = new HashMap[String, String]
@@ -124,13 +125,6 @@
       }
     }
 
-    private def read_symbols(path: String) = {
-      val file = new File(isabelle_system.platform_path(path))
-      if (file.canRead) {
-        for (line <- Source.fromFile(file).getLines) read_line(line)
-      }
-    }
-
 
     /* init tables */
 
@@ -154,9 +148,7 @@
 
     /* constructor */
 
-    read_symbols("$ISABELLE_HOME/etc/symbols")
-    read_symbols("$ISABELLE_HOME_USER/etc/symbols")
+    symbol_decls.foreach(read_decl)
     init_recoders()
   }
-
 }
--- a/src/Pure/General/symbol_pos.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/symbol_pos.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/General/symbol_pos.ML
-    ID:         $Id$
     Author:     Makarius
 
 Symbols with explicit position information.
--- a/src/Pure/General/table.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/table.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/General/table.ML
-    ID:         $Id$
     Author:     Markus Wenzel and Stefan Berghofer, TU Muenchen
 
 Generic tables.  Efficient purely functional implementation using
--- a/src/Pure/General/url.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/url.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/General/url.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Basic URLs, see RFC 1738 and RFC 2396.
--- a/src/Pure/General/xml.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/xml.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/General/xml.ML
-    ID:         $Id$
     Author:     David Aspinall, Stefan Berghofer and Markus Wenzel
 
 Basic support for XML.
--- a/src/Pure/General/yxml.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/yxml.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/General/yxml.ML
-    ID:         $Id$
     Author:     Makarius
 
 Efficient text representation of XML trees using extra characters X
--- a/src/Pure/General/yxml.scala	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/General/yxml.scala	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/IsaMakefile	Thu Jan 29 12:24:00 2009 +0000
@@ -41,7 +41,7 @@
   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/old_locale.ML	\
+  Isar/locale.ML Isar/method.ML Isar/net_rules.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		\
@@ -75,7 +75,7 @@
   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_header.ML Thy/thy_info.ML Thy/thy_load.ML Thy/thy_output.ML	\
-  Thy/thy_syntax.ML Tools/ROOT.ML Tools/invoke.ML			\
+  Thy/thy_syntax.ML Tools/ROOT.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	\
@@ -126,9 +126,9 @@
 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/outer_keyword.scala Thy/thy_header.scala				\
-  Tools/isabelle_process.scala Tools/isabelle_syntax.scala		\
-  Tools/isabelle_system.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
 
 
 SCALA_TARGET = $(ISABELLE_HOME)/lib/classes/Pure.jar
--- a/src/Pure/Isar/ROOT.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/ROOT.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -53,7 +53,6 @@
 (*local theories and targets*)
 use "local_theory.ML";
 use "overloading.ML";
-use "old_locale.ML";
 use "locale.ML";
 use "class_target.ML";
 use "theory_target.ML";
--- a/src/Pure/Isar/antiquote.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/antiquote.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/antiquote.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Text with antiquotations of inner items (terms, types etc.).
--- a/src/Pure/Isar/args.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/args.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/args.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Parsing with implicit value assigment.  Concrete argument syntax of
@@ -35,7 +34,7 @@
   val name_source: T list -> string * T list
   val name_source_position: T list -> (SymbolPos.text * Position.T) * T list
   val name: T list -> string * T list
-  val binding: T list -> Binding.T * T list
+  val binding: T list -> binding * T list
   val alt_name: T list -> string * T list
   val symbol: T list -> string * T list
   val liberal_name: T list -> string * T list
@@ -66,8 +65,8 @@
   val parse1: (string -> bool) -> OuterLex.token list -> T list * OuterLex.token list
   val attribs: (string -> string) -> T list -> src list * T list
   val opt_attribs: (string -> string) -> T list -> src list * T list
-  val thm_name: (string -> string) -> string -> T list -> (Binding.T * src list) * T list
-  val opt_thm_name: (string -> string) -> string -> T list -> (Binding.T * src list) * T list
+  val thm_name: (string -> string) -> string -> T list -> (binding * src list) * T list
+  val opt_thm_name: (string -> string) -> string -> T list -> (binding * src list) * T list
   val syntax: string -> ('b * T list -> 'a * ('b * T list)) -> src -> 'b -> 'a * 'b
   val context_syntax: string -> (Context.generic * T list -> 'a * (Context.generic * T list)) ->
     src -> Proof.context -> 'a * Proof.context
--- a/src/Pure/Isar/attrib.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/attrib.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -7,7 +7,7 @@
 signature ATTRIB =
 sig
   type src = Args.src
-  type binding = Binding.T * src list
+  type binding = binding * src list
   val empty_binding: binding
   val print_attributes: theory -> unit
   val intern: theory -> xstring -> string
@@ -54,7 +54,7 @@
 
 type src = Args.src;
 
-type binding = Binding.T * src list;
+type binding = binding * src list;
 val empty_binding: binding = (Binding.empty, []);
 
 
--- a/src/Pure/Isar/auto_bind.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/auto_bind.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/auto_bind.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Automatic bindings of Isar text elements.
--- a/src/Pure/Isar/calculation.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/calculation.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/calculation.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Generic calculational proofs.
@@ -98,8 +97,8 @@
     ("sym", sym_att, "declaration of symmetry rule"),
     ("symmetric", Attrib.no_args symmetric, "resolution with symmetry rule")] #>
   PureThy.add_thms
-   [(("", transitive_thm), [trans_add]),
-    (("", symmetric_thm), [sym_add])] #> snd));
+   [((Binding.empty, transitive_thm), [trans_add]),
+    ((Binding.empty, symmetric_thm), [sym_add])] #> snd));
 
 
 
--- a/src/Pure/Isar/class.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/class.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,7 +1,7 @@
 (*  Title:      Pure/Isar/ML
     Author:     Florian Haftmann, TU Muenchen
 
-Type classes derived from primitive axclasses and locales - interfaces
+Type classes derived from primitive axclasses and locales - interfaces.
 *)
 
 signature CLASS =
@@ -24,141 +24,207 @@
 
 open Class_Target;
 
-(** rule calculation **)
-
-fun calculate_axiom thy sups base_sort assm_axiom param_map class =
-  case Old_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 (fst o rules thy) sups
-          @ the_list assm_axiom;
-      in intro
-        |> instantiate thy [class]
-        |> (fn thm => thm OF axiom_premises)
-        |> Drule.standard'
-        |> Thm.close_derivation
-        |> SOME
-      end;
-
-fun raw_morphism thy class param_map some_axiom =
-  let
-    val ctxt = ProofContext.init thy;
-    val some_wit = case some_axiom
-     of SOME axiom => SOME (Element.prove_witness ctxt
-          (Logic.unvarify (Thm.prop_of axiom))
-            (ALLGOALS (ProofContext.fact_tac [axiom])))
-      | NONE => NONE;
-    val instT = Symtab.empty |> Symtab.update ("'a", TFree ("'a", [class]));
-    val inst = Symtab.make ((map o apsnd) Const param_map);
-  in case some_wit
-   of SOME wit => Element.inst_morphism thy (instT, inst)
-        $> Morphism.binding_morphism (Binding.add_prefix false (class_prefix class))
-        $> Element.satisfy_morphism [wit]
-    | NONE => Element.inst_morphism thy (instT, inst)
-        $> Morphism.binding_morphism (Binding.add_prefix false (class_prefix class))
-  end;
-
-fun calculate_rules thy morph 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 = Old_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 snd o rules thy) sups;
-    val locale_dests = map Drule.standard' (Old_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;
-
-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;
-
-
-(** define classes **)
+(** class definitions **)
 
 local
 
-fun gen_class_spec prep_class process_expr thy raw_supclasses raw_elems =
+(* calculating class-related rules including canonical interpretation *)
+
+fun calculate thy class sups base_sort param_map assm_axiom =
   let
-    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 empty_ctxt = ProofContext.init thy;
+
+    (* instantiation of canonical interpretation *)
+    val aT = TFree (Name.aT, base_sort);
+    val param_map_const = (map o apsnd) Const param_map;
+    val param_map_inst = (map o apsnd)
+      (Const o apsnd (map_atyps (K aT))) param_map;
+    val const_morph = Element.inst_morphism thy
+      (Symtab.empty, Symtab.make param_map_inst);
+    val (([props], [(_, inst_morph)], export_morph), _) = empty_ctxt
+      |> Expression.cert_goal_expression ([(class, (("", false),
+           Expression.Named param_map_const))], []);
+
+    (* 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;
+
+    (* 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;
+
+    (* assm_intro *)
+    fun prove_assm_intro thm = 
+      let
+        val ((_, [thm']), _) = Variable.import_thms true [thm] empty_ctxt;
+        val thm'' = Morphism.thm (const_morph $> eq_morph) thm';
+        val tac = ALLGOALS (ProofContext.fact_tac [thm'']);
+      in Goal.prove_global thy [] [] (Thm.prop_of thm'') (K tac) end;
+    val assm_intro = Option.map prove_assm_intro
+      (fst (Locale.intros_of thy class));
+
+    (* 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 const_morph
+          ((map_types o map_atyps) (K aT) prop), of_class_prop_concl);
+    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);
+
+  in (base_morph, morph, export_morph, axiom, assm_intro, of_class) end;
+
+
+(* reading and processing class specifications *)
+
+fun prep_class_elems prep_decl thy supexpr sups proto_base_sort raw_elems =
+  let
+
+    (* user space type system: only permits 'a type variable, improves towards 'a *)
+    val base_constraints = (map o apsnd)
+      (map_type_tfree (K (TVar ((Name.aT, 0), proto_base_sort))) o fst o snd)
+        (these_operations thy sups);
+    val reject_bcd_etc = (map o map_atyps) (fn T as TFree (v, sort) =>
+          if v = Name.aT then T
+          else error ("No type variable other than " ^ Name.aT ^ " allowed in class specification")
+      | T => T);
+    fun singleton_fixate thy algebra Ts =
+      let
+        fun extract f = (fold o fold_atyps) f Ts [];
+        val tfrees = extract
+          (fn TFree (v, sort) => insert (op =) (v, sort) | _ => I);
+        val inferred_sort = extract
+          (fn TVar (_, sort) => curry (Sorts.inter_sort algebra) sort | _ => I);
+        val fixate_sort = if null tfrees then inferred_sort
+          else let val a_sort = (snd o the_single) tfrees
+          in if Sorts.sort_le algebra (a_sort, inferred_sort)
+            then Sorts.inter_sort algebra (a_sort, inferred_sort)
+            else error ("Type inference imposes additional sort constraint "
+              ^ Syntax.string_of_sort_global thy inferred_sort
+              ^ " of type parameter " ^ Name.aT ^ " of sort "
+              ^ Syntax.string_of_sort_global thy a_sort)
+          end
+      in (map o map_atyps) (K (TFree (Name.aT, fixate_sort))) Ts end;
+    fun add_typ_check level name f = Context.proof_map (Syntax.add_typ_check level name (fn xs => fn ctxt =>
+      let val xs' = f xs in if eq_list (op =) (xs, xs') then NONE else SOME (xs', ctxt) end));
+
+    (* preproceesing elements, retrieving base sort from type-checked elements *)
+    val ((_, _, inferred_elems), _) = ProofContext.init thy
+      |> fold (ProofContext.add_const_constraint o apsnd SOME) base_constraints
+      |> redeclare_operations thy sups
+      |> add_typ_check 10 "reject_bcd_etc" reject_bcd_etc
+      |> add_typ_check ~10 "singleton_fixate" (singleton_fixate thy (Sign.classes_of thy))
+      |> prep_decl supexpr raw_elems;
+    fun fold_element_types f (Element.Fixes fxs) = fold (fn (_, SOME T, _) => f T) fxs
+      | fold_element_types f (Element.Constrains cnstrs) = fold (f o snd) cnstrs
+      | fold_element_types f (Element.Assumes assms) = fold (fold (fn (t, ts) =>
+          fold_types f t #> (fold o fold_types) f ts) o snd) assms
+      | fold_element_types f (Element.Defines _) =
+          error ("\"defines\" element not allowed in class specification.")
+      | fold_element_types f (Element.Notes _) =
+          error ("\"notes\" element not allowed in class specification.");
+    val base_sort = if null inferred_elems then proto_base_sort else
+      case (fold o fold_element_types) Term.add_tfreesT inferred_elems []
+       of [] => error "No type variable in class specification"
+        | [(_, sort)] => sort
+        | _ => error "Multiple type variables in class specification"
+
+  in (base_sort, inferred_elems) end;
+
+val cert_class_elems = prep_class_elems Expression.cert_declaration;
+val read_class_elems = prep_class_elems Expression.cert_read_declaration;
+
+fun prep_class_spec prep_class prep_class_elems thy raw_supclasses raw_elems =
+  let
+
+    (* prepare import *)
+    val inter_sort = curry (Sorts.inter_sort (Sign.classes_of thy));
+    val sups = map (prep_class thy) raw_supclasses
+      |> Sign.minimize_sort thy;
+    val _ = case filter_out (is_class thy) sups
+     of [] => ()
+      | no_classes => error ("No proper classes: " ^ commas (map quote no_classes));
+          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 thy) sups);
-    val suplocales = map Old_Locale.Locale sups;
-    val supexpr = Old_Locale.Merge suplocales;
-    val supparams = (map fst o Old_Locale.parameters_of_expr thy) supexpr;
-    val mergeexpr = Old_Locale.Merge suplocales;
-    val constrain = Element.Constrains ((map o apsnd o map_atyps)
-      (K (TFree (Name.aT, base_sort))) supparams);
+    val supexpr = (map (fn sup => (sup, (("", false), Expression.Positional [])))
+      sups, []);
+    val given_basesort = fold inter_sort (map (base_sort thy) sups) [];
+
+    (* infer types and base sort *)
+    val (base_sort, inferred_elems) = prep_class_elems thy supexpr sups
+      given_basesort raw_elems;
+    val sup_sort = inter_sort base_sort sups
+
+    (* process elements as class specification *)
+    val class_ctxt = begin sups base_sort (ProofContext.init thy)
+    val ((_, _, syntax_elems), _) = class_ctxt
+      (*#> fold (Variable.declare_constraints o Free) ((map o apsnd o map_atyps)
+          (K (TFree (Name.aT, base_sort))) supparams) FIXME really necessary?*)
+            (*FIXME should constraints be issued in begin?*)
+      |> Expression.cert_declaration supexpr inferred_elems;
+    fun check_vars e vs = if null vs
+      then error ("No type variable in part of specification element "
+        ^ (Pretty.output o Pretty.chunks) (Element.pretty_ctxt class_ctxt e))
+      else ();
+    fun check_element (e as Element.Fixes fxs) =
+          map (fn (_, SOME T, _) => check_vars e (Term.add_tfreesT T [])) fxs
+      | check_element (e as Element.Assumes assms) =
+          maps (fn (_, ts_pss) => map
+            (fn (t, _) => check_vars e (Term.add_tfrees t [])) ts_pss) assms
+      | check_element e = [()];
+    val _ = map check_element syntax_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
-      |> Old_Locale.cert_expr supexpr [constrain]
-      |> snd
-      |> begin sups base_sort
-      |> process_expr Old_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 [];
+    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 read_class_spec = gen_class_spec Sign.intern_class Old_Locale.read_expr;
-val check_class_spec = gen_class_spec (K I) Old_Locale.cert_expr;
+  in (((sups, supparam_names), (sup_sort, base_sort, supexpr)), (constrain :: elems, global_syntax)) end;
+
+val cert_class_spec = prep_class_spec (K I) cert_class_elems;
+val read_class_spec = prep_class_spec Sign.intern_class read_class_elems;
+
+
+(* class establishment *)
 
 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 (Old_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;
@@ -172,7 +238,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))
@@ -180,12 +246,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 = Old_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;
@@ -194,45 +261,38 @@
     |> 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
-    |> Old_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 thy => raw_morphism thy class param_map axiom)
-    #-> (fn morph =>
-        `(fn thy => calculate_rules thy morph sups base_sort param_map axiom class)
-    #-> (fn (assm_intro, of_class) =>
-        register class sups params base_sort inst
-          morph axiom assm_intro of_class
-    #> fold (note_assm_intro class) (the_list assm_intro)))))
+    |-> (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*)
 
@@ -244,42 +304,30 @@
 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"
+    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 _ = 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 (these_params thy [sub]);
-    val sup_params = map fst (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 =
-      Old_Locale.global_asms_of thy sup
-      |> maps snd
-      |> try the_single
-      |> Option.map (ObjectLogic.ensure_propT thy);
-    fun after_qed some_thm =
-      LocalTheory.theory (prove_subclass_relation (sub, sup) some_thm)
-      #> (TheoryTarget.init (SOME sub) o ProofContext.theory_of);
-  in
-    do_proof after_qed sublocale_prop lthy
-  end;
+    val (sub, sup) = AxClass.cert_classrel thy (proto_sub, proto_sup)
 
-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, [])]];
+    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 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;
+fun user_proof after_qed some_prop =
+  Element.witness_proof (after_qed o try the_single o the_single)
+    [the_list some_prop];
+
+fun tactic_proof tac after_qed some_prop ctxt =
+  after_qed (Option.map
+    (fn prop => Element.prove_witness ctxt prop tac) some_prop) ctxt;
 
 in
 
@@ -289,6 +337,4 @@
 
 end; (*local*)
 
-
 end;
-
--- a/src/Pure/Isar/class_target.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/class_target.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -8,8 +8,18 @@
 sig
   (*classes*)
   val register: class -> class list -> ((string * typ) * (string * typ)) list
-    -> sort -> term list -> morphism
-    -> thm option -> thm option -> thm -> theory -> theory
+    -> 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 these_operations: theory -> sort -> (string * (class * (typ * term))) list
+  val print_classes: theory -> unit
 
   val begin: class list -> sort -> Proof.context -> Proof.context
   val init: class -> theory -> Proof.context
@@ -17,21 +27,9 @@
     -> (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
-
-  val intro_classes_tac: thm list -> tactic
-  val default_intro_tac: Proof.context -> thm list -> tactic
-  val prove_class_interpretation: class -> term list -> (class * string) list
-    -> thm list -> thm list -> theory -> theory
-  val prove_subclass_relation: 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 these_defs: theory -> sort -> thm list
-  val base_sort: theory -> class -> sort
-  val rules: theory -> class -> thm option * thm
-  val print_classes: theory -> unit
+  val redeclare_operations: theory -> sort -> Proof.context -> Proof.context
 
   (*instances*)
   val init_instantiation: string list * (string * sort) list * sort
@@ -50,13 +48,14 @@
   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;
@@ -64,36 +63,6 @@
 structure Class_Target : CLASS_TARGET =
 struct
 
-(*temporary adaption code to mediate between old and new locale code*)
-
-structure Locale_Layer =
-struct
-
-val init = Old_Locale.init;
-val parameters_of = Old_Locale.parameters_of;
-val intros = Old_Locale.intros;
-val dests = Old_Locale.dests;
-val get_interpret_morph = Old_Locale.get_interpret_morph;
-val Locale = Old_Locale.Locale;
-val extern = Old_Locale.extern;
-val intro_locales_tac = Old_Locale.intro_locales_tac;
-
-fun prove_interpretation tac prfx_atts expr inst =
-  Old_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) =
-  Old_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;
-
-end;
-
-
 (** primitive axclass and instance commands **)
 
 fun axclass_cmd (class, raw_superclasses) raw_specs thy =
@@ -146,9 +115,7 @@
   consts: (string * string) list
     (*locale parameter ~> constant name*),
   base_sort: sort,
-  inst: term list
-    (*canonical interpretation*),
-  base_morph: Morphism.morphism
+  base_morph: morphism
     (*static part of canonical morphism*),
   assm_intro: thm option,
   of_class: thm,
@@ -160,22 +127,22 @@
 
 };
 
-fun rep_class_data (ClassData d) = d;
-fun mk_class_data ((consts, base_sort, inst, base_morph, assm_intro, of_class, axiom),
+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, inst = inst,
+  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, inst, base_morph, assm_intro,
+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, inst, base_morph, assm_intro, of_class, axiom),
+  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, inst = inst, base_morph = base_morph, assm_intro = assm_intro,
+    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 = _, inst = _, base_morph = _, assm_intro = _,
+  ClassData { consts = _, base_sort = _, base_morph = _, assm_intro = _,
     of_class = _, axiom = _, defs = defs2, operations = operations2 }) =
-  mk_class_data ((consts, base_sort, inst, base_morph, assm_intro, of_class, axiom),
+  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)));
 
@@ -200,8 +167,7 @@
 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;
+val heritage = Graph.all_preds o ClassData.get;
 
 fun these_params thy =
   let
@@ -215,34 +181,22 @@
       end;
   in maps params o ancestry thy end;
 
-fun these_assm_intros thy =
+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 (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;
-
-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 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 class_prefix = Logic.const_of_class o Sign.base_name;
-
-fun class_binding_morph class =
-  Binding.map_prefix (K (Binding.add_prefix false (class_prefix class)));
-
-fun morphism thy class =
-  let
-    val { base_morph, defs, ... } = the_class_data thy class;
-  in
-    base_morph 
-    $> Morphism.term_morphism (MetaSimplifier.rewrite_term thy defs [])
-    $> Morphism.thm_morphism (MetaSimplifier.rewrite_rule defs)
-  end;
+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
@@ -264,8 +218,6 @@
       (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_Layer.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,
@@ -282,17 +234,24 @@
 
 (* updaters *)
 
-fun register class superclasses params base_sort inst morph
+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,
-          inst, morph, assm_intro, of_class, axiom), ([], operations)))
-      #> fold (curry Graph.add_edge class) superclasses;
+          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 =
   let
     val base_sort = base_sort thy class;
@@ -307,91 +266,53 @@
       (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;
 
-
-(** tactics and methods **)
-
-fun prove_class_interpretation class inst consts hyp_facts def_facts thy =
+fun register_subclass (sub, sup) some_dep_morph some_wit export 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_Layer.intro_locales_tac false ctxt []) n));
-    val prfx = class_prefix class;
-  in
-    thy
-    |> fold2 add_constraint (map snd consts) no_constraints
-    |> Locale_Layer.prove_interpretation tac (class_binding_morph class) (Locale_Layer.Locale class)
-          (map SOME inst, map (pair (Attrib.empty_binding) o Thm.prop_of) def_facts)
-    |> snd
-    |> fold2 add_constraint (map snd consts) constraints
-  end;
-
-fun prove_subclass_relation (sub, sup) some_thm thy =
-  let
-    val of_class = (snd o rules 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 fst o rules thy) sub)
-      |> Thm.close_derivation;
+    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 = EVERY (map (TRYALL o 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
-    |> Locale_Layer.prove_interpretation_in (ALLGOALS (ProofContext.fact_tac (the_list some_thm)))
-         I (sub, Locale_Layer.Locale sup)
     |> ClassData.map (Graph.add_edge (sub, sup))
-  end;
-
-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
+    |> activate_defs sub (these_defs thy diff_sort)
+    |> fold (fn dep_morph => Locale.add_dependency sub (sup,
+        dep_morph $> Element.satisfy_morphism (the_list some_wit) $> export))
+          (the_list some_dep_morph)
+    |> (fn thy => fold_rev Locale.activate_global_facts
+      (Locale.get_registrations thy) thy)
   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")]));
-
 
 (** classes and class target **)
 
 (* class context syntax *)
 
-fun synchronize_class_syntax sups base_sort ctxt =
+fun these_unchecks thy =
+  map (fn (c, (_, (ty, t))) => (t, Const (c, ty))) o these_operations thy;
+
+fun redeclare_const thy c =
+  let val b = Sign.base_name c
+  in Sign.intern_const thy b = c ? Variable.declare_const (b, c) end;
+
+fun synchronize_class_syntax sort base_sort ctxt =
   let
     val thy = ProofContext.theory_of ctxt;
     val algebra = Sign.classes_of thy;
-    val operations = these_operations thy sups;
+    val operations = these_operations thy sort;
     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)
@@ -404,10 +325,10 @@
           | 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;
+    val unchecks = these_unchecks thy sort;
   in
     ctxt
-    |> fold declare_const primary_constraints
+    |> fold (redeclare_const thy o fst) primary_constraints
     |> Overloading.map_improvable_syntax (K (((primary_constraints, secondary_constraints),
         (((improve, subst), true), unchecks)), false))
     |> Overloading.set_primary_constraints
@@ -419,67 +340,64 @@
     val base_sort = base_sort thy class;
   in synchronize_class_syntax [class] base_sort ctxt end;
 
-fun begin sups base_sort ctxt =
+fun redeclare_operations thy sort =
+  fold (redeclare_const thy o fst) (these_operations thy sort);
+
+fun begin sort base_sort ctxt =
   ctxt
   |> Variable.declare_term
       (Logic.mk_type (TFree (Name.aT, base_sort)))
-  |> synchronize_class_syntax sups base_sort
+  |> synchronize_class_syntax sort base_sort
   |> Overloading.add_improvable_syntax;
 
 fun init class thy =
   thy
-  |> Locale_Layer.init class
+  |> 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;
-    val morph = 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 morph = morphism thy class;
+    val b = Morphism.binding morph (Binding.name c);
+    val b_def = Morphism.binding morph (Binding.name (c ^ "_dict"));
+    val c' = Sign.full_name thy b;
     val dict' = Morphism.term morph dict;
     val ty' = Term.fastype_of dict';
-    val ty'' = Type.strip_sorts ty';
-    val def_eq = Logic.mk_equals (Const (c', ty'), dict');
-    fun get_axiom thy = ((Thm.varifyT o Thm.axiom thy) c', thy);
+    val def_eq = Logic.mk_equals (Const (c', ty'), dict')
+      |> map_types Type.strip_sorts;
   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]
-      #> register_operation class (c', (dict', SOME (Thm.symmetric def')))
-      #> PureThy.store_thm (c ^ "_raw", def')
-      #> snd)
-    |> Sign.restore_naming thy
+    thy
+    |> Sign.declare_const pos ((b, Type.strip_sorts ty'), mx)
+    |> snd
+    |> Thm.add_def false false (b_def, def_eq)
+    |>> Thm.varifyT
+    |-> (fn def_thm => PureThy.store_thm (b_def, def_thm)
+      #> snd
+      #> register_operation class (c', (dict', SOME (Thm.symmetric def_thm))))
     |> 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 morph = morphism thy class;
+    val unchecks = these_unchecks thy [class];
+    val b = Morphism.binding morph (Binding.name c);
+    val c' = Sign.full_name thy b;
     val rhs' = Pattern.rewrite_term thy unchecks [] rhs;
-    val rhs'' = map_types Logic.varifyT rhs';
     val ty' = Term.fastype_of rhs';
+    val rhs'' = map_types ((*Type.strip_sorts o *)Logic.varifyT) rhs';
   in
-    thy'
-    |> Sign.add_abbrev (#1 prmode) pos (Binding.name c, map_types Type.strip_sorts rhs'') |> snd
+    thy
+    |> Sign.add_abbrev (#1 prmode) pos (b, 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;
 
 
@@ -557,7 +475,7 @@
 
 fun type_name "*" = "prod"
   | type_name "+" = "sum"
-  | type_name s = sanatize_name (NameSpace.base s); (*FIXME*)
+  | type_name s = sanatize_name (NameSpace.base s);
 
 fun resort_terms pp algebra consts constraints ts =
   let
@@ -576,7 +494,8 @@
   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)
+    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;
@@ -673,5 +592,34 @@
       (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 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/constdefs.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/constdefs.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -8,12 +8,12 @@
 
 signature CONSTDEFS =
 sig
-  val add_constdefs: (Binding.T * string option) list *
-    ((Binding.T * string option * mixfix) option *
+  val add_constdefs: (binding * string option) list *
+    ((binding * string option * mixfix) option *
       (Attrib.binding * string)) list -> theory -> theory
-  val add_constdefs_i: (Binding.T * typ option) list *
-    ((Binding.T * typ option * mixfix) option *
-      ((Binding.T * attribute list) * term)) list -> theory -> theory
+  val add_constdefs_i: (binding * typ option) list *
+    ((binding * typ option * mixfix) option *
+      ((binding * attribute list) * term)) list -> theory -> theory
 end;
 
 structure Constdefs: CONSTDEFS =
@@ -52,7 +52,7 @@
     val thy' =
       thy
       |> Sign.add_consts_i [(c, T, mx)]
-      |> PureThy.add_defs false [((name, def), atts)]
+      |> PureThy.add_defs false [((Binding.name name, def), atts)]
       |-> (fn [thm] => Code.add_default_eqn thm);
   in ((c, T), thy') end;
 
--- a/src/Pure/Isar/context_rules.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/context_rules.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/context_rules.ML
-    ID:         $Id$
     Author:     Stefan Berghofer and Markus Wenzel, TU Muenchen
 
 Declarations of intro/elim/dest rules in Pure (see also
@@ -199,7 +198,7 @@
 val dest_query  = rule_add elim_queryK Tactic.make_elim;
 
 val _ = Context.>> (Context.map_theory
-  (snd o PureThy.add_thms [(("", Drule.equal_intr_rule), [intro_query NONE])]));
+  (snd o PureThy.add_thms [((Binding.empty, Drule.equal_intr_rule), [intro_query NONE])]));
 
 
 (* concrete syntax *)
--- a/src/Pure/Isar/element.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/element.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -9,11 +9,11 @@
 sig
   datatype ('typ, 'term) stmt =
     Shows of (Attrib.binding * ('term * 'term list) list) list |
-    Obtains of (Binding.T * ((Binding.T * 'typ option) list * 'term list)) list
+    Obtains of (binding * ((binding * 'typ option) list * 'term list)) list
   type statement = (string, string) stmt
   type statement_i = (typ, term) stmt
   datatype ('typ, 'term, 'fact) ctxt =
-    Fixes of (Binding.T * 'typ option * mixfix) list |
+    Fixes of (binding * 'typ option * mixfix) list |
     Constrains of (string * 'typ) list |
     Assumes of (Attrib.binding * ('term * 'term list) list) list |
     Defines of (Attrib.binding * ('term * 'term list)) list |
@@ -23,46 +23,26 @@
   val facts_map: (('typ, 'term, 'fact) ctxt -> ('a, 'b, 'c) ctxt) ->
    (Attrib.binding * ('fact * Attrib.src list) list) list ->
    (Attrib.binding * ('c * Attrib.src list) list) list
-  val map_ctxt': {binding: Binding.T -> Binding.T,
-    var: Binding.T * mixfix -> Binding.T * mixfix,
-    typ: 'typ -> 'a, term: 'term -> 'b, pat: 'term -> 'b, fact: 'fact -> 'c,
-    attrib: Attrib.src -> Attrib.src} -> ('typ, 'term, 'fact) ctxt -> ('a, 'b, 'c) ctxt
-  val map_ctxt: {binding: Binding.T -> Binding.T,
-    var: Binding.T * mixfix -> Binding.T * mixfix,
-    typ: 'typ -> 'a, term: 'term -> 'b, fact: 'fact -> 'c,
-    attrib: Attrib.src -> Attrib.src} -> ('typ, 'term, 'fact) ctxt -> ('a, 'b, 'c) ctxt
+  val map_ctxt: {binding: binding -> binding, typ: 'typ -> 'a, term: 'term -> 'b,
+    pattern: 'term -> 'b, fact: 'fact -> 'c, attrib: Attrib.src -> Attrib.src} ->
+    ('typ, 'term, 'fact) ctxt -> ('a, 'b, 'c) ctxt
   val map_ctxt_attrib: (Attrib.src -> Attrib.src) ->
     ('typ, 'term, 'fact) ctxt -> ('typ, 'term, 'fact) ctxt
   val morph_ctxt: morphism -> context_i -> context_i
-  val params_of: context_i -> (string * typ) list
-  val prems_of: context_i -> term list
-  val facts_of: theory -> context_i -> (Attrib.binding * (thm list * Attrib.src list) list) list
   val pretty_stmt: Proof.context -> statement_i -> Pretty.T list
   val pretty_ctxt: Proof.context -> context_i -> Pretty.T list
   val pretty_statement: Proof.context -> string -> thm -> Pretty.T
   type witness
-  val map_witness: (term * thm -> term * thm) -> witness -> witness
+  val prove_witness: Proof.context -> term -> tactic -> witness
+  val witness_proof: (witness list list -> Proof.context -> Proof.context) ->
+    term list list -> Proof.context -> Proof.state
+  val witness_proof_eqs: (witness list list -> thm list -> Proof.context -> Proof.context) ->
+    term list list -> term list -> Proof.context -> Proof.state
+  val witness_local_proof: (witness list list -> Proof.state -> Proof.state) ->
+    string -> term list list -> Proof.context -> bool -> Proof.state -> Proof.state
   val morph_witness: morphism -> witness -> witness
-  val witness_prop: witness -> term
-  val witness_hyps: witness -> term list
-  val assume_witness: theory -> term -> witness
-  val prove_witness: Proof.context -> term -> tactic -> witness
-  val close_witness: witness -> witness
   val conclude_witness: witness -> thm
-  val mark_witness: term -> term
-  val make_witness: term -> thm -> witness
-  val dest_witness: witness -> term * thm
-  val transfer_witness: theory -> witness -> witness
-  val refine_witness: Proof.state -> Proof.state Seq.seq
   val pretty_witness: Proof.context -> witness -> Pretty.T
-  val rename: (string * (string * mixfix option)) list -> string -> string
-  val rename_var_name: (string * (string * mixfix option)) list ->
-    string * mixfix -> string * mixfix
-  val rename_var: (string * (string * mixfix option)) list ->
-    Binding.T * mixfix -> Binding.T * mixfix
-  val rename_term: (string * (string * mixfix option)) list -> term -> term
-  val rename_thm: (string * (string * mixfix option)) list -> thm -> thm
-  val rename_morphism: (string * (string * mixfix option)) list -> morphism
   val instT_type: typ Symtab.table -> typ -> typ
   val instT_term: typ Symtab.table -> term -> term
   val instT_thm: theory -> typ Symtab.table -> thm -> thm
@@ -78,6 +58,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
@@ -92,7 +73,7 @@
 
 datatype ('typ, 'term) stmt =
   Shows of (Attrib.binding * ('term * 'term list) list) list |
-  Obtains of (Binding.T * ((Binding.T * 'typ option) list * 'term list)) list;
+  Obtains of (binding * ((binding * 'typ option) list * 'term list)) list;
 
 type statement = (string, string) stmt;
 type statement_i = (typ, term) stmt;
@@ -101,7 +82,7 @@
 (* context *)
 
 datatype ('typ, 'term, 'fact) ctxt =
-  Fixes of (Binding.T * 'typ option * mixfix) list |
+  Fixes of (binding * 'typ option * mixfix) list |
   Constrains of (string * 'typ) list |
   Assumes of (Attrib.binding * ('term * 'term list) list) list |
   Defines of (Attrib.binding * ('term * 'term list)) list |
@@ -112,53 +93,29 @@
 
 fun facts_map f facts = Notes ("", facts) |> f |> (fn Notes (_, facts') => facts');
 
-fun map_ctxt' {binding, var, typ, term, pat, fact, attrib} =
-  fn Fixes fixes => Fixes (fixes |> map (fn (x, T, mx) =>
-       let val (x', mx') = var (x, mx) in (x', Option.map typ T, mx') end))
+fun map_ctxt {binding, typ, term, pattern, fact, attrib} =
+  fn Fixes fixes => Fixes (fixes |> map (fn (x, T, mx) => (binding x, Option.map typ T, mx)))
    | Constrains xs => Constrains (xs |> map (fn (x, T) =>
-       let val x' = Binding.base_name (#1 (var (Binding.name x, NoSyn))) in (x', typ T) end))
+      (Binding.base_name (binding (Binding.name x)), typ T)))
    | Assumes asms => Assumes (asms |> map (fn ((a, atts), propps) =>
-      ((binding a, map attrib atts), propps |> map (fn (t, ps) => (term t, map pat ps)))))
+      ((binding a, map attrib atts), propps |> map (fn (t, ps) => (term t, map pattern ps)))))
    | Defines defs => Defines (defs |> map (fn ((a, atts), (t, ps)) =>
-      ((binding a, map attrib atts), (term t, map pat ps))))
+      ((binding a, map attrib atts), (term t, map pattern ps))))
    | Notes (kind, facts) => Notes (kind, facts |> map (fn ((a, atts), bs) =>
       ((binding a, map attrib atts), bs |> map (fn (ths, btts) => (fact ths, map attrib btts)))));
 
-fun map_ctxt {binding, var, typ, term, fact, attrib} =
-  map_ctxt' {binding = binding, var = var, typ = typ, term = term, pat = term,
-    fact = fact, attrib = attrib}
-
 fun map_ctxt_attrib attrib =
-  map_ctxt {binding = I, var = I, typ = I, term = I, fact = I, attrib = attrib};
+  map_ctxt {binding = I, typ = I, term = I, pattern = I, fact = I, attrib = attrib};
 
 fun morph_ctxt phi = map_ctxt
  {binding = Morphism.binding phi,
-  var = Morphism.var phi,
   typ = Morphism.typ phi,
   term = Morphism.term phi,
+  pattern = Morphism.term phi,
   fact = Morphism.fact phi,
   attrib = Args.morph_values phi};
 
 
-(* logical content *)
-
-fun params_of (Fixes fixes) = fixes |> map
-    (fn (x, SOME T, _) => (Binding.base_name x, T)
-      | (x, _, _) => raise TERM ("Untyped context element parameter " ^ quote (Binding.display x), []))
-  | params_of _ = [];
-
-fun prems_of (Assumes asms) = maps (map fst o snd) asms
-  | prems_of (Defines defs) = map (fst o snd) defs
-  | prems_of _ = [];
-
-fun assume thy t = Assumption.assume (Thm.cterm_of thy t);
-
-fun facts_of thy (Assumes asms) = map (apsnd (map (fn (t, _) => ([assume thy t], [])))) asms
-  | facts_of thy (Defines defs) = map (apsnd (fn (t, _) => [([assume thy t], [])])) defs
-  | facts_of _ (Notes (_, facts)) = facts
-  | facts_of _ _ = [];
-
-
 
 (** pretty printing **)
 
@@ -168,9 +125,8 @@
         map (fn y => Pretty.block [Pretty.str "  ", Pretty.keyword sep, Pretty.brk 1, y]) ys;
 
 fun pretty_name_atts ctxt (b, atts) sep =
-  let
-    val name = Binding.display b;
-  in if name = "" andalso null atts then []
+  let val name = Binding.display b in
+    if name = "" andalso null atts then []
     else [Pretty.block
       (Pretty.breaks (Pretty.str name :: Attrib.pretty_attribs ctxt atts @ [Pretty.str sep]))]
   end;
@@ -299,24 +255,52 @@
 
 datatype witness = Witness of term * thm;
 
+val mark_witness = Logic.protect;
+fun witness_prop (Witness (t, _)) = t;
+fun witness_hyps (Witness (_, th)) = #hyps (Thm.rep_thm th);
 fun map_witness f (Witness witn) = Witness (f witn);
 
 fun morph_witness phi = map_witness (fn (t, th) => (Morphism.term phi t, Morphism.thm phi th));
 
-fun witness_prop (Witness (t, _)) = t;
-fun witness_hyps (Witness (_, th)) = #hyps (Thm.rep_thm th);
-
-fun assume_witness thy t =
-  Witness (t, Goal.protect (Thm.assume (Thm.cterm_of thy t)));
-
 fun prove_witness ctxt t tac =
-  Witness (t, Thm.close_derivation (Goal.prove ctxt [] [] (Logic.protect t) (fn _ =>
+  Witness (t, Thm.close_derivation (Goal.prove ctxt [] [] (mark_witness t) (fn _ =>
     Tactic.rtac Drule.protectI 1 THEN tac)));
 
-val close_witness = map_witness (fn (t, th) => (t, Thm.close_derivation th));
+
+local
+
+val refine_witness =
+  Proof.refine (Method.Basic (K (Method.RAW_METHOD
+    (K (ALLGOALS
+      (CONJUNCTS (ALLGOALS
+        (CONJUNCTS (TRYALL (Tactic.rtac Drule.protectI)))))))), Position.none));
+
+fun gen_witness_proof proof after_qed wit_propss eq_props =
+  let
+    val propss = (map o map) (fn prop => (mark_witness prop, [])) wit_propss
+      @ [map (rpair []) eq_props];
+    fun after_qed' thmss =
+      let val (wits, eqs) = split_last ((map o map) Thm.close_derivation thmss);
+      in after_qed ((map2 o map2) (curry Witness) wit_propss wits) eqs end;
+  in proof after_qed' propss #> refine_witness #> Seq.hd end;
 
-fun conclude_witness (Witness (_, th)) =
-  Thm.close_derivation (MetaSimplifier.norm_hhf_protect (Goal.conclude th));
+in
+
+fun witness_proof after_qed wit_propss =
+  gen_witness_proof (Proof.theorem_i NONE) (fn wits => fn _ => after_qed wits)
+    wit_propss [];
+
+val witness_proof_eqs = gen_witness_proof (Proof.theorem_i NONE);
+
+fun witness_local_proof after_qed cmd wit_propss goal_ctxt int =
+  gen_witness_proof (fn after_qed' => fn propss =>
+    Proof.map_context (K goal_ctxt)
+    #> Proof.local_goal (ProofDisplay.print_results int) (K I) ProofContext.bind_propp_i
+      cmd NONE after_qed' (map (pair (Binding.empty, [])) propss))
+    (fn wits => fn _ => after_qed wits) wit_propss [];
+
+end;
+
 
 fun compose_witness (Witness (_, th)) r =
   let
@@ -329,18 +313,8 @@
         (Thm.instantiate (Thm.match (Thm.cprop_of th', A)) th'))
   end;
 
-val mark_witness = Logic.protect;
-
-fun make_witness t th = Witness (t, th);
-fun dest_witness (Witness w) = w;
-
-fun transfer_witness thy (Witness (t, th)) = Witness (t, Thm.transfer thy th);
-
-val refine_witness =
-  Proof.refine (Method.Basic (K (Method.RAW_METHOD
-    (K (ALLGOALS
-      (CONJUNCTS (ALLGOALS
-        (CONJUNCTS (TRYALL (Tactic.rtac Drule.protectI)))))))), Position.none));
+fun conclude_witness (Witness (_, th)) =
+  Thm.close_derivation (MetaSimplifier.norm_hhf_protect (Goal.conclude th));
 
 fun pretty_witness ctxt witn =
   let val prt_term = Pretty.quote o Syntax.pretty_term ctxt in
@@ -384,50 +358,6 @@
   end;
 
 
-(* rename *)
-
-fun rename ren x =
-  (case AList.lookup (op =) ren (x: string) of
-    NONE => x
-  | SOME (x', _) => x');
-
-fun rename_var_name ren (x, mx) =
-  (case (AList.lookup (op =) ren x, mx) of
-    (NONE, _) => (x, mx)
-  | (SOME (x', NONE), Structure) => (x', mx)
-  | (SOME (x', SOME _), Structure) =>
-      error ("Attempt to change syntax of structure parameter " ^ quote x)
-  | (SOME (x', NONE), _) => (x', NoSyn)
-  | (SOME (x', SOME mx'), _) => (x', mx'));
-
-fun rename_var ren (b, mx) =
-  let
-    val x = Binding.base_name b;
-    val (x', mx') = rename_var_name ren (x, mx);
-  in (Binding.name x', mx') end;
-
-fun rename_term ren (Free (x, T)) = Free (rename ren x, T)
-  | rename_term ren (t $ u) = rename_term ren t $ rename_term ren u
-  | rename_term ren (Abs (x, T, t)) = Abs (x, T, rename_term ren t)
-  | rename_term _ a = a;
-
-fun rename_thm ren th =
-  let
-    val thy = Thm.theory_of_thm th;
-    val subst = (Thm.fold_terms o Term.fold_aterms)
-      (fn Free (x, T) =>
-        let val x' = rename ren x
-        in if x = x' then I else insert (eq_fst (op =)) ((x, T), Free (x', T)) end
-      | _ => I) th [];
-  in
-    if null subst then th
-    else th |> hyps_rule (instantiate_frees thy subst)
-  end;
-
-fun rename_morphism ren = Morphism.morphism
-  {binding = I, var = rename_var ren, typ = I, term = rename_term ren, fact = map (rename_thm ren)};
-
-
 (* instantiate types *)
 
 fun instT_type env =
@@ -453,7 +383,7 @@
 fun instT_morphism thy env =
   let val thy_ref = Theory.check_thy thy in
     Morphism.morphism
-     {binding = I, var = I,
+     {binding = I,
       typ = instT_type env,
       term = instT_term env,
       fact = map (fn th => instT_thm (Theory.deref thy_ref) env th)}
@@ -502,7 +432,7 @@
 fun inst_morphism thy envs =
   let val thy_ref = Theory.check_thy thy in
     Morphism.morphism
-     {binding = I, var = I,
+     {binding = I,
       typ = instT_type (#1 envs),
       term = inst_term envs,
       fact = map (fn th => inst_thm (Theory.deref thy_ref) envs th)}
@@ -516,9 +446,17 @@
       NONE => I
     | SOME w => Thm.implies_intr hyp #> compose_witness w)) (#hyps (Thm.crep_thm thm));
 
-fun satisfy_morphism witns = Morphism.thm_morphism (satisfy_thm witns);
+val satisfy_morphism = Morphism.thm_morphism o satisfy_thm;
+val satisfy_facts = facts_map o morph_ctxt o satisfy_morphism;
+
+
+(* rewriting with equalities *)
 
-fun satisfy_facts witns = facts_map (morph_ctxt (satisfy_morphism witns));
+fun eq_morphism thy thms = Morphism.morphism
+ {binding = I,
+  typ = I,
+  term = MetaSimplifier.rewrite_term thy thms [],
+  fact = map (MetaSimplifier.rewrite_rule thms)};
 
 
 (* generalize type/term parameters *)
@@ -533,18 +471,16 @@
     val exp_fact = map (Thm.adjust_maxidx_thm maxidx) #> Variable.export inner outer;
     val exp_term = Drule.term_rule thy (singleton exp_fact);
     val exp_typ = Logic.type_map exp_term;
-    val morphism =
-      Morphism.morphism {binding = I, var = I, typ = exp_typ, term = exp_term, fact = exp_fact};
+    val morphism = Morphism.morphism {binding = I, typ = exp_typ, term = exp_term, fact = exp_fact};
   in facts_map (morph_ctxt morphism) facts end;
 
 
 (* transfer to theory using closure *)
 
 fun transfer_morphism thy =
-  let val thy_ref = Theory.check_thy thy in
-    Morphism.morphism {binding = I, var = I, typ = I, term = I,
-      fact = map (fn th => transfer (Theory.deref thy_ref) th)}
-  end;
+  let val thy_ref = Theory.check_thy thy
+  in Morphism.thm_morphism (fn th => transfer (Theory.deref thy_ref) th) end;
+
 
 
 (** activate in context, return elements and facts **)
@@ -591,11 +527,14 @@
   if NameSpace.is_qualified name then error ("Illegal qualified name: " ^ quote name)
   else name;
 
-fun prep_facts prep_name get intern ctxt elem = elem |> map_ctxt
-     {var = I, typ = I, term = I,
-      binding = Binding.map_base prep_name,
-      fact = get ctxt,
-      attrib = intern (ProofContext.theory_of ctxt)};
+fun prep_facts prep_name get intern ctxt =
+  map_ctxt
+   {binding = Binding.map_base prep_name,
+    typ = I,
+    term = I,
+    pattern = I,
+    fact = get ctxt,
+    attrib = intern (ProofContext.theory_of ctxt)};
 
 in
 
--- a/src/Pure/Isar/expression.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/expression.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -6,35 +6,47 @@
 
 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_i = term expr * (Binding.T * typ option * mixfix) list;
-  type expression = string expr * (Binding.T * string 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 * typ option * mixfix) list
+  type expression = string expr * (binding * 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;
+    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;
+    Proof.context -> (term * term list) list list * Proof.context
 
   (* Declaring locales *)
+  val cert_declaration: expression_i -> Element.context_i list -> Proof.context ->
+    ((binding * 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 * 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 * 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;
+    theory -> string * local_theory
   val add_locale_cmd: bstring -> bstring -> expression -> Element.context list ->
-    theory -> string * local_theory;
+    theory -> string * local_theory
 
   (* Interpretation *)
   val cert_goal_expression: expression_i -> 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;
+    (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;
+    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;
+    theory -> Proof.state
+  val interpret: expression_i -> bool -> Proof.state -> Proof.state
+  val interpret_cmd: expression -> bool -> Proof.state -> Proof.state
 end;
 
 
@@ -52,8 +64,8 @@
 
 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;
+type expression = string expr * (binding * string option * mixfix) list;
+type expression_i = term expr * (binding * typ option * mixfix) list;
 
 
 (** Internalise locale names in expr **)
@@ -138,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
 
@@ -186,11 +198,14 @@
 
 (** Parsing **)
 
-fun parse_elem prep_typ prep_term ctxt elem =
-  Element.map_ctxt' {binding = I, var = I, typ = prep_typ ctxt,
-  term = prep_term (ProofContext.set_mode ProofContext.mode_schematic ctxt), (* FIXME ?? *)
-  pat = prep_term (ProofContext.set_mode ProofContext.mode_pattern ctxt),
-  fact = I, attrib = I} elem;
+fun parse_elem prep_typ prep_term ctxt =
+  Element.map_ctxt
+   {binding = I,
+    typ = prep_typ ctxt,
+    term = prep_term (ProofContext.set_mode ProofContext.mode_schematic ctxt),
+    pattern = prep_term (ProofContext.set_mode ProofContext.mode_pattern ctxt),
+    fact = I,
+    attrib = I};
 
 fun parse_concl prep_term ctxt concl =
   (map o map) (fn (t, ps) =>
@@ -313,7 +328,7 @@
   let
     val thy = ProofContext.theory_of ctxt;
     val (parm_names, parm_types) = Locale.params_of thy loc |>
-      map (fn (b, SOME T, _) => (Binding.base_name b, T)) |> split_list;
+      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;
 
@@ -335,7 +350,7 @@
 
 local
 
-fun prep_full_context_statement parse_typ parse_prop parse_inst prep_vars prep_expr
+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 ctxt1;
@@ -345,8 +360,9 @@
     fun prep_inst (loc, (prfx, inst)) (i, insts, ctxt) =
       let
         val (parm_names, parm_types) = Locale.params_of thy loc |>
-          map (fn (b, SOME T, _) => (Binding.base_name b, T)) |> split_list;
-        val inst' = parse_inst parm_names inst ctxt;
+          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';
@@ -357,43 +373,47 @@
         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 ctxt1 |> fst;
+    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 raw_elems (insts', [], ctxt3);
-    val (insts, elems, concl, ctxt5) =
-      prep_concl raw_concl (insts', elems'', ctxt4);
+    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 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 ctxt6 parms do_close insts elems;
+    val (deps, elems'') = finish ctxt6 parms do_close insts elems';
 
-  in ((fors', deps, elems', concl), (parms, ctxt6)) 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;
 
@@ -405,14 +425,16 @@
 fun prep_statement prep activate raw_elems raw_concl context =
   let
      val ((_, _, elems, concl), _) =
-       prep true false ([], []) raw_elems raw_concl context ;
-     val (_, context') = activate elems (ProofContext.set_stmt true context);
+       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;
 
@@ -429,13 +451,16 @@
     val context' = context |>
       ProofContext.add_fixes_i fixed |> snd |>
       fold Locale.activate_local_facts deps;
-    val (elems', _) = activate elems (ProofContext.set_stmt true context');
+    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;
 
@@ -468,14 +493,13 @@
     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;
     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};
+    val export' = Morphism.morphism {binding = I, typ = exp_typ, term = exp_term, fact = exp_fact};
   in ((propss, deps, export'), goal_ctxt) end;
     
 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;
 
@@ -618,7 +642,7 @@
       |> 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])];
+        [((Binding.name (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;
@@ -756,54 +780,36 @@
 
 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;
 
 
 (*** Interpretation ***)
 
-(** Witnesses and goals **)
-
-fun prep_propp propss = propss |> map (map (rpair [] o Element.mark_witness));
-
-val prep_result = map2 (fn props => fn thms =>
-  map2 Element.make_witness props (map Thm.close_derivation thms));
-
-
 (** Interpretation between locales: declaring sublocale relationships **)
 
 local
 
-fun gen_sublocale prep_expr intern
-    raw_target expression thy =
+fun gen_sublocale prep_expr intern raw_target expression thy =
   let
     val target = intern thy raw_target;
     val target_ctxt = Locale.init target thy;
 
     val ((propss, deps, export), goal_ctxt) = prep_expr expression target_ctxt;
     
-    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) #>
-        (* propagate registrations *)
-        (fn thy => fold_rev (fn reg => Locale.activate_global_facts reg)
+    fun after_qed witss = ProofContext.theory (
+      fold2 (fn (name, morph) => fn wits => Locale.add_dependency target
+        (name, morph $> Element.satisfy_morphism wits $> export)) deps witss #>
+      (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) |>
-      Element.refine_witness |> Seq.hd
-  end;
+  in Element.witness_proof after_qed propss goal_ctxt end;
 
 in
 
+fun sublocale x = gen_sublocale cert_goal_expression (K I) x;
 fun sublocale_cmd x = gen_sublocale read_goal_expression Locale.intern x;
-fun sublocale x = gen_sublocale cert_goal_expression (K I) x;
 
 end;
 
@@ -812,25 +818,21 @@
 
 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) = prep_expr expression ctxt;
+    val ((propss, regs, export), expr_ctxt) = ProofContext.init theory
+      |> prep_expr expression;
 
     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;
 
-    (*** alternative code -- unclear why it does not work yet ***)
-    fun store_reg ((name, morph), thms) thy =
+    fun store_reg ((name, morph), wits) thy =
       let
-        val thms' = map (Element.morph_witness export') thms;
-        val morph' = morph $> Element.satisfy_morphism thms';
+        val wits' = map (Element.morph_witness export') wits;
+        val morph' = morph $> Element.satisfy_morphism wits';
       in
         thy
         |> Locale.add_registration (name, (morph', export))
@@ -841,76 +843,32 @@
           thy
           |> fold (fn (name, morph) =>
                Locale.activate_global_facts (name, morph $> export)) regs
-      | store_eqns_activate regs thms thys =
+      | store_eqns_activate regs eqs thy =
           let
-            val thms' = thms |> map (Element.conclude_witness #>
-              Morphism.thm (export' $> export) #>
+            val eqs' = eqs |> map (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 eq_morph = Element.eq_morphism thy eqs';
             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')
+            |> PureThy.note_thmss Thm.lemmaK (eq_attns' ~~ map (fn eq => [([eq], [])]) eqs')
             |> snd
           end;
 
-    fun after_qed results =
-      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;
-    (*** alternative code end ***)
+    fun after_qed wits eqs = ProofContext.theory (fold_map store_reg (regs ~~ wits)
+        #-> (fn regs => store_eqns_activate regs eqs));
 
-    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 = Locale.add_registration (name, (morph', export));
-        in ((name, morph') :: regs, add thy) end
-      | store (Eqns [], []) (regs, thy) =
-        let val add = fold_rev (fn (name, morph) =>
-              Locale.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) =>
-              Locale.amend_registration eq_morph (name, morph) #>
-              Locale.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 after_qed results =
-      ProofContext.theory (fn thy =>
-        fold store (map Reg regs @ [Eqns eq_attns] ~~
-          prep_result (propss @ [eqns]) results) ([], thy) |> snd);
-  in
-    goal_ctxt |>
-      Proof.theorem_i NONE after_qed (prep_propp (propss @ [eqns])) |>
-      Element.refine_witness |> Seq.hd
-  end;
+  in Element.witness_proof_eqs after_qed propss eqns goal_ctxt end;
 
 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;
 
@@ -927,26 +885,22 @@
 
     val ((propss, regs, export), goal_ctxt) = prep_expr expression ctxt;
     
-    fun store_reg ((name, morph), thms) =
+    fun store_reg (name, morph) thms =
       let
         val morph' = morph $> Element.satisfy_morphism thms $> export;
-      in
-        Locale.activate_local_facts (name, morph')
-      end;
+      in Locale.activate_local_facts (name, morph') end;
 
-    fun after_qed results =
-      Proof.map_context (fold store_reg (regs ~~ prep_result propss results));
+    fun after_qed wits =
+      Proof.map_context (fold2 store_reg regs wits);
   in
-    state |> Proof.map_context (K goal_ctxt) |>
-      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
+    state
+    |> Element.witness_local_proof after_qed "interpret" propss goal_ctxt int
   end;
 
 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/isar_cmd.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/isar_cmd.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -13,8 +13,8 @@
   val typed_print_translation: bool * (string * Position.T) -> theory -> theory
   val print_ast_translation: bool * (string * Position.T) -> theory -> theory
   val oracle: bstring -> SymbolPos.text * Position.T -> theory -> theory
-  val add_axioms: ((Binding.T * string) * Attrib.src list) list -> theory -> theory
-  val add_defs: (bool * bool) * ((Binding.T * string) * Attrib.src list) list -> theory -> theory
+  val add_axioms: ((binding * string) * Attrib.src list) list -> theory -> theory
+  val add_defs: (bool * bool) * ((binding * string) * Attrib.src list) list -> theory -> theory
   val declaration: string * Position.T -> local_theory -> local_theory
   val simproc_setup: string -> string list -> string * Position.T -> string list ->
     local_theory -> local_theory
@@ -53,7 +53,6 @@
   val print_theorems: Toplevel.transition -> Toplevel.transition
   val print_locales: Toplevel.transition -> Toplevel.transition
   val print_locale: bool * xstring -> Toplevel.transition -> Toplevel.transition
-  val print_registrations: bool -> string -> Toplevel.transition -> Toplevel.transition
   val print_attributes: Toplevel.transition -> Toplevel.transition
   val print_simpset: Toplevel.transition -> Toplevel.transition
   val print_rules: Toplevel.transition -> Toplevel.transition
@@ -359,12 +358,6 @@
   Toplevel.keep (fn state =>
     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 (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);
 
--- a/src/Pure/Isar/isar_document.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/isar_document.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -11,8 +11,8 @@
   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
-  val end_document: document_id -> unit
 end;
 
 structure IsarDocument: ISAR_DOCUMENT =
@@ -20,50 +20,14 @@
 
 (* unique identifiers *)
 
-type document_id = string;
+type state_id = string;
 type command_id = string;
-type state_id = string;
-
-
-(** execution states **)
-
-(* command status *)
-
-datatype status =
-  Unprocessed |
-  Running of Toplevel.state option future |
-  Failed |
-  Finished of Toplevel.state;
-
-fun status_markup Unprocessed = Markup.unprocessed
-  | status_markup (Running future) = Markup.running (Task_Queue.str_of_task (Future.task_of future))
-  | status_markup Failed = Markup.failed
-  | status_markup (Finished _) = Markup.finished;
+type document_id = string;
 
-fun status_update tr state status =
-  (case status of
-    Unprocessed => SOME (Running (Future.fork_pri 1 (fn () =>
-      (case Toplevel.transition true tr state of
-        NONE => (Toplevel.error_msg tr (SYS_ERROR "Cannot terminate here!", ""); NONE)
-      | SOME (_, SOME err) => (Toplevel.error_msg tr err; NONE)
-      | SOME (state', NONE) => SOME state'))))
-  | Running future =>
-      (case Future.peek future of
-        NONE => NONE
-      | SOME (Exn.Result NONE) => SOME Failed
-      | SOME (Exn.Result (SOME state')) => SOME (Finished state')
-      | SOME (Exn.Exn exn) => (Toplevel.error_msg tr (exn, ""); SOME Failed))
-  | _ => NONE);
+fun make_id () = "isabelle:" ^ serial_string ();
 
-
-(* state *)
-
-datatype state = State of
- {prev: state_id option,
-  command: command_id,
-  status: status};
-
-fun make_state prev command status = State {prev = prev, command = command, status = status};
+fun err_dup kind id = error ("Duplicate " ^ kind ^ ": " ^ quote id);
+fun err_undef kind id = error ("Unknown " ^ kind ^ ": " ^ quote id);
 
 
 
@@ -71,12 +35,22 @@
 
 (* command entries *)
 
-datatype entry = Entry of
- {prev: command_id option,
-  next: command_id option,
-  state: state_id};
+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 make_entry prev next state = Entry {prev = prev, next = next, state = state};
+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 *)
@@ -84,19 +58,61 @@
 datatype document = Document of
  {dir: Path.T,                    (*master directory*)
   name: string,                   (*theory name*)
-  commands: entry Symtab.table};  (*unique command entries indexed by command_id*)
+  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 make_document dir name commands = Document {dir = dir, name = name, commands = commands};
+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 **)
 
-fun err_dup kind id = sys_error ("Duplicate " ^ kind ^ ": " ^ quote id);
-fun err_undef kind id = sys_error ("Unknown " ^ kind ^ ": " ^ quote id);
-
 local
-  val global_states = ref (Symtab.empty: state Symtab.table);
+  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
@@ -118,27 +134,36 @@
 end;
 
 
-fun define_state (id: state_id) state =
-  change_states (Symtab.update_new (id, state))
+(* 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 {status as Finished state, ...}) => state
-  | _ => sys_error ("Unfinished 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 " ^ Toplevel.name_of tr) dup;
+    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 cmd => cmd);
+  | 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;
@@ -146,26 +171,103 @@
 fun the_document (id: document_id) =
   (case Symtab.lookup (get_documents ()) id of
     NONE => err_undef "document" id
-  | SOME doc => doc);
+  | 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 (make_state NONE id (Finished Toplevel.toplevel));
-    val commands = Symtab.make [(id, make_entry NONE NONE id)];
-    val _ = define_document id (make_document dir name commands);
+    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;
 
-fun edit_document (id: document_id) (new_id: document_id) edits =
+
+(* 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 Document {dir, name, commands} = the_document id;
-    val commands' = sys_error "FIXME";
-    val _ = define_document new_id (make_document dir name commands');
+    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;
 
-fun end_document (id: document_id) = sys_error "FIXME";
+end;
 
 
 
@@ -185,14 +287,15 @@
       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 () => end_document id new_id edits)));
-
-val _ =
-  OuterSyntax.internal_command "Isar.end_document"
-    (P.string >> (fn id => Toplevel.imperative (fn () => end_document id)));
+      >> (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	Thu Jan 29 12:24:00 2009 +0000
@@ -0,0 +1,64 @@
+/*  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
+}
+
+trait IsarDocument extends IsabelleProcess
+{
+  import IsarDocument._
+
+
+  /* commands */
+
+  def define_command(id: Command_ID, text: String) {
+    output_sync("Isar.define_command " + IsabelleSyntax.encode_string(id) + " " +
+      IsabelleSyntax.encode_string(text))
+  }
+
+
+  /* documents */
+
+  def begin_document(id: Document_ID, path: String) {
+    output_sync("Isar.begin_document " + IsabelleSyntax.encode_string(id) + " " +
+       IsabelleSyntax.encode_string(path))
+  }
+
+  def end_document(id: Document_ID) {
+    output_sync("Isar.end_document " + IsabelleSyntax.encode_string(id))
+  }
+
+  def edit_document(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)
+    output_sync(buf.toString)
+  }
+}
--- a/src/Pure/Isar/isar_syn.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/isar_syn.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -418,45 +418,6 @@
         >> (fn expr => Toplevel.print o
             Toplevel.proof' (fn int => Expression.interpret_cmd expr int)));
 
-local
-
-val opt_prefix = Scan.optional (P.binding --| P.$$$ ":") Binding.empty;
-
-in
-
-val locale_val =
-  SpecParse.locale_expr --
-    Scan.optional (P.$$$ "+" |-- P.!!! (Scan.repeat1 SpecParse.context_element)) [] ||
-  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) (Old_Locale.empty, []) -- P.opt_begin
-      >> (fn ((name, (expr, elems)), begin) =>
-          (begin ? Toplevel.print) o Toplevel.begin_local_theory 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 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
-              (Old_Locale.interpretation_cmd (Binding.base_name name) expr insts)));
-
-val _ =
-  OuterSyntax.command "class_interpret"
-    "prove and register interpretation of locale expression in proof context"
-    (K.tag_proof K.prf_goal)
-    (opt_prefix -- SpecParse.locale_expr -- SpecParse.locale_insts
-      >> (fn ((name, expr), insts) => Toplevel.print o
-          Toplevel.proof'
-            (fn int => Old_Locale.interpret_cmd (Binding.base_name name) expr insts int)));
-
-end;
-
 
 (* classes *)
 
@@ -467,7 +428,7 @@
 
 val _ =
   OuterSyntax.command "class" "define type class" K.thy_decl
-   (P.name -- (P.$$$ "=" |-- class_val) -- P.opt_begin
+   (P.name -- Scan.optional (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 #> snd)));
@@ -857,12 +818,6 @@
     (opt_bang -- P.xname >> (Toplevel.no_timing oo IsarCmd.print_locale));
 
 val _ =
-  OuterSyntax.improper_command "print_interps"
-    "print interpretations of named locale" K.diag
-    (Scan.optional (P.$$$ "!" >> K true) false -- P.xname
-      >> (Toplevel.no_timing oo uncurry IsarCmd.print_registrations));
-
-val _ =
   OuterSyntax.improper_command "print_attributes" "print attributes of this theory" K.diag
     (Scan.succeed (Toplevel.no_timing o IsarCmd.print_attributes));
 
--- a/src/Pure/Isar/local_defs.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/local_defs.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -11,10 +11,10 @@
   val mk_def: Proof.context -> (string * term) list -> term list
   val expand: cterm list -> thm -> thm
   val def_export: Assumption.export
-  val add_defs: ((Binding.T * mixfix) * ((Binding.T * attribute list) * term)) list ->
+  val add_defs: ((binding * mixfix) * ((binding * attribute list) * term)) list ->
     Proof.context -> (term * (string * thm)) list * Proof.context
-  val add_def: (Binding.T * mixfix) * term -> Proof.context -> (term * thm) * Proof.context
-  val fixed_abbrev: (Binding.T * mixfix) * term -> Proof.context ->
+  val add_def: (binding * mixfix) * term -> Proof.context -> (term * thm) * Proof.context
+  val fixed_abbrev: (binding * mixfix) * term -> Proof.context ->
     (term * term) * Proof.context
   val export: Proof.context -> Proof.context -> thm -> thm list * thm
   val export_cterm: Proof.context -> Proof.context -> cterm -> cterm * thm
--- a/src/Pure/Isar/local_syntax.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/local_syntax.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/local_syntax.ML
-    ID:         $Id$
     Author:     Makarius
 
 Local syntax depending on theory syntax.
--- a/src/Pure/Isar/local_theory.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/local_theory.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -18,16 +18,16 @@
   val raw_theory: (theory -> theory) -> local_theory -> local_theory
   val checkpoint: local_theory -> local_theory
   val full_naming: local_theory -> NameSpace.naming
-  val full_name: local_theory -> Binding.T -> string
+  val full_name: local_theory -> binding -> string
   val theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory
   val theory: (theory -> theory) -> local_theory -> local_theory
   val target_result: (Proof.context -> 'a * Proof.context) -> local_theory -> 'a * local_theory
   val target: (Proof.context -> Proof.context) -> local_theory -> local_theory
   val affirm: local_theory -> local_theory
   val pretty: local_theory -> Pretty.T list
-  val abbrev: Syntax.mode -> (Binding.T * mixfix) * term -> local_theory ->
+  val abbrev: Syntax.mode -> (binding * mixfix) * term -> local_theory ->
     (term * term) * local_theory
-  val define: string -> (Binding.T * mixfix) * (Attrib.binding * term) -> local_theory ->
+  val define: string -> (binding * mixfix) * (Attrib.binding * term) -> local_theory ->
     (term * (string * thm)) * local_theory
   val note: string -> Attrib.binding * thm list -> local_theory -> (string * thm list) * local_theory
   val notes: string -> (Attrib.binding * (thm list * Attrib.src list) list) list ->
@@ -55,10 +55,10 @@
 
 type operations =
  {pretty: local_theory -> Pretty.T list,
-  abbrev: Syntax.mode -> (Binding.T * mixfix) * term -> local_theory ->
+  abbrev: Syntax.mode -> (binding * mixfix) * term -> local_theory ->
     (term * term) * local_theory,
   define: string ->
-    (Binding.T * mixfix) * (Attrib.binding * term) -> local_theory ->
+    (binding * mixfix) * (Attrib.binding * term) -> local_theory ->
     (term * (string * thm)) * local_theory,
   notes: string ->
     (Attrib.binding * (thm list * Attrib.src list) list) list ->
--- a/src/Pure/Isar/locale.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/locale.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -29,28 +29,24 @@
 
 signature LOCALE =
 sig
-  type locale
-
+  (* Locale specification *)
   val register_locale: bstring ->
-    (string * sort) list * (Binding.T * typ option * mixfix) list ->
+    (string * sort) list * (binding * 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.morphism) * stamp) list -> theory -> theory
-
-  (* Locale name space *)
+    ((string * morphism) * stamp) list -> theory -> theory
   val intern: theory -> xstring -> string
   val extern: theory -> string -> xstring
-
-  (* Specification *)
   val defined: theory -> string -> bool
-  val params_of: theory -> string -> (Binding.T * typ option * mixfix) list
+  val params_of: theory -> string -> (binding * 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.morphism -> term 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 add_thmss: string -> string -> (Attrib.binding * (thm list * Attrib.src list) list) list ->
@@ -58,13 +54,13 @@
   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
+  val add_dependency: string -> string * morphism -> theory -> theory
 
   (* Activation *)
-  val activate_declarations: theory -> string * Morphism.morphism ->
+  val activate_declarations: theory -> string * 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 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 *)
@@ -74,11 +70,11 @@
   val intro_locales_tac: bool -> Proof.context -> thm list -> tactic
 
   (* Registrations *)
-  val add_registration: (string * (Morphism.morphism * Morphism.morphism)) ->
+  val add_registration: string * (morphism * morphism) ->
     theory -> theory
-  val amend_registration: Morphism.morphism -> (string * Morphism.morphism) ->
+  val amend_registration: morphism -> string * morphism ->
     theory -> theory
-  val get_registrations: theory -> (string * Morphism.morphism) list
+  val get_registrations: theory -> (string * morphism) list
 
   (* Diagnostic *)
   val print_locales: theory -> unit
@@ -111,13 +107,25 @@
 
 datatype ctxt = datatype Element.ctxt;
 
+fun global_note_qualified kind facts thy = (*FIXME*)
+  thy
+  |> Sign.qualified_names
+  |> PureThy.note_thmss kind facts
+  ||> Sign.restore_naming thy;
+
+fun local_note_qualified kind facts ctxt = (*FIXME*)
+  ctxt
+  |> ProofContext.qualified_names
+  |> ProofContext.note_thmss_i kind facts
+  ||> ProofContext.restore_naming ctxt;
+
 
 
 (*** Theory data ***)
 
 datatype locale = Loc of {
   (** static part **)
-  parameters: (string * sort) list * (Binding.T * typ option * mixfix) list,
+  parameters: (string * sort) list * (binding * typ option * mixfix) list,
     (* type and term parameters *)
   spec: term option * term list,
     (* assumptions (as a single predicate expression) and defines *)
@@ -128,7 +136,7 @@
     (* 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
+  dependencies: ((string * morphism) * stamp) list
     (* locale dependencies (sublocale relation) *)
 };
 
@@ -193,6 +201,9 @@
 fun declarations_of thy name = the_locale thy name |>
   #decls |> apfst (map fst) |> apsnd (map fst);
 
+fun dependencies_of thy name = the_locale thy name |>
+  #dependencies |> map fst;
+
 
 (*** Activate context elements of locale ***)
 
@@ -326,7 +337,7 @@
 fun init_global_elem (Notes (kind, facts)) thy =
   let
     val facts' = Attrib.map_facts (Attrib.attribute_i thy) facts
-  in Old_Locale.global_note_qualified kind facts' thy |> snd end
+  in global_note_qualified kind facts' thy |> snd end
 
 fun init_local_elem (Fixes fixes) ctxt = ctxt |>
       ProofContext.add_fixes_i fixes |> snd
@@ -348,7 +359,7 @@
   | 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
+      in local_note_qualified kind facts' ctxt |> snd end
 
 fun cons_elem false (Notes notes) elems = elems
   | cons_elem _ elem elems = elem :: elems
@@ -356,23 +367,20 @@
 in
 
 fun activate_declarations thy dep ctxt =
-  roundup thy activate_decls dep (get_local_idents ctxt, ctxt) |> uncurry put_local_idents;
+  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) |>
-  uncurry put_global_idents;
+    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) |>
-  uncurry put_local_idents;
+    (get_local_idents ctxt, ctxt) |-> 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;
+    (empty, ProofContext.init thy) |-> put_local_idents;
 
 fun print_locale thy show_facts name =
   let
@@ -392,7 +400,7 @@
 
 structure RegistrationsData = TheoryDataFun
 (
-  type T = ((string * (Morphism.morphism * Morphism.morphism)) * stamp) list;
+  type T = ((string * (morphism * morphism)) * stamp) list;
     (* FIXME mixins need to be stamped *)
     (* registrations, in reverse order of declaration *)
   val empty = [];
@@ -408,8 +416,8 @@
 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 ?? uncurry put_global_idents *);
+    (name, base_morph) (get_global_idents thy, thy) |> snd
+    (* FIXME |-> put_global_idents ?*);
 
 fun amend_registration morph (name, base_morph) thy =
   let
@@ -428,6 +436,7 @@
   end;
 
 
+
 (*** Storing results ***)
 
 (* Theorems *)
@@ -443,7 +452,7 @@
             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)
+            in global_note_qualified kind args'' #> snd end)
         (get_registrations thy |> filter (fn (name, _) => name = loc)) thy))
   in ctxt'' end;
 
@@ -494,12 +503,10 @@
 val _ = Context.>> (Context.map_theory
   (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)),
+      Method.ctxt_args (fn ctxt => Method.METHOD (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)),
+      Method.ctxt_args (fn ctxt => Method.METHOD (intro_locales_tac true ctxt)),
       "back-chain all introduction rules of locales")]));
 
 end;
--- a/src/Pure/Isar/net_rules.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/net_rules.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/net_rules.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Efficient storage of rules: preserves order, prefers later entries.
--- a/src/Pure/Isar/object_logic.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/object_logic.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/object_logic.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Specifics about common object-logics.
--- a/src/Pure/Isar/obtain.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/obtain.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -39,16 +39,16 @@
 signature OBTAIN =
 sig
   val thatN: string
-  val obtain: string -> (Binding.T * string option * mixfix) list ->
+  val obtain: string -> (binding * string option * mixfix) list ->
     (Attrib.binding * (string * string list) list) list ->
     bool -> Proof.state -> Proof.state
-  val obtain_i: string -> (Binding.T * typ option * mixfix) list ->
-    ((Binding.T * attribute list) * (term * term list) list) list ->
+  val obtain_i: string -> (binding * typ option * mixfix) list ->
+    ((binding * attribute list) * (term * term list) list) list ->
     bool -> Proof.state -> Proof.state
   val result: (Proof.context -> tactic) -> thm list -> Proof.context ->
     (cterm list * thm list) * Proof.context
-  val guess: (Binding.T * string option * mixfix) list -> bool -> Proof.state -> Proof.state
-  val guess_i: (Binding.T * typ option * mixfix) list -> bool -> Proof.state -> Proof.state
+  val guess: (binding * string option * mixfix) list -> bool -> Proof.state -> Proof.state
+  val guess_i: (binding * typ option * mixfix) list -> bool -> Proof.state -> Proof.state
 end;
 
 structure Obtain: OBTAIN =
--- a/src/Pure/Isar/old_locale.ML	Thu Jan 29 12:05:19 2009 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,2485 +0,0 @@
-(*  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_lex.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/outer_lex.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/outer_lex.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Outer lexical syntax for Isabelle/Isar.
--- a/src/Pure/Isar/outer_parse.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/outer_parse.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -61,12 +61,12 @@
   val list: 'a parser -> 'a list parser
   val list1: 'a parser -> 'a list parser
   val name: bstring parser
-  val binding: Binding.T parser
+  val binding: binding parser
   val xname: xstring parser
   val text: string parser
   val path: Path.T parser
   val parname: string parser
-  val parbinding: Binding.T parser
+  val parbinding: binding parser
   val sort: string parser
   val arity: (string * string list * string) parser
   val multi_arity: (string list * string list * string) parser
@@ -81,11 +81,11 @@
   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 params: (binding * string option) list parser
+  val simple_fixes: (binding * string option) list parser
+  val fixes: (binding * string option * mixfix) list parser
+  val for_fixes: (binding * string option * mixfix) list parser
+  val for_simple_fixes: (binding * 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
--- a/src/Pure/Isar/overloading.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/overloading.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/overloading.ML
-    ID:         $Id$
     Author:     Florian Haftmann, TU Muenchen
 
 Overloaded definitions without any discipline.
@@ -134,8 +133,8 @@
 
 fun declare c_ty = pair (Const c_ty);
 
-fun define checked name (c, t) =
-  Thm.add_def (not checked) true (name, Logic.mk_equals (Const (c, Term.fastype_of t), t));
+fun define checked name (c, t) = Thm.add_def (not checked) true (Binding.name name,
+  Logic.mk_equals (Const (c, Term.fastype_of t), t));
 
 
 (* target *)
--- a/src/Pure/Isar/proof.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/proof.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -43,27 +43,27 @@
   val match_bind_i: (term list * term) list -> state -> state
   val let_bind: (string list * string) list -> state -> state
   val let_bind_i: (term list * term) list -> state -> state
-  val fix: (Binding.T * string option * mixfix) list -> state -> state
-  val fix_i: (Binding.T * typ option * mixfix) list -> state -> state
+  val fix: (binding * string option * mixfix) list -> state -> state
+  val fix_i: (binding * typ option * mixfix) list -> state -> state
   val assm: Assumption.export ->
     (Attrib.binding * (string * string list) list) list -> state -> state
   val assm_i: Assumption.export ->
-    ((Binding.T * attribute list) * (term * term list) list) list -> state -> state
+    ((binding * attribute list) * (term * term list) list) list -> state -> state
   val assume: (Attrib.binding * (string * string list) list) list -> state -> state
-  val assume_i: ((Binding.T * attribute list) * (term * term list) list) list ->
+  val assume_i: ((binding * attribute list) * (term * term list) list) list ->
     state -> state
   val presume: (Attrib.binding * (string * string list) list) list -> state -> state
-  val presume_i: ((Binding.T * attribute list) * (term * term list) list) list ->
+  val presume_i: ((binding * attribute list) * (term * term list) list) list ->
     state -> state
-  val def: (Attrib.binding * ((Binding.T * mixfix) * (string * string list))) list ->
+  val def: (Attrib.binding * ((binding * mixfix) * (string * string list))) list ->
     state -> state
-  val def_i: ((Binding.T * attribute list) *
-    ((Binding.T * mixfix) * (term * term list))) list -> state -> state
+  val def_i: ((binding * attribute list) *
+    ((binding * mixfix) * (term * term list))) list -> state -> state
   val chain: state -> state
   val chain_facts: thm list -> state -> state
   val get_thmss: state -> (Facts.ref * Attrib.src list) list -> thm list
   val note_thmss: (Attrib.binding * (Facts.ref * Attrib.src list) list) list -> state -> state
-  val note_thmss_i: ((Binding.T * attribute list) *
+  val note_thmss_i: ((binding * attribute list) *
     (thm list * attribute list) list) list -> state -> state
   val from_thmss: ((Facts.ref * Attrib.src list) list) list -> state -> state
   val from_thmss_i: ((thm list * attribute list) list) list -> state -> state
@@ -87,7 +87,7 @@
     (theory -> 'a -> attribute) ->
     (context * 'b list -> context * (term list list * (context -> context))) ->
     string -> Method.text option -> (thm list list -> state -> state) ->
-    ((Binding.T * 'a list) * 'b) list -> state -> state
+    ((binding * 'a list) * 'b) list -> state -> state
   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
@@ -107,11 +107,11 @@
   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) ->
-    ((Binding.T * attribute list) * (term * term list) list) list -> bool -> state -> state
+    ((binding * attribute list) * (term * term list) list) list -> bool -> state -> state
   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) ->
-    ((Binding.T * attribute list) * (term * term list) list) list -> bool -> state -> state
+    ((binding * attribute list) * (term * term list) list) list -> bool -> state -> state
   val schematic_goal: state -> bool
   val is_relevant: state -> bool
   val local_future_proof: (state -> ('a * state) Future.future) ->
--- a/src/Pure/Isar/proof_context.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/proof_context.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/proof_context.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 The key concept of Isar proof contexts: elevates primitive local
@@ -23,7 +22,7 @@
   val abbrev_mode: Proof.context -> bool
   val set_stmt: bool -> Proof.context -> Proof.context
   val naming_of: Proof.context -> NameSpace.naming
-  val full_name: Proof.context -> Binding.T -> string
+  val full_name: Proof.context -> binding -> string
   val full_bname: Proof.context -> bstring -> string
   val consts_of: Proof.context -> Consts.T
   val const_syntax_name: Proof.context -> string -> string
@@ -105,27 +104,27 @@
   val restore_naming: Proof.context -> Proof.context -> Proof.context
   val reset_naming: Proof.context -> Proof.context
   val note_thmss: string ->
-    ((Binding.T * attribute list) * (Facts.ref * attribute list) list) list ->
+    ((binding * attribute list) * (Facts.ref * attribute list) list) list ->
       Proof.context -> (string * thm list) list * Proof.context
   val note_thmss_i: string ->
-    ((Binding.T * attribute list) * (thm list * attribute list) list) list ->
+    ((binding * attribute list) * (thm list * attribute list) list) list ->
       Proof.context -> (string * thm list) list * Proof.context
   val put_thms: bool -> string * thm list option -> Proof.context -> Proof.context
-  val read_vars: (Binding.T * string option * mixfix) list -> Proof.context ->
-    (Binding.T * typ option * mixfix) list * Proof.context
-  val cert_vars: (Binding.T * typ option * mixfix) list -> Proof.context ->
-    (Binding.T * typ option * mixfix) list * Proof.context
-  val add_fixes: (Binding.T * string option * mixfix) list ->
+  val read_vars: (binding * string option * mixfix) list -> Proof.context ->
+    (binding * typ option * mixfix) list * Proof.context
+  val cert_vars: (binding * typ option * mixfix) list -> Proof.context ->
+    (binding * typ option * mixfix) list * Proof.context
+  val add_fixes: (binding * string option * mixfix) list ->
     Proof.context -> string list * Proof.context
-  val add_fixes_i: (Binding.T * typ option * mixfix) list ->
+  val add_fixes_i: (binding * typ option * mixfix) list ->
     Proof.context -> string list * Proof.context
   val auto_fixes: Proof.context * (term list list * 'a) -> Proof.context * (term list list * 'a)
   val bind_fixes: string list -> Proof.context -> (term -> term) * Proof.context
   val add_assms: Assumption.export ->
-    ((Binding.T * attribute list) * (string * string list) list) list ->
+    ((binding * attribute list) * (string * string list) list) list ->
     Proof.context -> (string * thm list) list * Proof.context
   val add_assms_i: Assumption.export ->
-    ((Binding.T * attribute list) * (term * term list) list) list ->
+    ((binding * attribute list) * (term * term list) list) list ->
     Proof.context -> (string * thm list) list * Proof.context
   val add_cases: bool -> (string * RuleCases.T option) list -> Proof.context -> Proof.context
   val apply_case: RuleCases.T -> Proof.context -> (string * term list) list * Proof.context
@@ -135,7 +134,7 @@
     Context.generic -> Context.generic
   val add_const_constraint: string * typ option -> Proof.context -> Proof.context
   val add_abbrev: string -> Properties.T ->
-    Binding.T * term -> Proof.context -> (term * term) * Proof.context
+    binding * term -> Proof.context -> (term * term) * Proof.context
   val revert_abbrev: string -> string -> Proof.context -> Proof.context
   val verbose: bool ref
   val setmp_verbose: ('a -> 'b) -> 'a -> 'b
--- a/src/Pure/Isar/proof_display.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/proof_display.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/proof_display.ML
-    ID:         $Id$
     Author:     Makarius
 
 Printing of theorems, goals, results etc.
--- a/src/Pure/Isar/proof_node.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/proof_node.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/proof_node.ML
-    ID:         $Id$
     Author:     Makarius
 
 Proof nodes with linear position and backtracking.
--- a/src/Pure/Isar/rule_insts.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/rule_insts.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/rule_insts.ML
-    ID:         $Id$
     Author:     Makarius
 
 Rule instantiations -- operations within a rule/subgoal context.
--- a/src/Pure/Isar/skip_proof.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/skip_proof.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/skip_proof.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Skipping proofs -- quick_and_dirty mode.
--- a/src/Pure/Isar/spec_parse.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/spec_parse.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -15,16 +15,15 @@
   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 spec_name: ((binding * string) * Attrib.src list) parser
+  val spec_opt_name: ((binding * 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_fixes: (binding * 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
@@ -32,7 +31,7 @@
   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
+    (binding * ((Attrib.binding * string list) list * (binding * string option) list)) list parser
 end;
 
 structure SpecParse: SPEC_PARSE =
@@ -115,13 +114,6 @@
 
 val class_expr = plus1_unless locale_keyword P.xname;
 
-val locale_expr =
-  let
-    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 =
   let
     fun expr2 x = P.xname x;
--- a/src/Pure/Isar/specification.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/specification.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/specification.ML
-    ID:         $Id$
     Author:     Makarius
 
 Derived local theory specifications --- with type-inference and
@@ -9,33 +8,33 @@
 signature SPECIFICATION =
 sig
   val print_consts: local_theory -> (string * typ -> bool) -> (string * typ) list -> unit
-  val check_specification: (Binding.T * typ option * mixfix) list ->
+  val check_specification: (binding * typ option * mixfix) list ->
     (Attrib.binding * term list) list list -> Proof.context ->
-    (((Binding.T * typ) * mixfix) list * (Attrib.binding * term list) list) * Proof.context
-  val read_specification: (Binding.T * string option * mixfix) list ->
+    (((binding * typ) * mixfix) list * (Attrib.binding * term list) list) * Proof.context
+  val read_specification: (binding * string option * mixfix) list ->
     (Attrib.binding * string list) list list -> Proof.context ->
-    (((Binding.T * typ) * mixfix) list * (Attrib.binding * term list) list) * Proof.context
-  val check_free_specification: (Binding.T * typ option * mixfix) list ->
+    (((binding * typ) * mixfix) list * (Attrib.binding * term list) list) * Proof.context
+  val check_free_specification: (binding * typ option * mixfix) list ->
     (Attrib.binding * term list) list -> Proof.context ->
-    (((Binding.T * typ) * mixfix) list * (Attrib.binding * term list) list) * Proof.context
-  val read_free_specification: (Binding.T * string option * mixfix) list ->
+    (((binding * typ) * mixfix) list * (Attrib.binding * term list) list) * Proof.context
+  val read_free_specification: (binding * string option * mixfix) list ->
     (Attrib.binding * string list) list -> Proof.context ->
-    (((Binding.T * typ) * mixfix) list * (Attrib.binding * term list) list) * Proof.context
-  val axiomatization: (Binding.T * typ option * mixfix) list ->
+    (((binding * typ) * mixfix) list * (Attrib.binding * term list) list) * Proof.context
+  val axiomatization: (binding * typ option * mixfix) list ->
     (Attrib.binding * term list) list -> theory ->
     (term list * (string * thm list) list) * theory
-  val axiomatization_cmd: (Binding.T * string option * mixfix) list ->
+  val axiomatization_cmd: (binding * string option * mixfix) list ->
     (Attrib.binding * string list) list -> theory ->
     (term list * (string * thm list) list) * theory
   val definition:
-    (Binding.T * typ option * mixfix) option * (Attrib.binding * term) ->
+    (binding * typ option * mixfix) option * (Attrib.binding * term) ->
     local_theory -> (term * (string * thm)) * local_theory
   val definition_cmd:
-    (Binding.T * string option * mixfix) option * (Attrib.binding * string) ->
+    (binding * string option * mixfix) option * (Attrib.binding * string) ->
     local_theory -> (term * (string * thm)) * local_theory
-  val abbreviation: Syntax.mode -> (Binding.T * typ option * mixfix) option * term ->
+  val abbreviation: Syntax.mode -> (binding * typ option * mixfix) option * term ->
     local_theory -> local_theory
-  val abbreviation_cmd: Syntax.mode -> (Binding.T * string option * mixfix) option * string ->
+  val abbreviation_cmd: Syntax.mode -> (binding * string option * mixfix) option * string ->
     local_theory -> local_theory
   val notation: bool -> Syntax.mode -> (term * mixfix) list -> local_theory -> local_theory
   val notation_cmd: bool -> Syntax.mode -> (string * mixfix) list -> local_theory -> local_theory
@@ -149,7 +148,8 @@
 
     (*axioms*)
     val (axioms, axioms_thy) = consts_thy |> fold_map (fn ((b, atts), props) =>
-        fold_map Thm.add_axiom (PureThy.name_multi (Binding.base_name b) (map subst props))
+        fold_map Thm.add_axiom
+          ((map o apfst) Binding.name (PureThy.name_multi (Binding.base_name b) (map subst props)))
         #>> (fn ths => ((b, atts), [(map Drule.standard' ths, [])]))) specs;
     val (facts, thy') = axioms_thy |> PureThy.note_thmss Thm.axiomK
       (Attrib.map_facts (Attrib.attribute_i axioms_thy) axioms);
--- a/src/Pure/Isar/theory_target.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/theory_target.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -6,7 +6,7 @@
 
 signature THEORY_TARGET =
 sig
-  val peek: local_theory -> {target: string, new_locale: bool, is_locale: bool,
+  val peek: local_theory -> {target: string, is_locale: bool,
     is_class: bool, instantiation: string list * (string * sort) list * sort,
     overloading: (string * (string * typ) * bool) list}
   val init: string option -> theory -> local_theory
@@ -21,34 +21,17 @@
 structure TheoryTarget: THEORY_TARGET =
 struct
 
-(* new locales *)
-
-fun locale_extern new_locale 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 Locale.add_type_syntax x else Old_Locale.add_type_syntax x;
-fun locale_add_term_syntax new_locale 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 Locale.add_declaration x else Old_Locale.add_declaration x;
-fun locale_add_thmss new_locale 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 Locale.init x else Old_Locale.init x;
-fun locale_intern new_locale x =
-  if new_locale then Locale.intern x else Old_Locale.intern x;
-
 (* context data *)
 
-datatype target = Target of {target: string, new_locale: bool, is_locale: bool,
+datatype target = Target of {target: string, is_locale: bool,
   is_class: bool, instantiation: string list * (string * sort) list * sort,
   overloading: (string * (string * typ) * bool) list};
 
-fun make_target target new_locale is_locale is_class instantiation overloading =
-  Target {target = target, new_locale = new_locale, is_locale = is_locale,
+fun make_target target is_locale is_class instantiation overloading =
+  Target {target = target, is_locale = is_locale,
     is_class = is_class, instantiation = instantiation, overloading = overloading};
 
-val global_target = make_target "" false false false ([], [], []) [];
+val global_target = make_target "" false false ([], [], []) [];
 
 structure Data = ProofDataFun
 (
@@ -64,7 +47,7 @@
 fun pretty_thy ctxt target is_locale is_class =
   let
     val thy = ProofContext.theory_of ctxt;
-    val target_name = (if is_class then "class " else "locale ") ^ locale_extern is_class thy target;
+    val target_name = (if is_class then "class " else "locale ") ^ Locale.extern thy target;
     val fixes = map (fn (x, T) => (Binding.name x, SOME T, NoSyn))
       (#1 (ProofContext.inferred_fixes ctxt));
     val assumes = map (fn A => (Attrib.empty_binding, [(Thm.term_of A, [])]))
@@ -89,7 +72,7 @@
 
 (* target declarations *)
 
-fun target_decl add (Target {target, new_locale, ...}) d lthy =
+fun target_decl add (Target {target, ...}) d lthy =
   let
     val d' = Morphism.transform (LocalTheory.target_morphism lthy) d;
     val d0 = Morphism.form d';
@@ -100,12 +83,12 @@
       |> LocalTheory.target (Context.proof_map d0)
     else
       lthy
-      |> LocalTheory.target (add new_locale target d')
+      |> LocalTheory.target (add target d')
   end;
 
-val type_syntax = target_decl locale_add_type_syntax;
-val term_syntax = target_decl locale_add_term_syntax;
-val declaration = target_decl locale_add_declaration;
+val type_syntax = target_decl Locale.add_type_syntax;
+val term_syntax = target_decl Locale.add_term_syntax;
+val declaration = target_decl Locale.add_declaration;
 
 fun class_target (Target {target, ...}) f =
   LocalTheory.raw_theory f #>
@@ -166,7 +149,7 @@
   |> ProofContext.note_thmss_i kind facts
   ||> ProofContext.restore_naming ctxt;
 
-fun notes (Target {target, is_locale, new_locale, ...}) kind facts lthy =
+fun notes (Target {target, is_locale, ...}) kind facts lthy =
   let
     val thy = ProofContext.theory_of lthy;
     val facts' = facts
@@ -185,7 +168,7 @@
         #> PureThy.note_thmss_grouped kind (LocalTheory.group_of lthy) global_facts #> snd
         #> Sign.restore_naming thy)
     |> not is_locale ? LocalTheory.target (note_local kind global_facts #> snd)
-    |> is_locale ? LocalTheory.target (locale_add_thmss new_locale target kind target_facts)
+    |> is_locale ? LocalTheory.target (Locale.add_thmss target kind target_facts)
     |> note_local kind local_facts
   end;
 
@@ -313,7 +296,7 @@
           (fn name => fn (Const (c, _), rhs) => Overloading.define checked name (c, rhs))
       | NONE =>
           if is_none (Class_Target.instantiation_param lthy c)
-          then (fn name => fn eq => Thm.add_def false false (name, Logic.mk_equals eq))
+          then (fn name => fn eq => Thm.add_def false false (Binding.name name, Logic.mk_equals eq))
           else (fn name => fn (Const (c, _), rhs) => AxClass.define_overloaded name (c, rhs)));
     val (global_def, lthy3) = lthy2
       |> LocalTheory.theory_result (define_const (Binding.base_name name') (lhs', rhs'));
@@ -335,13 +318,13 @@
 fun init_target _ NONE = global_target
   | init_target thy (SOME target) =
       make_target target (Locale.defined thy (Locale.intern thy target))
-      true (Class_Target.is_class thy target) ([], [], []) [];
+      (Class_Target.is_class thy target) ([], [], []) [];
 
-fun init_ctxt (Target {target, new_locale, is_locale, is_class, instantiation, overloading}) =
+fun init_ctxt (Target {target, is_locale, is_class, instantiation, overloading}) =
   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 if not is_class then Locale.init target
   else Class_Target.init target;
 
 fun init_lthy (ta as Target {target, instantiation, overloading, ...}) =
@@ -375,7 +358,7 @@
     val ctxt = ProofContext.init thy;
     val ops = raw_ops |> map (fn (name, const, checked) =>
       (name, Term.dest_Const (prep_const ctxt const), checked));
-  in thy |> init_lthy_ctxt (make_target "" false false false ([], [], []) ops) end;
+  in thy |> init_lthy_ctxt (make_target "" false false ([], [], []) ops) end;
 
 in
 
@@ -383,10 +366,9 @@
 fun begin target ctxt = init_lthy (init_target (ProofContext.theory_of ctxt) (SOME target)) ctxt;
 
 fun context "-" thy = init NONE thy
-  | context target thy = init (SOME (locale_intern
-      (Locale.defined thy (Locale.intern thy target)) thy target)) thy;
+  | context target thy = init (SOME (Locale.intern thy target)) thy;
 
-fun instantiation arities = init_lthy_ctxt (make_target "" false false false arities []);
+fun instantiation arities = init_lthy_ctxt (make_target "" false false arities []);
 fun instantiation_cmd raw_arities thy =
   instantiation (read_multi_arity thy raw_arities) thy;
 
--- a/src/Pure/Isar/toplevel.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Isar/toplevel.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -96,6 +96,7 @@
   val transition: bool -> transition -> state -> (state * (exn * string) option) option
   val commit_exit: Position.T -> transition
   val command: transition -> state -> state
+  val run_command: string -> transition -> state -> state option
   val excursion: (transition * transition list) list -> (transition * state) list lazy
 end;
 
@@ -693,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 *)
 
@@ -702,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
--- a/src/Pure/ML-Systems/alice.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ML-Systems/alice.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ML-Systems/exn.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ML-Systems/install_pp_polyml.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ML-Systems/ml_name_space.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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/multithreading.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ML-Systems/multithreading.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ML-Systems/multithreading_polyml.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ML-Systems/overloading_smlnj.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ML-Systems/polyml-4.1.3.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ML-Systems/polyml-4.1.4.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ML-Systems/polyml-4.2.0.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ML-Systems/polyml-5.0.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML-Systems/polyml-5.0.ML
-    ID:         $Id$
 
 Compatibility wrapper for Poly/ML 5.0.
 *)
@@ -11,3 +10,5 @@
 
 val pointer_eq = PolyML.pointerEq;
 
+fun share_common_data () = PolyML.shareCommonData PolyML.rootFunction;
+
--- a/src/Pure/ML-Systems/polyml-5.1.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ML-Systems/polyml-5.1.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML-Systems/polyml-5.1.ML
-    ID:         $Id$
 
 Compatibility wrapper for Poly/ML 5.1.
 *)
@@ -9,3 +8,6 @@
 use "ML-Systems/polyml_old_compiler5.ML";
 
 val pointer_eq = PolyML.pointerEq;
+
+fun share_common_data () = PolyML.shareCommonData PolyML.rootFunction;
+
--- a/src/Pure/ML-Systems/polyml.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ML-Systems/polyml.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML-Systems/polyml.ML
-    ID:         $Id$
 
 Compatibility wrapper for Poly/ML 5.2 or later.
 *)
@@ -13,6 +12,8 @@
 
 val pointer_eq = PolyML.pointerEq;
 
+fun share_common_data () = PolyML.shareCommonData PolyML.rootFunction;
+
 
 (* runtime compilation *)
 
--- a/src/Pure/ML-Systems/polyml_common.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ML-Systems/polyml_common.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ML-Systems/polyml_old_basis.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ML-Systems/polyml_old_compiler4.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ML-Systems/polyml_old_compiler5.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ML-Systems/proper_int.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ML-Systems/smlnj.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ML-Systems/system_shell.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ML-Systems/thread_dummy.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ML-Systems/time_limit.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ML-Systems/universal.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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/ML/ml_antiquote.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ML/ml_antiquote.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML/ml_antiquote.ML
-    ID:         $Id$
     Author:     Makarius
 
 Common ML antiquotations.
--- a/src/Pure/ML/ml_context.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ML/ml_context.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML/ml_context.ML
-    ID:         $Id$
     Author:     Makarius
 
 ML context and antiquotations.
@@ -126,7 +125,8 @@
 
 fun ml_store sel (name, ths) =
   let
-    val ths' = Context.>>> (Context.map_theory_result (PureThy.store_thms (name, ths)));
+    val ths' = Context.>>> (Context.map_theory_result
+      (PureThy.store_thms (Binding.name name, ths)));
     val _ =
       if name = "" then ()
       else if not (ML_Syntax.is_identifier name) then
--- a/src/Pure/ML/ml_lex.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ML/ml_lex.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML/ml_lex.ML
-    ID:         $Id$
     Author:     Makarius
 
 Lexical syntax for SML.
--- a/src/Pure/ML/ml_parse.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ML/ml_parse.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML/ml_parse.ML
-    ID:         $Id$
     Author:     Makarius
 
 Minimal parsing for SML -- fixing integer numerals.
--- a/src/Pure/ML/ml_syntax.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ML/ml_syntax.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML/ml_syntax.ML
-    ID:         $Id$
     Author:     Makarius
 
 Basic ML syntax operations.
--- a/src/Pure/ML/ml_thms.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ML/ml_thms.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ML/ml_thms.ML
-    ID:         $Id$
     Author:     Makarius
 
 Isar theorem values within ML.
--- a/src/Pure/Proof/extraction.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Proof/extraction.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -546,7 +546,7 @@
 
       | corr d defs vs ts Ts hs (prf0 as PThm (_, ((name, prop, SOME Ts'), body))) _ _ =
           let
-            val prf = force_proof body;
+            val prf = join_proof body;
             val (vs', tye) = find_inst prop Ts ts vs;
             val tye' = (map fst (OldTerm.term_tvars prop) ~~ Ts') @ tye;
             val T = etype_of thy' vs' [] prop;
@@ -570,7 +570,7 @@
                       (proof_combt
                          (PThm (serial (),
                           ((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))
+                            Future.value (make_proof_body corr_prf))), vfs_of corr_prop))
                       (map (get_var_type corr_prop) (vfs_of prop))
                   in
                     ((name, (vs', ((nullt, nullt), (corr_prf, corr_prf')))) :: defs'',
@@ -636,7 +636,7 @@
 
       | extr d defs vs ts Ts hs (prf0 as PThm (_, ((s, prop, SOME Ts'), body))) =
           let
-            val prf = force_proof body;
+            val prf = join_proof body;
             val (vs', tye) = find_inst prop Ts ts vs;
             val tye' = (map fst (OldTerm.term_tvars prop) ~~ Ts') @ tye
           in
@@ -681,7 +681,7 @@
                       (proof_combt
                         (PThm (serial (),
                          ((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))
+                           Future.value (make_proof_body corr_prf'))), vfs_of corr_prop))
                       (map (get_var_type corr_prop) (vfs_of prop));
                   in
                     ((s, (vs', ((t', u), (corr_prf', corr_prf'')))) :: defs'',
@@ -733,11 +733,11 @@
              val (def_thms, thy') = if t = nullt then ([], thy) else
                thy
                |> Sign.add_consts_i [(extr_name s vs, fastype_of ft, NoSyn)]
-               |> PureThy.add_defs false [((extr_name s vs ^ "_def",
+               |> PureThy.add_defs false [((Binding.name (extr_name s vs ^ "_def"),
                     Logic.mk_equals (head_of (strip_abs_body fu), ft)), [])]
            in
              thy'
-             |> PureThy.store_thm (corr_name s vs,
+             |> PureThy.store_thm (Binding.name (corr_name s vs),
                   Thm.varifyT (funpow (length (OldTerm.term_vars corr_prop))
                     (Thm.forall_elim_var 0) (forall_intr_frees
                       (ProofChecker.thm_of_proof thy'
--- a/src/Pure/Proof/proof_syntax.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Proof/proof_syntax.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Proof/proof_syntax.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Function for parsing and printing proof terms.
@@ -229,7 +228,7 @@
     val prop = Thm.full_prop_of thm;
     val prf = Thm.proof_of thm;
     val prf' = (case strip_combt (fst (strip_combP prf)) of
-        (PThm (_, ((_, prop', _), body)), _) => if prop = prop' then force_proof body else prf
+        (PThm (_, ((_, prop', _), body)), _) => if prop = prop' then join_proof body else prf
       | _ => prf)
   in if full then Reconstruct.reconstruct_proof thy prop prf' else prf' end;
 
--- a/src/Pure/Proof/reconstruct.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Proof/reconstruct.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -358,7 +358,7 @@
                     val _ = message ("Reconstructing proof of " ^ a);
                     val _ = message (Syntax.string_of_term_global thy prop);
                     val prf' = forall_intr_vfs_prf prop
-                      (reconstruct_proof thy prop (force_proof body));
+                      (reconstruct_proof thy prop (join_proof body));
                     val (maxidx', prfs', prf) = expand
                       (maxidx_proof prf' ~1) prfs prf'
                   in (maxidx' + maxidx + 1, inc (maxidx + 1) prf,
--- a/src/Pure/ProofGeneral/ROOT.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ProofGeneral/ROOT.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ProofGeneral/ROOT.ML
-    ID:         $Id$
     Author:     David Aspinall
 
 Proof General interface for Isabelle, both the traditional Emacs version,
--- a/src/Pure/ProofGeneral/pgip.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ProofGeneral/pgip.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ProofGeneral/pgip.ML
-    ID:         $Id$
     Author:     David Aspinall
 
 Prover-side PGIP abstraction.  
--- a/src/Pure/ProofGeneral/pgip_input.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ProofGeneral/pgip_input.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ProofGeneral/pgip_input.ML
-    ID:         $Id$
     Author:     David Aspinall
 
 PGIP abstraction: input commands.
--- a/src/Pure/ProofGeneral/pgip_isabelle.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ProofGeneral/pgip_isabelle.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ProofGeneral/pgip_isabelle.ML
-    ID:         $Id$
     Author:     David Aspinall
 
 Prover-side PGIP abstraction: Isabelle configuration and mapping to Isabelle types.
--- a/src/Pure/ProofGeneral/pgip_markup.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ProofGeneral/pgip_markup.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ProofGeneral/pgip_markup.ML
-    ID:         $Id$
     Author:     David Aspinall
 
 PGIP abstraction: document markup for proof scripts (in progress).
--- a/src/Pure/ProofGeneral/pgip_output.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ProofGeneral/pgip_output.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ProofGeneral/pgip_output.ML
-    ID:         $Id$
     Author:     David Aspinall
 
 PGIP abstraction: output commands.
--- a/src/Pure/ProofGeneral/pgip_parser.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ProofGeneral/pgip_parser.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ProofGeneral/pgip_parser.ML
-    ID:         $Id$
     Author:     David Aspinall and Makarius
 
 Parsing theory sources without execution (via keyword classification).
--- a/src/Pure/ProofGeneral/pgip_tests.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ProofGeneral/pgip_tests.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ProofGeneral/pgip_tests.ML
-    ID:         $Id$
     Author:     David Aspinall
 
 A test suite for the PGIP abstraction code (in progress).
--- a/src/Pure/ProofGeneral/pgip_types.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ProofGeneral/pgip_types.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ProofGeneral/pgip_types.ML
-    ID:         $Id$
     Author:     David Aspinall
 
 PGIP abstraction: types and conversions.
--- a/src/Pure/ProofGeneral/pgml.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ProofGeneral/pgml.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ProofGeneral/pgml.ML
-    ID:         $Id$
     Author:     David Aspinall
 
 PGIP abstraction: PGML
--- a/src/Pure/ProofGeneral/proof_general_keywords.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ProofGeneral/proof_general_keywords.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ProofGeneral/proof_general_keywords.ML
-    ID:         $Id$
     Author:     Makarius
 
 Dummy session with outer syntax keyword initialization.
--- a/src/Pure/ProofGeneral/proof_general_pgip.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/ProofGeneral/proof_general_pgip.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/ProofGeneral/proof_general_pgip.ML
-    ID:         $Id$
     Author:     David Aspinall and Markus Wenzel
 
 Isabelle configuration for Proof General using PGIP protocol.
@@ -257,7 +256,7 @@
   (case Thm.proof_body_of th of 
     PBody {proof = PThm (_, ((name, _, _), body)), ...} =>
       if Thm.has_name_hint th andalso Thm.get_name_hint th = name
-      then add_proof_body (Lazy.force body)
+      then add_proof_body (Future.join body)
       else I
   | body => add_proof_body body);
 
--- a/src/Pure/Pure.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Pure.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,6 +1,3 @@
-(*  Title:      Pure/Pure.thy
-    ID:         $Id$
-*)
 
 section {* Further content for the Pure theory *}
 
--- a/src/Pure/Syntax/ast.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Syntax/ast.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Syntax/lexicon.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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).
--- a/src/Pure/Syntax/mixfix.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Syntax/mixfix.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Syntax/parser.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Syntax/printer.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Syntax/simple_syntax.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Syntax/syn_ext.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Syntax/syn_ext.ML
-    ID:         $Id$
     Author:     Markus Wenzel and Carsten Clasohm, TU Muenchen
 
 Syntax extension (internal interface).
--- a/src/Pure/Syntax/type_ext.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Syntax/type_ext.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -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/html.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Thy/html.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Thy/html.ML
-    ID:         $Id$
     Author:     Markus Wenzel and Stefan Berghofer, TU Muenchen
 
 HTML presentation elements.
--- a/src/Pure/Thy/latex.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Thy/latex.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Thy/latex.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 LaTeX presentation elements -- based on outer lexical syntax.
--- a/src/Pure/Thy/present.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Thy/present.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Thy/present.ML
-    ID:         $Id$
     Author:     Markus Wenzel and Stefan Berghofer, TU Muenchen
 
 Theory presentation: HTML, graph files, (PDF)LaTeX documents.
@@ -467,7 +466,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/term_style.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Thy/term_style.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Thy/term_style.ML
-    ID:         $Id$
     Author:     Florian Haftmann, TU Muenchen
 
 Styles for terms, to use with the "term_style" and "thm_style"
--- a/src/Pure/Thy/thm_deps.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Thy/thm_deps.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Thy/thm_deps.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Visualize dependencies of theorems.
--- a/src/Pure/Thy/thy_header.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Thy/thy_header.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Thy/thy_header.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Theory headers -- independent of outer syntax.
--- a/src/Pure/Thy/thy_load.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Thy/thy_load.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -25,6 +25,7 @@
   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
@@ -95,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 *)
 
@@ -104,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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Thy/thy_output.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Thy/thy_output.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Theory document output.
@@ -11,6 +10,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
--- a/src/Pure/Tools/ROOT.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Tools/ROOT.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -9,8 +9,5 @@
 (*basic XML support*)
 use "xml_syntax.ML";
 
-(*derived theory and proof elements*)
-use "invoke.ML";
-
 (*quickcheck needed here because of pg preferences*)
 use "../../Tools/quickcheck.ML"
--- a/src/Pure/Tools/invoke.ML	Thu Jan 29 12:05:19 2009 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,127 +0,0 @@
-(*  Title:      Pure/Tools/invoke.ML
-    Author:     Makarius
-
-Schematic invocation of locale expression in proof context.
-*)
-
-signature INVOKE =
-sig
-  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 -> Old_Locale.expr -> term option list ->
-    (Binding.T * typ option * mixfix) list -> bool -> Proof.state -> Proof.state
-end;
-
-structure Invoke: INVOKE =
-struct
-
-
-(* invoke *)
-
-local
-
-fun gen_invoke prep_att prep_expr parse_term add_fixes
-    (prfx, raw_atts) raw_expr raw_insts fixes int state =
-  let
-    val thy = Proof.theory_of state;
-    val _ = Proof.assert_forward_or_chain state;
-    val chain_facts = if can Proof.assert_chain state then Proof.the_facts state else [];
-
-    val more_atts = map (prep_att thy) raw_atts;
-    val (elems, _) = prep_expr raw_expr [] (ProofContext.init thy);
-
-    val prems = maps Element.prems_of elems;
-    val params = maps Element.params_of elems;
-    val types = rev (fold Term.add_tfrees prems (fold (Term.add_tfreesT o #2) params []));
-
-    val prems' = map Logic.varify prems;
-    val params' = map (Logic.varify o Free) params;
-    val types' = map (Logic.varifyT o TFree) types;
-
-    val state' = state
-      |> Proof.enter_forward
-      |> Proof.begin_block
-      |> Proof.map_context (snd o add_fixes fixes);
-    val ctxt' = Proof.context_of state';
-
-    val raw_insts' = zip_options params' raw_insts
-      handle Library.UnequalLengths => error "Too many instantiations";
-
-    fun prep_inst (t, u) =
-      TypeInfer.constrain (TypeInfer.paramify_vars (Term.fastype_of t)) (parse_term ctxt' u);
-    val insts = map #1 raw_insts' ~~
-      Variable.polymorphic ctxt' (Syntax.check_terms ctxt' (map prep_inst raw_insts'));
-    val inst_rules =
-      replicate (length types') Drule.termI @
-      map (fn t =>
-        (case AList.lookup (op =) insts t of
-          SOME u => Drule.mk_term (Thm.cterm_of thy u)
-        | NONE => Drule.termI)) params';
-
-    val propp =
-      [((Binding.empty, []), map (rpair [] o Logic.mk_term o Logic.mk_type) types'),
-       ((Binding.empty, []), map (rpair [] o Logic.mk_term) params'),
-       ((Binding.empty, []), map (rpair [] o Element.mark_witness) prems')];
-    fun after_qed results =
-      Proof.end_block #>
-      Proof.map_context (fn ctxt =>
-        let
-          val ([res_types, res_params, res_prems], ctxt'') =
-            fold_burrow (apfst snd oo Variable.import_thms false) results ctxt';
-
-          val types'' = map (Logic.dest_type o Thm.term_of o Drule.dest_term) res_types;
-          val params'' = map (Thm.term_of o Drule.dest_term) res_params;
-          val inst = Element.morph_ctxt (Element.inst_morphism thy
-            (Symtab.make (map #1 types ~~ types''), Symtab.make (map #1 params ~~ params'')));
-          val elems' = map inst elems;
-          val prems'' = map2 Element.make_witness (maps Element.prems_of elems') res_prems;
-          val notes =
-            maps (Element.facts_of thy) elems'
-            |> Element.satisfy_facts prems''
-            |> Element.generalize_facts ctxt'' ctxt
-            |> Attrib.map_facts (Attrib.attribute_i thy)
-            |> map (fn ((a, atts), bs) => ((a, atts @ more_atts), bs));
-        in
-          ctxt
-          |> ProofContext.sticky_prefix prfx
-          |> ProofContext.qualified_names
-          |> (snd o ProofContext.note_thmss_i "" notes)
-          |> ProofContext.restore_naming ctxt
-        end) #>
-      Proof.put_facts NONE;
-  in
-    state'
-    |> Proof.chain_facts chain_facts
-    |> Proof.local_goal (K (K ())) (K I) ProofContext.bind_propp_schematic_i
-      "invoke" NONE after_qed propp
-    |> Element.refine_witness
-    |> Seq.hd
-    |> Proof.refine (Method.Basic (K (Method.METHOD (K (HEADGOAL (RANGE (map rtac inst_rules))))),
-      Position.none))
-    |> Seq.hd
-  end;
-
-in
-
-fun invoke 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;
-
-
-(* concrete syntax *)
-
-local structure P = OuterParse and K = OuterKeyword in
-
-val _ =
-  OuterSyntax.command "invoke"
-    "schematic invocation of locale expression in proof context"
-    (K.tag_proof K.prf_goal)
-    (SpecParse.opt_thm_name ":" -- SpecParse.locale_expr -- SpecParse.locale_insts -- P.for_fixes
-      >> (fn ((((name, atts), expr), (insts, _)), fixes) =>
-          Toplevel.print o Toplevel.proof' (invoke (Binding.base_name name, atts) expr insts fixes)));
-
-end;
-
-end;
--- a/src/Pure/Tools/isabelle_process.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Tools/isabelle_process.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -55,9 +55,6 @@
   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 body =
-  YXML.element Markup.messageN [(Markup.classN, class)] [clean_string [Symbol.STX] body];
-
 fun message_pos trees = trees |> get_first
   (fn XML.Elem (name, atts, ts) =>
         if name = Markup.positionN then SOME (Position.of_properties atts)
@@ -69,21 +66,21 @@
 
 in
 
-fun message _ _ _ "" = ()
-  | message out_stream ch class body =
+fun message _ _ "" = ()
+  | message out_stream ch body =
       let
         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 = message_text class body;
+        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 (Session.welcome ());
+    val text = Session.welcome ();
   in output out_stream (special "A" ^ message_props [pid, session] ^ text ^ special_end) end;
 
 end;
@@ -116,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	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Tools/isabelle_process.scala	Thu Jan 29 12:24:00 2009 +0000
@@ -7,7 +7,6 @@
 
 package isabelle
 
-import java.util.Properties
 import java.util.concurrent.LinkedBlockingQueue
 import java.io.{BufferedReader, BufferedWriter, InputStreamReader, OutputStreamWriter,
   InputStream, OutputStream, IOException}
@@ -50,6 +49,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
@@ -65,17 +79,28 @@
       kind == STATUS
   }
 
-  class Result(val kind: Kind.Value, val props: Properties, val result: String) {
+  class Result(val kind: Kind.Value, val props: List[(String, String)], val result: String) {
     override def toString = {
-      val res = XML.content(YXML.parse_failsafe(result)).mkString
-      if (props == null) kind.toString + " [[" + res + "]]"
-      else kind.toString + " " + props.toString + " [[" + res + "]]"
+      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.isEmpty)
+        kind.toString + " [[" + res + "]]"
+      else
+        kind.toString + " " +
+          (for ((x, y) <- props) yield x + "=" + y).mkString("{", ",", "}") + " [[" + res + "]]"
     }
     def is_raw = Kind.is_raw(kind)
     def is_control = Kind.is_control(kind)
     def is_system = Kind.is_system(kind)
   }
 
+  def parse_message(isabelle_system: IsabelleSystem, result: Result) =
+  {
+    XML.Elem(Markup.MESSAGE, (Markup.CLASS, Kind.markup(result.kind)) :: result.props,
+      YXML.parse_body_failsafe(isabelle_system.symbols.decode(result.result)))
+  }
 }
 
 
@@ -93,21 +118,26 @@
 
   /* process information */
 
-  private var proc: Process = null
-  private var closing = false
-  private var pid: String = null
-  private var the_session: String = null
-  def session() = the_session
+  @volatile private var proc: Process = null
+  @volatile private var closing = false
+  @volatile private var pid: String = null
+  @volatile private var the_session: String = null
+  def session = the_session
 
 
   /* results */
 
+  def parse_message(result: Result): XML.Tree =
+    IsabelleProcess.parse_message(isabelle_system, result)
+
   private val result_queue = new LinkedBlockingQueue[Result]
 
-  private def put_result(kind: Kind.Value, props: Properties, result: String) {
-    if (kind == Kind.INIT && props != null) {
-      pid = props.getProperty(Markup.PID)
-      the_session = props.getProperty(Markup.SESSION)
+  private def put_result(kind: Kind.Value, props: List[(String, String)], result: String)
+  {
+    if (kind == Kind.INIT) {
+      val map = Map(props: _*)
+      if (map.isDefinedAt(Markup.PID)) pid = map(Markup.PID)
+      if (map.isDefinedAt(Markup.SESSION)) the_session = map(Markup.SESSION)
     }
     result_queue.put(new Result(kind, props, result))
   }
@@ -121,7 +151,7 @@
           catch { case _: NullPointerException => null }
 
         if (result != null) {
-          results.event(result)  // FIXME try/catch (!??)
+          results.event(result)
           if (result.kind == Kind.EXIT) finished = true
         }
         else finished = true
@@ -134,13 +164,13 @@
 
   def interrupt() = synchronized {
     if (proc == null) error("Cannot interrupt Isabelle: no process")
-    if (pid == null) put_result(Kind.SYSTEM, null, "Cannot interrupt: unknown pid")
+    if (pid == null) put_result(Kind.SYSTEM, Nil, "Cannot interrupt: unknown pid")
     else {
       try {
         if (isabelle_system.execute(true, "kill", "-INT", pid).waitFor == 0)
-          put_result(Kind.SIGNAL, null, "INT")
+          put_result(Kind.SIGNAL, Nil, "INT")
         else
-          put_result(Kind.SYSTEM, null, "Cannot interrupt: kill command failed")
+          put_result(Kind.SYSTEM, Nil, "Cannot interrupt: kill command failed")
       }
       catch { case e: IOException => error("Cannot interrupt Isabelle: " + e.getMessage) }
     }
@@ -151,7 +181,7 @@
     else {
       try_close()
       Thread.sleep(500)
-      put_result(Kind.SIGNAL, null, "KILL")
+      put_result(Kind.SIGNAL, Nil, "KILL")
       proc.destroy
       proc = null
       pid = null
@@ -176,7 +206,7 @@
   def command(text: String) =
     output_sync("Isabelle.command " + IsabelleSyntax.encode_string(text))
 
-  def command(props: Properties, text: String) =
+  def command(props: List[(String, String)], text: String) =
     output_sync("Isabelle.command " + IsabelleSyntax.encode_properties(props) + " " +
       IsabelleSyntax.encode_string(text))
 
@@ -211,17 +241,17 @@
             finished = true
           }
           else {
-            put_result(Kind.STDIN, null, s)
+            put_result(Kind.STDIN, Nil, s)
             writer.write(s)
             writer.flush
           }
           //}}}
         }
         catch {
-          case e: IOException => put_result(Kind.SYSTEM, null, "Stdin thread: " + e.getMessage)
+          case e: IOException => put_result(Kind.SYSTEM, Nil, "Stdin thread: " + e.getMessage)
         }
       }
-      put_result(Kind.SYSTEM, null, "Stdin thread terminated")
+      put_result(Kind.SYSTEM, Nil, "Stdin thread terminated")
     }
   }
 
@@ -245,7 +275,7 @@
             else done = true
           }
           if (result.length > 0) {
-            put_result(Kind.STDOUT, null, result.toString)
+            put_result(Kind.STDOUT, Nil, result.toString)
             result.length = 0
           }
           else {
@@ -256,10 +286,10 @@
           //}}}
         }
         catch {
-          case e: IOException => put_result(Kind.SYSTEM, null, "Stdout thread: " + e.getMessage)
+          case e: IOException => put_result(Kind.SYSTEM, Nil, "Stdout thread: " + e.getMessage)
         }
       }
-      put_result(Kind.SYSTEM, null, "Stdout thread terminated")
+      put_result(Kind.SYSTEM, Nil, "Stdout thread terminated")
     }
   }
 
@@ -270,7 +300,7 @@
     override def run() = {
       val reader = isabelle_system.fifo_reader(fifo)
       var kind: Kind.Value = null
-      var props: Properties = null
+      var props: List[(String, String)] = Nil
       var result = new StringBuilder
 
       var finished = false
@@ -285,7 +315,7 @@
             } while (c >= 0 && c != 2)
 
             if (result.length > 0) {
-              put_result(Kind.SYSTEM, null, "Malformed message:\n" + result.toString)
+              put_result(Kind.SYSTEM, Nil, "Malformed message:\n" + result.toString)
               result.length = 0
             }
             if (c < 0) {
@@ -297,7 +327,6 @@
               c = reader.read
               if (Kind.code.isDefinedAt(c)) kind = Kind.code(c)
               else kind = null
-              props = null
             }
             //}}}
           }
@@ -317,16 +346,16 @@
                 if (i > 0) {
                   val name = line.substring(0, i)
                   val value = line.substring(i + 1, len - 2)
-                  if (props == null) props = new Properties
-                  if (!props.containsKey(name)) props.setProperty(name, value)
+                  props = (name, value) :: props
                 }
               }
               // last text line
               else if (line.endsWith("\u0002.")) {
                 result.append(line.substring(0, len - 2))
-                put_result(kind, props, result.toString)
+                put_result(kind, props.reverse, result.toString)
+                kind = null
+                props = Nil
                 result.length = 0
-                kind = null
               }
               // text line
               else {
@@ -338,10 +367,10 @@
           }
         }
         catch {
-          case e: IOException => put_result(Kind.SYSTEM, null, "Message thread: " + e.getMessage)
+          case e: IOException => put_result(Kind.SYSTEM, Nil, "Message thread: " + e.getMessage)
         }
       }
-      put_result(Kind.SYSTEM, null, "Message thread terminated")
+      put_result(Kind.SYSTEM, Nil, "Message thread terminated")
     }
   }
 
@@ -355,7 +384,7 @@
     {
       val (msg, rc) = isabelle_system.isabelle_tool("version")
       if (rc != 0) error("Version check failed -- bad Isabelle installation:\n" + msg)
-      put_result(Kind.SYSTEM, null, msg)
+      put_result(Kind.SYSTEM, Nil, msg)
     }
 
 
@@ -396,8 +425,8 @@
       override def run() = {
         val rc = proc.waitFor()
         Thread.sleep(300)
-        put_result(Kind.SYSTEM, null, "Exit thread terminated")
-        put_result(Kind.EXIT, null, Integer.toString(rc))
+        put_result(Kind.SYSTEM, Nil, "Exit thread terminated")
+        put_result(Kind.EXIT, Nil, Integer.toString(rc))
         rm_fifo()
       }
     }.start
--- a/src/Pure/Tools/isabelle_syntax.scala	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Tools/isabelle_syntax.scala	Thu Jan 29 12:24:00 2009 +0000
@@ -6,14 +6,13 @@
 
 package isabelle
 
-import java.util.{Properties, Enumeration}
-
 
 object IsabelleSyntax {
 
   /* 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,30 +26,49 @@
     result.append("\"")
   }
 
-  def encode_string(str: String) = {
-    val result = new StringBuilder(str.length + 20)
+  def encode_string(str: String) =
+  {
+    val result = new StringBuilder(str.length + 10)
     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
+    append_list(append_elem, elems, result)
+    result.toString
+  }
+
+
   /* properties */
 
-  def append_properties(props: Properties, result: StringBuilder) = {
-    result.append("(")
-    val names = props.propertyNames.asInstanceOf[Enumeration[String]]
-    while (names.hasMoreElements) {
-      val name = names.nextElement; val value = props.getProperty(name)
-      append_string(name, result); result.append(" = "); append_string(value, result)
-      if (names.hasMoreElements) result.append(", ")
-    }
-    result.append(")")
+  def append_properties(props: List[(String, String)], result: StringBuilder)
+  {
+    append_list((p: (String, String), res) =>
+      { append_string(p._1, res); res.append(" = "); append_string(p._2, res) }, props, result)
   }
 
-  def encode_properties(props: Properties) = {
-    val result = new StringBuilder(40)
+  def encode_properties(props: List[(String, String)]) =
+  {
+    val result = new StringBuilder
     append_properties(props, result)
     result.toString
   }
-
 }
--- a/src/Pure/Tools/isabelle_system.scala	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Tools/isabelle_system.scala	Thu Jan 29 12:24:00 2009 +0000
@@ -143,4 +143,16 @@
     }
     logics.toList.sort(_ < _)
   }
+
+
+  /* symbols */
+
+  private def read_symbols(path: String) = {
+    val file = new File(platform_path(path))
+    if (file.canRead) Source.fromFile(file).getLines
+    else Iterator.empty
+  }
+  val symbols = new Symbol.Interpretation(
+    read_symbols("$ISABELLE_HOME/etc/symbols") ++
+    read_symbols("$ISABELLE_HOME_USER/etc/symbols"))
 }
--- a/src/Pure/Tools/named_thms.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Tools/named_thms.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Tools/named_thms.ML
-    ID:         $Id$
     Author:     Makarius
 
 Named collections of theorems in canonical order.
@@ -36,6 +35,6 @@
 
 val setup =
   Attrib.add_attributes [(name, Attrib.add_del_args add del, "declaration of " ^ description)] #>
-  PureThy.add_thms_dynamic (name, Data.get);
+  PureThy.add_thms_dynamic (Binding.name name, Data.get);
 
 end;
--- a/src/Pure/Tools/xml_syntax.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/Tools/xml_syntax.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Tools/xml_syntax.ML
-    ID:         $Id$
     Author:     Stefan Berghofer, TU Muenchen
 
 Input and output of types, terms, and proofs in XML format.
@@ -159,7 +158,7 @@
   | proof_of_xml (XML.Elem ("PThm", [("name", s)], term :: opttypes)) =
       (* FIXME? *)
       PThm (serial (), ((s, term_of_xml term, opttypes_of_xml opttypes),
-        Lazy.value (Proofterm.make_proof_body MinProof)))
+        Future.value (Proofterm.make_proof_body MinProof)))
   | proof_of_xml (XML.Elem ("PAxm", [("name", s)], term :: opttypes)) =
       PAxm (s, term_of_xml term, opttypes_of_xml opttypes)
   | proof_of_xml (XML.Elem ("Oracle", [("name", s)], term :: opttypes)) =
--- a/src/Pure/assumption.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/assumption.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/assumption.ML
-    ID:         $Id$
     Author:     Makarius
 
 Local assumptions, parameterized by export rules.
@@ -79,7 +78,7 @@
 (* named prems -- legacy feature *)
 
 val _ = Context.>>
-  (Context.map_theory (PureThy.add_thms_dynamic ("prems",
+  (Context.map_theory (PureThy.add_thms_dynamic (Binding.name "prems",
     fn Context.Theory _ => [] | Context.Proof ctxt => prems_of ctxt)));
 
 
@@ -120,6 +119,6 @@
     val thm = export false inner outer;
     val term = export_term inner outer;
     val typ = Logic.type_map term;
-  in Morphism.morphism {binding = I, var = I, typ = typ, term = term, fact = map thm} end;
+  in Morphism.morphism {binding = I, typ = typ, term = term, fact = map thm} end;
 
 end;
--- a/src/Pure/axclass.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/axclass.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/axclass.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Type classes defined as predicates, associated with a record of
@@ -9,7 +8,7 @@
 signature AX_CLASS =
 sig
   val define_class: bstring * class list -> string list ->
-    ((Binding.T * attribute list) * term list) list -> theory -> class * theory
+    ((binding * attribute list) * term list) list -> theory -> class * theory
   val add_classrel: thm -> theory -> theory
   val add_arity: thm -> theory -> theory
   val prove_classrel: class * class -> tactic -> theory -> theory
@@ -295,7 +294,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;
 
@@ -329,7 +328,8 @@
         quote (Syntax.string_of_classrel ctxt [c1, c2]));
   in
     thy
-    |> PureThy.add_thms [((prefix classrel_prefix (Logic.name_classrel (c1, c2)), th), [])]
+    |> PureThy.add_thms [((Binding.name
+        (prefix classrel_prefix (Logic.name_classrel (c1, c2))), th), [])]
     |-> (fn [th'] => add_classrel th')
   end;
 
@@ -345,7 +345,7 @@
           quote (Syntax.string_of_arity ctxt arity));
   in
     thy
-    |> PureThy.add_thms (map (rpair []) (names ~~ ths))
+    |> PureThy.add_thms (map (rpair []) (map Binding.name names ~~ ths))
     |-> fold add_arity
   end;
 
@@ -372,10 +372,10 @@
     |> Sign.no_base_names
     |> Sign.declare_const [] ((Binding.name c', T'), NoSyn)
     |-> (fn const' as Const (c'', _) => Thm.add_def false true
-          (Thm.def_name c', Logic.mk_equals (Const (c, T'), const'))
+          (Binding.name (Thm.def_name c'), Logic.mk_equals (Const (c, T'), const'))
     #>> Thm.varifyT
     #-> (fn thm => add_inst_param (c, tyco) (c'', thm)
-    #> PureThy.add_thm ((c', thm), [Thm.kind_internal])
+    #> PureThy.add_thm ((Binding.name c', thm), [Thm.kind_internal])
     #> snd
     #> Sign.restore_naming thy
     #> pair (Const (c, T))))
@@ -391,7 +391,7 @@
       (NameSpace.base c ^ "_" ^ NameSpace.base tyco) name;
   in
     thy
-    |> Thm.add_def false false (name', prop)
+    |> Thm.add_def false false (Binding.name name', prop)
     |>> (fn thm =>  Drule.transitive_thm OF [eq, thm])
   end;
 
@@ -469,7 +469,7 @@
     val ([def], def_thy) =
       thy
       |> Sign.primitive_class (bclass, super)
-      |> PureThy.add_defs false [((Thm.def_name bconst, class_eq), [])];
+      |> PureThy.add_defs false [((Binding.name (Thm.def_name bconst), class_eq), [])];
     val (raw_intro, (raw_classrel, raw_axioms)) =
       split_defined (length conjs) def ||> chop (length super);
 
@@ -515,7 +515,11 @@
     val args = prep thy raw_args;
     val specs = mk args;
     val names = name args;
-  in thy |> PureThy.add_axioms (map (rpair []) (names ~~ specs)) |-> fold add end;
+  in
+    thy
+    |> PureThy.add_axioms (map (rpair []) (map Binding.name names ~~ specs))
+    |-> fold add
+  end;
 
 fun ax_classrel prep =
   axiomatize (map o prep) (map Logic.mk_classrel)
--- a/src/Pure/config.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/config.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/config.ML
-    ID:         $Id$
     Author:     Makarius
 
 Configuration options as values within the local context.
--- a/src/Pure/conjunction.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/conjunction.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/conjunction.ML
-    ID:         $Id$
     Author:     Makarius
 
 Meta-level conjunction.
--- a/src/Pure/consts.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/consts.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/consts.ML
-    ID:         $Id$
     Author:     Makarius
 
 Polymorphic constants: declarations, abbreviations, additional type
@@ -30,10 +29,10 @@
   val certify: Pretty.pp -> Type.tsig -> bool -> T -> term -> term  (*exception TYPE*)
   val typargs: T -> string * typ -> typ list
   val instance: T -> string * typ list -> typ
-  val declare: bool -> NameSpace.naming -> Properties.T -> (Binding.T * typ) -> T -> T
+  val declare: bool -> NameSpace.naming -> Properties.T -> (binding * typ) -> T -> T
   val constrain: string * typ option -> T -> T
   val abbreviate: Pretty.pp -> Type.tsig -> NameSpace.naming -> string -> Properties.T ->
-    Binding.T * term -> T -> (term * term) * T
+    binding * term -> T -> (term * term) * T
   val revert_abbrev: string -> string -> T -> T
   val hide: bool -> string -> T -> T
   val empty: T
--- a/src/Pure/context_position.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/context_position.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/context_position.ML
-    ID:         $Id$
     Author:     Makarius
 
 Context position visibility flag.
--- a/src/Pure/conv.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/conv.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/conv.ML
-    ID:         $Id$
     Author:     Amine Chaieb and Makarius
 
 Conversions: primitive equality reasoning.
--- a/src/Pure/defs.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/defs.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/defs.ML
-    ID:         $Id$
     Author:     Makarius
 
 Global well-formedness checks for constant definitions.  Covers plain
--- a/src/Pure/display.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/display.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/display.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1993  University of Cambridge
 
--- a/src/Pure/drule.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/drule.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -460,10 +460,10 @@
 val read_prop = certify o SimpleSyntax.read_prop;
 
 fun store_thm name th =
-  Context.>>> (Context.map_theory_result (PureThy.store_thm (name, th)));
+  Context.>>> (Context.map_theory_result (PureThy.store_thm (Binding.name name, th)));
 
 fun store_thm_open name th =
-  Context.>>> (Context.map_theory_result (PureThy.store_thm_open (name, th)));
+  Context.>>> (Context.map_theory_result (PureThy.store_thm_open (Binding.name name, th)));
 
 fun store_standard_thm name th = store_thm name (standard th);
 fun store_standard_thm_open name thm = store_thm_open name (standard' thm);
--- a/src/Pure/facts.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/facts.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -30,9 +30,9 @@
   val props: T -> thm list
   val could_unify: T -> term -> thm list
   val merge: T * T -> T
-  val add_global: NameSpace.naming -> Binding.T * thm list -> T -> string * T
-  val add_local: bool -> NameSpace.naming -> Binding.T * thm list -> T -> string * T
-  val add_dynamic: NameSpace.naming -> Binding.T * (Context.generic -> thm list) -> T -> string * T
+  val add_global: NameSpace.naming -> binding * thm list -> T -> string * T
+  val add_local: bool -> NameSpace.naming -> binding * thm list -> T -> string * T
+  val add_dynamic: NameSpace.naming -> binding * (Context.generic -> thm list) -> T -> string * T
   val del: string -> T -> T
   val hide: bool -> string -> T -> T
 end;
--- a/src/Pure/interpretation.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/interpretation.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/interpretation.ML
-    ID:         $Id$
     Author:     Florian Haftmann and Makarius
 
 Generic interpretation of theory data.
--- a/src/Pure/more_thm.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/more_thm.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -38,8 +38,8 @@
   val forall_elim_vars: int -> thm -> thm
   val unvarify: thm -> thm
   val close_derivation: thm -> thm
-  val add_axiom: bstring * term -> theory -> thm * theory
-  val add_def: bool -> bool -> bstring * term -> theory -> thm * theory
+  val add_axiom: binding * term -> theory -> thm * theory
+  val add_def: bool -> bool -> binding * term -> theory -> thm * theory
   val rule_attribute: (Context.generic -> thm -> thm) -> attribute
   val declaration_attribute: (thm -> Context.generic -> Context.generic) -> attribute
   val theory_attributes: attribute list -> theory * thm -> theory * thm
@@ -276,14 +276,15 @@
 
 (** specification primitives **)
 
-fun add_axiom (name, prop) thy =
+fun add_axiom (b, prop) thy =
   let
-    val name' = if name = "" then "axiom_" ^ serial_string () else name;
-    val thy' = thy |> Theory.add_axioms_i [(name', prop)];
-    val axm = unvarify (Thm.axiom thy' (Sign.full_bname thy' name'));
+    val b' = if Binding.is_empty b
+      then Binding.name ("axiom_" ^ serial_string ()) else b;
+    val thy' = thy |> Theory.add_axioms_i [(b', prop)];
+    val axm = unvarify (Thm.axiom thy' (Sign.full_name thy' b'));
   in (axm, thy') end;
 
-fun add_def unchecked overloaded (name, prop) thy =
+fun add_def unchecked overloaded (b, prop) thy =
   let
     val tfrees = rev (map TFree (Term.add_tfrees prop []));
     val tfrees' = map (fn a => TFree (a, [])) (Name.invents Name.context Name.aT (length tfrees));
@@ -291,8 +292,8 @@
     val recover_sorts = map (pairself (Thm.ctyp_of thy o Logic.varifyT)) (tfrees' ~~ tfrees);
 
     val prop' = Term.map_types (Term.map_atyps (perhaps (AList.lookup (op =) strip_sorts))) prop;
-    val thy' = Theory.add_defs_i unchecked overloaded [(name, prop')] thy;
-    val axm' = Thm.axiom thy' (Sign.full_bname thy' name);
+    val thy' = Theory.add_defs_i unchecked overloaded [(b, prop')] thy;
+    val axm' = Thm.axiom thy' (Sign.full_name thy' b);
     val thm = unvarify (Thm.instantiate (recover_sorts, []) axm');
   in (thm, thy') end;
 
--- a/src/Pure/morphism.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/morphism.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -16,21 +16,18 @@
 signature MORPHISM =
 sig
   include BASIC_MORPHISM
-  val var: morphism -> Binding.T * mixfix -> Binding.T * mixfix
-  val binding: morphism -> Binding.T -> Binding.T
+  val binding: morphism -> binding -> binding
   val typ: morphism -> typ -> typ
   val term: morphism -> term -> term
   val fact: morphism -> thm list -> thm list
   val thm: morphism -> thm -> thm
   val cterm: morphism -> cterm -> cterm
   val morphism:
-   {binding: Binding.T -> Binding.T,
-    var: Binding.T * mixfix -> Binding.T * mixfix,
+   {binding: binding -> binding,
     typ: typ -> typ,
     term: term -> term,
     fact: thm list -> thm list} -> morphism
-  val binding_morphism: (Binding.T -> Binding.T) -> morphism
-  val var_morphism: (Binding.T * mixfix -> Binding.T * mixfix) -> morphism
+  val binding_morphism: (binding -> binding) -> morphism
   val typ_morphism: (typ -> typ) -> morphism
   val term_morphism: (term -> term) -> morphism
   val fact_morphism: (thm list -> thm list) -> morphism
@@ -45,8 +42,7 @@
 struct
 
 datatype morphism = Morphism of
- {binding: Binding.T -> Binding.T,
-  var: Binding.T * mixfix -> Binding.T * mixfix,
+ {binding: binding -> binding,
   typ: typ -> typ,
   term: term -> term,
   fact: thm list -> thm list};
@@ -54,7 +50,6 @@
 type declaration = morphism -> Context.generic -> Context.generic;
 
 fun binding (Morphism {binding, ...}) = binding;
-fun var (Morphism {var, ...}) = var;
 fun typ (Morphism {typ, ...}) = typ;
 fun term (Morphism {term, ...}) = term;
 fun fact (Morphism {fact, ...}) = fact;
@@ -63,20 +58,19 @@
 
 val morphism = Morphism;
 
-fun binding_morphism binding = morphism {binding = binding, var = I, typ = I, term = I, fact = I};
-fun var_morphism var = morphism {binding = I, var = var, typ = I, term = I, fact = I};
-fun typ_morphism typ = morphism {binding = I, var = I, typ = typ, term = I, fact = I};
-fun term_morphism term = morphism {binding = I, var = I, typ = I, term = term, fact = I};
-fun fact_morphism fact = morphism {binding = I, var = I, typ = I, term = I, fact = fact};
-fun thm_morphism thm = morphism {binding = I, var = I, typ = I, term = I, fact = map thm};
+fun binding_morphism binding = morphism {binding = binding, typ = I, term = I, fact = I};
+fun typ_morphism typ = morphism {binding = I, typ = typ, term = I, fact = I};
+fun term_morphism term = morphism {binding = I, typ = I, term = term, fact = I};
+fun fact_morphism fact = morphism {binding = I, typ = I, term = I, fact = fact};
+fun thm_morphism thm = morphism {binding = I, typ = I, term = I, fact = map thm};
 
-val identity = morphism {binding = I, var = I, typ = I, term = I, fact = I};
+val identity = morphism {binding = I, typ = I, term = I, fact = I};
 
 fun compose
-    (Morphism {binding = binding1, var = var1, typ = typ1, term = term1, fact = fact1})
-    (Morphism {binding = binding2, var = var2, typ = typ2, term = term2, fact = fact2}) =
-  morphism {binding = binding1 o binding2, var = var1 o var2,
-    typ = typ1 o typ2, term = term1 o term2, fact = fact1 o fact2};
+    (Morphism {binding = binding1, typ = typ1, term = term1, fact = fact1})
+    (Morphism {binding = binding2, typ = typ2, term = term2, fact = fact2}) =
+  morphism {binding = binding1 o binding2, typ = typ1 o typ2,
+    term = term1 o term2, fact = fact1 o fact2};
 
 fun phi1 $> phi2 = compose phi2 phi1;
 
--- a/src/Pure/net.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/net.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/net.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1993  University of Cambridge
 
--- a/src/Pure/old_goals.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/old_goals.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/old_goals.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1993  University of Cambridge
 
--- a/src/Pure/primitive_defs.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/primitive_defs.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/primitive_defs.ML
-    ID:         $Id$
     Author:     Makarius
 
 Primitive definition forms.
--- a/src/Pure/proofterm.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/proofterm.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -21,10 +21,10 @@
    | PAxm of string * term * typ list option
    | Oracle of string * term * typ list option
    | Promise of serial * term * typ list
-   | PThm of serial * ((string * term * typ list option) * proof_body lazy)
+   | PThm of serial * ((string * term * typ list option) * proof_body future)
   and proof_body = PBody of
     {oracles: (string * term) OrdList.T,
-     thms: (serial * (string * term * proof_body lazy)) OrdList.T,
+     thms: (serial * (string * term * proof_body future)) OrdList.T,
      proof: proof}
 
   val %> : proof * term -> proof
@@ -35,10 +35,10 @@
   include BASIC_PROOFTERM
 
   type oracle = string * term
-  type pthm = serial * (string * term * proof_body lazy)
-  val force_body: proof_body lazy ->
+  type pthm = serial * (string * term * proof_body future)
+  val join_body: proof_body future ->
     {oracles: oracle OrdList.T, thms: pthm OrdList.T, proof: proof}
-  val force_proof: proof_body lazy -> proof
+  val join_proof: proof_body future -> proof
   val proof_of: proof_body -> proof
   val fold_body_thms: (string * term * proof_body -> 'a -> 'a) -> proof_body list -> 'a -> 'a
   val fold_proof_atoms: bool -> (proof -> 'a -> 'a) -> proof list -> 'a -> 'a
@@ -110,7 +110,7 @@
   val promise_proof: theory -> serial -> term -> proof
   val fulfill_proof: theory -> (serial * proof) list -> proof_body -> proof_body
   val thm_proof: theory -> string -> term list -> term ->
-    (serial * proof) list lazy -> proof_body -> pthm * proof
+    (serial * proof future) list -> proof_body -> pthm * proof
   val get_name: term list -> term -> proof -> string
 
   (** rewriting on proof terms **)
@@ -142,17 +142,17 @@
  | PAxm of string * term * typ list option
  | Oracle of string * term * typ list option
  | Promise of serial * term * typ list
- | PThm of serial * ((string * term * typ list option) * proof_body lazy)
+ | PThm of serial * ((string * term * typ list option) * proof_body future)
 and proof_body = PBody of
   {oracles: (string * term) OrdList.T,
-   thms: (serial * (string * term * proof_body lazy)) OrdList.T,
+   thms: (serial * (string * term * proof_body future)) OrdList.T,
    proof: proof};
 
 type oracle = string * term;
-type pthm = serial * (string * term * proof_body lazy);
+type pthm = serial * (string * term * proof_body future);
 
-val force_body = Lazy.force #> (fn PBody args => args);
-val force_proof = #proof o force_body;
+val join_body = Future.join #> (fn PBody args => args);
+val join_proof = #proof o join_body;
 
 fun proof_of (PBody {proof, ...}) = proof;
 
@@ -165,7 +165,7 @@
       if Inttab.defined seen i then (x, seen)
       else
         let
-          val body' = Lazy.force body;
+          val body' = Future.join body;
           val (x', seen') = app body' (x, Inttab.update (i, ()) seen);
         in (f (name, prop, body') x', seen') end);
   in fn bodies => fn x => #1 (fold app bodies (x, Inttab.empty)) end;
@@ -180,7 +180,7 @@
           if Inttab.defined seen i then (x, seen)
           else
             let val (x', seen') =
-              (if all then app (force_proof body) else I) (x, Inttab.update (i, ()) seen)
+              (if all then app (join_proof body) else I) (x, Inttab.update (i, ()) seen)
             in (f prf x', seen') end)
       | app prf = (fn (x, seen) => (f prf x, seen));
   in fn prfs => fn x => #1 (fold app prfs (x, Inttab.empty)) end;
@@ -233,7 +233,7 @@
 
 fun strip_thm (body as PBody {proof, ...}) =
   (case strip_combt (fst (strip_combP proof)) of
-    (PThm (_, (_, body')), _) => Lazy.force body'
+    (PThm (_, (_, body')), _) => Future.join body'
   | _ => body);
 
 val mk_Abst = fold_rev (fn (s, T:typ) => fn prf => Abst (s, NONE, prf));
@@ -1227,6 +1227,11 @@
         val proof = rewrite_prf fst (rules, K fill :: procs) proof0;
       in PBody {oracles = oracles, thms = thms, proof = proof} end;
 
+fun fulfill_proof_future _ [] body = Future.value body
+  | fulfill_proof_future thy promises body =
+      Future.fork_deps (map snd promises) (fn () =>
+        fulfill_proof thy (map (apsnd Future.join) promises) body);
+
 
 (***** theorems *****)
 
@@ -1243,11 +1248,9 @@
       if ! proofs = 2 then
         #4 (shrink_proof thy [] 0 (rew_proof thy (fold_rev implies_intr_proof hyps prf)))
       else MinProof;
+    val body0 = PBody {oracles = oracles0, thms = thms0, proof = proof0};
 
-    fun new_prf () = (serial (), name, prop,
-      Lazy.lazy (fn () => fulfill_proof thy (Lazy.force promises)
-        (PBody {oracles = oracles0, thms = thms0, proof = proof0})));
-
+    fun new_prf () = (serial (), name, prop, fulfill_proof_future thy promises body0);
     val (i, name, prop, body') =
       (case strip_combt (fst (strip_combP prf)) of
         (PThm (i, ((old_name, prop', NONE), body')), args') =>
--- a/src/Pure/pure_setup.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/pure_setup.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -33,7 +33,7 @@
   map (fn (x, y) => Pretty.str (x ^ "=" ^ y)) o Position.properties_of));
 install_pp (make_pp ["Thm", "thm"] ProofDisplay.pprint_thm);
 install_pp (make_pp ["Thm", "cterm"] ProofDisplay.pprint_cterm);
-install_pp (make_pp ["Binding", "T"] (Pretty.pprint o Pretty.str o Binding.display));
+install_pp (make_pp ["Binding", "binding"] (Pretty.pprint o Pretty.str o Binding.display));
 install_pp (make_pp ["Thm", "ctyp"] ProofDisplay.pprint_ctyp);
 install_pp (make_pp ["Context", "theory"] Context.pprint_thy);
 install_pp (make_pp ["Context", "theory_ref"] Context.pprint_thy_ref);
--- a/src/Pure/pure_thy.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/pure_thy.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -24,27 +24,27 @@
   val name_thm: bool -> bool -> Position.T -> string -> thm -> thm
   val name_thms: bool -> bool -> Position.T -> string -> thm list -> thm list
   val name_thmss: bool -> Position.T -> string -> (thm list * 'a) list -> (thm list * 'a) list
-  val store_thms: bstring * thm list -> theory -> thm list * theory
-  val store_thm: bstring * thm -> theory -> thm * theory
-  val store_thm_open: bstring * thm -> theory -> thm * theory
-  val add_thms: ((bstring * thm) * attribute list) list -> theory -> thm list * theory
-  val add_thm: (bstring * thm) * attribute list -> theory -> thm * theory
-  val add_thmss: ((bstring * thm list) * attribute list) list -> theory -> thm list list * theory
-  val add_thms_dynamic: bstring * (Context.generic -> thm list) -> theory -> theory
-  val note_thmss: string -> ((Binding.T * attribute list) *
+  val store_thms: binding * thm list -> theory -> thm list * theory
+  val store_thm: binding * thm -> theory -> thm * theory
+  val store_thm_open: binding * thm -> theory -> thm * theory
+  val add_thms: ((binding * thm) * attribute list) list -> theory -> thm list * theory
+  val add_thm: (binding * thm) * attribute list -> theory -> thm * theory
+  val add_thmss: ((binding * thm list) * attribute list) list -> theory -> thm list list * theory
+  val add_thms_dynamic: binding * (Context.generic -> thm list) -> theory -> theory
+  val note_thmss: string -> ((binding * attribute list) *
     (thm list * attribute list) list) list -> theory -> (string * thm list) list * theory
-  val note_thmss_grouped: string -> string -> ((Binding.T * attribute list) *
+  val note_thmss_grouped: string -> string -> ((binding * attribute list) *
     (thm list * attribute list) list) list -> theory -> (string * thm list) list * theory
-  val add_axioms: ((bstring * term) * attribute list) list -> theory -> thm list * theory
+  val add_axioms: ((binding * term) * attribute list) list -> theory -> thm list * theory
   val add_axioms_cmd: ((bstring * string) * attribute list) list -> theory -> thm list * theory
-  val add_defs: bool -> ((bstring * term) * attribute list) list ->
+  val add_defs: bool -> ((binding * term) * attribute list) list ->
     theory -> thm list * theory
-  val add_defs_unchecked: bool -> ((bstring * term) * attribute list) list ->
+  val add_defs_unchecked: bool -> ((binding * term) * attribute list) list ->
+    theory -> thm list * theory
+  val add_defs_cmd: bool -> ((bstring * string) * attribute list) list ->
     theory -> thm list * theory
   val add_defs_unchecked_cmd: bool -> ((bstring * string) * attribute list) list ->
     theory -> thm list * theory
-  val add_defs_cmd: bool -> ((bstring * string) * attribute list) list ->
-    theory -> thm list * theory
   val old_appl_syntax: theory -> bool
   val old_appl_syntax_setup: theory -> theory
 end;
@@ -163,21 +163,21 @@
 
 (* store_thm(s) *)
 
-fun store_thms (bname, thms) = enter_thms (name_thms true true Position.none)
-  (name_thms false true Position.none) I (Binding.name bname, thms);
+fun store_thms (b, thms) = enter_thms (name_thms true true Position.none)
+  (name_thms false true Position.none) I (b, thms);
 
-fun store_thm (bname, th) = store_thms (bname, [th]) #>> the_single;
+fun store_thm (b, th) = store_thms (b, [th]) #>> the_single;
 
-fun store_thm_open (bname, th) =
+fun store_thm_open (b, th) =
   enter_thms (name_thms true false Position.none) (name_thms false false Position.none) I
-    (Binding.name bname, [th]) #>> the_single;
+    (b, [th]) #>> the_single;
 
 
 (* add_thms(s) *)
 
-fun add_thms_atts pre_name ((bname, thms), atts) =
+fun add_thms_atts pre_name ((b, thms), atts) =
   enter_thms pre_name (name_thms false true Position.none)
-    (foldl_map (Thm.theory_attributes atts)) (Binding.name bname, thms);
+    (foldl_map (Thm.theory_attributes atts)) (b, thms);
 
 fun gen_add_thmss pre_name =
   fold_map (add_thms_atts pre_name);
@@ -192,9 +192,9 @@
 
 (* add_thms_dynamic *)
 
-fun add_thms_dynamic (bname, f) thy = thy
+fun add_thms_dynamic (b, f) thy = thy
   |> (FactsData.map o apfst)
-      (Facts.add_dynamic (Sign.naming_of thy) (Binding.name bname, f) #> snd);
+      (Facts.add_dynamic (Sign.naming_of thy) (b, f) #> snd);
 
 
 (* note_thmss *)
@@ -224,21 +224,21 @@
 (* store axioms as theorems *)
 
 local
-  fun get_ax thy (name, _) = Thm.axiom thy (Sign.full_bname thy name);
+  fun get_ax thy (b, _) = Thm.axiom thy (Sign.full_name thy b);
   fun get_axs thy named_axs = map (Thm.forall_elim_vars 0 o get_ax thy) named_axs;
-  fun add_axm add = fold_map (fn ((name, ax), atts) => fn thy =>
+  fun add_axm prep_b add = fold_map (fn ((b, ax), atts) => fn thy =>
     let
-      val named_ax = [(name, ax)];
+      val named_ax = [(b, ax)];
       val thy' = add named_ax thy;
-      val thm = hd (get_axs thy' named_ax);
-    in apfst hd (gen_add_thms (K I) [((name, thm), atts)] thy') end);
+      val thm = hd (get_axs thy' ((map o apfst) prep_b named_ax));
+    in apfst hd (gen_add_thms (K I) [((prep_b b, thm), atts)] thy') end);
 in
-  val add_defs               = add_axm o Theory.add_defs_i false;
-  val add_defs_unchecked     = add_axm o Theory.add_defs_i true;
-  val add_axioms             = add_axm Theory.add_axioms_i;
-  val add_defs_cmd           = add_axm o Theory.add_defs false;
-  val add_defs_unchecked_cmd = add_axm o Theory.add_defs true;
-  val add_axioms_cmd         = add_axm Theory.add_axioms;
+  val add_defs               = add_axm I o Theory.add_defs_i false;
+  val add_defs_unchecked     = add_axm I o Theory.add_defs_i true;
+  val add_axioms             = add_axm I Theory.add_axioms_i;
+  val add_defs_cmd           = add_axm Binding.name o Theory.add_defs false;
+  val add_defs_unchecked_cmd = add_axm Binding.name o Theory.add_defs true;
+  val add_axioms_cmd         = add_axm Binding.name Theory.add_axioms;
 end;
 
 
@@ -378,16 +378,16 @@
     ("sort_constraint", typ "'a itself => prop", NoSyn),
     ("conjunction", typ "prop => prop => prop", NoSyn)]
   #> (add_defs false o map Thm.no_attributes)
-   [("prop_def", prop "(CONST prop :: prop => prop) (A::prop) == A::prop"),
-    ("term_def", prop "(CONST Pure.term :: 'a => prop) (x::'a) == (!!A::prop. A ==> A)"),
-    ("sort_constraint_def",
+   [(Binding.name "prop_def", prop "(CONST prop :: prop => prop) (A::prop) == A::prop"),
+    (Binding.name "term_def", prop "(CONST Pure.term :: 'a => prop) (x::'a) == (!!A::prop. A ==> A)"),
+    (Binding.name "sort_constraint_def",
       prop "(CONST Pure.sort_constraint :: 'a itself => prop) (CONST TYPE :: 'a itself) ==\
       \ (CONST Pure.term :: 'a itself => prop) (CONST TYPE :: 'a itself)"),
-    ("conjunction_def", prop "(A &&& B) == (!!C::prop. (A ==> B ==> C) ==> C)")] #> snd
+    (Binding.name "conjunction_def", prop "(A &&& B) == (!!C::prop. (A ==> B ==> C) ==> C)")] #> snd
   #> Sign.hide_const false "Pure.term"
   #> Sign.hide_const false "Pure.sort_constraint"
   #> Sign.hide_const false "Pure.conjunction"
-  #> add_thmss [(("nothing", []), [])] #> snd
-  #> Theory.add_axioms_i Proofterm.equality_axms));
+  #> add_thmss [((Binding.name "nothing", []), [])] #> snd
+  #> Theory.add_axioms_i ((map o apfst) Binding.name Proofterm.equality_axms)));
 
 end;
--- a/src/Pure/sign.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/sign.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/sign.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson and Markus Wenzel
 
 Logical signature content: naming conventions, concrete syntax, type
@@ -14,7 +13,7 @@
     tsig: Type.tsig,
     consts: Consts.T}
   val naming_of: theory -> NameSpace.naming
-  val full_name: theory -> Binding.T -> string
+  val full_name: theory -> binding -> string
   val base_name: string -> bstring
   val full_bname: theory -> bstring -> string
   val full_bname_path: theory -> string -> bstring -> string
@@ -91,10 +90,10 @@
   val del_modesyntax: Syntax.mode -> (bstring * string * mixfix) list -> theory -> theory
   val del_modesyntax_i: Syntax.mode -> (bstring * typ * mixfix) list -> theory -> theory
   val notation: bool -> Syntax.mode -> (term * mixfix) list -> theory -> theory
-  val declare_const: Properties.T -> (Binding.T * typ) * mixfix -> theory -> term * theory
+  val declare_const: Properties.T -> (binding * typ) * mixfix -> theory -> term * theory
   val add_consts: (bstring * string * mixfix) list -> theory -> theory
   val add_consts_i: (bstring * typ * mixfix) list -> theory -> theory
-  val add_abbrev: string -> Properties.T -> Binding.T * term -> theory -> (term * term) * theory
+  val add_abbrev: string -> Properties.T -> binding * term -> theory -> (term * term) * theory
   val revert_abbrev: string -> string -> theory -> theory
   val add_const_constraint: string * typ option -> theory -> theory
   val primitive_class: string * class list -> theory -> theory
--- a/src/Pure/simplifier.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/simplifier.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/simplifier.ML
-    ID:         $Id$
     Author:     Tobias Nipkow and Markus Wenzel, TU Muenchen
 
 Generic simplifier, suitable for most logics (see also
--- a/src/Pure/subgoal.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/subgoal.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/subgoal.ML
-    ID:         $Id$
     Author:     Makarius
 
 Tactical operations depending on local subgoal structure.
--- a/src/Pure/theory.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/theory.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/theory.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson and Markus Wenzel
 
 Logical theory content: axioms, definitions, and begin/end wrappers.
@@ -29,14 +28,14 @@
   val at_end: (theory -> theory option) -> theory -> theory
   val begin_theory: string -> theory list -> theory
   val end_theory: theory -> theory
+  val add_axioms_i: (binding * term) list -> theory -> theory
   val add_axioms: (bstring * string) list -> theory -> theory
-  val add_axioms_i: (bstring * term) list -> theory -> theory
   val add_deps: string -> string * typ -> (string * typ) list -> theory -> theory
+  val add_defs_i: bool -> bool -> (binding * term) list -> theory -> theory
   val add_defs: bool -> bool -> (bstring * string) list -> theory -> theory
-  val add_defs_i: bool -> bool -> (bstring * term) list -> theory -> theory
+  val add_finals_i: bool -> term list -> theory -> theory
   val add_finals: bool -> string list -> theory -> theory
-  val add_finals_i: bool -> term list -> theory -> theory
-  val specify_const: Properties.T -> (Binding.T * typ) * mixfix -> theory -> term * theory
+  val specify_const: Properties.T -> (binding * typ) * mixfix -> theory -> term * theory
 end
 
 structure Theory: THEORY =
@@ -157,19 +156,19 @@
 fun err_in_axm msg name =
   cat_error msg ("The error(s) above occurred in axiom " ^ quote name);
 
-fun cert_axm thy (name, raw_tm) =
+fun cert_axm thy (b, raw_tm) =
   let
     val (t, T, _) = Sign.certify_prop thy raw_tm
       handle TYPE (msg, _, _) => error msg
         | TERM (msg, _) => error msg;
   in
     Term.no_dummy_patterns t handle TERM (msg, _) => error msg;
-    (name, Sign.no_vars (Syntax.pp_global thy) t)
+    (b, Sign.no_vars (Syntax.pp_global thy) t)
   end;
 
-fun read_axm thy (name, str) =
-  cert_axm thy (name, Syntax.read_prop_global thy str)
-    handle ERROR msg => err_in_axm msg name;
+fun read_axm thy (bname, str) =
+  cert_axm thy (Binding.name bname, Syntax.read_prop_global thy str)
+    handle ERROR msg => err_in_axm msg bname;
 
 
 (* add_axioms(_i) *)
@@ -178,15 +177,15 @@
 
 fun gen_add_axioms prep_axm raw_axms thy = thy |> map_axioms (fn axioms =>
   let
-    val axms = map (apfst Binding.name o apsnd Logic.varify o prep_axm thy) raw_axms;
+    val axms = map (apsnd Logic.varify o prep_axm thy) raw_axms;
     val axioms' = fold (snd oo NameSpace.bind (Sign.naming_of thy)) axms axioms
       handle Symtab.DUP dup => err_dup_axm dup;
   in axioms' end);
 
 in
 
+val add_axioms_i = gen_add_axioms cert_axm;
 val add_axioms = gen_add_axioms read_axm;
-val add_axioms_i = gen_add_axioms cert_axm;
 
 end;
 
@@ -250,16 +249,16 @@
 
 (* check_def *)
 
-fun check_def thy unchecked overloaded (bname, tm) defs =
+fun check_def thy unchecked overloaded (b, tm) defs =
   let
     val ctxt = ProofContext.init thy;
-    val name = Sign.full_bname thy bname;
+    val name = Sign.full_name thy b;
     val (lhs_const, rhs) = Sign.cert_def ctxt tm;
     val rhs_consts = fold_aterms (fn Const const => insert (op =) const | _ => I) rhs [];
     val _ = check_overloading thy overloaded lhs_const;
   in defs |> dependencies thy unchecked true name lhs_const rhs_consts end
   handle ERROR msg => cat_error msg (Pretty.string_of (Pretty.block
-   [Pretty.str ("The error(s) above occurred in definition " ^ quote bname ^ ":"),
+   [Pretty.str ("The error(s) above occurred in definition " ^ quote (Binding.display b) ^ ":"),
     Pretty.fbrk, Pretty.quote (Syntax.pretty_term_global thy tm)]));
 
 
@@ -298,8 +297,8 @@
 
 in
 
+val add_finals_i = gen_add_finals (K I);
 val add_finals = gen_add_finals Syntax.read_term_global;
-val add_finals_i = gen_add_finals (K I);
 
 end;
 
--- a/src/Pure/thm.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/thm.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1658,7 +1658,7 @@
     val {thy_ref, hyps, prop, tpairs, ...} = args;
     val _ = null tpairs orelse raise THM ("put_name: unsolved flex-flex constraints", 0, [thm]);
 
-    val ps = Lazy.lazy (fn () => map (apsnd (proof_of o Future.join)) promises);
+    val ps = map (apsnd (Future.map proof_of)) promises;
     val thy = Theory.deref thy_ref;
     val (pthm, proof) = Pt.thm_proof thy name hyps prop ps body;
 
--- a/src/Pure/type_infer.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/type_infer.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      Pure/type_infer.ML
-    ID:         $Id$
     Author:     Stefan Berghofer and Markus Wenzel, TU Muenchen
 
 Simple type inference.
--- a/src/Pure/variable.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Pure/variable.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -397,7 +397,7 @@
     val fact = export inner outer;
     val term = singleton (export_terms inner outer);
     val typ = Logic.type_map term;
-  in Morphism.morphism {binding = I, var = I, typ = typ, term = term, fact = fact} end;
+  in Morphism.morphism {binding = I, typ = typ, term = term, fact = fact} end;
 
 
 
--- a/src/Tools/induct.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/Tools/induct.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -50,7 +50,7 @@
   val setN: string
   (*proof methods*)
   val fix_tac: Proof.context -> int -> (string * typ) list -> int -> tactic
-  val add_defs: (Binding.T option * term) option list -> Proof.context ->
+  val add_defs: (binding option * term) option list -> Proof.context ->
     (term option list * thm list) * Proof.context
   val atomize_term: theory -> term -> term
   val atomize_tac: int -> tactic
@@ -62,7 +62,7 @@
   val cases_tac: Proof.context -> term option list list -> thm option ->
     thm list -> int -> cases_tactic
   val get_inductT: Proof.context -> term option list list -> thm list list
-  val induct_tac: Proof.context -> (Binding.T option * term) option list list ->
+  val induct_tac: Proof.context -> (binding option * term) option list list ->
     (string * typ) list list -> term option list -> thm list option ->
     thm list -> int -> cases_tactic
   val coinduct_tac: Proof.context -> term option list -> term option list -> thm option ->
--- a/src/ZF/Inductive_ZF.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/ZF/Inductive_ZF.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,4 @@
 (*  Title:      ZF/Inductive_ZF.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1993  University of Cambridge
 
--- a/src/ZF/Main_ZF.thy	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/ZF/Main_ZF.thy	Thu Jan 29 12:24:00 2009 +0000
@@ -1,5 +1,3 @@
-(*$Id$*)
-
 header{*Theory Main: Everything Except AC*}
 
 theory Main_ZF imports List_ZF IntDiv_ZF CardinalArith begin
--- a/src/ZF/Tools/datatype_package.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/ZF/Tools/datatype_package.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -247,7 +247,7 @@
       if need_recursor then
            thy |> Sign.add_consts_i
                     [(recursor_base_name, recursor_typ, NoSyn)]
-               |> (snd o PureThy.add_defs false [Thm.no_attributes recursor_def])
+               |> (snd o PureThy.add_defs false [(Thm.no_attributes o apfst Binding.name) recursor_def])
       else thy;
 
   val (con_defs, thy0) = thy_path
@@ -255,7 +255,7 @@
                  ((case_base_name, case_typ, NoSyn) ::
                   map #1 (List.concat con_ty_lists))
              |> PureThy.add_defs false
-                 (map Thm.no_attributes
+                 (map (Thm.no_attributes o apfst Binding.name)
                   (case_def ::
                    List.concat (ListPair.map mk_con_defs
                          (1 upto npart, con_ty_lists))))
@@ -383,13 +383,13 @@
   (*Updating theory components: simprules and datatype info*)
   (thy1 |> Sign.add_path big_rec_base_name
         |> PureThy.add_thmss
-         [(("simps", simps), [Simplifier.simp_add]),
-          (("", intrs), [Classical.safe_intro NONE]),
-          (("con_defs", con_defs), []),
-          (("case_eqns", case_eqns), []),
-          (("recursor_eqns", recursor_eqns), []),
-          (("free_iffs", free_iffs), []),
-          (("free_elims", free_SEs), [])] |> snd
+         [((Binding.name "simps", simps), [Simplifier.simp_add]),
+          ((Binding.empty , intrs), [Classical.safe_intro NONE]),
+          ((Binding.name "con_defs", con_defs), []),
+          ((Binding.name "case_eqns", case_eqns), []),
+          ((Binding.name "recursor_eqns", recursor_eqns), []),
+          ((Binding.name "free_iffs", free_iffs), []),
+          ((Binding.name "free_elims", free_SEs), [])] |> snd
         |> DatatypesData.map (Symtab.update (big_rec_name, dt_info))
         |> ConstructorsData.map (fold Symtab.update con_pairs)
         |> Sign.parent_path,
--- a/src/ZF/Tools/induct_tacs.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/ZF/Tools/induct_tacs.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -158,7 +158,7 @@
   in
     thy
     |> Sign.add_path (Sign.base_name big_rec_name)
-    |> PureThy.add_thmss [(("simps", simps), [Simplifier.simp_add])] |> snd
+    |> PureThy.add_thmss [((Binding.name "simps", simps), [Simplifier.simp_add])] |> snd
     |> DatatypesData.put (Symtab.update (big_rec_name, dt_info) (DatatypesData.get thy))
     |> ConstructorsData.put (fold_rev Symtab.update con_pairs (ConstructorsData.get thy))
     |> Sign.parent_path
--- a/src/ZF/Tools/inductive_package.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/ZF/Tools/inductive_package.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -27,10 +27,10 @@
   (*Insert definitions for the recursive sets, which
      must *already* be declared as constants in parent theory!*)
   val add_inductive_i: bool -> term list * term ->
-    ((Binding.T * term) * attribute list) list ->
+    ((binding * term) * attribute list) list ->
     thm list * thm list * thm list * thm list -> theory -> theory * inductive_result
   val add_inductive: string list * string ->
-    ((Binding.T * string) * Attrib.src list) list ->
+    ((binding * string) * Attrib.src list) list ->
     (Facts.ref * Attrib.src list) list * (Facts.ref * Attrib.src list) list *
     (Facts.ref * Attrib.src list) list * (Facts.ref * Attrib.src list) list ->
     theory -> theory * inductive_result
@@ -173,7 +173,7 @@
   val (_, thy1) =
     thy
     |> Sign.add_path big_rec_base_name
-    |> PureThy.add_defs false (map Thm.no_attributes axpairs);
+    |> PureThy.add_defs false (map (Thm.no_attributes o apfst Binding.name) axpairs);
 
   val ctxt1 = ProofContext.init thy1;
 
@@ -510,9 +510,9 @@
 
      val ([induct', mutual_induct'], thy') =
        thy
-       |> PureThy.add_thms [((co_prefix ^ "induct", induct),
+       |> PureThy.add_thms [((Binding.name (co_prefix ^ "induct"), induct),
              [case_names, Induct.induct_pred big_rec_name]),
-           (("mutual_induct", mutual_induct), [case_names])];
+           ((Binding.name "mutual_induct", mutual_induct), [case_names])];
     in ((thy', induct'), mutual_induct')
     end;  (*of induction_rules*)
 
@@ -522,7 +522,7 @@
     if not coind then induction_rules raw_induct thy1
     else
       (thy1
-      |> PureThy.add_thms [((co_prefix ^ "induct", raw_induct), [])]
+      |> PureThy.add_thms [((Binding.name (co_prefix ^ "induct"), raw_induct), [])]
       |> apfst hd |> Library.swap, TrueI)
   and defs = big_rec_def :: part_rec_defs
 
@@ -531,15 +531,15 @@
     thy2
     |> IndCases.declare big_rec_name make_cases
     |> PureThy.add_thms
-      [(("bnd_mono", bnd_mono), []),
-       (("dom_subset", dom_subset), []),
-       (("cases", elim), [case_names, Induct.cases_pred big_rec_name])]
+      [((Binding.name "bnd_mono", bnd_mono), []),
+       ((Binding.name "dom_subset", dom_subset), []),
+       ((Binding.name "cases", elim), [case_names, Induct.cases_pred big_rec_name])]
     ||>> (PureThy.add_thmss o map Thm.no_attributes)
-        [("defs", defs),
-         ("intros", intrs)];
+        [(Binding.name "defs", defs),
+         (Binding.name "intros", intrs)];
   val (intrs'', thy4) =
     thy3
-    |> PureThy.add_thms ((intr_names ~~ intrs') ~~ map #2 intr_specs)
+    |> PureThy.add_thms ((map Binding.name intr_names ~~ intrs') ~~ map #2 intr_specs)
     ||> Sign.parent_path;
   in
     (thy4,
--- a/src/ZF/Tools/primrec_package.ML	Thu Jan 29 12:05:19 2009 +0000
+++ b/src/ZF/Tools/primrec_package.ML	Thu Jan 29 12:24:00 2009 +0000
@@ -8,8 +8,8 @@
 
 signature PRIMREC_PACKAGE =
 sig
-  val add_primrec: ((Binding.T * string) * Attrib.src list) list -> theory -> theory * thm list
-  val add_primrec_i: ((Binding.T * term) * attribute list) list -> theory -> theory * thm list
+  val add_primrec: ((binding * string) * Attrib.src list) list -> theory -> theory * thm list
+  val add_primrec_i: ((binding * term) * attribute list) list -> theory -> theory * thm list
 end;
 
 structure PrimrecPackage : PRIMREC_PACKAGE =
@@ -169,7 +169,7 @@
 
     val ([def_thm], thy1) = thy
       |> Sign.add_path (Sign.base_name fname)
-      |> (PureThy.add_defs false o map Thm.no_attributes) [def];
+      |> PureThy.add_defs false [Thm.no_attributes (apfst Binding.name def)];
 
     val rewrites = def_thm :: map mk_meta_eq (#rec_rewrites con_info)
     val eqn_thms =
@@ -179,10 +179,10 @@
 
     val (eqn_thms', thy2) =
       thy1
-      |> PureThy.add_thms ((map Binding.base_name eqn_names ~~ eqn_thms) ~~ eqn_atts);
+      |> PureThy.add_thms ((eqn_names ~~ eqn_thms) ~~ eqn_atts);
     val (_, thy3) =
       thy2
-      |> PureThy.add_thmss [(("simps", eqn_thms'), [Simplifier.simp_add])]
+      |> PureThy.add_thmss [((Binding.name "simps", eqn_thms'), [Simplifier.simp_add])]
       ||> Sign.parent_path;
   in (thy3, eqn_thms') end;