Merged.
authorballarin
Thu, 01 Oct 2009 20:52:18 +0200
changeset 32847 88b58880d52c
parent 32846 29941e925c82 (current diff)
parent 32837 3a2fa964ad08 (diff)
child 32848 484863ae9b98
Merged.
src/Pure/Concurrent/par_list_dummy.ML
src/Pure/Concurrent/synchronized_dummy.ML
src/Pure/General/lazy.ML
src/Pure/Isar/expression.ML
--- a/doc-src/Codegen/Thy/Setup.thy	Thu Oct 01 20:49:46 2009 +0200
+++ b/doc-src/Codegen/Thy/Setup.thy	Thu Oct 01 20:52:18 2009 +0200
@@ -10,6 +10,6 @@
    "~~/src/HOL/Decision_Procs/Ferrack"] *}
 
 ML_command {* Code_Target.code_width := 74 *}
-ML_command {* reset unique_names *}
+ML_command {* Unsynchronized.reset unique_names *}
 
 end
--- a/doc-src/IsarImplementation/Thy/Integration.thy	Thu Oct 01 20:49:46 2009 +0200
+++ b/doc-src/IsarImplementation/Thy/Integration.thy	Thu Oct 01 20:52:18 2009 +0200
@@ -59,9 +59,9 @@
   @{index_ML Toplevel.is_toplevel: "Toplevel.state -> bool"} \\
   @{index_ML Toplevel.theory_of: "Toplevel.state -> theory"} \\
   @{index_ML Toplevel.proof_of: "Toplevel.state -> Proof.state"} \\
-  @{index_ML Toplevel.debug: "bool ref"} \\
-  @{index_ML Toplevel.timing: "bool ref"} \\
-  @{index_ML Toplevel.profiling: "int ref"} \\
+  @{index_ML Toplevel.debug: "bool Unsynchronized.ref"} \\
+  @{index_ML Toplevel.timing: "bool Unsynchronized.ref"} \\
+  @{index_ML Toplevel.profiling: "int Unsynchronized.ref"} \\
   \end{mldecls}
 
   \begin{description}
@@ -85,11 +85,11 @@
   \item @{ML Toplevel.proof_of}~@{text "state"} selects the Isar proof
   state if available, otherwise raises @{ML Toplevel.UNDEF}.
 
-  \item @{ML "set Toplevel.debug"} makes the toplevel print further
+  \item @{ML "Toplevel.debug := true"} makes the toplevel print further
   details about internal error conditions, exceptions being raised
   etc.
 
-  \item @{ML "set Toplevel.timing"} makes the toplevel print timing
+  \item @{ML "Toplevel.timing := true"} makes the toplevel print timing
   information for each Isar command being executed.
 
   \item @{ML Toplevel.profiling}~@{verbatim ":="}~@{text "n"} controls
--- a/doc-src/IsarImplementation/Thy/Logic.thy	Thu Oct 01 20:49:46 2009 +0200
+++ b/doc-src/IsarImplementation/Thy/Logic.thy	Thu Oct 01 20:52:18 2009 +0200
@@ -547,7 +547,7 @@
   \end{mldecls}
   \begin{mldecls}
   @{index_ML_type thm} \\
-  @{index_ML proofs: "int ref"} \\
+  @{index_ML proofs: "int Unsynchronized.ref"} \\
   @{index_ML Thm.assume: "cterm -> thm"} \\
   @{index_ML Thm.forall_intr: "cterm -> thm -> thm"} \\
   @{index_ML Thm.forall_elim: "cterm -> thm -> thm"} \\
--- a/doc-src/IsarImplementation/Thy/ML.thy	Thu Oct 01 20:49:46 2009 +0200
+++ b/doc-src/IsarImplementation/Thy/ML.thy	Thu Oct 01 20:52:18 2009 +0200
@@ -258,7 +258,7 @@
   \begin{mldecls}
   @{index_ML NAMED_CRITICAL: "string -> (unit -> 'a) -> 'a"} \\
   @{index_ML CRITICAL: "(unit -> 'a) -> 'a"} \\
-  @{index_ML setmp: "'a ref -> 'a -> ('b -> 'c) -> 'b -> 'c"} \\
+  @{index_ML setmp: "'a Unsynchronized.ref -> 'a -> ('b -> 'c) -> 'b -> 'c"} \\
   \end{mldecls}
 
   \begin{description}
--- a/doc-src/IsarImplementation/Thy/document/Integration.tex	Thu Oct 01 20:49:46 2009 +0200
+++ b/doc-src/IsarImplementation/Thy/document/Integration.tex	Thu Oct 01 20:52:18 2009 +0200
@@ -86,9 +86,9 @@
   \indexdef{}{ML}{Toplevel.is\_toplevel}\verb|Toplevel.is_toplevel: Toplevel.state -> bool| \\
   \indexdef{}{ML}{Toplevel.theory\_of}\verb|Toplevel.theory_of: Toplevel.state -> theory| \\
   \indexdef{}{ML}{Toplevel.proof\_of}\verb|Toplevel.proof_of: Toplevel.state -> Proof.state| \\
-  \indexdef{}{ML}{Toplevel.debug}\verb|Toplevel.debug: bool ref| \\
-  \indexdef{}{ML}{Toplevel.timing}\verb|Toplevel.timing: bool ref| \\
-  \indexdef{}{ML}{Toplevel.profiling}\verb|Toplevel.profiling: int ref| \\
+  \indexdef{}{ML}{Toplevel.debug}\verb|Toplevel.debug: bool Unsynchronized.ref| \\
+  \indexdef{}{ML}{Toplevel.timing}\verb|Toplevel.timing: bool Unsynchronized.ref| \\
+  \indexdef{}{ML}{Toplevel.profiling}\verb|Toplevel.profiling: int Unsynchronized.ref| \\
   \end{mldecls}
 
   \begin{description}
@@ -112,11 +112,11 @@
   \item \verb|Toplevel.proof_of|~\isa{state} selects the Isar proof
   state if available, otherwise raises \verb|Toplevel.UNDEF|.
 
-  \item \verb|set Toplevel.debug| makes the toplevel print further
+  \item \verb|Toplevel.debug := true| makes the toplevel print further
   details about internal error conditions, exceptions being raised
   etc.
 
-  \item \verb|set Toplevel.timing| makes the toplevel print timing
+  \item \verb|Toplevel.timing := true| makes the toplevel print timing
   information for each Isar command being executed.
 
   \item \verb|Toplevel.profiling|~\verb|:=|~\isa{n} controls
--- a/doc-src/IsarImplementation/Thy/document/Logic.tex	Thu Oct 01 20:49:46 2009 +0200
+++ b/doc-src/IsarImplementation/Thy/document/Logic.tex	Thu Oct 01 20:52:18 2009 +0200
@@ -546,7 +546,7 @@
   \end{mldecls}
   \begin{mldecls}
   \indexdef{}{ML type}{thm}\verb|type thm| \\
-  \indexdef{}{ML}{proofs}\verb|proofs: int ref| \\
+  \indexdef{}{ML}{proofs}\verb|proofs: int Unsynchronized.ref| \\
   \indexdef{}{ML}{Thm.assume}\verb|Thm.assume: cterm -> thm| \\
   \indexdef{}{ML}{Thm.forall\_intr}\verb|Thm.forall_intr: cterm -> thm -> thm| \\
   \indexdef{}{ML}{Thm.forall\_elim}\verb|Thm.forall_elim: cterm -> thm -> thm| \\
--- a/doc-src/IsarImplementation/Thy/document/ML.tex	Thu Oct 01 20:49:46 2009 +0200
+++ b/doc-src/IsarImplementation/Thy/document/ML.tex	Thu Oct 01 20:52:18 2009 +0200
@@ -277,7 +277,7 @@
 \begin{mldecls}
   \indexdef{}{ML}{NAMED\_CRITICAL}\verb|NAMED_CRITICAL: string -> (unit -> 'a) -> 'a| \\
   \indexdef{}{ML}{CRITICAL}\verb|CRITICAL: (unit -> 'a) -> 'a| \\
-  \indexdef{}{ML}{setmp}\verb|setmp: 'a ref -> 'a -> ('b -> 'c) -> 'b -> 'c| \\
+  \indexdef{}{ML}{setmp}\verb|setmp: 'a Unsynchronized.ref -> 'a -> ('b -> 'c) -> 'b -> 'c| \\
   \end{mldecls}
 
   \begin{description}
--- a/doc-src/IsarOverview/Isar/Logic.thy	Thu Oct 01 20:49:46 2009 +0200
+++ b/doc-src/IsarOverview/Isar/Logic.thy	Thu Oct 01 20:52:18 2009 +0200
@@ -526,7 +526,6 @@
 tendency to use the default introduction and elimination rules to
 decompose goals and facts. This can lead to very tedious proofs:
 *}
-(*<*)ML"set quick_and_dirty"(*>*)
 lemma "\<forall>x y. A x y \<and> B x y \<longrightarrow> C x y"
 proof
   fix x show "\<forall>y. A x y \<and> B x y \<longrightarrow> C x y"
--- a/doc-src/IsarOverview/Isar/ROOT.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/doc-src/IsarOverview/Isar/ROOT.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -1,2 +1,3 @@
-use_thy "Logic";
-use_thy "Induction"
+Unsynchronized.set quick_and_dirty;
+
+use_thys ["Logic", "Induction"];
--- a/doc-src/IsarOverview/Isar/document/Logic.tex	Thu Oct 01 20:49:46 2009 +0200
+++ b/doc-src/IsarOverview/Isar/document/Logic.tex	Thu Oct 01 20:52:18 2009 +0200
@@ -1228,19 +1228,6 @@
 decompose goals and facts. This can lead to very tedious proofs:%
 \end{isamarkuptext}%
 \isamarkuptrue%
-%
-\isadelimML
-%
-\endisadelimML
-%
-\isatagML
-%
-\endisatagML
-{\isafoldML}%
-%
-\isadelimML
-%
-\endisadelimML
 \isacommand{lemma}\isamarkupfalse%
 \ {\isachardoublequoteopen}{\isasymforall}x\ y{\isachardot}\ A\ x\ y\ {\isasymand}\ B\ x\ y\ {\isasymlongrightarrow}\ C\ x\ y{\isachardoublequoteclose}\isanewline
 %
--- a/doc-src/IsarRef/Thy/Inner_Syntax.thy	Thu Oct 01 20:49:46 2009 +0200
+++ b/doc-src/IsarRef/Thy/Inner_Syntax.thy	Thu Oct 01 20:52:18 2009 +0200
@@ -96,19 +96,19 @@
 
 text {*
   \begin{mldecls} 
-    @{index_ML show_types: "bool ref"} & default @{ML false} \\
-    @{index_ML show_sorts: "bool ref"} & default @{ML false} \\
-    @{index_ML show_consts: "bool ref"} & default @{ML false} \\
-    @{index_ML long_names: "bool ref"} & default @{ML false} \\
-    @{index_ML short_names: "bool ref"} & default @{ML false} \\
-    @{index_ML unique_names: "bool ref"} & default @{ML true} \\
-    @{index_ML show_brackets: "bool ref"} & default @{ML false} \\
-    @{index_ML eta_contract: "bool ref"} & default @{ML true} \\
-    @{index_ML goals_limit: "int ref"} & default @{ML 10} \\
-    @{index_ML Proof.show_main_goal: "bool ref"} & default @{ML false} \\
-    @{index_ML show_hyps: "bool ref"} & default @{ML false} \\
-    @{index_ML show_tags: "bool ref"} & default @{ML false} \\
-    @{index_ML show_question_marks: "bool ref"} & default @{ML true} \\
+    @{index_ML show_types: "bool Unsynchronized.ref"} & default @{ML false} \\
+    @{index_ML show_sorts: "bool Unsynchronized.ref"} & default @{ML false} \\
+    @{index_ML show_consts: "bool Unsynchronized.ref"} & default @{ML false} \\
+    @{index_ML long_names: "bool Unsynchronized.ref"} & default @{ML false} \\
+    @{index_ML short_names: "bool Unsynchronized.ref"} & default @{ML false} \\
+    @{index_ML unique_names: "bool Unsynchronized.ref"} & default @{ML true} \\
+    @{index_ML show_brackets: "bool Unsynchronized.ref"} & default @{ML false} \\
+    @{index_ML eta_contract: "bool Unsynchronized.ref"} & default @{ML true} \\
+    @{index_ML goals_limit: "int Unsynchronized.ref"} & default @{ML 10} \\
+    @{index_ML Proof.show_main_goal: "bool Unsynchronized.ref"} & default @{ML false} \\
+    @{index_ML show_hyps: "bool Unsynchronized.ref"} & default @{ML false} \\
+    @{index_ML show_tags: "bool Unsynchronized.ref"} & default @{ML false} \\
+    @{index_ML show_question_marks: "bool Unsynchronized.ref"} & default @{ML true} \\
   \end{mldecls}
 
   These global ML variables control the detail of information that is
--- a/doc-src/IsarRef/Thy/ROOT-HOLCF.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/doc-src/IsarRef/Thy/ROOT-HOLCF.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -1,4 +1,4 @@
-set ThyOutput.source;
+Unsynchronized.set ThyOutput.source;
 use "../../antiquote_setup.ML";
 
 use_thy "HOLCF_Specific";
--- a/doc-src/IsarRef/Thy/ROOT-ZF.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/doc-src/IsarRef/Thy/ROOT-ZF.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -1,4 +1,4 @@
-set ThyOutput.source;
+Unsynchronized.set ThyOutput.source;
 use "../../antiquote_setup.ML";
 
 use_thy "ZF_Specific";
--- a/doc-src/IsarRef/Thy/ROOT.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/doc-src/IsarRef/Thy/ROOT.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -1,5 +1,5 @@
-set quick_and_dirty;
-set ThyOutput.source;
+Unsynchronized.set quick_and_dirty;
+Unsynchronized.set ThyOutput.source;
 use "../../antiquote_setup.ML";
 
 use_thys [
--- a/doc-src/IsarRef/Thy/document/Inner_Syntax.tex	Thu Oct 01 20:49:46 2009 +0200
+++ b/doc-src/IsarRef/Thy/document/Inner_Syntax.tex	Thu Oct 01 20:52:18 2009 +0200
@@ -118,19 +118,19 @@
 %
 \begin{isamarkuptext}%
 \begin{mldecls} 
-    \indexdef{}{ML}{show\_types}\verb|show_types: bool ref| & default \verb|false| \\
-    \indexdef{}{ML}{show\_sorts}\verb|show_sorts: bool ref| & default \verb|false| \\
-    \indexdef{}{ML}{show\_consts}\verb|show_consts: bool ref| & default \verb|false| \\
-    \indexdef{}{ML}{long\_names}\verb|long_names: bool ref| & default \verb|false| \\
-    \indexdef{}{ML}{short\_names}\verb|short_names: bool ref| & default \verb|false| \\
-    \indexdef{}{ML}{unique\_names}\verb|unique_names: bool ref| & default \verb|true| \\
-    \indexdef{}{ML}{show\_brackets}\verb|show_brackets: bool ref| & default \verb|false| \\
-    \indexdef{}{ML}{eta\_contract}\verb|eta_contract: bool ref| & default \verb|true| \\
-    \indexdef{}{ML}{goals\_limit}\verb|goals_limit: int ref| & default \verb|10| \\
-    \indexdef{}{ML}{Proof.show\_main\_goal}\verb|Proof.show_main_goal: bool ref| & default \verb|false| \\
-    \indexdef{}{ML}{show\_hyps}\verb|show_hyps: bool ref| & default \verb|false| \\
-    \indexdef{}{ML}{show\_tags}\verb|show_tags: bool ref| & default \verb|false| \\
-    \indexdef{}{ML}{show\_question\_marks}\verb|show_question_marks: bool ref| & default \verb|true| \\
+    \indexdef{}{ML}{show\_types}\verb|show_types: bool Unsynchronized.ref| & default \verb|false| \\
+    \indexdef{}{ML}{show\_sorts}\verb|show_sorts: bool Unsynchronized.ref| & default \verb|false| \\
+    \indexdef{}{ML}{show\_consts}\verb|show_consts: bool Unsynchronized.ref| & default \verb|false| \\
+    \indexdef{}{ML}{long\_names}\verb|long_names: bool Unsynchronized.ref| & default \verb|false| \\
+    \indexdef{}{ML}{short\_names}\verb|short_names: bool Unsynchronized.ref| & default \verb|false| \\
+    \indexdef{}{ML}{unique\_names}\verb|unique_names: bool Unsynchronized.ref| & default \verb|true| \\
+    \indexdef{}{ML}{show\_brackets}\verb|show_brackets: bool Unsynchronized.ref| & default \verb|false| \\
+    \indexdef{}{ML}{eta\_contract}\verb|eta_contract: bool Unsynchronized.ref| & default \verb|true| \\
+    \indexdef{}{ML}{goals\_limit}\verb|goals_limit: int Unsynchronized.ref| & default \verb|10| \\
+    \indexdef{}{ML}{Proof.show\_main\_goal}\verb|Proof.show_main_goal: bool Unsynchronized.ref| & default \verb|false| \\
+    \indexdef{}{ML}{show\_hyps}\verb|show_hyps: bool Unsynchronized.ref| & default \verb|false| \\
+    \indexdef{}{ML}{show\_tags}\verb|show_tags: bool Unsynchronized.ref| & default \verb|false| \\
+    \indexdef{}{ML}{show\_question\_marks}\verb|show_question_marks: bool Unsynchronized.ref| & default \verb|true| \\
   \end{mldecls}
 
   These global ML variables control the detail of information that is
--- a/doc-src/LaTeXsugar/IsaMakefile	Thu Oct 01 20:49:46 2009 +0200
+++ b/doc-src/LaTeXsugar/IsaMakefile	Thu Oct 01 20:52:18 2009 +0200
@@ -14,7 +14,7 @@
 OUT = $(ISABELLE_OUTPUT)
 LOG = $(OUT)/log
 
-USEDIR = $(ISABELLE_TOOL) usedir -v true -i false -g false -d false -D document
+USEDIR = $(ISABELLE_TOOL) usedir -v true -i false -g false -d false -D document -M 1
 
 
 ## Sugar
--- a/doc-src/LaTeXsugar/Sugar/Sugar.thy	Thu Oct 01 20:49:46 2009 +0200
+++ b/doc-src/LaTeXsugar/Sugar/Sugar.thy	Thu Oct 01 20:52:18 2009 +0200
@@ -143,7 +143,7 @@
 internal index. This can be avoided by turning the last digit into a
 subscript: write \verb!x\<^isub>1! and obtain the much nicer @{text"x\<^isub>1"}. *}
 
-(*<*)ML"reset show_question_marks"(*>*)
+(*<*)ML"Unsynchronized.reset show_question_marks"(*>*)
 
 subsection {*Qualified names*}
 
--- a/doc-src/Locales/Locales/document/Examples3.tex	Thu Oct 01 20:49:46 2009 +0200
+++ b/doc-src/Locales/Locales/document/Examples3.tex	Thu Oct 01 20:52:18 2009 +0200
@@ -42,7 +42,7 @@
 \isatagvisible
 \isacommand{interpretation}\isamarkupfalse%
 \ nat{\isacharcolon}\ partial{\isacharunderscore}order\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
-\ \ \isakeyword{where}\ nat{\isacharunderscore}less{\isacharunderscore}eq{\isacharcolon}\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{where}\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
 \isacommand{proof}\isamarkupfalse%
 \ {\isacharminus}\isanewline
 \ \ \isacommand{show}\isamarkupfalse%
@@ -101,13 +101,12 @@
 \isatagvisible
 \isacommand{interpretation}\isamarkupfalse%
 \ nat{\isacharcolon}\ lattice\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
-\ \ \isakeyword{where}\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isakeyword{and}\ nat{\isacharunderscore}meet{\isacharunderscore}eq{\isacharcolon}\ {\isachardoublequoteopen}lattice{\isachardot}meet\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ min\ x\ y{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isakeyword{and}\ nat{\isacharunderscore}join{\isacharunderscore}eq{\isacharcolon}\ {\isachardoublequoteopen}lattice{\isachardot}join\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ max\ x\ y{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{where}\ {\isachardoublequoteopen}lattice{\isachardot}meet\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ min\ x\ y{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isakeyword{and}\ {\isachardoublequoteopen}lattice{\isachardot}join\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ max\ x\ y{\isachardoublequoteclose}\isanewline
 \isacommand{proof}\isamarkupfalse%
 \ {\isacharminus}\isanewline
 \ \ \isacommand{show}\isamarkupfalse%
-\ lattice{\isacharcolon}\ {\isachardoublequoteopen}lattice\ {\isacharparenleft}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}{\isachardoublequoteclose}%
+\ {\isachardoublequoteopen}lattice\ {\isacharparenleft}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}{\isachardoublequoteclose}%
 \begin{isamarkuptxt}%
 We have already shown that this is a partial order,%
 \end{isamarkuptxt}%
@@ -134,21 +133,12 @@
 \ \ \ \ \isacommand{by}\isamarkupfalse%
 \ arith{\isacharplus}%
 \begin{isamarkuptxt}%
-For the first of the equations, we refer to the theorem
-  shown in the previous interpretation.%
-\end{isamarkuptxt}%
-\isamarkuptrue%
-\ \ \isacommand{show}\isamarkupfalse%
-\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isacommand{by}\isamarkupfalse%
-\ {\isacharparenleft}rule\ nat{\isacharunderscore}less{\isacharunderscore}eq{\isacharparenright}%
-\begin{isamarkuptxt}%
-In order to show the remaining equations, we put ourselves in a
+In order to show the equations, we put ourselves in a
     situation where the lattice theorems can be used in a convenient way.%
 \end{isamarkuptxt}%
 \isamarkuptrue%
-\ \ \isacommand{from}\isamarkupfalse%
-\ lattice\ \isacommand{interpret}\isamarkupfalse%
+\ \ \isacommand{then}\isamarkupfalse%
+\ \isacommand{interpret}\isamarkupfalse%
 \ nat{\isacharcolon}\ lattice\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{show}\isamarkupfalse%
@@ -180,45 +170,8 @@
 \isatagvisible
 \isacommand{interpretation}\isamarkupfalse%
 \ nat{\isacharcolon}\ total{\isacharunderscore}order\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
-\ \ \isakeyword{where}\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isakeyword{and}\ {\isachardoublequoteopen}lattice{\isachardot}meet\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ min\ x\ y{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isakeyword{and}\ {\isachardoublequoteopen}lattice{\isachardot}join\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ max\ x\ y{\isachardoublequoteclose}\isanewline
-\isacommand{proof}\isamarkupfalse%
-\ {\isacharminus}\isanewline
-\ \ \isacommand{show}\isamarkupfalse%
-\ {\isachardoublequoteopen}total{\isacharunderscore}order\ {\isacharparenleft}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isacommand{by}\isamarkupfalse%
-\ unfold{\isacharunderscore}locales\ arith\isanewline
-\isacommand{qed}\isamarkupfalse%
-\ {\isacharparenleft}rule\ nat{\isacharunderscore}less{\isacharunderscore}eq\ nat{\isacharunderscore}meet{\isacharunderscore}eq\ nat{\isacharunderscore}join{\isacharunderscore}eq{\isacharparenright}{\isacharplus}%
-\endisatagvisible
-{\isafoldvisible}%
-%
-\isadelimvisible
-%
-\endisadelimvisible
-%
-\begin{isamarkuptext}%
-Since the locale hierarchy reflects that total
-  orders are distributive lattices, an explicit interpretation of
-  distributive lattices for the order relation on natural numbers is
-  only necessary for mapping the definitions to the right operators on
-  \isa{nat}.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimvisible
-%
-\endisadelimvisible
-%
-\isatagvisible
-\isacommand{interpretation}\isamarkupfalse%
-\ nat{\isacharcolon}\ distrib{\isacharunderscore}lattice\ {\isachardoublequoteopen}op\ {\isasymle}\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
-\ \ \isakeyword{where}\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isakeyword{and}\ {\isachardoublequoteopen}lattice{\isachardot}meet\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ min\ x\ y{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isakeyword{and}\ {\isachardoublequoteopen}lattice{\isachardot}join\ op\ {\isasymle}\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ max\ x\ y{\isachardoublequoteclose}\isanewline
 \ \ \isacommand{by}\isamarkupfalse%
-\ unfold{\isacharunderscore}locales\ {\isacharbrackleft}{\isadigit{1}}{\isacharbrackright}\ {\isacharparenleft}rule\ nat{\isacharunderscore}less{\isacharunderscore}eq\ nat{\isacharunderscore}meet{\isacharunderscore}eq\ nat{\isacharunderscore}join{\isacharunderscore}eq{\isacharparenright}{\isacharplus}%
+\ unfold{\isacharunderscore}locales\ arith%
 \endisatagvisible
 {\isafoldvisible}%
 %
@@ -248,7 +201,30 @@
 \hrule
 \caption{Interpreted theorems for \isa{{\isasymle}} on the natural numbers.}
 \label{tab:nat-lattice}
-\end{table}%
+\end{table}
+
+  Note that since the locale hierarchy reflects that total orders are
+  distributive lattices, an explicit interpretation of distributive
+  lattices for the order relation on natural numbers is not neccessary.
+
+  Why not push this idea further and just give the last interpretation
+  as a single interpretation instead of the sequence of three?  The
+  reasons for this are twofold:
+\begin{itemize}
+\item
+  Often it is easier to work in an incremental fashion, because later
+  interpretations require theorems provided by earlier
+  interpretations.
+\item
+  Assume that a definition is made in some locale $l_1$, and that $l_2$
+  imports $l_1$.  Let an equation for the definition be
+  proved in an interpretation of $l_2$.  The equation will be unfolded
+  in interpretations of theorems added to $l_2$ or below in the import
+  hierarchy, but not for theorems added above $l_2$.
+  Hence, an equation interpreting a definition should always be given in
+  an interpretation of the locale where the definition is made, not in
+  an interpretation of a locale further down the hierarchy.
+\end{itemize}%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -264,8 +240,7 @@
 \isamarkuptrue%
 \isacommand{interpretation}\isamarkupfalse%
 \ nat{\isacharunderscore}dvd{\isacharcolon}\ partial{\isacharunderscore}order\ {\isachardoublequoteopen}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
-\ \ \isakeyword{where}\ nat{\isacharunderscore}dvd{\isacharunderscore}less{\isacharunderscore}eq{\isacharcolon}\isanewline
-\ \ \ \ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ dvd\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ dvd\ y\ {\isasymand}\ x\ {\isasymnoteq}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{where}\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ dvd\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ dvd\ y\ {\isasymand}\ x\ {\isasymnoteq}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
 %
 \isadelimproof
 %
@@ -306,8 +281,7 @@
 \isamarkuptrue%
 \isacommand{interpretation}\isamarkupfalse%
 \ nat{\isacharunderscore}dvd{\isacharcolon}\ lattice\ {\isachardoublequoteopen}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
-\ \ \isakeyword{where}\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ dvd\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ dvd\ y\ {\isasymand}\ x\ {\isasymnoteq}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isakeyword{and}\ nat{\isacharunderscore}dvd{\isacharunderscore}meet{\isacharunderscore}eq{\isacharcolon}\ {\isachardoublequoteopen}lattice{\isachardot}meet\ {\isacharparenleft}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}\ {\isacharequal}\ gcd{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{where}\ nat{\isacharunderscore}dvd{\isacharunderscore}meet{\isacharunderscore}eq{\isacharcolon}\ {\isachardoublequoteopen}lattice{\isachardot}meet\ {\isacharparenleft}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}\ {\isacharequal}\ gcd{\isachardoublequoteclose}\isanewline
 \ \ \ \ \isakeyword{and}\ nat{\isacharunderscore}dvd{\isacharunderscore}join{\isacharunderscore}eq{\isacharcolon}\ {\isachardoublequoteopen}lattice{\isachardot}join\ {\isacharparenleft}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}\ {\isacharequal}\ lcm{\isachardoublequoteclose}\isanewline
 %
 \isadelimproof
@@ -338,10 +312,6 @@
 \ nat{\isacharunderscore}dvd{\isacharcolon}\ lattice\ {\isachardoublequoteopen}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse%
 \isanewline
 \ \ \isacommand{show}\isamarkupfalse%
-\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ dvd\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ dvd\ y\ {\isasymand}\ x\ {\isasymnoteq}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isacommand{by}\isamarkupfalse%
-\ {\isacharparenleft}rule\ nat{\isacharunderscore}dvd{\isacharunderscore}less{\isacharunderscore}eq{\isacharparenright}\isanewline
-\ \ \isacommand{show}\isamarkupfalse%
 \ {\isachardoublequoteopen}lattice{\isachardot}meet\ {\isacharparenleft}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}\ {\isacharequal}\ gcd{\isachardoublequoteclose}\isanewline
 \ \ \ \ \isacommand{apply}\isamarkupfalse%
 \ {\isacharparenleft}auto\ simp\ add{\isacharcolon}\ expand{\isacharunderscore}fun{\isacharunderscore}eq{\isacharparenright}\isanewline
@@ -408,14 +378,7 @@
 \isacommand{interpretation}\isamarkupfalse%
 \ nat{\isacharunderscore}dvd{\isacharcolon}\isanewline
 \ \ distrib{\isacharunderscore}lattice\ {\isachardoublequoteopen}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
-\ \ \isakeyword{where}\ {\isachardoublequoteopen}partial{\isacharunderscore}order{\isachardot}less\ op\ dvd\ {\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ y\ {\isacharequal}\ {\isacharparenleft}x\ dvd\ y\ {\isasymand}\ x\ {\isasymnoteq}\ y{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isakeyword{and}\ {\isachardoublequoteopen}lattice{\isachardot}meet\ {\isacharparenleft}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}\ {\isacharequal}\ gcd{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isakeyword{and}\ {\isachardoublequoteopen}lattice{\isachardot}join\ {\isacharparenleft}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}\ {\isacharequal}\ lcm{\isachardoublequoteclose}\isanewline
-\isacommand{proof}\isamarkupfalse%
-\ {\isacharminus}\isanewline
-\ \ \isacommand{show}\isamarkupfalse%
-\ {\isachardoublequoteopen}distrib{\isacharunderscore}lattice\ {\isacharparenleft}op\ dvd\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ \ \isacommand{apply}\isamarkupfalse%
 \ unfold{\isacharunderscore}locales%
 \begin{isamarkuptxt}%
 \begin{isabelle}%
@@ -426,7 +389,7 @@
 \end{isabelle}%
 \end{isamarkuptxt}%
 \isamarkuptrue%
-\ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ \ \isacommand{apply}\isamarkupfalse%
 \ {\isacharparenleft}unfold\ nat{\isacharunderscore}dvd{\isacharunderscore}meet{\isacharunderscore}eq\ nat{\isacharunderscore}dvd{\isacharunderscore}join{\isacharunderscore}eq{\isacharparenright}%
 \begin{isamarkuptxt}%
 \begin{isabelle}%
@@ -434,12 +397,9 @@
 \end{isabelle}%
 \end{isamarkuptxt}%
 \isamarkuptrue%
-\ \ \ \ \isacommand{apply}\isamarkupfalse%
-\ {\isacharparenleft}rule\ gcd{\isacharunderscore}lcm{\isacharunderscore}distr{\isacharparenright}\isanewline
-\ \ \ \ \isacommand{done}\isamarkupfalse%
-\isanewline
-\isacommand{qed}\isamarkupfalse%
-\ {\isacharparenleft}rule\ nat{\isacharunderscore}dvd{\isacharunderscore}less{\isacharunderscore}eq\ nat{\isacharunderscore}dvd{\isacharunderscore}meet{\isacharunderscore}eq\ nat{\isacharunderscore}dvd{\isacharunderscore}join{\isacharunderscore}eq{\isacharparenright}{\isacharplus}%
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}rule\ gcd{\isacharunderscore}lcm{\isacharunderscore}distr{\isacharparenright}\ \isacommand{done}\isamarkupfalse%
+%
 \endisatagvisible
 {\isafoldvisible}%
 %
--- a/doc-src/Main/Docs/document/Main_Doc.tex	Thu Oct 01 20:49:46 2009 +0200
+++ b/doc-src/Main/Docs/document/Main_Doc.tex	Thu Oct 01 20:52:18 2009 +0200
@@ -519,7 +519,7 @@
 \isa{takeWhile} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ bool{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list}\\
 \isa{tl} & \isa{{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list}\\
 \isa{upt} & \isa{nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat\ list}\\
-\isa{upto} & \isa{{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ list}\\
+\isa{upto} & \isa{int\ {\isasymRightarrow}\ int\ {\isasymRightarrow}\ int\ list}\\
 \isa{zip} & \isa{{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}b\ list\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}b{\isacharparenright}\ list}\\
 \end{supertabular}
 
--- a/doc-src/System/Thy/ROOT.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/doc-src/System/Thy/ROOT.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -1,7 +1,4 @@
-set ThyOutput.source;
+Unsynchronized.set ThyOutput.source;
 use "../../antiquote_setup.ML";
 
-use_thy "Basics";
-use_thy "Interfaces";
-use_thy "Presentation";
-use_thy "Misc";
+use_thys ["Basics", "Interfaces", "Presentation", "Misc"];
--- a/doc-src/TutorialI/Inductive/document/Advanced.tex	Thu Oct 01 20:49:46 2009 +0200
+++ b/doc-src/TutorialI/Inductive/document/Advanced.tex	Thu Oct 01 20:52:18 2009 +0200
@@ -362,7 +362,7 @@
 subgoal may look uninviting, but fortunately 
 \isa{lists} distributes over intersection:
 \begin{isabelle}%
-lists\ {\isacharparenleft}A\ {\isasyminter}\ B{\isacharparenright}\ {\isacharequal}\ lists\ A\ {\isasyminter}\ lists\ B\rulename{lists{\isacharunderscore}Int{\isacharunderscore}eq}%
+listsp\ {\isacharparenleft}{\isacharparenleft}{\isasymlambda}x{\isachardot}\ x\ {\isasymin}\ A{\isacharparenright}\ {\isasyminter}\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ x\ {\isasymin}\ B{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ x\ {\isasymin}\ lists\ A{\isacharparenright}\ {\isasyminter}\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ x\ {\isasymin}\ lists\ B{\isacharparenright}\rulename{lists{\isacharunderscore}Int{\isacharunderscore}eq}%
 \end{isabelle}
 Thanks to this default simplification rule, the induction hypothesis 
 is quickly replaced by its two parts:
--- a/doc-src/TutorialI/IsaMakefile	Thu Oct 01 20:49:46 2009 +0200
+++ b/doc-src/TutorialI/IsaMakefile	Thu Oct 01 20:52:18 2009 +0200
@@ -17,7 +17,7 @@
 SRC = $(ISABELLE_HOME)/src
 OUT = $(ISABELLE_OUTPUT)
 LOG = $(OUT)/log
-OPTIONS = -m brackets -i true -d "" -D document
+OPTIONS = -m brackets -i true -d "" -D document -M 1
 USEDIR = @$(ISABELLE_TOOL) usedir $(OPTIONS) $(OUT)/HOL
 
 
--- a/doc-src/TutorialI/Misc/Itrev.thy	Thu Oct 01 20:49:46 2009 +0200
+++ b/doc-src/TutorialI/Misc/Itrev.thy	Thu Oct 01 20:52:18 2009 +0200
@@ -2,7 +2,7 @@
 theory Itrev
 imports Main
 begin
-ML"reset NameSpace.unique_names"
+ML"Unsynchronized.reset NameSpace.unique_names"
 (*>*)
 
 section{*Induction Heuristics*}
@@ -141,6 +141,6 @@
 \index{induction heuristics|)}
 *}
 (*<*)
-ML"set NameSpace.unique_names"
+ML"Unsynchronized.set NameSpace.unique_names"
 end
 (*>*)
--- a/doc-src/TutorialI/Rules/Basic.thy	Thu Oct 01 20:49:46 2009 +0200
+++ b/doc-src/TutorialI/Rules/Basic.thy	Thu Oct 01 20:52:18 2009 +0200
@@ -188,7 +188,7 @@
 
 text{*unification failure trace *}
 
-ML "set trace_unify_fail"
+ML "Unsynchronized.set trace_unify_fail"
 
 lemma "P(a, f(b, g(e,a), b), a) \<Longrightarrow> P(a, f(b, g(c,a), b), a)"
 txt{*
@@ -213,7 +213,7 @@
 *}
 oops
 
-ML "reset trace_unify_fail"
+ML "Unsynchronized.reset trace_unify_fail"
 
 
 text{*Quantifiers*}
--- a/doc-src/TutorialI/Sets/Examples.thy	Thu Oct 01 20:49:46 2009 +0200
+++ b/doc-src/TutorialI/Sets/Examples.thy	Thu Oct 01 20:52:18 2009 +0200
@@ -1,7 +1,7 @@
 (* ID:         $Id$ *)
 theory Examples imports Main Binomial begin
 
-ML "reset eta_contract"
+ML "Unsynchronized.reset eta_contract"
 ML "Pretty.setmargin 64"
 
 text{*membership, intersection *}
--- a/doc-src/TutorialI/Types/Numbers.thy	Thu Oct 01 20:49:46 2009 +0200
+++ b/doc-src/TutorialI/Types/Numbers.thy	Thu Oct 01 20:52:18 2009 +0200
@@ -252,18 +252,13 @@
 \rulename{mult_cancel_left}
 *}
 
-ML{*set show_sorts*}
-
 text{*
 effect of show sorts on the above
 
-@{thm[display] mult_cancel_left[no_vars]}
+@{thm[display,show_sorts] mult_cancel_left[no_vars]}
 \rulename{mult_cancel_left}
 *}
 
-ML{*reset show_sorts*}
-
-
 text{*
 absolute value
 
--- a/doc-src/TutorialI/Types/document/Numbers.tex	Thu Oct 01 20:49:46 2009 +0200
+++ b/doc-src/TutorialI/Types/document/Numbers.tex	Thu Oct 01 20:52:18 2009 +0200
@@ -550,44 +550,18 @@
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
-\isadelimML
-%
-\endisadelimML
-%
-\isatagML
-\isacommand{ML}\isamarkupfalse%
-{\isacharverbatimopen}set\ show{\isacharunderscore}sorts{\isacharverbatimclose}%
-\endisatagML
-{\isafoldML}%
-%
-\isadelimML
-%
-\endisadelimML
-%
 \begin{isamarkuptext}%
 effect of show sorts on the above
 
 \begin{isabelle}%
-{\isacharparenleft}c\ {\isacharasterisk}\ a\ {\isacharequal}\ c\ {\isacharasterisk}\ b{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}c\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isasymor}\ a\ {\isacharequal}\ b{\isacharparenright}%
+{\isacharparenleft}{\isacharparenleft}c{\isasymColon}{\isacharprime}a{\isasymColon}ring{\isacharunderscore}no{\isacharunderscore}zero{\isacharunderscore}divisors{\isacharparenright}\ {\isacharasterisk}\ {\isacharparenleft}a{\isasymColon}{\isacharprime}a{\isasymColon}ring{\isacharunderscore}no{\isacharunderscore}zero{\isacharunderscore}divisors{\isacharparenright}\ {\isacharequal}\isanewline
+\isaindent{{\isacharparenleft}}c\ {\isacharasterisk}\ {\isacharparenleft}b{\isasymColon}{\isacharprime}a{\isasymColon}ring{\isacharunderscore}no{\isacharunderscore}zero{\isacharunderscore}divisors{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline
+{\isacharparenleft}c\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}{\isasymColon}{\isacharprime}a{\isasymColon}ring{\isacharunderscore}no{\isacharunderscore}zero{\isacharunderscore}divisors{\isacharparenright}\ {\isasymor}\ a\ {\isacharequal}\ b{\isacharparenright}%
 \end{isabelle}
 \rulename{mult_cancel_left}%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
-\isadelimML
-%
-\endisadelimML
-%
-\isatagML
-\isacommand{ML}\isamarkupfalse%
-{\isacharverbatimopen}reset\ show{\isacharunderscore}sorts{\isacharverbatimclose}%
-\endisatagML
-{\isafoldML}%
-%
-\isadelimML
-%
-\endisadelimML
-%
 \begin{isamarkuptext}%
 absolute value
 
--- a/src/HOL/IsaMakefile	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/HOL/IsaMakefile	Thu Oct 01 20:52:18 2009 +0200
@@ -192,6 +192,7 @@
   Tools/sat_funcs.ML \
   Tools/sat_solver.ML \
   Tools/split_rule.ML \
+  Tools/transfer.ML \
   Tools/typecopy.ML \
   Tools/typedef_codegen.ML \
   Tools/typedef.ML \
@@ -258,12 +259,12 @@
   Tools/Qelim/presburger.ML \
   Tools/Qelim/qelim.ML \
   Tools/recdef.ML \
+  Tools/choice_specification.ML \
   Tools/res_atp.ML \
   Tools/res_axioms.ML \
   Tools/res_clause.ML \
   Tools/res_hol_clause.ML \
   Tools/res_reconstruct.ML \
-  Tools/choice_specification.ML \
   Tools/string_code.ML \
   Tools/string_syntax.ML \
   Tools/TFL/casesplit.ML \
@@ -308,7 +309,6 @@
   Taylor.thy \
   Transcendental.thy \
   Tools/float_syntax.ML \
-  Tools/transfer.ML \
   Tools/Qelim/ferrante_rackoff_data.ML \
   Tools/Qelim/ferrante_rackoff.ML \
   Tools/Qelim/langford_data.ML \
--- a/src/HOL/Library/Nat_Int_Bij.thy	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/HOL/Library/Nat_Int_Bij.thy	Thu Oct 01 20:52:18 2009 +0200
@@ -128,6 +128,9 @@
   thus "\<forall>y. \<exists>x. y = nat2_to_nat x"  by fast
 qed
 
+lemma nat_to_nat2_inj: "inj nat_to_nat2"
+  by (simp add: nat_to_nat2_def surj_imp_inj_inv nat2_to_nat_surj) 
+
 
 subsection{*  A bijection between @{text "\<nat>"} and @{text "\<int>"} *}
 
--- a/src/HOL/Library/Sum_Of_Squares/positivstellensatz_tools.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/HOL/Library/Sum_Of_Squares/positivstellensatz_tools.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -16,8 +16,6 @@
 structure PositivstellensatzTools : POSITIVSTELLENSATZ_TOOLS =
 struct
 
-open RealArith FuncUtil
-
 (*** certificate generation ***)
 
 fun string_of_rat r =
@@ -41,42 +39,42 @@
   end
 
 fun string_of_monomial m = 
- if Ctermfunc.is_undefined m then "1" 
+ if FuncUtil.Ctermfunc.is_empty m then "1" 
  else 
   let 
-   val m' = dest_monomial m
+   val m' = FuncUtil.dest_monomial m
    val vps = fold_rev (fn (x,k) => cons (string_of_varpow x k)) m' [] 
   in foldr1 (fn (s, t) => s ^ "*" ^ t) vps
   end
 
 fun string_of_cmonomial (m,c) =
-  if Ctermfunc.is_undefined m then string_of_rat c
+  if FuncUtil.Ctermfunc.is_empty m then string_of_rat c
   else if c = Rat.one then string_of_monomial m
   else (string_of_rat c) ^ "*" ^ (string_of_monomial m);
 
 fun string_of_poly p = 
- if Monomialfunc.is_undefined p then "0" 
+ if FuncUtil.Monomialfunc.is_empty p then "0" 
  else
   let 
    val cms = map string_of_cmonomial
-     (sort (prod_ord monomial_order (K EQUAL)) (Monomialfunc.graph p))
+     (sort (prod_ord FuncUtil.monomial_order (K EQUAL)) (FuncUtil.Monomialfunc.dest p))
   in foldr1 (fn (t1, t2) => t1 ^ " + " ^ t2) cms
   end;
 
-fun pss_to_cert (Axiom_eq i) = "A=" ^ string_of_int i
-  | pss_to_cert (Axiom_le i) = "A<=" ^ string_of_int i
-  | pss_to_cert (Axiom_lt i) = "A<" ^ string_of_int i
-  | pss_to_cert (Rational_eq r) = "R=" ^ string_of_rat r
-  | pss_to_cert (Rational_le r) = "R<=" ^ string_of_rat r
-  | pss_to_cert (Rational_lt r) = "R<" ^ string_of_rat r
-  | pss_to_cert (Square p) = "[" ^ string_of_poly p ^ "]^2"
-  | pss_to_cert (Eqmul (p, pss)) = "([" ^ string_of_poly p ^ "] * " ^ pss_to_cert pss ^ ")"
-  | pss_to_cert (Sum (pss1, pss2)) = "(" ^ pss_to_cert pss1 ^ " + " ^ pss_to_cert pss2 ^ ")"
-  | pss_to_cert (Product (pss1, pss2)) = "(" ^ pss_to_cert pss1 ^ " * " ^ pss_to_cert pss2 ^ ")"
+fun pss_to_cert (RealArith.Axiom_eq i) = "A=" ^ string_of_int i
+  | pss_to_cert (RealArith.Axiom_le i) = "A<=" ^ string_of_int i
+  | pss_to_cert (RealArith.Axiom_lt i) = "A<" ^ string_of_int i
+  | pss_to_cert (RealArith.Rational_eq r) = "R=" ^ string_of_rat r
+  | pss_to_cert (RealArith.Rational_le r) = "R<=" ^ string_of_rat r
+  | pss_to_cert (RealArith.Rational_lt r) = "R<" ^ string_of_rat r
+  | pss_to_cert (RealArith.Square p) = "[" ^ string_of_poly p ^ "]^2"
+  | pss_to_cert (RealArith.Eqmul (p, pss)) = "([" ^ string_of_poly p ^ "] * " ^ pss_to_cert pss ^ ")"
+  | pss_to_cert (RealArith.Sum (pss1, pss2)) = "(" ^ pss_to_cert pss1 ^ " + " ^ pss_to_cert pss2 ^ ")"
+  | pss_to_cert (RealArith.Product (pss1, pss2)) = "(" ^ pss_to_cert pss1 ^ " * " ^ pss_to_cert pss2 ^ ")"
 
-fun pss_tree_to_cert Trivial = "()"
-  | pss_tree_to_cert (Cert pss) = "(" ^ pss_to_cert pss ^ ")"
-  | pss_tree_to_cert (Branch (t1, t2)) = "(" ^ pss_tree_to_cert t1 ^ " & " ^ pss_tree_to_cert t2 ^ ")"
+fun pss_tree_to_cert RealArith.Trivial = "()"
+  | pss_tree_to_cert (RealArith.Cert pss) = "(" ^ pss_to_cert pss ^ ")"
+  | pss_tree_to_cert (RealArith.Branch (t1, t2)) = "(" ^ pss_tree_to_cert t1 ^ " & " ^ pss_tree_to_cert t2 ^ ")"
 
 (*** certificate parsing ***)
 
@@ -103,27 +101,27 @@
   (fn (x, k) => (cterm_of (Context.theory_of_proof ctxt) (Free (x, @{typ real})), k)) 
 
 fun parse_monomial ctxt = repeat_sep "*" (parse_varpow ctxt) >>
-  foldl (uncurry Ctermfunc.update) Ctermfunc.undefined
+  foldl (uncurry FuncUtil.Ctermfunc.update) FuncUtil.Ctermfunc.empty
 
 fun parse_cmonomial ctxt =
   rat_int --| str "*" -- (parse_monomial ctxt) >> swap ||
   (parse_monomial ctxt) >> (fn m => (m, Rat.one)) ||
-  rat_int >> (fn r => (Ctermfunc.undefined, r))
+  rat_int >> (fn r => (FuncUtil.Ctermfunc.empty, r))
 
 fun parse_poly ctxt = repeat_sep "+" (parse_cmonomial ctxt) >>
-  foldl (uncurry Monomialfunc.update) Monomialfunc.undefined
+  foldl (uncurry FuncUtil.Monomialfunc.update) FuncUtil.Monomialfunc.empty
 
 (* positivstellensatz parser *)
 
 val parse_axiom =
-  (str "A=" |-- int >> Axiom_eq) ||
-  (str "A<=" |-- int >> Axiom_le) ||
-  (str "A<" |-- int >> Axiom_lt)
+  (str "A=" |-- int >> RealArith.Axiom_eq) ||
+  (str "A<=" |-- int >> RealArith.Axiom_le) ||
+  (str "A<" |-- int >> RealArith.Axiom_lt)
 
 val parse_rational =
-  (str "R=" |-- rat_int >> Rational_eq) ||
-  (str "R<=" |-- rat_int >> Rational_le) ||
-  (str "R<" |-- rat_int >> Rational_lt)
+  (str "R=" |-- rat_int >> RealArith.Rational_eq) ||
+  (str "R<=" |-- rat_int >> RealArith.Rational_le) ||
+  (str "R<" |-- rat_int >> RealArith.Rational_lt)
 
 fun parse_cert ctxt input =
   let
@@ -132,10 +130,10 @@
   in
   (parse_axiom ||
    parse_rational ||
-   str "[" |-- pp --| str "]^2" >> Square ||
-   str "([" |-- pp --| str "]*" -- pc --| str ")" >> Eqmul ||
-   str "(" |-- pc --| str "*" -- pc --| str ")" >> Product ||
-   str "(" |-- pc --| str "+" -- pc --| str ")" >> Sum) input
+   str "[" |-- pp --| str "]^2" >> RealArith.Square ||
+   str "([" |-- pp --| str "]*" -- pc --| str ")" >> RealArith.Eqmul ||
+   str "(" |-- pc --| str "*" -- pc --| str ")" >> RealArith.Product ||
+   str "(" |-- pc --| str "+" -- pc --| str ")" >> RealArith.Sum) input
   end
 
 fun parse_cert_tree ctxt input =
@@ -143,9 +141,9 @@
     val pc = parse_cert ctxt
     val pt = parse_cert_tree ctxt
   in
-  (str "()" >> K Trivial ||
-   str "(" |-- pc --| str ")" >> Cert ||
-   str "(" |-- pt --| str "&" -- pt --| str ")" >> Branch) input
+  (str "()" >> K RealArith.Trivial ||
+   str "(" |-- pc --| str ")" >> RealArith.Cert ||
+   str "(" |-- pt --| str "&" -- pt --| str ")" >> RealArith.Branch) input
   end
 
 (* scanner *)
--- a/src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -23,8 +23,6 @@
 structure Sos : SOS = 
 struct
 
-open FuncUtil;
-
 val rat_0 = Rat.zero;
 val rat_1 = Rat.one;
 val rat_2 = Rat.two;
@@ -104,9 +102,9 @@
 
 (* The main types.                                                           *)
 
-type vector = int* Rat.rat Intfunc.T;
+type vector = int* Rat.rat FuncUtil.Intfunc.table;
 
-type matrix = (int*int)*(Rat.rat Intpairfunc.T);
+type matrix = (int*int)*(Rat.rat FuncUtil.Intpairfunc.table);
 
 fun iszero (k,r) = r =/ rat_0;
 
@@ -118,29 +116,29 @@
  
 (* Vectors. Conventionally indexed 1..n.                                     *)
 
-fun vector_0 n = (n,Intfunc.undefined):vector;
+fun vector_0 n = (n,FuncUtil.Intfunc.empty):vector;
 
 fun dim (v:vector) = fst v;
 
 fun vector_const c n =
   if c =/ rat_0 then vector_0 n
-  else (n,fold_rev (fn k => Intfunc.update (k,c)) (1 upto n) Intfunc.undefined) :vector;
+  else (n,fold_rev (fn k => FuncUtil.Intfunc.update (k,c)) (1 upto n) FuncUtil.Intfunc.empty) :vector;
 
 val vector_1 = vector_const rat_1;
 
 fun vector_cmul c (v:vector) =
  let val n = dim v 
  in if c =/ rat_0 then vector_0 n
-    else (n,Intfunc.mapf (fn x => c */ x) (snd v))
+    else (n,FuncUtil.Intfunc.map (fn x => c */ x) (snd v))
  end;
 
-fun vector_neg (v:vector) = (fst v,Intfunc.mapf Rat.neg (snd v)) :vector;
+fun vector_neg (v:vector) = (fst v,FuncUtil.Intfunc.map Rat.neg (snd v)) :vector;
 
 fun vector_add (v1:vector) (v2:vector) =
  let val m = dim v1  
      val n = dim v2 
  in if m <> n then error "vector_add: incompatible dimensions"
-    else (n,Intfunc.combine (curry op +/) (fn x => x =/ rat_0) (snd v1) (snd v2)) :vector 
+    else (n,FuncUtil.Intfunc.combine (curry op +/) (fn x => x =/ rat_0) (snd v1) (snd v2)) :vector 
  end;
 
 fun vector_sub v1 v2 = vector_add v1 (vector_neg v2);
@@ -149,43 +147,43 @@
  let val m = dim v1 
      val n = dim v2 
  in if m <> n then error "vector_dot: incompatible dimensions" 
-    else Intfunc.fold (fn (i,x) => fn a => x +/ a) 
-        (Intfunc.combine (curry op */) (fn x => x =/ rat_0) (snd v1) (snd v2)) rat_0
+    else FuncUtil.Intfunc.fold (fn (i,x) => fn a => x +/ a) 
+        (FuncUtil.Intfunc.combine (curry op */) (fn x => x =/ rat_0) (snd v1) (snd v2)) rat_0
  end;
 
 fun vector_of_list l =
  let val n = length l 
- in (n,fold_rev2 (curry Intfunc.update) (1 upto n) l Intfunc.undefined) :vector
+ in (n,fold_rev2 (curry FuncUtil.Intfunc.update) (1 upto n) l FuncUtil.Intfunc.empty) :vector
  end;
 
 (* Matrices; again rows and columns indexed from 1.                          *)
 
-fun matrix_0 (m,n) = ((m,n),Intpairfunc.undefined):matrix;
+fun matrix_0 (m,n) = ((m,n),FuncUtil.Intpairfunc.empty):matrix;
 
 fun dimensions (m:matrix) = fst m;
 
 fun matrix_const c (mn as (m,n)) =
   if m <> n then error "matrix_const: needs to be square"
   else if c =/ rat_0 then matrix_0 mn
-  else (mn,fold_rev (fn k => Intpairfunc.update ((k,k), c)) (1 upto n) Intpairfunc.undefined) :matrix;;
+  else (mn,fold_rev (fn k => FuncUtil.Intpairfunc.update ((k,k), c)) (1 upto n) FuncUtil.Intpairfunc.empty) :matrix;;
 
 val matrix_1 = matrix_const rat_1;
 
 fun matrix_cmul c (m:matrix) =
  let val (i,j) = dimensions m 
  in if c =/ rat_0 then matrix_0 (i,j)
-    else ((i,j),Intpairfunc.mapf (fn x => c */ x) (snd m))
+    else ((i,j),FuncUtil.Intpairfunc.map (fn x => c */ x) (snd m))
  end;
 
 fun matrix_neg (m:matrix) = 
-  (dimensions m, Intpairfunc.mapf Rat.neg (snd m)) :matrix;
+  (dimensions m, FuncUtil.Intpairfunc.map Rat.neg (snd m)) :matrix;
 
 fun matrix_add (m1:matrix) (m2:matrix) =
  let val d1 = dimensions m1 
      val d2 = dimensions m2 
  in if d1 <> d2 
      then error "matrix_add: incompatible dimensions"
-    else (d1,Intpairfunc.combine (curry op +/) (fn x => x =/ rat_0) (snd m1) (snd m2)) :matrix
+    else (d1,FuncUtil.Intpairfunc.combine (curry op +/) (fn x => x =/ rat_0) (snd m1) (snd m2)) :matrix
  end;;
 
 fun matrix_sub m1 m2 = matrix_add m1 (matrix_neg m2);
@@ -193,112 +191,112 @@
 fun row k (m:matrix) =
  let val (i,j) = dimensions m 
  in (j,
-   Intpairfunc.fold (fn ((i,j), c) => fn a => if i = k then Intfunc.update (j,c) a else a) (snd m) Intfunc.undefined ) : vector
+   FuncUtil.Intpairfunc.fold (fn ((i,j), c) => fn a => if i = k then FuncUtil.Intfunc.update (j,c) a else a) (snd m) FuncUtil.Intfunc.empty ) : vector
  end;
 
 fun column k (m:matrix) =
   let val (i,j) = dimensions m 
   in (i,
-   Intpairfunc.fold (fn ((i,j), c) => fn a => if j = k then Intfunc.update (i,c) a else a) (snd m)  Intfunc.undefined)
+   FuncUtil.Intpairfunc.fold (fn ((i,j), c) => fn a => if j = k then FuncUtil.Intfunc.update (i,c) a else a) (snd m)  FuncUtil.Intfunc.empty)
    : vector
  end;
 
 fun transp (m:matrix) =
   let val (i,j) = dimensions m 
   in
-  ((j,i),Intpairfunc.fold (fn ((i,j), c) => fn a => Intpairfunc.update ((j,i), c) a) (snd m) Intpairfunc.undefined) :matrix
+  ((j,i),FuncUtil.Intpairfunc.fold (fn ((i,j), c) => fn a => FuncUtil.Intpairfunc.update ((j,i), c) a) (snd m) FuncUtil.Intpairfunc.empty) :matrix
  end;
 
 fun diagonal (v:vector) =
  let val n = dim v 
- in ((n,n),Intfunc.fold (fn (i, c) => fn a => Intpairfunc.update ((i,i), c) a) (snd v) Intpairfunc.undefined) : matrix
+ in ((n,n),FuncUtil.Intfunc.fold (fn (i, c) => fn a => FuncUtil.Intpairfunc.update ((i,i), c) a) (snd v) FuncUtil.Intpairfunc.empty) : matrix
  end;
 
 fun matrix_of_list l =
  let val m = length l 
  in if m = 0 then matrix_0 (0,0) else
    let val n = length (hd l) 
-   in ((m,n),itern 1 l (fn v => fn i => itern 1 v (fn c => fn j => Intpairfunc.update ((i,j), c))) Intpairfunc.undefined)
+   in ((m,n),itern 1 l (fn v => fn i => itern 1 v (fn c => fn j => FuncUtil.Intpairfunc.update ((i,j), c))) FuncUtil.Intpairfunc.empty)
    end
  end;
 
 (* Monomials.                                                                *)
 
-fun monomial_eval assig (m:monomial) =
-  Ctermfunc.fold (fn (x, k) => fn a => a */ rat_pow (Ctermfunc.apply assig x) k)
+fun monomial_eval assig m =
+  FuncUtil.Ctermfunc.fold (fn (x, k) => fn a => a */ rat_pow (FuncUtil.Ctermfunc.apply assig x) k)
         m rat_1;
-val monomial_1 = (Ctermfunc.undefined:monomial);
+val monomial_1 = FuncUtil.Ctermfunc.empty;
 
-fun monomial_var x = Ctermfunc.onefunc (x, 1) :monomial;
+fun monomial_var x = FuncUtil.Ctermfunc.onefunc (x, 1);
 
-val (monomial_mul:monomial->monomial->monomial) =
-  Ctermfunc.combine (curry op +) (K false);
+val monomial_mul =
+  FuncUtil.Ctermfunc.combine (curry op +) (K false);
 
-fun monomial_pow (m:monomial) k =
+fun monomial_pow m k =
   if k = 0 then monomial_1
-  else Ctermfunc.mapf (fn x => k * x) m;
+  else FuncUtil.Ctermfunc.map (fn x => k * x) m;
 
-fun monomial_divides (m1:monomial) (m2:monomial) =
-  Ctermfunc.fold (fn (x, k) => fn a => Ctermfunc.tryapplyd m2 x 0 >= k andalso a) m1 true;;
+fun monomial_divides m1 m2 =
+  FuncUtil.Ctermfunc.fold (fn (x, k) => fn a => FuncUtil.Ctermfunc.tryapplyd m2 x 0 >= k andalso a) m1 true;;
 
-fun monomial_div (m1:monomial) (m2:monomial) =
- let val m = Ctermfunc.combine (curry op +) 
-   (fn x => x = 0) m1 (Ctermfunc.mapf (fn x => ~ x) m2) 
- in if Ctermfunc.fold (fn (x, k) => fn a => k >= 0 andalso a) m true then m
+fun monomial_div m1 m2 =
+ let val m = FuncUtil.Ctermfunc.combine (curry op +) 
+   (fn x => x = 0) m1 (FuncUtil.Ctermfunc.map (fn x => ~ x) m2) 
+ in if FuncUtil.Ctermfunc.fold (fn (x, k) => fn a => k >= 0 andalso a) m true then m
   else error "monomial_div: non-divisible"
  end;
 
-fun monomial_degree x (m:monomial) = 
-  Ctermfunc.tryapplyd m x 0;;
+fun monomial_degree x m = 
+  FuncUtil.Ctermfunc.tryapplyd m x 0;;
 
-fun monomial_lcm (m1:monomial) (m2:monomial) =
-  fold_rev (fn x => Ctermfunc.update (x, max (monomial_degree x m1) (monomial_degree x m2)))
-          (gen_union (is_equal o  cterm_ord) (Ctermfunc.dom m1, Ctermfunc.dom m2)) (Ctermfunc.undefined :monomial);
+fun monomial_lcm m1 m2 =
+  fold_rev (fn x => FuncUtil.Ctermfunc.update (x, max (monomial_degree x m1) (monomial_degree x m2)))
+          (gen_union (is_equal o  FuncUtil.cterm_ord) (FuncUtil.Ctermfunc.dom m1, FuncUtil.Ctermfunc.dom m2)) (FuncUtil.Ctermfunc.empty);
 
-fun monomial_multidegree (m:monomial) = 
- Ctermfunc.fold (fn (x, k) => fn a => k + a) m 0;;
+fun monomial_multidegree m = 
+ FuncUtil.Ctermfunc.fold (fn (x, k) => fn a => k + a) m 0;;
 
-fun monomial_variables m = Ctermfunc.dom m;;
+fun monomial_variables m = FuncUtil.Ctermfunc.dom m;;
 
 (* Polynomials.                                                              *)
 
-fun eval assig (p:poly) =
-  Monomialfunc.fold (fn (m, c) => fn a => a +/ c */ monomial_eval assig m) p rat_0;
+fun eval assig p =
+  FuncUtil.Monomialfunc.fold (fn (m, c) => fn a => a +/ c */ monomial_eval assig m) p rat_0;
 
-val poly_0 = (Monomialfunc.undefined:poly);
+val poly_0 = FuncUtil.Monomialfunc.empty;
 
-fun poly_isconst (p:poly) = 
-  Monomialfunc.fold (fn (m, c) => fn a => Ctermfunc.is_undefined m andalso a) p true;
+fun poly_isconst p = 
+  FuncUtil.Monomialfunc.fold (fn (m, c) => fn a => FuncUtil.Ctermfunc.is_empty m andalso a) p true;
 
-fun poly_var x = Monomialfunc.onefunc (monomial_var x,rat_1) :poly;
+fun poly_var x = FuncUtil.Monomialfunc.onefunc (monomial_var x,rat_1);
 
 fun poly_const c =
-  if c =/ rat_0 then poly_0 else Monomialfunc.onefunc(monomial_1, c);
+  if c =/ rat_0 then poly_0 else FuncUtil.Monomialfunc.onefunc(monomial_1, c);
 
-fun poly_cmul c (p:poly) =
+fun poly_cmul c p =
   if c =/ rat_0 then poly_0
-  else Monomialfunc.mapf (fn x => c */ x) p;
+  else FuncUtil.Monomialfunc.map (fn x => c */ x) p;
 
-fun poly_neg (p:poly) = (Monomialfunc.mapf Rat.neg p :poly);;
+fun poly_neg p = FuncUtil.Monomialfunc.map Rat.neg p;;
 
-fun poly_add (p1:poly) (p2:poly) =
-  (Monomialfunc.combine (curry op +/) (fn x => x =/ rat_0) p1 p2 :poly);
+fun poly_add p1 p2 =
+  FuncUtil.Monomialfunc.combine (curry op +/) (fn x => x =/ rat_0) p1 p2;
 
 fun poly_sub p1 p2 = poly_add p1 (poly_neg p2);
 
-fun poly_cmmul (c,m) (p:poly) =
+fun poly_cmmul (c,m) p =
  if c =/ rat_0 then poly_0
- else if Ctermfunc.is_undefined m 
-      then Monomialfunc.mapf (fn d => c */ d) p
-      else Monomialfunc.fold (fn (m', d) => fn a => (Monomialfunc.update (monomial_mul m m', c */ d) a)) p poly_0;
+ else if FuncUtil.Ctermfunc.is_empty m 
+      then FuncUtil.Monomialfunc.map (fn d => c */ d) p
+      else FuncUtil.Monomialfunc.fold (fn (m', d) => fn a => (FuncUtil.Monomialfunc.update (monomial_mul m m', c */ d) a)) p poly_0;
 
-fun poly_mul (p1:poly) (p2:poly) =
-  Monomialfunc.fold (fn (m, c) => fn a => poly_add (poly_cmmul (c,m) p2) a) p1 poly_0;
+fun poly_mul p1 p2 =
+  FuncUtil.Monomialfunc.fold (fn (m, c) => fn a => poly_add (poly_cmmul (c,m) p2) a) p1 poly_0;
 
-fun poly_div (p1:poly) (p2:poly) =
+fun poly_div p1 p2 =
  if not(poly_isconst p2) 
  then error "poly_div: non-constant" else
- let val c = eval Ctermfunc.undefined p2 
+ let val c = eval FuncUtil.Ctermfunc.empty p2 
  in if c =/ rat_0 then error "poly_div: division by zero"
     else poly_cmul (Rat.inv c) p1
  end;
@@ -314,22 +312,20 @@
 fun poly_exp p1 p2 =
   if not(poly_isconst p2) 
   then error "poly_exp: not a constant" 
-  else poly_pow p1 (int_of_rat (eval Ctermfunc.undefined p2));
+  else poly_pow p1 (int_of_rat (eval FuncUtil.Ctermfunc.empty p2));
 
-fun degree x (p:poly) = 
- Monomialfunc.fold (fn (m,c) => fn a => max (monomial_degree x m) a) p 0;
+fun degree x p = 
+ FuncUtil.Monomialfunc.fold (fn (m,c) => fn a => max (monomial_degree x m) a) p 0;
 
-fun multidegree (p:poly) =
-  Monomialfunc.fold (fn (m, c) => fn a => max (monomial_multidegree m) a) p 0;
+fun multidegree p =
+  FuncUtil.Monomialfunc.fold (fn (m, c) => fn a => max (monomial_multidegree m) a) p 0;
 
-fun poly_variables (p:poly) =
-  sort cterm_ord (Monomialfunc.fold_rev (fn (m, c) => curry (gen_union (is_equal o  cterm_ord)) (monomial_variables m)) p []);;
+fun poly_variables p =
+  sort FuncUtil.cterm_ord (FuncUtil.Monomialfunc.fold_rev (fn (m, c) => curry (gen_union (is_equal o FuncUtil.cterm_ord)) (monomial_variables m)) p []);;
 
 (* Order monomials for human presentation.                                   *)
 
-fun cterm_ord (t,t') = TermOrd.fast_term_ord (term_of t, term_of t');
-
-val humanorder_varpow = prod_ord cterm_ord (rev_order o int_ord);
+val humanorder_varpow = prod_ord FuncUtil.cterm_ord (rev_order o int_ord);
 
 local
  fun ord (l1,l2) = case (l1,l2) of
@@ -341,8 +337,8 @@
    | EQUAL => ord (t1,t2)
    | GREATER => GREATER)
 in fun humanorder_monomial m1 m2 = 
- ord (sort humanorder_varpow (Ctermfunc.graph m1),
-  sort humanorder_varpow (Ctermfunc.graph m2))
+ ord (sort humanorder_varpow (FuncUtil.Ctermfunc.dest m1),
+  sort humanorder_varpow (FuncUtil.Ctermfunc.dest m2))
 end;
 
 (* Conversions to strings.                                                   *)
@@ -352,8 +348,8 @@
  in if n_raw = 0 then "[]" else
   let 
    val n = max min_size (min n_raw max_size) 
-   val xs = map (Rat.string_of_rat o (fn i => Intfunc.tryapplyd (snd v) i rat_0)) (1 upto n) 
-  in "[" ^ foldr1 (fn (s, t) => s ^ ", " ^ t) xs ^
+   val xs = map (Rat.string_of_rat o (fn i => FuncUtil.Intfunc.tryapplyd (snd v) i rat_0)) (1 upto n) 
+  in "[" ^ space_implode ", " xs ^
   (if n_raw > max_size then ", ...]" else "]")
   end
  end;
@@ -364,7 +360,7 @@
   val i = min max_size i_raw 
   val j = min max_size j_raw
   val rstr = map (fn k => string_of_vector j j (row k m)) (1 upto i) 
- in "["^ foldr1 (fn (s, t) => s^";\n "^t) rstr ^
+ in "["^ space_implode ";\n " rstr ^
   (if j > max_size then "\n ...]" else "]")
  end;
 
@@ -387,21 +383,21 @@
   else string_of_cterm x^"^"^string_of_int k;
 
 fun string_of_monomial m =
- if Ctermfunc.is_undefined m then "1" else
+ if FuncUtil.Ctermfunc.is_empty m then "1" else
  let val vps = fold_rev (fn (x,k) => fn a => string_of_varpow x k :: a)
-  (sort humanorder_varpow (Ctermfunc.graph m)) [] 
- in foldr1 (fn (s, t) => s^"*"^t) vps
+  (sort humanorder_varpow (FuncUtil.Ctermfunc.dest m)) [] 
+ in space_implode "*" vps
  end;
 
 fun string_of_cmonomial (c,m) =
- if Ctermfunc.is_undefined m then Rat.string_of_rat c
+ if FuncUtil.Ctermfunc.is_empty m then Rat.string_of_rat c
  else if c =/ rat_1 then string_of_monomial m
  else Rat.string_of_rat c ^ "*" ^ string_of_monomial m;;
 
 fun string_of_poly p =
- if Monomialfunc.is_undefined p then "<<0>>" else
+ if FuncUtil.Monomialfunc.is_empty p then "<<0>>" else
  let 
-  val cms = sort (fn ((m1,_),(m2,_)) => humanorder_monomial m1  m2) (Monomialfunc.graph p)
+  val cms = sort (fn ((m1,_),(m2,_)) => humanorder_monomial m1  m2) (FuncUtil.Monomialfunc.dest p)
   val s = fold (fn (m,c) => fn a =>
              if c </ rat_0 then a ^ " - " ^ string_of_cmonomial(Rat.neg c,m)
              else a ^ " + " ^ string_of_cmonomial(c,m))
@@ -434,7 +430,7 @@
       else if lop aconvc inv_tm then
        let val p = poly_of_term r 
        in if poly_isconst p 
-          then poly_const(Rat.inv (eval Ctermfunc.undefined p))
+          then poly_const(Rat.inv (eval FuncUtil.Ctermfunc.empty p))
           else error "poly_of_term: inverse of non-constant polyomial"
        end
    else (let val (opr,l) = Thm.dest_comb lop
@@ -451,7 +447,7 @@
            then let 
                   val p = poly_of_term l 
                   val q = poly_of_term r 
-                in if poly_isconst q then poly_cmul (Rat.inv (eval Ctermfunc.undefined q)) p
+                in if poly_isconst q then poly_cmul (Rat.inv (eval FuncUtil.Ctermfunc.empty q)) p
                    else error "poly_of_term: division by non-constant polynomial"
                 end
           else poly_var tm
@@ -471,8 +467,8 @@
 fun sdpa_of_vector (v:vector) =
  let 
   val n = dim v
-  val strs = map (decimalize 20 o (fn i => Intfunc.tryapplyd (snd v) i rat_0)) (1 upto n) 
- in foldr1 (fn (x, y) => x ^ " " ^ y) strs ^ "\n"
+  val strs = map (decimalize 20 o (fn i => FuncUtil.Intfunc.tryapplyd (snd v) i rat_0)) (1 upto n) 
+ in space_implode " " strs ^ "\n"
  end;
 
 fun triple_int_ord ((a,b,c),(a',b',c')) = 
@@ -487,7 +483,7 @@
   val pfx = string_of_int k ^" "
   val ents =
     Inttriplefunc.fold (fn ((b,i,j), c) => fn a => if i > j then a else ((b,i,j),c)::a) m []
-  val entss = sort (increasing fst triple_int_ord ) ents
+  val entss = sort (triple_int_ord o pairself fst) ents
  in  fold_rev (fn ((b,i,j),c) => fn a =>
      pfx ^ string_of_int b ^ " " ^ string_of_int i ^ " " ^ string_of_int j ^
      " " ^ decimalize 20 c ^ "\n" ^ a) entss ""
@@ -498,8 +494,8 @@
 fun sdpa_of_matrix k (m:matrix) =
  let 
   val pfx = string_of_int k ^ " 1 "
-  val ms = Intpairfunc.fold (fn ((i,j), c) => fn  a => if i > j then a else ((i,j),c)::a) (snd m) [] 
-  val mss = sort (increasing fst (prod_ord int_ord int_ord)) ms 
+  val ms = FuncUtil.Intpairfunc.fold (fn ((i,j), c) => fn  a => if i > j then a else ((i,j),c)::a) (snd m) [] 
+  val mss = sort ((prod_ord int_ord int_ord) o pairself fst) ms 
  in fold_rev (fn ((i,j),c) => fn a =>
      pfx ^ string_of_int i ^ " " ^ string_of_int j ^
      " " ^ decimalize 20 c ^ "\n" ^ a) mss ""
@@ -544,18 +540,15 @@
 
 (* More parser basics.                                                       *)
 
-local
- open Scan
-in 
- val word = this_string
+ val word = Scan.this_string
  fun token s =
-  repeat ($$ " ") |-- word s --| repeat ($$ " ")
- val numeral = one isnum
- val decimalint = bulk numeral >> (rat_of_string o implode)
- val decimalfrac = bulk numeral
+  Scan.repeat ($$ " ") |-- word s --| Scan.repeat ($$ " ")
+ val numeral = Scan.one isnum
+ val decimalint = Scan.bulk numeral >> (rat_of_string o implode)
+ val decimalfrac = Scan.bulk numeral
     >> (fn s => rat_of_string(implode s) // pow10 (length s))
  val decimalsig =
-    decimalint -- option (Scan.$$ "." |-- decimalfrac)
+    decimalint -- Scan.option (Scan.$$ "." |-- decimalfrac)
     >> (fn (h,NONE) => h | (h,SOME x) => h +/ x)
  fun signed prs =
        $$ "-" |-- prs >> Rat.neg 
@@ -568,7 +561,6 @@
 
  val decimal = signed decimalsig -- (emptyin rat_0|| exponent)
     >> (fn (h, x) => h */ pow10 (int_of_rat x));
-end;
 
  fun mkparser p s =
   let val (x,rst) = p (explode s) 
@@ -605,15 +597,15 @@
 
 fun pi_scale_then solver (obj:vector)  mats =
  let 
-  val cd1 = fold_rev (common_denominator Intpairfunc.fold) mats (rat_1)
-  val cd2 = common_denominator Intfunc.fold (snd obj)  (rat_1) 
-  val mats' = map (Intpairfunc.mapf (fn x => cd1 */ x)) mats
+  val cd1 = fold_rev (common_denominator FuncUtil.Intpairfunc.fold) mats (rat_1)
+  val cd2 = common_denominator FuncUtil.Intfunc.fold (snd obj)  (rat_1) 
+  val mats' = map (FuncUtil.Intpairfunc.map (fn x => cd1 */ x)) mats
   val obj' = vector_cmul cd2 obj
-  val max1 = fold_rev (maximal_element Intpairfunc.fold) mats' (rat_0)
-  val max2 = maximal_element Intfunc.fold (snd obj') (rat_0) 
+  val max1 = fold_rev (maximal_element FuncUtil.Intpairfunc.fold) mats' (rat_0)
+  val max2 = maximal_element FuncUtil.Intfunc.fold (snd obj') (rat_0) 
   val scal1 = pow2 (20 - trunc(Math.ln (float_of_rat max1) / Math.ln 2.0))
   val scal2 = pow2 (20 - trunc(Math.ln (float_of_rat max2) / Math.ln 2.0)) 
-  val mats'' = map (Intpairfunc.mapf (fn x => x */ scal1)) mats'
+  val mats'' = map (FuncUtil.Intpairfunc.map (fn x => x */ scal1)) mats'
   val obj'' = vector_cmul scal2 obj' 
  in solver obj'' mats''
   end
@@ -639,14 +631,14 @@
 fun tri_scale_then solver (obj:vector)  mats =
  let 
   val cd1 = fold_rev (common_denominator Inttriplefunc.fold) mats (rat_1)
-  val cd2 = common_denominator Intfunc.fold (snd obj)  (rat_1) 
-  val mats' = map (Inttriplefunc.mapf (fn x => cd1 */ x)) mats
+  val cd2 = common_denominator FuncUtil.Intfunc.fold (snd obj)  (rat_1) 
+  val mats' = map (Inttriplefunc.map (fn x => cd1 */ x)) mats
   val obj' = vector_cmul cd2 obj
   val max1 = fold_rev (maximal_element Inttriplefunc.fold) mats' (rat_0)
-  val max2 = maximal_element Intfunc.fold (snd obj') (rat_0) 
+  val max2 = maximal_element FuncUtil.Intfunc.fold (snd obj') (rat_0) 
   val scal1 = pow2 (20 - int_of_float(Math.ln (float_of_rat max1) / Math.ln 2.0))
   val scal2 = pow2 (20 - int_of_float(Math.ln (float_of_rat max2) / Math.ln 2.0)) 
-  val mats'' = map (Inttriplefunc.mapf (fn x => x */ scal1)) mats'
+  val mats'' = map (Inttriplefunc.map (fn x => x */ scal1)) mats'
   val obj'' = vector_cmul scal2 obj' 
  in solver obj'' mats''
   end
@@ -656,17 +648,17 @@
 
 fun nice_rational n x = round_rat (n */ x) // n;;
 fun nice_vector n ((d,v) : vector) = 
- (d, Intfunc.fold (fn (i,c) => fn a => 
+ (d, FuncUtil.Intfunc.fold (fn (i,c) => fn a => 
    let val y = nice_rational n c 
    in if c =/ rat_0 then a 
-      else Intfunc.update (i,y) a end) v Intfunc.undefined):vector
+      else FuncUtil.Intfunc.update (i,y) a end) v FuncUtil.Intfunc.empty):vector
 
 fun dest_ord f x = is_equal (f x);
 
 (* Stuff for "equations" ((int*int*int)->num functions).                         *)
 
 fun tri_equation_cmul c eq =
-  if c =/ rat_0 then Inttriplefunc.undefined else Inttriplefunc.mapf (fn d => c */ d) eq;
+  if c =/ rat_0 then Inttriplefunc.empty else Inttriplefunc.map (fn d => c */ d) eq;
 
 fun tri_equation_add eq1 eq2 = Inttriplefunc.combine (curry op +/) (fn x => x =/ rat_0) eq1 eq2;
 
@@ -685,25 +677,25 @@
  | h::t => if p h then (h,t) else
           let val (k,s) = extract_first p t in (k,h::s) end
 fun eliminate vars dun eqs = case vars of 
-  [] => if forall Inttriplefunc.is_undefined eqs then dun
+  [] => if forall Inttriplefunc.is_empty eqs then dun
         else raise Unsolvable
  | v::vs =>
   ((let 
     val (eq,oeqs) = extract_first (fn e => Inttriplefunc.defined e v) eqs 
     val a = Inttriplefunc.apply eq v
-    val eq' = tri_equation_cmul ((Rat.neg rat_1) // a) (Inttriplefunc.undefine v eq)
+    val eq' = tri_equation_cmul ((Rat.neg rat_1) // a) (Inttriplefunc.delete_safe v eq)
     fun elim e =
      let val b = Inttriplefunc.tryapplyd e v rat_0 
      in if b =/ rat_0 then e else
         tri_equation_add e (tri_equation_cmul (Rat.neg b // a) eq)
      end
-   in eliminate vs (Inttriplefunc.update (v,eq') (Inttriplefunc.mapf elim dun)) (map elim oeqs)
+   in eliminate vs (Inttriplefunc.update (v,eq') (Inttriplefunc.map elim dun)) (map elim oeqs)
    end)
   handle Failure _ => eliminate vs dun eqs)
 in
 fun tri_eliminate_equations one vars eqs =
  let 
-  val assig = eliminate vars Inttriplefunc.undefined eqs
+  val assig = eliminate vars Inttriplefunc.empty eqs
   val vs = Inttriplefunc.fold (fn (x, f) => fn a => remove (dest_ord triple_int_ord) one (Inttriplefunc.dom f) @ a) assig []
   in (distinct (dest_ord triple_int_ord) vs, assig)
   end
@@ -716,8 +708,8 @@
   fun choose_variable eq =
    let val (v,_) = Inttriplefunc.choose eq 
    in if is_equal (triple_int_ord(v,one)) then
-      let val eq' = Inttriplefunc.undefine v eq 
-      in if Inttriplefunc.is_undefined eq' then error "choose_variable" 
+      let val eq' = Inttriplefunc.delete_safe v eq 
+      in if Inttriplefunc.is_empty eq' then error "choose_variable" 
          else fst (Inttriplefunc.choose eq')
       end
     else v 
@@ -725,22 +717,22 @@
   fun eliminate dun eqs = case eqs of 
     [] => dun
   | eq::oeqs =>
-    if Inttriplefunc.is_undefined eq then eliminate dun oeqs else
+    if Inttriplefunc.is_empty eq then eliminate dun oeqs else
     let val v = choose_variable eq
         val a = Inttriplefunc.apply eq v
         val eq' = tri_equation_cmul ((Rat.rat_of_int ~1) // a) 
-                   (Inttriplefunc.undefine v eq)
+                   (Inttriplefunc.delete_safe v eq)
         fun elim e =
          let val b = Inttriplefunc.tryapplyd e v rat_0 
          in if b =/ rat_0 then e 
             else tri_equation_add e (tri_equation_cmul (Rat.neg b // a) eq)
          end
-    in eliminate (Inttriplefunc.update(v, eq') (Inttriplefunc.mapf elim dun)) 
+    in eliminate (Inttriplefunc.update(v, eq') (Inttriplefunc.map elim dun)) 
                  (map elim oeqs) 
     end
 in fn eqs =>
  let 
-  val assig = eliminate Inttriplefunc.undefined eqs
+  val assig = eliminate Inttriplefunc.empty eqs
   val vs = Inttriplefunc.fold (fn (x, f) => fn a => remove (dest_ord triple_int_ord) one (Inttriplefunc.dom f) @ a) assig []
  in (distinct (dest_ord triple_int_ord) vs,assig)
  end
@@ -755,36 +747,36 @@
             (Inttriplefunc.onefunc(one, Rat.rat_of_int ~1))
   val ass =
     Inttriplefunc.combine (curry op +/) (K false) 
-    (Inttriplefunc.mapf (tri_equation_eval vfn) assigs) vfn 
+    (Inttriplefunc.map (tri_equation_eval vfn) assigs) vfn 
  in if forall (fn e => tri_equation_eval ass e =/ rat_0) eqs
-    then Inttriplefunc.undefine one ass else raise Sanity
+    then Inttriplefunc.delete_safe one ass else raise Sanity
  end;
 
 (* Multiply equation-parametrized poly by regular poly and add accumulator.  *)
 
 fun tri_epoly_pmul p q acc =
- Monomialfunc.fold (fn (m1, c) => fn a =>
-  Monomialfunc.fold (fn (m2,e) => fn b =>
+ FuncUtil.Monomialfunc.fold (fn (m1, c) => fn a =>
+  FuncUtil.Monomialfunc.fold (fn (m2,e) => fn b =>
    let val m =  monomial_mul m1 m2
-       val es = Monomialfunc.tryapplyd b m Inttriplefunc.undefined 
-   in Monomialfunc.update (m,tri_equation_add (tri_equation_cmul c e) es) b 
+       val es = FuncUtil.Monomialfunc.tryapplyd b m Inttriplefunc.empty 
+   in FuncUtil.Monomialfunc.update (m,tri_equation_add (tri_equation_cmul c e) es) b 
    end) q a) p acc ;
 
 (* Usual operations on equation-parametrized poly.                           *)
 
 fun tri_epoly_cmul c l =
-  if c =/ rat_0 then Inttriplefunc.undefined else Inttriplefunc.mapf (tri_equation_cmul c) l;;
+  if c =/ rat_0 then Inttriplefunc.empty else Inttriplefunc.map (tri_equation_cmul c) l;;
 
 val tri_epoly_neg = tri_epoly_cmul (Rat.rat_of_int ~1);
 
-val tri_epoly_add = Inttriplefunc.combine tri_equation_add Inttriplefunc.is_undefined;
+val tri_epoly_add = Inttriplefunc.combine tri_equation_add Inttriplefunc.is_empty;
 
 fun tri_epoly_sub p q = tri_epoly_add p (tri_epoly_neg q);;
 
 (* Stuff for "equations" ((int*int)->num functions).                         *)
 
 fun pi_equation_cmul c eq =
-  if c =/ rat_0 then Inttriplefunc.undefined else Inttriplefunc.mapf (fn d => c */ d) eq;
+  if c =/ rat_0 then Inttriplefunc.empty else Inttriplefunc.map (fn d => c */ d) eq;
 
 fun pi_equation_add eq1 eq2 = Inttriplefunc.combine (curry op +/) (fn x => x =/ rat_0) eq1 eq2;
 
@@ -803,25 +795,25 @@
  | h::t => if p h then (h,t) else
           let val (k,s) = extract_first p t in (k,h::s) end
 fun eliminate vars dun eqs = case vars of 
-  [] => if forall Inttriplefunc.is_undefined eqs then dun
+  [] => if forall Inttriplefunc.is_empty eqs then dun
         else raise Unsolvable
  | v::vs =>
    let 
     val (eq,oeqs) = extract_first (fn e => Inttriplefunc.defined e v) eqs 
     val a = Inttriplefunc.apply eq v
-    val eq' = pi_equation_cmul ((Rat.neg rat_1) // a) (Inttriplefunc.undefine v eq)
+    val eq' = pi_equation_cmul ((Rat.neg rat_1) // a) (Inttriplefunc.delete_safe v eq)
     fun elim e =
      let val b = Inttriplefunc.tryapplyd e v rat_0 
      in if b =/ rat_0 then e else
         pi_equation_add e (pi_equation_cmul (Rat.neg b // a) eq)
      end
-   in eliminate vs (Inttriplefunc.update (v,eq') (Inttriplefunc.mapf elim dun)) (map elim oeqs)
+   in eliminate vs (Inttriplefunc.update (v,eq') (Inttriplefunc.map elim dun)) (map elim oeqs)
    end
   handle Failure _ => eliminate vs dun eqs
 in
 fun pi_eliminate_equations one vars eqs =
  let 
-  val assig = eliminate vars Inttriplefunc.undefined eqs
+  val assig = eliminate vars Inttriplefunc.empty eqs
   val vs = Inttriplefunc.fold (fn (x, f) => fn a => remove (dest_ord triple_int_ord) one (Inttriplefunc.dom f) @ a) assig []
   in (distinct (dest_ord triple_int_ord) vs, assig)
   end
@@ -834,8 +826,8 @@
   fun choose_variable eq =
    let val (v,_) = Inttriplefunc.choose eq 
    in if is_equal (triple_int_ord(v,one)) then
-      let val eq' = Inttriplefunc.undefine v eq 
-      in if Inttriplefunc.is_undefined eq' then error "choose_variable" 
+      let val eq' = Inttriplefunc.delete_safe v eq 
+      in if Inttriplefunc.is_empty eq' then error "choose_variable" 
          else fst (Inttriplefunc.choose eq')
       end
     else v 
@@ -843,22 +835,22 @@
   fun eliminate dun eqs = case eqs of 
     [] => dun
   | eq::oeqs =>
-    if Inttriplefunc.is_undefined eq then eliminate dun oeqs else
+    if Inttriplefunc.is_empty eq then eliminate dun oeqs else
     let val v = choose_variable eq
         val a = Inttriplefunc.apply eq v
         val eq' = pi_equation_cmul ((Rat.rat_of_int ~1) // a) 
-                   (Inttriplefunc.undefine v eq)
+                   (Inttriplefunc.delete_safe v eq)
         fun elim e =
          let val b = Inttriplefunc.tryapplyd e v rat_0 
          in if b =/ rat_0 then e 
             else pi_equation_add e (pi_equation_cmul (Rat.neg b // a) eq)
          end
-    in eliminate (Inttriplefunc.update(v, eq') (Inttriplefunc.mapf elim dun)) 
+    in eliminate (Inttriplefunc.update(v, eq') (Inttriplefunc.map elim dun)) 
                  (map elim oeqs) 
     end
 in fn eqs =>
  let 
-  val assig = eliminate Inttriplefunc.undefined eqs
+  val assig = eliminate Inttriplefunc.empty eqs
   val vs = Inttriplefunc.fold (fn (x, f) => fn a => remove (dest_ord triple_int_ord) one (Inttriplefunc.dom f) @ a) assig []
  in (distinct (dest_ord triple_int_ord) vs,assig)
  end
@@ -873,29 +865,29 @@
             (Inttriplefunc.onefunc(one, Rat.rat_of_int ~1))
   val ass =
     Inttriplefunc.combine (curry op +/) (K false) 
-    (Inttriplefunc.mapf (pi_equation_eval vfn) assigs) vfn 
+    (Inttriplefunc.map (pi_equation_eval vfn) assigs) vfn 
  in if forall (fn e => pi_equation_eval ass e =/ rat_0) eqs
-    then Inttriplefunc.undefine one ass else raise Sanity
+    then Inttriplefunc.delete_safe one ass else raise Sanity
  end;
 
 (* Multiply equation-parametrized poly by regular poly and add accumulator.  *)
 
 fun pi_epoly_pmul p q acc =
- Monomialfunc.fold (fn (m1, c) => fn a =>
-  Monomialfunc.fold (fn (m2,e) => fn b =>
+ FuncUtil.Monomialfunc.fold (fn (m1, c) => fn a =>
+  FuncUtil.Monomialfunc.fold (fn (m2,e) => fn b =>
    let val m =  monomial_mul m1 m2
-       val es = Monomialfunc.tryapplyd b m Inttriplefunc.undefined 
-   in Monomialfunc.update (m,pi_equation_add (pi_equation_cmul c e) es) b 
+       val es = FuncUtil.Monomialfunc.tryapplyd b m Inttriplefunc.empty 
+   in FuncUtil.Monomialfunc.update (m,pi_equation_add (pi_equation_cmul c e) es) b 
    end) q a) p acc ;
 
 (* Usual operations on equation-parametrized poly.                           *)
 
 fun pi_epoly_cmul c l =
-  if c =/ rat_0 then Inttriplefunc.undefined else Inttriplefunc.mapf (pi_equation_cmul c) l;;
+  if c =/ rat_0 then Inttriplefunc.empty else Inttriplefunc.map (pi_equation_cmul c) l;;
 
 val pi_epoly_neg = pi_epoly_cmul (Rat.rat_of_int ~1);
 
-val pi_epoly_add = Inttriplefunc.combine pi_equation_add Inttriplefunc.is_undefined;
+val pi_epoly_add = Inttriplefunc.combine pi_equation_add Inttriplefunc.is_empty;
 
 fun pi_epoly_sub p q = pi_epoly_add p (pi_epoly_neg q);;
 
@@ -914,27 +906,27 @@
 
 local
 fun diagonalize n i m =
- if Intpairfunc.is_undefined (snd m) then [] 
+ if FuncUtil.Intpairfunc.is_empty (snd m) then [] 
  else
-  let val a11 = Intpairfunc.tryapplyd (snd m) (i,i) rat_0 
+  let val a11 = FuncUtil.Intpairfunc.tryapplyd (snd m) (i,i) rat_0 
   in if a11 </ rat_0 then raise Failure "diagonalize: not PSD"
     else if a11 =/ rat_0 then
-          if Intfunc.is_undefined (snd (row i m)) then diagonalize n (i + 1) m
+          if FuncUtil.Intfunc.is_empty (snd (row i m)) then diagonalize n (i + 1) m
           else raise Failure "diagonalize: not PSD ___ "
     else
      let 
       val v = row i m
-      val v' = (fst v, Intfunc.fold (fn (i, c) => fn a => 
+      val v' = (fst v, FuncUtil.Intfunc.fold (fn (i, c) => fn a => 
        let val y = c // a11 
-       in if y = rat_0 then a else Intfunc.update (i,y) a 
-       end)  (snd v) Intfunc.undefined)
-      fun upt0 x y a = if y = rat_0 then a else Intpairfunc.update (x,y) a
+       in if y = rat_0 then a else FuncUtil.Intfunc.update (i,y) a 
+       end)  (snd v) FuncUtil.Intfunc.empty)
+      fun upt0 x y a = if y = rat_0 then a else FuncUtil.Intpairfunc.update (x,y) a
       val m' =
       ((n,n),
       iter (i+1,n) (fn j =>
           iter (i+1,n) (fn k =>
-              (upt0 (j,k) (Intpairfunc.tryapplyd (snd m) (j,k) rat_0 -/ Intfunc.tryapplyd (snd v) j rat_0 */ Intfunc.tryapplyd (snd v') k rat_0))))
-          Intpairfunc.undefined)
+              (upt0 (j,k) (FuncUtil.Intpairfunc.tryapplyd (snd m) (j,k) rat_0 -/ FuncUtil.Intfunc.tryapplyd (snd v) j rat_0 */ FuncUtil.Intfunc.tryapplyd (snd v') k rat_0))))
+          FuncUtil.Intpairfunc.empty)
      in (a11,v')::diagonalize n (i + 1) m' 
      end
   end
@@ -953,14 +945,14 @@
 (* Adjust a diagonalization to collect rationals at the start.               *)
   (* FIXME : Potentially polymorphic keys, but here only: integers!! *)
 local
- fun upd0 x y a = if y =/ rat_0 then a else Intfunc.update(x,y) a;
+ fun upd0 x y a = if y =/ rat_0 then a else FuncUtil.Intfunc.update(x,y) a;
  fun mapa f (d,v) = 
-  (d, Intfunc.fold (fn (i,c) => fn a => upd0 i (f c) a) v Intfunc.undefined)
+  (d, FuncUtil.Intfunc.fold (fn (i,c) => fn a => upd0 i (f c) a) v FuncUtil.Intfunc.empty)
  fun adj (c,l) =
  let val a = 
-  Intfunc.fold (fn (i,c) => fn a => lcm_rat a (denominator_rat c)) 
+  FuncUtil.Intfunc.fold (fn (i,c) => fn a => lcm_rat a (denominator_rat c)) 
     (snd l) rat_1 //
-  Intfunc.fold (fn (i,c) => fn a => gcd_rat a (numerator_rat c)) 
+  FuncUtil.Intfunc.fold (fn (i,c) => fn a => gcd_rat a (numerator_rat c)) 
     (snd l) rat_0
   in ((c // (a */ a)),mapa (fn x => a */ x) l)
   end
@@ -977,39 +969,35 @@
 
 fun enumerate_monomials d vars = 
  if d < 0 then []
- else if d = 0 then [Ctermfunc.undefined]
+ else if d = 0 then [FuncUtil.Ctermfunc.empty]
  else if null vars then [monomial_1] else
  let val alts =
   map (fn k => let val oths = enumerate_monomials (d - k) (tl vars) 
-               in map (fn ks => if k = 0 then ks else Ctermfunc.update (hd vars, k) ks) oths end) (0 upto d) 
- in foldr1 op @ alts
+               in map (fn ks => if k = 0 then ks else FuncUtil.Ctermfunc.update (hd vars, k) ks) oths end) (0 upto d) 
+ in flat alts
  end;
 
 (* Enumerate products of distinct input polys with degree <= d.              *)
 (* We ignore any constant input polynomials.                                 *)
 (* Give the output polynomial and a record of how it was derived.            *)
 
-local
- open RealArith
-in
 fun enumerate_products d pols =
-if d = 0 then [(poly_const rat_1,Rational_lt rat_1)] 
+if d = 0 then [(poly_const rat_1,RealArith.Rational_lt rat_1)] 
 else if d < 0 then [] else
 case pols of 
-   [] => [(poly_const rat_1,Rational_lt rat_1)]
+   [] => [(poly_const rat_1,RealArith.Rational_lt rat_1)]
  | (p,b)::ps => 
     let val e = multidegree p 
     in if e = 0 then enumerate_products d ps else
        enumerate_products d ps @
-       map (fn (q,c) => (poly_mul p q,Product(b,c)))
+       map (fn (q,c) => (poly_mul p q,RealArith.Product(b,c)))
          (enumerate_products (d - e) ps)
     end
-end;
 
 (* Convert regular polynomial. Note that we treat (0,0,0) as -1.             *)
 
 fun epoly_of_poly p =
-  Monomialfunc.fold (fn (m,c) => fn a => Monomialfunc.update (m, Inttriplefunc.onefunc ((0,0,0), Rat.neg c)) a) p Monomialfunc.undefined;
+  FuncUtil.Monomialfunc.fold (fn (m,c) => fn a => FuncUtil.Monomialfunc.update (m, Inttriplefunc.onefunc ((0,0,0), Rat.neg c)) a) p FuncUtil.Monomialfunc.empty;
 
 (* String for block diagonal matrix numbered k.                              *)
 
@@ -1020,7 +1008,7 @@
     Inttriplefunc.fold 
       (fn ((b,i,j),c) => fn a => if i > j then a else ((b,i,j),c)::a) 
       m [] 
-  val entss = sort (increasing fst triple_int_ord) ents 
+  val entss = sort (triple_int_ord o pairself fst) ents 
  in fold_rev (fn ((b,i,j),c) => fn a =>
      pfx ^ string_of_int b ^ " " ^ string_of_int i ^ " " ^ string_of_int j ^
      " " ^ decimalize 20 c ^ "\n" ^ a) entss ""
@@ -1033,7 +1021,7 @@
  in
   string_of_int m ^ "\n" ^
   string_of_int nblocks ^ "\n" ^
-  (foldr1 (fn (s, t) => s^" "^t) (map string_of_int blocksizes)) ^
+  (space_implode " " (map string_of_int blocksizes)) ^
   "\n" ^
   sdpa_of_vector obj ^
   fold_rev2 (fn k => fn m => fn a => sdpa_of_blockdiagonal (k - 1) m ^ a)
@@ -1049,8 +1037,8 @@
 
 val bmatrix_add = Inttriplefunc.combine (curry op +/) (fn x => x =/ rat_0);
 fun bmatrix_cmul c bm =
-  if c =/ rat_0 then Inttriplefunc.undefined
-  else Inttriplefunc.mapf (fn x => c */ x) bm;
+  if c =/ rat_0 then Inttriplefunc.empty
+  else Inttriplefunc.map (fn x => c */ x) bm;
 
 val bmatrix_neg = bmatrix_cmul (Rat.rat_of_int ~1);
 fun bmatrix_sub m1 m2 = bmatrix_add m1 (bmatrix_neg m2);;
@@ -1060,8 +1048,8 @@
 fun blocks blocksizes bm =
  map (fn (bs,b0) =>
       let val m = Inttriplefunc.fold
-          (fn ((b,i,j),c) => fn a => if b = b0 then Intpairfunc.update ((i,j),c) a else a) bm Intpairfunc.undefined
-          val d = Intpairfunc.fold (fn ((i,j),c) => fn a => max a (max i j)) m 0 
+          (fn ((b,i,j),c) => fn a => if b = b0 then FuncUtil.Intpairfunc.update ((i,j),c) a else a) bm FuncUtil.Intpairfunc.empty
+          val d = FuncUtil.Intpairfunc.fold (fn ((i,j),c) => fn a => max a (max i j)) m 0 
       in (((bs,bs),m):matrix) end)
  (blocksizes ~~ (1 upto length blocksizes));;
 
@@ -1076,15 +1064,12 @@
 (* Positiv- and Nullstellensatz. Flag "linf" forces a linear representation. *)
 
  
-local
- open RealArith
-in
 fun real_positivnullstellensatz_general prover linf d eqs leqs pol =
 let 
  val vars = fold_rev (curry (gen_union (op aconvc)) o poly_variables) 
               (pol::eqs @ map fst leqs) []
  val monoid = if linf then 
-      (poly_const rat_1,Rational_lt rat_1)::
+      (poly_const rat_1,RealArith.Rational_lt rat_1)::
       (filter (fn (p,c) => multidegree p <= d) leqs)
     else enumerate_products d leqs
  val nblocks = length monoid
@@ -1094,7 +1079,7 @@
    val mons = enumerate_monomials e vars
    val nons = mons ~~ (1 upto length mons) 
   in (mons,
-      fold_rev (fn (m,n) => Monomialfunc.update(m,Inttriplefunc.onefunc((~k,~n,n),rat_1))) nons Monomialfunc.undefined)
+      fold_rev (fn (m,n) => FuncUtil.Monomialfunc.update(m,Inttriplefunc.onefunc((~k,~n,n),rat_1))) nons FuncUtil.Monomialfunc.empty)
   end
 
  fun mk_sqmultiplier k (p,c) =
@@ -1108,11 +1093,11 @@
         let val m = monomial_mul m1 m2 
         in if n1 > n2 then a else
           let val c = if n1 = n2 then rat_1 else rat_2
-              val e = Monomialfunc.tryapplyd a m Inttriplefunc.undefined 
-          in Monomialfunc.update(m, tri_equation_add (Inttriplefunc.onefunc((k,n1,n2), c)) e) a
+              val e = FuncUtil.Monomialfunc.tryapplyd a m Inttriplefunc.empty 
+          in FuncUtil.Monomialfunc.update(m, tri_equation_add (Inttriplefunc.onefunc((k,n1,n2), c)) e) a
           end
         end)  nons)
-       nons Monomialfunc.undefined)
+       nons FuncUtil.Monomialfunc.empty)
   end
 
   val (sqmonlist,sqs) = split_list (map2 mk_sqmultiplier (1 upto length monoid) monoid)
@@ -1122,7 +1107,7 @@
     fold_rev2 (fn p => fn q => fn a => tri_epoly_pmul p q a) eqs ids
             (fold_rev2 (fn (p,c) => fn s => fn a => tri_epoly_pmul p s a) monoid sqs
                      (epoly_of_poly(poly_neg pol)))
-  val eqns = Monomialfunc.fold (fn (m,e) => fn a => e::a) bigsum []
+  val eqns = FuncUtil.Monomialfunc.fold (fn (m,e) => fn a => e::a) bigsum []
   val (pvs,assig) = tri_eliminate_all_equations (0,0,0) eqns
   val qvars = (0,0,0)::pvs
   val allassig = fold_rev (fn v => Inttriplefunc.update(v,(Inttriplefunc.onefunc(v,rat_1)))) pvs assig
@@ -1133,19 +1118,19 @@
          in if c = rat_0 then m else
             Inttriplefunc.update ((b,j,i), c) (Inttriplefunc.update ((b,i,j), c) m)
          end)
-          allassig Inttriplefunc.undefined
+          allassig Inttriplefunc.empty
   val diagents = Inttriplefunc.fold
     (fn ((b,i,j), e) => fn a => if b > 0 andalso i = j then tri_equation_add e a else a)
-    allassig Inttriplefunc.undefined
+    allassig Inttriplefunc.empty
 
   val mats = map mk_matrix qvars
   val obj = (length pvs,
-            itern 1 pvs (fn v => fn i => Intfunc.updatep iszero (i,Inttriplefunc.tryapplyd diagents v rat_0))
-                        Intfunc.undefined)
+            itern 1 pvs (fn v => fn i => FuncUtil.Intfunc.updatep iszero (i,Inttriplefunc.tryapplyd diagents v rat_0))
+                        FuncUtil.Intfunc.empty)
   val raw_vec = if null pvs then vector_0 0
                 else tri_scale_then (run_blockproblem prover nblocks blocksizes) obj mats
-  fun int_element (d,v) i = Intfunc.tryapplyd v i rat_0
-  fun cterm_element (d,v) i = Ctermfunc.tryapplyd v i rat_0
+  fun int_element (d,v) i = FuncUtil.Intfunc.tryapplyd v i rat_0
+  fun cterm_element (d,v) i = FuncUtil.Ctermfunc.tryapplyd v i rat_0
 
   fun find_rounding d =
    let 
@@ -1169,12 +1154,12 @@
   val finalassigs =
     Inttriplefunc.fold (fn (v,e) => fn a => Inttriplefunc.update(v, tri_equation_eval newassigs e) a) allassig newassigs
   fun poly_of_epoly p =
-    Monomialfunc.fold (fn (v,e) => fn a => Monomialfunc.updatep iszero (v,tri_equation_eval finalassigs e) a)
-          p Monomialfunc.undefined
+    FuncUtil.Monomialfunc.fold (fn (v,e) => fn a => FuncUtil.Monomialfunc.updatep iszero (v,tri_equation_eval finalassigs e) a)
+          p FuncUtil.Monomialfunc.empty
   fun  mk_sos mons =
    let fun mk_sq (c,m) =
-    (c,fold_rev (fn k=> fn a => Monomialfunc.updatep iszero (nth mons (k - 1), int_element m k) a)
-                 (1 upto length mons) Monomialfunc.undefined)
+    (c,fold_rev (fn k=> fn a => FuncUtil.Monomialfunc.updatep iszero (nth mons (k - 1), int_element m k) a)
+                 (1 upto length mons) FuncUtil.Monomialfunc.empty)
    in map mk_sq
    end
   val sqs = map2 mk_sos sqmonlist ratdias
@@ -1186,13 +1171,11 @@
            (fold_rev2 (fn p => fn q => poly_add (poly_mul p q)) cfs eqs
                     (poly_neg pol))
 
-in if not(Monomialfunc.is_undefined sanity) then raise Sanity else
+in if not(FuncUtil.Monomialfunc.is_empty sanity) then raise Sanity else
   (cfs,map (fn (a,b) => (snd a,b)) msq)
  end
 
 
-end;
-
 (* Iterative deepening.                                                      *)
 
 fun deepen f n = 
@@ -1201,21 +1184,15 @@
 
 (* Map back polynomials and their composites to a positivstellensatz.        *)
 
-local
- open Thm Numeral RealArith
-in
-
-fun cterm_of_sqterm (c,p) = Product(Rational_lt c,Square p);
+fun cterm_of_sqterm (c,p) = RealArith.Product(RealArith.Rational_lt c,RealArith.Square p);
 
 fun cterm_of_sos (pr,sqs) = if null sqs then pr
-  else Product(pr,foldr1 (fn (a, b) => Sum(a,b)) (map cterm_of_sqterm sqs));
-
-end
+  else RealArith.Product(pr,foldr1 RealArith.Sum (map cterm_of_sqterm sqs));
 
 (* Interface to HOL.                                                         *)
 local
-  open Thm Conv RealArith
-  val concl = dest_arg o cprop_of
+  open Conv
+  val concl = Thm.dest_arg o cprop_of
   fun simple_cterm_ord t u = TermOrd.fast_term_ord (term_of t, term_of u) = LESS
 in
   (* FIXME: Replace tryfind by get_first !! *)
@@ -1228,37 +1205,37 @@
        real_poly_pow_conv,real_poly_sub_conv,real_poly_conv) = (add,mul,neg,pow,sub,main)
   fun mainf cert_choice translator (eqs,les,lts) = 
   let 
-   val eq0 = map (poly_of_term o dest_arg1 o concl) eqs
-   val le0 = map (poly_of_term o dest_arg o concl) les
-   val lt0 = map (poly_of_term o dest_arg o concl) lts
-   val eqp0 = map (fn (t,i) => (t,Axiom_eq i)) (eq0 ~~ (0 upto (length eq0 - 1)))
-   val lep0 = map (fn (t,i) => (t,Axiom_le i)) (le0 ~~ (0 upto (length le0 - 1)))
-   val ltp0 = map (fn (t,i) => (t,Axiom_lt i)) (lt0 ~~ (0 upto (length lt0 - 1)))
+   val eq0 = map (poly_of_term o Thm.dest_arg1 o concl) eqs
+   val le0 = map (poly_of_term o Thm.dest_arg o concl) les
+   val lt0 = map (poly_of_term o Thm.dest_arg o concl) lts
+   val eqp0 = map (fn (t,i) => (t,RealArith.Axiom_eq i)) (eq0 ~~ (0 upto (length eq0 - 1)))
+   val lep0 = map (fn (t,i) => (t,RealArith.Axiom_le i)) (le0 ~~ (0 upto (length le0 - 1)))
+   val ltp0 = map (fn (t,i) => (t,RealArith.Axiom_lt i)) (lt0 ~~ (0 upto (length lt0 - 1)))
    val (keq,eq) = List.partition (fn (p,_) => multidegree p = 0) eqp0
    val (klep,lep) = List.partition (fn (p,_) => multidegree p = 0) lep0
    val (kltp,ltp) = List.partition (fn (p,_) => multidegree p = 0) ltp0
    fun trivial_axiom (p,ax) =
     case ax of
-       Axiom_eq n => if eval Ctermfunc.undefined p <>/ Rat.zero then nth eqs n 
+       RealArith.Axiom_eq n => if eval FuncUtil.Ctermfunc.empty p <>/ Rat.zero then nth eqs n 
                      else raise Failure "trivial_axiom: Not a trivial axiom"
-     | Axiom_le n => if eval Ctermfunc.undefined p </ Rat.zero then nth les n 
+     | RealArith.Axiom_le n => if eval FuncUtil.Ctermfunc.empty p </ Rat.zero then nth les n 
                      else raise Failure "trivial_axiom: Not a trivial axiom"
-     | Axiom_lt n => if eval Ctermfunc.undefined p <=/ Rat.zero then nth lts n 
+     | RealArith.Axiom_lt n => if eval FuncUtil.Ctermfunc.empty p <=/ Rat.zero then nth lts n 
                      else raise Failure "trivial_axiom: Not a trivial axiom"
      | _ => error "trivial_axiom: Not a trivial axiom"
    in 
   (let val th = tryfind trivial_axiom (keq @ klep @ kltp)
    in
-    (fconv_rule (arg_conv (arg1_conv real_poly_conv) then_conv field_comp_conv) th, Trivial)
+    (fconv_rule (arg_conv (arg1_conv real_poly_conv) then_conv field_comp_conv) th, RealArith.Trivial)
    end)
    handle Failure _ => 
      (let val proof =
        (case proof_method of Certificate certs =>
          (* choose certificate *)
          let
-           fun chose_cert [] (Cert c) = c
-             | chose_cert (Left::s) (Branch (l, _)) = chose_cert s l
-             | chose_cert (Right::s) (Branch (_, r)) = chose_cert s r
+           fun chose_cert [] (RealArith.Cert c) = c
+             | chose_cert (RealArith.Left::s) (RealArith.Branch (l, _)) = chose_cert s l
+             | chose_cert (RealArith.Right::s) (RealArith.Branch (_, r)) = chose_cert s r
              | chose_cert _ _ = error "certificate tree in invalid form"
          in
            chose_cert cert_choice certs
@@ -1278,17 +1255,17 @@
            end
          val (d,i,(cert_ideal,cert_cone)) = deepen tryall 0
          val proofs_ideal =
-           map2 (fn q => fn (p,ax) => Eqmul(q,ax)) cert_ideal eq
+           map2 (fn q => fn (p,ax) => RealArith.Eqmul(q,ax)) cert_ideal eq
          val proofs_cone = map cterm_of_sos cert_cone
-         val proof_ne = if null ltp then Rational_lt Rat.one else
-           let val p = foldr1 (fn (s, t) => Product(s,t)) (map snd ltp) 
-           in  funpow i (fn q => Product(p,q)) (Rational_lt Rat.one)
+         val proof_ne = if null ltp then RealArith.Rational_lt Rat.one else
+           let val p = foldr1 RealArith.Product (map snd ltp) 
+           in  funpow i (fn q => RealArith.Product(p,q)) (RealArith.Rational_lt Rat.one)
            end
          in 
-           foldr1 (fn (s, t) => Sum(s,t)) (proof_ne :: proofs_ideal @ proofs_cone) 
+           foldr1 RealArith.Sum (proof_ne :: proofs_ideal @ proofs_cone) 
          end)
      in
-        (translator (eqs,les,lts) proof, Cert proof)
+        (translator (eqs,les,lts) proof, RealArith.Cert proof)
      end)
    end
  in mainf end
@@ -1305,9 +1282,9 @@
 (* A wrapper that tries to substitute away variables first.                  *)
 
 local
- open Thm Conv RealArith
+ open Conv
   fun simple_cterm_ord t u = TermOrd.fast_term_ord (term_of t, term_of u) = LESS
- val concl = dest_arg o cprop_of
+ val concl = Thm.dest_arg o cprop_of
  val shuffle1 = 
    fconv_rule (rewr_conv @{lemma "(a + x == y) == (x == y - (a::real))" by (atomize (full)) (simp add: ring_simps) })
  val shuffle2 =
@@ -1316,19 +1293,19 @@
     Free(_,@{typ real}) => if not (member (op aconvc) fvs tm) then (Rat.one,tm) 
                            else raise Failure "substitutable_monomial"
   | @{term "op * :: real => _"}$c$(t as Free _ ) => 
-     if is_ratconst (dest_arg1 tm) andalso not (member (op aconvc) fvs (dest_arg tm))
-         then (dest_ratconst (dest_arg1 tm),dest_arg tm) else raise Failure "substitutable_monomial"
+     if RealArith.is_ratconst (Thm.dest_arg1 tm) andalso not (member (op aconvc) fvs (Thm.dest_arg tm))
+         then (RealArith.dest_ratconst (Thm.dest_arg1 tm),Thm.dest_arg tm) else raise Failure "substitutable_monomial"
   | @{term "op + :: real => _"}$s$t => 
-       (substitutable_monomial (add_cterm_frees (dest_arg tm) fvs) (dest_arg1 tm)
-        handle Failure _ => substitutable_monomial (add_cterm_frees (dest_arg1 tm) fvs) (dest_arg tm))
+       (substitutable_monomial (Thm.add_cterm_frees (Thm.dest_arg tm) fvs) (Thm.dest_arg1 tm)
+        handle Failure _ => substitutable_monomial (Thm.add_cterm_frees (Thm.dest_arg1 tm) fvs) (Thm.dest_arg tm))
   | _ => raise Failure "substitutable_monomial"
 
   fun isolate_variable v th = 
-   let val w = dest_arg1 (cprop_of th)
+   let val w = Thm.dest_arg1 (cprop_of th)
    in if v aconvc w then th
       else case term_of w of
            @{term "op + :: real => _"}$s$t => 
-              if dest_arg1 w aconvc v then shuffle2 th 
+              if Thm.dest_arg1 w aconvc v then shuffle2 th 
               else isolate_variable v (shuffle1 th)
           | _ => error "isolate variable : This should not happen?"
    end 
@@ -1345,8 +1322,8 @@
 
   fun make_substitution th =
    let 
-    val (c,v) = substitutable_monomial [] (dest_arg1(concl th))
-    val th1 = Drule.arg_cong_rule (capply @{cterm "op * :: real => _"} (cterm_of_rat (Rat.inv c))) (mk_meta_eq th)
+    val (c,v) = substitutable_monomial [] (Thm.dest_arg1(concl th))
+    val th1 = Drule.arg_cong_rule (Thm.capply @{cterm "op * :: real => _"} (RealArith.cterm_of_rat (Rat.inv c))) (mk_meta_eq th)
     val th2 = fconv_rule (binop_conv real_poly_mul_conv) th1
    in fconv_rule (arg_conv real_poly_conv) (isolate_variable v th2)
    end
@@ -1378,18 +1355,9 @@
 (* Overall function. *)
 
 fun real_sos prover ctxt =
-  gen_prover_real_arith ctxt (real_nonlinear_subst_prover prover ctxt)
+  RealArith.gen_prover_real_arith ctxt (real_nonlinear_subst_prover prover ctxt)
 end;
 
-(* A tactic *)
-fun strip_all ct = 
- case term_of ct of 
-  Const("all",_) $ Abs (xn,xT,p) => 
-   let val (a,(v,t')) = (apsnd (Thm.dest_abs (SOME xn)) o Thm.dest_comb) ct
-   in apfst (cons v) (strip_all t')
-   end
-| _ => ([],ct)
-
 val known_sos_constants = 
   [@{term "op ==>"}, @{term "Trueprop"}, 
    @{term "op -->"}, @{term "op &"}, @{term "op |"}, 
@@ -1424,13 +1392,17 @@
               else error ("SOSO: Unknown constants in Subgoal:" ^ commas (map fst ukcs))
 in () end
 
-fun core_sos_tac print_cert prover ctxt = CSUBGOAL (fn (ct, i) => 
-  let val _ = check_sos known_sos_constants ct
-      val (avs, p) = strip_all ct
-      val (ths, certificates) = real_sos prover ctxt (Thm.dest_arg p)
-      val th = standard (fold_rev forall_intr avs ths)
-      val _ = print_cert certificates
-  in rtac th i end);
+fun core_sos_tac print_cert prover ctxt i = 
+  let
+    fun core_tac {concl, context, ...} =
+      let
+        val _ = check_sos known_sos_constants concl
+        val (ths, certificates) = real_sos prover context (Thm.dest_arg concl)
+        val _ = print_cert certificates
+      in rtac ths i end
+  in
+    SUBPROOF core_tac ctxt i
+  end
 
 fun default_SOME f NONE v = SOME v
   | default_SOME f (SOME v) _ = SOME v;
--- a/src/HOL/Library/normarith.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/HOL/Library/normarith.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -15,7 +15,7 @@
 structure NormArith : NORM_ARITH = 
 struct
 
- open Conv Thm FuncUtil;
+ open Conv;
  val bool_eq = op = : bool *bool -> bool
   fun dest_ratconst t = case term_of t of
    Const(@{const_name divide}, _)$a$b => Rat.rat_of_quotient(HOLogic.dest_number a |> snd, HOLogic.dest_number b |> snd)
@@ -23,50 +23,50 @@
  | _ => Rat.rat_of_int (HOLogic.dest_number (term_of t) |> snd)
  fun is_ratconst t = can dest_ratconst t
  fun augment_norm b t acc = case term_of t of 
-     Const(@{const_name norm}, _) $ _ => insert (eq_pair bool_eq (op aconvc)) (b,dest_arg t) acc
+     Const(@{const_name norm}, _) $ _ => insert (eq_pair bool_eq (op aconvc)) (b,Thm.dest_arg t) acc
    | _ => acc
  fun find_normedterms t acc = case term_of t of
     @{term "op + :: real => _"}$_$_ =>
-            find_normedterms (dest_arg1 t) (find_normedterms (dest_arg t) acc)
+            find_normedterms (Thm.dest_arg1 t) (find_normedterms (Thm.dest_arg t) acc)
       | @{term "op * :: real => _"}$_$n =>
-            if not (is_ratconst (dest_arg1 t)) then acc else
-            augment_norm (dest_ratconst (dest_arg1 t) >=/ Rat.zero) 
-                      (dest_arg t) acc
+            if not (is_ratconst (Thm.dest_arg1 t)) then acc else
+            augment_norm (dest_ratconst (Thm.dest_arg1 t) >=/ Rat.zero) 
+                      (Thm.dest_arg t) acc
       | _ => augment_norm true t acc 
 
- val cterm_lincomb_neg = Ctermfunc.mapf Rat.neg
+ val cterm_lincomb_neg = FuncUtil.Ctermfunc.map Rat.neg
  fun cterm_lincomb_cmul c t = 
-    if c =/ Rat.zero then Ctermfunc.undefined else Ctermfunc.mapf (fn x => x */ c) t
- fun cterm_lincomb_add l r = Ctermfunc.combine (curry op +/) (fn x => x =/ Rat.zero) l r
+    if c =/ Rat.zero then FuncUtil.Ctermfunc.empty else FuncUtil.Ctermfunc.map (fn x => x */ c) t
+ fun cterm_lincomb_add l r = FuncUtil.Ctermfunc.combine (curry op +/) (fn x => x =/ Rat.zero) l r
  fun cterm_lincomb_sub l r = cterm_lincomb_add l (cterm_lincomb_neg r)
- fun cterm_lincomb_eq l r = Ctermfunc.is_undefined (cterm_lincomb_sub l r)
+ fun cterm_lincomb_eq l r = FuncUtil.Ctermfunc.is_empty (cterm_lincomb_sub l r)
 
- val int_lincomb_neg = Intfunc.mapf Rat.neg
+ val int_lincomb_neg = FuncUtil.Intfunc.map Rat.neg
  fun int_lincomb_cmul c t = 
-    if c =/ Rat.zero then Intfunc.undefined else Intfunc.mapf (fn x => x */ c) t
- fun int_lincomb_add l r = Intfunc.combine (curry op +/) (fn x => x =/ Rat.zero) l r
+    if c =/ Rat.zero then FuncUtil.Intfunc.empty else FuncUtil.Intfunc.map (fn x => x */ c) t
+ fun int_lincomb_add l r = FuncUtil.Intfunc.combine (curry op +/) (fn x => x =/ Rat.zero) l r
  fun int_lincomb_sub l r = int_lincomb_add l (int_lincomb_neg r)
- fun int_lincomb_eq l r = Intfunc.is_undefined (int_lincomb_sub l r)
+ fun int_lincomb_eq l r = FuncUtil.Intfunc.is_empty (int_lincomb_sub l r)
 
 fun vector_lincomb t = case term_of t of 
    Const(@{const_name plus}, _) $ _ $ _ =>
-    cterm_lincomb_add (vector_lincomb (dest_arg1 t)) (vector_lincomb (dest_arg t))
+    cterm_lincomb_add (vector_lincomb (Thm.dest_arg1 t)) (vector_lincomb (Thm.dest_arg t))
  | Const(@{const_name minus}, _) $ _ $ _ =>
-    cterm_lincomb_sub (vector_lincomb (dest_arg1 t)) (vector_lincomb (dest_arg t))
+    cterm_lincomb_sub (vector_lincomb (Thm.dest_arg1 t)) (vector_lincomb (Thm.dest_arg t))
  | Const(@{const_name scaleR}, _)$_$_ =>
-    cterm_lincomb_cmul (dest_ratconst (dest_arg1 t)) (vector_lincomb (dest_arg t))
+    cterm_lincomb_cmul (dest_ratconst (Thm.dest_arg1 t)) (vector_lincomb (Thm.dest_arg t))
  | Const(@{const_name uminus}, _)$_ =>
-     cterm_lincomb_neg (vector_lincomb (dest_arg t))
+     cterm_lincomb_neg (vector_lincomb (Thm.dest_arg t))
 (* FIXME: how should we handle numerals?
  | Const(@ {const_name vec},_)$_ => 
    let 
-     val b = ((snd o HOLogic.dest_number o term_of o dest_arg) t = 0 
+     val b = ((snd o HOLogic.dest_number o term_of o Thm.dest_arg) t = 0 
                handle TERM _=> false)
-   in if b then Ctermfunc.onefunc (t,Rat.one)
-      else Ctermfunc.undefined
+   in if b then FuncUtil.Ctermfunc.onefunc (t,Rat.one)
+      else FuncUtil.Ctermfunc.empty
    end
 *)
- | _ => Ctermfunc.onefunc (t,Rat.one)
+ | _ => FuncUtil.Ctermfunc.onefunc (t,Rat.one)
 
  fun vector_lincombs ts =
   fold_rev 
@@ -82,35 +82,35 @@
 fun replacenegnorms cv t = case term_of t of 
   @{term "op + :: real => _"}$_$_ => binop_conv (replacenegnorms cv) t
 | @{term "op * :: real => _"}$_$_ => 
-    if dest_ratconst (dest_arg1 t) </ Rat.zero then arg_conv cv t else reflexive t
+    if dest_ratconst (Thm.dest_arg1 t) </ Rat.zero then arg_conv cv t else reflexive t
 | _ => reflexive t
 fun flip v eq = 
-  if Ctermfunc.defined eq v 
-  then Ctermfunc.update (v, Rat.neg (Ctermfunc.apply eq v)) eq else eq
+  if FuncUtil.Ctermfunc.defined eq v 
+  then FuncUtil.Ctermfunc.update (v, Rat.neg (FuncUtil.Ctermfunc.apply eq v)) eq else eq
 fun allsubsets s = case s of 
   [] => [[]]
 |(a::t) => let val res = allsubsets t in
                map (cons a) res @ res end
 fun evaluate env lin =
- Intfunc.fold (fn (x,c) => fn s => s +/ c */ (Intfunc.apply env x)) 
+ FuncUtil.Intfunc.fold (fn (x,c) => fn s => s +/ c */ (FuncUtil.Intfunc.apply env x)) 
    lin Rat.zero
 
 fun solve (vs,eqs) = case (vs,eqs) of
-  ([],[]) => SOME (Intfunc.onefunc (0,Rat.one))
+  ([],[]) => SOME (FuncUtil.Intfunc.onefunc (0,Rat.one))
  |(_,eq::oeqs) => 
-   (case filter (member (op =) vs) (Intfunc.dom eq) of (*FIXME use find_first here*)
+   (case filter (member (op =) vs) (FuncUtil.Intfunc.dom eq) of (*FIXME use find_first here*)
      [] => NONE
     | v::_ => 
-       if Intfunc.defined eq v 
+       if FuncUtil.Intfunc.defined eq v 
        then 
         let 
-         val c = Intfunc.apply eq v
+         val c = FuncUtil.Intfunc.apply eq v
          val vdef = int_lincomb_cmul (Rat.neg (Rat.inv c)) eq
-         fun eliminate eqn = if not (Intfunc.defined eqn v) then eqn 
-                             else int_lincomb_add (int_lincomb_cmul (Intfunc.apply eqn v) vdef) eqn
+         fun eliminate eqn = if not (FuncUtil.Intfunc.defined eqn v) then eqn 
+                             else int_lincomb_add (int_lincomb_cmul (FuncUtil.Intfunc.apply eqn v) vdef) eqn
         in (case solve (vs \ v,map eliminate oeqs) of
             NONE => NONE
-          | SOME soln => SOME (Intfunc.update (v, evaluate soln (Intfunc.undefine v vdef)) soln))
+          | SOME soln => SOME (FuncUtil.Intfunc.update (v, evaluate soln (FuncUtil.Intfunc.delete_safe v vdef)) soln))
         end
        else NONE)
 
@@ -130,7 +130,7 @@
  let 
   fun vertex cmb = case solve(vs,cmb) of
     NONE => NONE
-   | SOME soln => SOME (map (fn v => Intfunc.tryapplyd soln v Rat.zero) vs)
+   | SOME soln => SOME (map (fn v => FuncUtil.Intfunc.tryapplyd soln v Rat.zero) vs)
   val rawvs = map_filter vertex (combinations (length vs) eqs)
   val unset = filter (forall (fn c => c >=/ Rat.zero)) rawvs 
  in fold_rev (insert (uncurry (forall2 (curry op =/)))) unset [] 
@@ -265,28 +265,28 @@
  | fold_rev2 f _ _ _ = raise UnequalLengths;
 
 fun int_flip v eq = 
-  if Intfunc.defined eq v 
-  then Intfunc.update (v, Rat.neg (Intfunc.apply eq v)) eq else eq;
+  if FuncUtil.Intfunc.defined eq v 
+  then FuncUtil.Intfunc.update (v, Rat.neg (FuncUtil.Intfunc.apply eq v)) eq else eq;
 
 local
  val pth_zero = @{thm norm_zero}
- val tv_n = (ctyp_of_term o dest_arg o dest_arg1 o dest_arg o cprop_of)
+ val tv_n = (ctyp_of_term o Thm.dest_arg o Thm.dest_arg1 o Thm.dest_arg o cprop_of)
              pth_zero
- val concl = dest_arg o cprop_of 
+ val concl = Thm.dest_arg o cprop_of 
  fun real_vector_combo_prover ctxt translator (nubs,ges,gts) = 
   let 
    (* FIXME: Should be computed statically!!*)
    val real_poly_conv = 
       Normalizer.semiring_normalize_wrapper ctxt
        (valOf (NormalizerData.match ctxt @{cterm "(0::real) + 1"}))
-   val sources = map (dest_arg o dest_arg1 o concl) nubs
-   val rawdests = fold_rev (find_normedterms o dest_arg o concl) (ges @ gts) [] 
+   val sources = map (Thm.dest_arg o Thm.dest_arg1 o concl) nubs
+   val rawdests = fold_rev (find_normedterms o Thm.dest_arg o concl) (ges @ gts) [] 
    val _ = if not (forall fst rawdests) then error "real_vector_combo_prover: Sanity check" 
            else ()
    val dests = distinct (op aconvc) (map snd rawdests)
    val srcfuns = map vector_lincomb sources
    val destfuns = map vector_lincomb dests 
-   val vvs = fold_rev (curry (gen_union op aconvc) o Ctermfunc.dom) (srcfuns @ destfuns) []
+   val vvs = fold_rev (curry (gen_union op aconvc) o FuncUtil.Ctermfunc.dom) (srcfuns @ destfuns) []
    val n = length srcfuns
    val nvs = 1 upto n
    val srccombs = srcfuns ~~ nvs
@@ -294,12 +294,12 @@
     let 
      fun coefficients x =
       let 
-       val inp = if Ctermfunc.defined d x then Intfunc.onefunc (0, Rat.neg(Ctermfunc.apply d x))
-                      else Intfunc.undefined 
-      in fold_rev (fn (f,v) => fn g => if Ctermfunc.defined f x then Intfunc.update (v, Ctermfunc.apply f x) g else g) srccombs inp 
+       val inp = if FuncUtil.Ctermfunc.defined d x then FuncUtil.Intfunc.onefunc (0, Rat.neg(FuncUtil.Ctermfunc.apply d x))
+                      else FuncUtil.Intfunc.empty 
+      in fold_rev (fn (f,v) => fn g => if FuncUtil.Ctermfunc.defined f x then FuncUtil.Intfunc.update (v, FuncUtil.Ctermfunc.apply f x) g else g) srccombs inp 
       end
      val equations = map coefficients vvs
-     val inequalities = map (fn n => Intfunc.onefunc (n,Rat.one)) nvs
+     val inequalities = map (fn n => FuncUtil.Intfunc.onefunc (n,Rat.one)) nvs
      fun plausiblevertices f =
       let 
        val flippedequations = map (fold_rev int_flip f) equations
@@ -307,7 +307,7 @@
        val rawverts = vertices nvs constraints
        fun check_solution v =
         let 
-          val f = fold_rev2 (curry Intfunc.update) nvs v (Intfunc.onefunc (0, Rat.one))
+          val f = fold_rev2 (curry FuncUtil.Intfunc.update) nvs v (FuncUtil.Intfunc.onefunc (0, Rat.one))
         in forall (fn e => evaluate f e =/ Rat.zero) flippedequations
         end
        val goodverts = filter check_solution rawverts
@@ -328,7 +328,7 @@
    val ges' = map_filter (try compute_ineq) (fold_rev (append o consider) destfuns []) @
                  map (inequality_canon_rule ctxt) nubs @ ges
    val zerodests = filter
-        (fn t => null (Ctermfunc.dom (vector_lincomb t))) (map snd rawdests)
+        (fn t => null (FuncUtil.Ctermfunc.dom (vector_lincomb t))) (map snd rawdests)
 
   in fst (RealArith.real_linear_prover translator
         (map (fn t => instantiate ([(tv_n, ctyp_of_term t)],[]) pth_zero)
@@ -344,19 +344,19 @@
 local
  val pth = @{thm norm_imp_pos_and_ge}
  val norm_mp = match_mp pth
- val concl = dest_arg o cprop_of
+ val concl = Thm.dest_arg o cprop_of
  fun conjunct1 th = th RS @{thm conjunct1}
  fun conjunct2 th = th RS @{thm conjunct2}
  fun C f x y = f y x
 fun real_vector_ineq_prover ctxt translator (ges,gts) = 
  let 
 (*   val _ = error "real_vector_ineq_prover: pause" *)
-  val ntms = fold_rev find_normedterms (map (dest_arg o concl) (ges @ gts)) []
+  val ntms = fold_rev find_normedterms (map (Thm.dest_arg o concl) (ges @ gts)) []
   val lctab = vector_lincombs (map snd (filter (not o fst) ntms))
   val (fxns, ctxt') = Variable.variant_fixes (replicate (length lctab) "x") ctxt
   fun instantiate_cterm' ty tms = Drule.cterm_rule (Drule.instantiate' ty tms)
-  fun mk_norm t = capply (instantiate_cterm' [SOME (ctyp_of_term t)] [] @{cpat "norm :: (?'a :: real_normed_vector) => real"}) t
-  fun mk_equals l r = capply (capply (instantiate_cterm' [SOME (ctyp_of_term l)] [] @{cpat "op == :: ?'a =>_"}) l) r
+  fun mk_norm t = Thm.capply (instantiate_cterm' [SOME (ctyp_of_term t)] [] @{cpat "norm :: (?'a :: real_normed_vector) => real"}) t
+  fun mk_equals l r = Thm.capply (Thm.capply (instantiate_cterm' [SOME (ctyp_of_term l)] [] @{cpat "op == :: ?'a =>_"}) l) r
   val asl = map2 (fn (t,_) => fn n => assume (mk_equals (mk_norm t) (cterm_of (ProofContext.theory_of ctxt') (Free(n,@{typ real}))))) lctab fxns
   val replace_conv = try_conv (rewrs_conv asl)
   val replace_rule = fconv_rule (funpow 2 arg_conv (replacenegnorms replace_conv))
@@ -368,7 +368,7 @@
   val th1 = real_vector_combo_prover ctxt' translator (nubs,ges',gts')
   val shs = filter (member (fn (t,th) => t aconvc cprop_of th) asl) (#hyps (crep_thm th1)) 
   val th11 = hd (Variable.export ctxt' ctxt [fold implies_intr shs th1])
-  val cps = map (swap o dest_equals) (cprems_of th11)
+  val cps = map (swap o Thm.dest_equals) (cprems_of th11)
   val th12 = instantiate ([], cps) th11
   val th13 = fold (C implies_elim) (map (reflexive o snd) cps) th12;
  in hd (Variable.export ctxt' ctxt [th13])
@@ -406,7 +406,7 @@
    val ctxt' = Variable.declare_term (term_of ct) ctxt
    val th = init_conv ctxt' ct
   in equal_elim (Drule.arg_cong_rule @{cterm Trueprop} (symmetric th)) 
-                (pure ctxt' (rhs_of th))
+                (pure ctxt' (Thm.rhs_of th))
  end
 
  fun norm_arith_tac ctxt = 
--- a/src/HOL/Library/positivstellensatz.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/HOL/Library/positivstellensatz.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -8,41 +8,24 @@
 
 signature FUNC = 
 sig
- type 'a T
- type key
- val apply : 'a T -> key -> 'a
- val applyd :'a T -> (key -> 'a) -> key -> 'a
- val combine : ('a -> 'a -> 'a) -> ('a -> bool) -> 'a T -> 'a T -> 'a T
- val defined : 'a T -> key -> bool
- val dom : 'a T -> key list
- val fold : (key * 'a -> 'b -> 'b) -> 'a T -> 'b -> 'b
- val fold_rev : (key * 'a -> 'b -> 'b) -> 'a T -> 'b -> 'b
- val graph : 'a T -> (key * 'a) list
- val is_undefined : 'a T -> bool
- val mapf : ('a -> 'b) -> 'a T -> 'b T
- val tryapplyd : 'a T -> key -> 'a -> 'a
- val undefine :  key -> 'a T -> 'a T
- val undefined : 'a T
- val update : key * 'a -> 'a T -> 'a T
- val updatep : (key * 'a -> bool) -> key * 'a -> 'a T -> 'a T
- val choose : 'a T -> key * 'a
- val onefunc : key * 'a -> 'a T
- val get_first: (key*'a -> 'a option) -> 'a T -> 'a option
+ include TABLE
+ val apply : 'a table -> key -> 'a
+ val applyd :'a table -> (key -> 'a) -> key -> 'a
+ val combine : ('a -> 'a -> 'a) -> ('a -> bool) -> 'a table -> 'a table -> 'a table
+ val dom : 'a table -> key list
+ val tryapplyd : 'a table -> key -> 'a -> 'a
+ val updatep : (key * 'a -> bool) -> key * 'a -> 'a table -> 'a table
+ val choose : 'a table -> key * 'a
+ val onefunc : key * 'a -> 'a table
 end;
 
 functor FuncFun(Key: KEY) : FUNC=
 struct
 
-type key = Key.key;
 structure Tab = Table(Key);
-type 'a T = 'a Tab.table;
 
-val undefined = Tab.empty;
-val is_undefined = Tab.is_empty;
-val mapf = Tab.map;
-val fold = Tab.fold;
-val fold_rev = Tab.fold_rev;
-val graph = Tab.dest;
+open Tab;
+
 fun dom a = sort Key.ord (Tab.keys a);
 fun applyd f d x = case Tab.lookup f x of 
    SOME y => y
@@ -50,9 +33,6 @@
 
 fun apply f x = applyd f (fn _ => raise Tab.UNDEF x) x;
 fun tryapplyd f a d = applyd f (K d) a;
-val defined = Tab.defined;
-fun undefine x t = (Tab.delete x t handle UNDEF => t);
-val update = Tab.update;
 fun updatep p (k,v) t = if p (k, v) then t else update (k,v) t
 fun combine f z a b = 
  let
@@ -64,16 +44,10 @@
 
 fun choose f = case Tab.min_key f of 
    SOME k => (k,valOf (Tab.lookup f k))
- | NONE => error "FuncFun.choose : Completely undefined function"
-
-fun onefunc kv = update kv undefined
+ | NONE => error "FuncFun.choose : Completely empty function"
 
-local
-fun  find f (k,v) NONE = f (k,v)
-   | find f (k,v) r = r
-in
-fun get_first f t = fold (find f) t NONE
-end
+fun onefunc kv = update kv empty
+
 end;
 
 (* Some standard functors and utility functions for them *)
@@ -81,33 +55,31 @@
 structure FuncUtil =
 struct
 
-fun increasing f ord (x,y) = ord (f x, f y);
-
 structure Intfunc = FuncFun(type key = int val ord = int_ord);
 structure Ratfunc = FuncFun(type key = Rat.rat val ord = Rat.ord);
 structure Intpairfunc = FuncFun(type key = int*int val ord = prod_ord int_ord int_ord);
 structure Symfunc = FuncFun(type key = string val ord = fast_string_ord);
 structure Termfunc = FuncFun(type key = term val ord = TermOrd.fast_term_ord);
 
-val cterm_ord = (fn (s,t) => TermOrd.fast_term_ord(term_of s, term_of t))
+val cterm_ord = TermOrd.fast_term_ord o pairself term_of
 
 structure Ctermfunc = FuncFun(type key = cterm val ord = cterm_ord);
 
-type monomial = int Ctermfunc.T;
+type monomial = int Ctermfunc.table;
 
-fun monomial_ord (m1,m2) = list_ord (prod_ord cterm_ord int_ord) (Ctermfunc.graph m1, Ctermfunc.graph m2)
+val monomial_ord = list_ord (prod_ord cterm_ord int_ord) o pairself Ctermfunc.dest
 
 structure Monomialfunc = FuncFun(type key = monomial val ord = monomial_ord)
 
-type poly = Rat.rat Monomialfunc.T;
+type poly = Rat.rat Monomialfunc.table;
 
 (* The ordering so we can create canonical HOL polynomials.                  *)
 
-fun dest_monomial mon = sort (increasing fst cterm_ord) (Ctermfunc.graph mon);
+fun dest_monomial mon = sort (cterm_ord o pairself fst) (Ctermfunc.dest mon);
 
 fun monomial_order (m1,m2) =
- if Ctermfunc.is_undefined m2 then LESS 
- else if Ctermfunc.is_undefined m1 then GREATER 
+ if Ctermfunc.is_empty m2 then LESS 
+ else if Ctermfunc.is_empty m1 then GREATER 
  else
   let val mon1 = dest_monomial m1 
       val mon2 = dest_monomial m2
@@ -165,7 +137,7 @@
 structure RealArith : REAL_ARITH =
 struct
 
- open Conv Thm FuncUtil;;
+ open Conv
 (* ------------------------------------------------------------------------- *)
 (* Data structure for Positivstellensatz refutations.                        *)
 (* ------------------------------------------------------------------------- *)
@@ -353,36 +325,31 @@
 
 (* Map back polynomials to HOL.                         *)
 
-local
- open Thm Numeral
-in
-
-fun cterm_of_varpow x k = if k = 1 then x else capply (capply @{cterm "op ^ :: real => _"} x) 
-  (mk_cnumber @{ctyp nat} k)
+fun cterm_of_varpow x k = if k = 1 then x else Thm.capply (Thm.capply @{cterm "op ^ :: real => _"} x) 
+  (Numeral.mk_cnumber @{ctyp nat} k)
 
 fun cterm_of_monomial m = 
- if Ctermfunc.is_undefined m then @{cterm "1::real"} 
+ if FuncUtil.Ctermfunc.is_empty m then @{cterm "1::real"} 
  else 
   let 
-   val m' = dest_monomial m
+   val m' = FuncUtil.dest_monomial m
    val vps = fold_rev (fn (x,k) => cons (cterm_of_varpow x k)) m' [] 
-  in foldr1 (fn (s, t) => capply (capply @{cterm "op * :: real => _"} s) t) vps
+  in foldr1 (fn (s, t) => Thm.capply (Thm.capply @{cterm "op * :: real => _"} s) t) vps
   end
 
-fun cterm_of_cmonomial (m,c) = if Ctermfunc.is_undefined m then cterm_of_rat c
+fun cterm_of_cmonomial (m,c) = if FuncUtil.Ctermfunc.is_empty m then cterm_of_rat c
     else if c = Rat.one then cterm_of_monomial m
-    else capply (capply @{cterm "op *::real => _"} (cterm_of_rat c)) (cterm_of_monomial m);
+    else Thm.capply (Thm.capply @{cterm "op *::real => _"} (cterm_of_rat c)) (cterm_of_monomial m);
 
 fun cterm_of_poly p = 
- if Monomialfunc.is_undefined p then @{cterm "0::real"} 
+ if FuncUtil.Monomialfunc.is_empty p then @{cterm "0::real"} 
  else
   let 
    val cms = map cterm_of_cmonomial
-     (sort (prod_ord monomial_order (K EQUAL)) (Monomialfunc.graph p))
-  in foldr1 (fn (t1, t2) => capply(capply @{cterm "op + :: real => _"} t1) t2) cms
+     (sort (prod_ord FuncUtil.monomial_order (K EQUAL)) (FuncUtil.Monomialfunc.dest p))
+  in foldr1 (fn (t1, t2) => Thm.capply(Thm.capply @{cterm "op + :: real => _"} t1) t2) cms
   end;
 
-end;
     (* A general real arithmetic prover *)
 
 fun gen_gen_real_arith ctxt (mk_numeric,
@@ -390,7 +357,6 @@
        poly_conv,poly_neg_conv,poly_add_conv,poly_mul_conv,
        absconv1,absconv2,prover) = 
 let
- open Conv Thm;
  val _ = my_context := ctxt 
  val _ = (my_mk_numeric := mk_numeric ; my_numeric_eq_conv := numeric_eq_conv ; 
           my_numeric_ge_conv := numeric_ge_conv; my_numeric_gt_conv := numeric_gt_conv ;
@@ -414,7 +380,7 @@
 
  fun real_ineq_conv th ct =
   let
-   val th' = (instantiate (match (lhs_of th, ct)) th 
+   val th' = (Thm.instantiate (Thm.match (Thm.lhs_of th, ct)) th 
       handle MATCH => raise CTERM ("real_ineq_conv", [ct]))
   in transitive th' (oprconv poly_conv (Thm.rhs_of th'))
   end 
@@ -442,14 +408,14 @@
         Axiom_eq n => nth eqs n
       | Axiom_le n => nth les n
       | Axiom_lt n => nth lts n
-      | Rational_eq x => eqT_elim(numeric_eq_conv(capply @{cterm Trueprop} 
-                          (capply (capply @{cterm "op =::real => _"} (mk_numeric x)) 
+      | Rational_eq x => eqT_elim(numeric_eq_conv(Thm.capply @{cterm Trueprop} 
+                          (Thm.capply (Thm.capply @{cterm "op =::real => _"} (mk_numeric x)) 
                                @{cterm "0::real"})))
-      | Rational_le x => eqT_elim(numeric_ge_conv(capply @{cterm Trueprop} 
-                          (capply (capply @{cterm "op <=::real => _"} 
+      | Rational_le x => eqT_elim(numeric_ge_conv(Thm.capply @{cterm Trueprop} 
+                          (Thm.capply (Thm.capply @{cterm "op <=::real => _"} 
                                      @{cterm "0::real"}) (mk_numeric x))))
-      | Rational_lt x => eqT_elim(numeric_gt_conv(capply @{cterm Trueprop} 
-                      (capply (capply @{cterm "op <::real => _"} @{cterm "0::real"})
+      | Rational_lt x => eqT_elim(numeric_gt_conv(Thm.capply @{cterm Trueprop} 
+                      (Thm.capply (Thm.capply @{cterm "op <::real => _"} @{cterm "0::real"})
                         (mk_numeric x))))
       | Square pt => square_rule (cterm_of_poly pt)
       | Eqmul(pt,p) => emul_rule (cterm_of_poly pt) (translate p)
@@ -463,8 +429,8 @@
       nnf_conv then_conv skolemize_conv then_conv prenex_conv then_conv
       weak_dnf_conv
 
-  val concl = dest_arg o cprop_of
-  fun is_binop opr ct = (dest_fun2 ct aconvc opr handle CTERM _ => false)
+  val concl = Thm.dest_arg o cprop_of
+  fun is_binop opr ct = (Thm.dest_fun2 ct aconvc opr handle CTERM _ => false)
   val is_req = is_binop @{cterm "op =:: real => _"}
   val is_ge = is_binop @{cterm "op <=:: real => _"}
   val is_gt = is_binop @{cterm "op <:: real => _"}
@@ -472,10 +438,13 @@
   val is_disj = is_binop @{cterm "op |"}
   fun conj_pair th = (th RS @{thm conjunct1}, th RS @{thm conjunct2})
   fun disj_cases th th1 th2 = 
-   let val (p,q) = dest_binop (concl th)
+   let val (p,q) = Thm.dest_binop (concl th)
        val c = concl th1
        val _ = if c aconvc (concl th2) then () else error "disj_cases : conclusions not alpha convertible"
-   in implies_elim (implies_elim (implies_elim (instantiate' [] (map SOME [p,q,c]) @{thm disjE}) th) (implies_intr (capply @{cterm Trueprop} p) th1)) (implies_intr (capply @{cterm Trueprop} q) th2)
+   in implies_elim (implies_elim
+          (implies_elim (instantiate' [] (map SOME [p,q,c]) @{thm disjE}) th)
+          (implies_intr (Thm.capply @{cterm Trueprop} p) th1))
+        (implies_intr (Thm.capply @{cterm Trueprop} q) th2)
    end
  fun overall cert_choice dun ths = case ths of
   [] =>
@@ -494,37 +463,37 @@
       overall cert_choice dun (th1::th2::oths) end
     else if is_disj ct then
       let 
-       val (th1, cert1) = overall (Left::cert_choice) dun (assume (capply @{cterm Trueprop} (dest_arg1 ct))::oths)
-       val (th2, cert2) = overall (Right::cert_choice) dun (assume (capply @{cterm Trueprop} (dest_arg ct))::oths)
+       val (th1, cert1) = overall (Left::cert_choice) dun (assume (Thm.capply @{cterm Trueprop} (Thm.dest_arg1 ct))::oths)
+       val (th2, cert2) = overall (Right::cert_choice) dun (assume (Thm.capply @{cterm Trueprop} (Thm.dest_arg ct))::oths)
       in (disj_cases th th1 th2, Branch (cert1, cert2)) end
    else overall cert_choice (th::dun) oths
   end
-  fun dest_binary b ct = if is_binop b ct then dest_binop ct 
+  fun dest_binary b ct = if is_binop b ct then Thm.dest_binop ct 
                          else raise CTERM ("dest_binary",[b,ct])
   val dest_eq = dest_binary @{cterm "op = :: real => _"}
   val neq_th = nth pth 5
   fun real_not_eq_conv ct = 
    let 
-    val (l,r) = dest_eq (dest_arg ct)
-    val th = instantiate ([],[(@{cpat "?x::real"},l),(@{cpat "?y::real"},r)]) neq_th
-    val th_p = poly_conv(dest_arg(dest_arg1(rhs_of th)))
+    val (l,r) = dest_eq (Thm.dest_arg ct)
+    val th = Thm.instantiate ([],[(@{cpat "?x::real"},l),(@{cpat "?y::real"},r)]) neq_th
+    val th_p = poly_conv(Thm.dest_arg(Thm.dest_arg1(Thm.rhs_of th)))
     val th_x = Drule.arg_cong_rule @{cterm "uminus :: real => _"} th_p
     val th_n = fconv_rule (arg_conv poly_neg_conv) th_x
     val th' = Drule.binop_cong_rule @{cterm "op |"} 
-     (Drule.arg_cong_rule (capply @{cterm "op <::real=>_"} @{cterm "0::real"}) th_p)
-     (Drule.arg_cong_rule (capply @{cterm "op <::real=>_"} @{cterm "0::real"}) th_n)
+     (Drule.arg_cong_rule (Thm.capply @{cterm "op <::real=>_"} @{cterm "0::real"}) th_p)
+     (Drule.arg_cong_rule (Thm.capply @{cterm "op <::real=>_"} @{cterm "0::real"}) th_n)
     in transitive th th' 
   end
  fun equal_implies_1_rule PQ = 
   let 
-   val P = lhs_of PQ
+   val P = Thm.lhs_of PQ
   in implies_intr P (equal_elim PQ (assume P))
   end
  (* FIXME!!! Copied from groebner.ml *)
  val strip_exists =
   let fun h (acc, t) =
    case (term_of t) of
-    Const("Ex",_)$Abs(x,T,p) => h (dest_abs NONE (dest_arg t) |>> (fn v => v::acc))
+    Const("Ex",_)$Abs(x,T,p) => h (Thm.dest_abs NONE (Thm.dest_arg t) |>> (fn v => v::acc))
   | _ => (acc,t)
   in fn t => h ([],t)
   end
@@ -559,7 +528,7 @@
  val strip_forall =
   let fun h (acc, t) =
    case (term_of t) of
-    Const("All",_)$Abs(x,T,p) => h (dest_abs NONE (dest_arg t) |>> (fn v => v::acc))
+    Const("All",_)$Abs(x,T,p) => h (Thm.dest_abs NONE (Thm.dest_arg t) |>> (fn v => v::acc))
   | _ => (acc,t)
   in fn t => h ([],t)
   end
@@ -576,16 +545,16 @@
   fun absremover ct = (literals_conv [@{term "op &"}, @{term "op |"}] [] 
                   (try_conv (absconv1 then_conv binop_conv (arg_conv poly_conv))) then_conv 
         try_conv (absconv2 then_conv nnf_norm_conv' then_conv binop_conv absremover)) ct
-  val nct = capply @{cterm Trueprop} (capply @{cterm "Not"} ct)
+  val nct = Thm.capply @{cterm Trueprop} (Thm.capply @{cterm "Not"} ct)
   val th0 = (init_conv then_conv arg_conv nnf_norm_conv') nct
-  val tm0 = dest_arg (rhs_of th0)
+  val tm0 = Thm.dest_arg (Thm.rhs_of th0)
   val (th, certificates) = if tm0 aconvc @{cterm False} then (equal_implies_1_rule th0, Trivial) else
    let 
     val (evs,bod) = strip_exists tm0
     val (avs,ibod) = strip_forall bod
     val th1 = Drule.arg_cong_rule @{cterm Trueprop} (fold mk_forall avs (absremover ibod))
-    val (th2, certs) = overall [] [] [specl avs (assume (rhs_of th1))]
-    val th3 = fold simple_choose evs (prove_hyp (equal_elim th1 (assume (capply @{cterm Trueprop} bod))) th2)
+    val (th2, certs) = overall [] [] [specl avs (assume (Thm.rhs_of th1))]
+    val th3 = fold simple_choose evs (prove_hyp (equal_elim th1 (assume (Thm.capply @{cterm Trueprop} bod))) th2)
    in (Drule.implies_intr_hyps (prove_hyp (equal_elim th0 (assume nct)) th3), certs)
    end
   in (implies_elim (instantiate' [] [SOME ct] pth_final) th, certificates)
@@ -595,11 +564,12 @@
 
 (* A linear arithmetic prover *)
 local
-  val linear_add = Ctermfunc.combine (curry op +/) (fn z => z =/ Rat.zero)
-  fun linear_cmul c = Ctermfunc.mapf (fn x => c */ x)
+  val linear_add = FuncUtil.Ctermfunc.combine (curry op +/) (fn z => z =/ Rat.zero)
+  fun linear_cmul c = FuncUtil.Ctermfunc.map (fn x => c */ x)
   val one_tm = @{cterm "1::real"}
-  fun contradictory p (e,_) = ((Ctermfunc.is_undefined e) andalso not(p Rat.zero)) orelse
-     ((gen_eq_set (op aconvc) (Ctermfunc.dom e, [one_tm])) andalso not(p(Ctermfunc.apply e one_tm)))
+  fun contradictory p (e,_) = ((FuncUtil.Ctermfunc.is_empty e) andalso not(p Rat.zero)) orelse
+     ((gen_eq_set (op aconvc) (FuncUtil.Ctermfunc.dom e, [one_tm])) andalso
+       not(p(FuncUtil.Ctermfunc.apply e one_tm)))
 
   fun linear_ineqs vars (les,lts) = 
    case find_first (contradictory (fn x => x >/ Rat.zero)) lts of
@@ -612,15 +582,15 @@
      let 
       val ineqs = les @ lts
       fun blowup v =
-       length(filter (fn (e,_) => Ctermfunc.tryapplyd e v Rat.zero =/ Rat.zero) ineqs) +
-       length(filter (fn (e,_) => Ctermfunc.tryapplyd e v Rat.zero >/ Rat.zero) ineqs) *
-       length(filter (fn (e,_) => Ctermfunc.tryapplyd e v Rat.zero </ Rat.zero) ineqs)
+       length(filter (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v Rat.zero =/ Rat.zero) ineqs) +
+       length(filter (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v Rat.zero >/ Rat.zero) ineqs) *
+       length(filter (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v Rat.zero </ Rat.zero) ineqs)
       val  v = fst(hd(sort (fn ((_,i),(_,j)) => int_ord (i,j))
                  (map (fn v => (v,blowup v)) vars)))
       fun addup (e1,p1) (e2,p2) acc =
        let 
-        val c1 = Ctermfunc.tryapplyd e1 v Rat.zero 
-        val c2 = Ctermfunc.tryapplyd e2 v Rat.zero
+        val c1 = FuncUtil.Ctermfunc.tryapplyd e1 v Rat.zero 
+        val c2 = FuncUtil.Ctermfunc.tryapplyd e2 v Rat.zero
        in if c1 */ c2 >=/ Rat.zero then acc else
         let 
          val e1' = linear_cmul (Rat.abs c2) e1
@@ -631,13 +601,13 @@
         end
        end
       val (les0,les1) = 
-         List.partition (fn (e,_) => Ctermfunc.tryapplyd e v Rat.zero =/ Rat.zero) les
+         List.partition (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v Rat.zero =/ Rat.zero) les
       val (lts0,lts1) = 
-         List.partition (fn (e,_) => Ctermfunc.tryapplyd e v Rat.zero =/ Rat.zero) lts
+         List.partition (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v Rat.zero =/ Rat.zero) lts
       val (lesp,lesn) = 
-         List.partition (fn (e,_) => Ctermfunc.tryapplyd e v Rat.zero >/ Rat.zero) les1
+         List.partition (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v Rat.zero >/ Rat.zero) les1
       val (ltsp,ltsn) = 
-         List.partition (fn (e,_) => Ctermfunc.tryapplyd e v Rat.zero >/ Rat.zero) lts1
+         List.partition (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v Rat.zero >/ Rat.zero) lts1
       val les' = fold_rev (fn ep1 => fold_rev (addup ep1) lesp) lesn les0
       val lts' = fold_rev (fn ep1 => fold_rev (addup ep1) (lesp@ltsp)) ltsn
                       (fold_rev (fn ep1 => fold_rev (addup ep1) (lesn@ltsn)) ltsp lts0)
@@ -650,20 +620,20 @@
   | NONE => (case eqs of 
     [] => 
      let val vars = remove (op aconvc) one_tm 
-           (fold_rev (curry (gen_union (op aconvc)) o Ctermfunc.dom o fst) (les@lts) []) 
+           (fold_rev (curry (gen_union (op aconvc)) o FuncUtil.Ctermfunc.dom o fst) (les@lts) []) 
      in linear_ineqs vars (les,lts) end
    | (e,p)::es => 
-     if Ctermfunc.is_undefined e then linear_eqs (es,les,lts) else
+     if FuncUtil.Ctermfunc.is_empty e then linear_eqs (es,les,lts) else
      let 
-      val (x,c) = Ctermfunc.choose (Ctermfunc.undefine one_tm e)
+      val (x,c) = FuncUtil.Ctermfunc.choose (FuncUtil.Ctermfunc.delete_safe one_tm e)
       fun xform (inp as (t,q)) =
-       let val d = Ctermfunc.tryapplyd t x Rat.zero in
+       let val d = FuncUtil.Ctermfunc.tryapplyd t x Rat.zero in
         if d =/ Rat.zero then inp else
         let 
          val k = (Rat.neg d) */ Rat.abs c // c
          val e' = linear_cmul k e
          val t' = linear_cmul (Rat.abs c) t
-         val p' = Eqmul(Monomialfunc.onefunc (Ctermfunc.undefined, k),p)
+         val p' = Eqmul(FuncUtil.Monomialfunc.onefunc (FuncUtil.Ctermfunc.empty, k),p)
          val q' = Product(Rational_lt(Rat.abs c),q) 
         in (linear_add e' t',Sum(p',q')) 
         end 
@@ -680,19 +650,19 @@
    end 
   
   fun lin_of_hol ct = 
-   if ct aconvc @{cterm "0::real"} then Ctermfunc.undefined
-   else if not (is_comb ct) then Ctermfunc.onefunc (ct, Rat.one)
-   else if is_ratconst ct then Ctermfunc.onefunc (one_tm, dest_ratconst ct)
+   if ct aconvc @{cterm "0::real"} then FuncUtil.Ctermfunc.empty
+   else if not (is_comb ct) then FuncUtil.Ctermfunc.onefunc (ct, Rat.one)
+   else if is_ratconst ct then FuncUtil.Ctermfunc.onefunc (one_tm, dest_ratconst ct)
    else
     let val (lop,r) = Thm.dest_comb ct 
-    in if not (is_comb lop) then Ctermfunc.onefunc (ct, Rat.one)
+    in if not (is_comb lop) then FuncUtil.Ctermfunc.onefunc (ct, Rat.one)
        else
         let val (opr,l) = Thm.dest_comb lop 
         in if opr aconvc @{cterm "op + :: real =>_"} 
            then linear_add (lin_of_hol l) (lin_of_hol r)
            else if opr aconvc @{cterm "op * :: real =>_"} 
-                   andalso is_ratconst l then Ctermfunc.onefunc (r, dest_ratconst l)
-           else Ctermfunc.onefunc (ct, Rat.one)
+                   andalso is_ratconst l then FuncUtil.Ctermfunc.onefunc (r, dest_ratconst l)
+           else FuncUtil.Ctermfunc.onefunc (ct, Rat.one)
         end
     end
 
@@ -700,21 +670,20 @@
    Const(@{const_name "real"}, _)$ n => 
      if can HOLogic.dest_number n then false else true
   | _ => false
- open Thm
 in 
 fun real_linear_prover translator (eq,le,lt) = 
  let 
-  val lhs = lin_of_hol o dest_arg1 o dest_arg o cprop_of
-  val rhs = lin_of_hol o dest_arg o dest_arg o cprop_of
+  val lhs = lin_of_hol o Thm.dest_arg1 o Thm.dest_arg o cprop_of
+  val rhs = lin_of_hol o Thm.dest_arg o Thm.dest_arg o cprop_of
   val eq_pols = map lhs eq
   val le_pols = map rhs le
   val lt_pols = map rhs lt 
   val aliens =  filter is_alien
-      (fold_rev (curry (gen_union (op aconvc)) o Ctermfunc.dom) 
+      (fold_rev (curry (gen_union (op aconvc)) o FuncUtil.Ctermfunc.dom) 
           (eq_pols @ le_pols @ lt_pols) [])
-  val le_pols' = le_pols @ map (fn v => Ctermfunc.onefunc (v,Rat.one)) aliens
+  val le_pols' = le_pols @ map (fn v => FuncUtil.Ctermfunc.onefunc (v,Rat.one)) aliens
   val (_,proof) = linear_prover (eq_pols,le_pols',lt_pols)
-  val le' = le @ map (fn a => instantiate' [] [SOME (dest_arg a)] @{thm real_of_nat_ge_zero}) aliens 
+  val le' = le @ map (fn a => instantiate' [] [SOME (Thm.dest_arg a)] @{thm real_of_nat_ge_zero}) aliens 
  in ((translator (eq,le',lt) proof), Trivial)
  end
 end;
@@ -737,28 +706,28 @@
    val y_tm = @{cpat "?y::real"}
    val is_max = is_binop @{cterm "max :: real => _"}
    val is_min = is_binop @{cterm "min :: real => _"} 
-   fun is_abs t = is_comb t andalso dest_fun t aconvc abs_tm
+   fun is_abs t = is_comb t andalso Thm.dest_fun t aconvc abs_tm
    fun eliminate_construct p c tm =
     let 
      val t = find_cterm p tm
-     val th0 = (symmetric o beta_conversion false) (capply (cabs t tm) t)
-     val (p,ax) = (dest_comb o rhs_of) th0
+     val th0 = (symmetric o beta_conversion false) (Thm.capply (Thm.cabs t tm) t)
+     val (p,ax) = (Thm.dest_comb o Thm.rhs_of) th0
     in fconv_rule(arg_conv(binop_conv (arg_conv (beta_conversion false))))
                (transitive th0 (c p ax))
    end
 
    val elim_abs = eliminate_construct is_abs
     (fn p => fn ax => 
-       instantiate ([], [(p_tm,p), (x_tm, dest_arg ax)]) pth_abs)
+       Thm.instantiate ([], [(p_tm,p), (x_tm, Thm.dest_arg ax)]) pth_abs)
    val elim_max = eliminate_construct is_max
     (fn p => fn ax => 
-      let val (ax,y) = dest_comb ax 
-      in  instantiate ([], [(p_tm,p), (x_tm, dest_arg ax), (y_tm,y)]) 
+      let val (ax,y) = Thm.dest_comb ax 
+      in  Thm.instantiate ([], [(p_tm,p), (x_tm, Thm.dest_arg ax), (y_tm,y)]) 
       pth_max end)
    val elim_min = eliminate_construct is_min
     (fn p => fn ax => 
-      let val (ax,y) = dest_comb ax 
-      in  instantiate ([], [(p_tm,p), (x_tm, dest_arg ax), (y_tm,y)]) 
+      let val (ax,y) = Thm.dest_comb ax 
+      in  Thm.instantiate ([], [(p_tm,p), (x_tm, Thm.dest_arg ax), (y_tm,y)]) 
       pth_min end)
    in first_conv [elim_abs, elim_max, elim_min, all_conv]
   end;
--- a/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -24,6 +24,7 @@
   calls: int,
   success: int,
   lemmas: int,
+  max_lems: int,
   time_isa: int,
   time_atp: int,
   time_atp_fail: int}
@@ -35,6 +36,7 @@
   time: int,
   timeout: int,
   lemmas: int,
+  max_lems: int,
   posns: Position.T list
   }
 
@@ -49,125 +51,141 @@
 *)
 datatype data = Data of sh_data * me_data * min_data * me_data
 
-fun make_sh_data (calls,success,lemmas,time_isa,time_atp,time_atp_fail) =
-  ShData{calls=calls, success=success, lemmas=lemmas, time_isa=time_isa,
-    time_atp=time_atp, time_atp_fail=time_atp_fail}
+fun make_sh_data
+      (calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail) =
+  ShData{calls=calls, success=success, lemmas=lemmas, max_lems=max_lems,
+         time_isa=time_isa, time_atp=time_atp, time_atp_fail=time_atp_fail}
 
 fun make_min_data (succs, ab_ratios, it_ratios) =
   MinData{succs=succs, ab_ratios=ab_ratios, it_ratios=it_ratios}
 
-fun make_me_data (calls, success, proofs, time, timeout, lemmas, posns) =
-  MeData{calls=calls, success=success, proofs=proofs, time=time, timeout=timeout, lemmas=lemmas, posns=posns}
+fun make_me_data (calls,success,proofs,time,timeout,lemmas,max_lems,posns) =
+  MeData{calls=calls, success=success, proofs=proofs, time=time,
+         timeout=timeout, lemmas=lemmas, max_lems=max_lems, posns=posns}
 
 val empty_data =
-  Data(make_sh_data (0, 0, 0, 0, 0, 0),
-       make_me_data(0, 0, 0, 0, 0, 0, []),
+  Data(make_sh_data (0, 0, 0, 0, 0, 0, 0),
+       make_me_data(0, 0, 0, 0, 0, 0, 0, []),
        MinData{succs=0, ab_ratios=0, it_ratios=0},
-       make_me_data(0, 0, 0, 0, 0, 0, []))
+       make_me_data(0, 0, 0, 0, 0, 0, 0, []))
 
 fun map_sh_data f
-  (Data (ShData{calls, success, lemmas, time_isa, time_atp, time_atp_fail}, meda0, minda, meda)) =
-  Data (make_sh_data (f (calls, success, lemmas, time_isa, time_atp, time_atp_fail)),
+    (Data(ShData{calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail},
+          meda0, minda, meda)) =
+  Data (make_sh_data (f (calls,success,lemmas,max_lems,
+                         time_isa,time_atp,time_atp_fail)),
         meda0, minda, meda)
 
 fun map_min_data f
   (Data(shda, meda0, MinData{succs,ab_ratios,it_ratios}, meda)) =
   Data(shda, meda0, make_min_data(f(succs,ab_ratios,it_ratios)), meda)
 
-fun map_me_data0 f (Data (shda, MeData{calls,success,proofs,time,timeout,lemmas,posns}, minda, meda)) =
-  Data(shda, make_me_data(f (calls,success,proofs,time,timeout,lemmas,posns)), minda, meda)
+fun map_me_data0 f (Data (shda, MeData{calls,success,proofs,time,timeout,lemmas,max_lems,posns}, minda, meda)) =
+  Data(shda, make_me_data(f (calls,success,proofs,time,timeout,lemmas,max_lems,posns)), minda, meda)
 
-fun map_me_data f (Data (shda, meda0, minda, MeData{calls,success,proofs,time,timeout,lemmas,posns})) =
-  Data(shda, meda0, minda, make_me_data(f (calls,success,proofs,time,timeout,lemmas,posns)))
+fun map_me_data f (Data (shda, meda0, minda, MeData{calls,success,proofs,time,timeout,lemmas,max_lems,posns})) =
+  Data(shda, meda0, minda, make_me_data(f (calls,success,proofs,time,timeout,lemmas,max_lems,posns)))
 
-val inc_sh_calls =
-  map_sh_data (fn (calls, success, lemmas, time_isa, time_atp, time_atp_fail)
-    => (calls + 1, success, lemmas, time_isa, time_atp, time_atp_fail))
+val inc_sh_calls =  map_sh_data
+  (fn (calls, success, lemmas,max_lems, time_isa, time_atp, time_atp_fail)
+    => (calls + 1, success, lemmas,max_lems, time_isa, time_atp, time_atp_fail))
 
-val inc_sh_success =
-  map_sh_data (fn (calls, success, lemmas, time_isa, time_atp, time_atp_fail)
-    => (calls, success + 1, lemmas, time_isa, time_atp, time_atp_fail))
+val inc_sh_success = map_sh_data
+  (fn (calls, success, lemmas,max_lems, time_isa, time_atp, time_atp_fail)
+    => (calls, success + 1, lemmas,max_lems, time_isa, time_atp, time_atp_fail))
 
-fun inc_sh_lemmas n =
-  map_sh_data (fn (calls, success, lemmas, time_isa, time_atp, time_atp_fail)
-    => (calls, success, lemmas + n, time_isa, time_atp, time_atp_fail))
+fun inc_sh_lemmas n = map_sh_data
+  (fn (calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail)
+    => (calls,success,lemmas+n,max_lems,time_isa,time_atp,time_atp_fail))
 
-fun inc_sh_time_isa t =
-  map_sh_data (fn (calls, success, lemmas, time_isa, time_atp, time_atp_fail)
-    => (calls, success, lemmas, time_isa + t, time_atp, time_atp_fail))
+fun inc_sh_max_lems n = map_sh_data
+  (fn (calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail)
+    => (calls,success,lemmas,Int.max(max_lems,n),time_isa,time_atp,time_atp_fail))
 
-fun inc_sh_time_atp t =
-  map_sh_data (fn (calls, success, lemmas, time_isa, time_atp, time_atp_fail)
-    => (calls, success, lemmas, time_isa, time_atp + t, time_atp_fail))
+fun inc_sh_time_isa t = map_sh_data
+  (fn (calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail)
+    => (calls,success,lemmas,max_lems,time_isa + t,time_atp,time_atp_fail))
+
+fun inc_sh_time_atp t = map_sh_data
+  (fn (calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail)
+    => (calls,success,lemmas,max_lems,time_isa,time_atp + t,time_atp_fail))
 
-fun inc_sh_time_atp_fail t =
-  map_sh_data (fn (calls, success, lemmas, time_isa, time_atp, time_atp_fail)
-    => (calls, success, lemmas, time_isa, time_atp, time_atp_fail + t))
+fun inc_sh_time_atp_fail t = map_sh_data
+  (fn (calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail)
+    => (calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail + t))
 
-val inc_min_succs =
-  map_min_data (fn (succs,ab_ratios,it_ratios) => (succs+1, ab_ratios, it_ratios))
+val inc_min_succs = map_min_data
+  (fn (succs,ab_ratios,it_ratios) => (succs+1, ab_ratios, it_ratios))
 
-fun inc_min_ab_ratios r =
-  map_min_data (fn (succs, ab_ratios, it_ratios) => (succs, ab_ratios+r, it_ratios))
+fun inc_min_ab_ratios r = map_min_data
+  (fn (succs, ab_ratios, it_ratios) => (succs, ab_ratios+r, it_ratios))
 
-fun inc_min_it_ratios r =
-  map_min_data (fn (succs, ab_ratios, it_ratios) => (succs, ab_ratios, it_ratios+r))
+fun inc_min_it_ratios r = map_min_data
+  (fn (succs, ab_ratios, it_ratios) => (succs, ab_ratios, it_ratios+r))
 
 val inc_metis_calls = map_me_data
- (fn (calls,success,proofs,time,timeout,lemmas,posns)
-  => (calls + 1, success, proofs, time, timeout, lemmas,posns))
+  (fn (calls,success,proofs,time,timeout,lemmas,max_lems,posns)
+    => (calls + 1, success, proofs, time, timeout, lemmas,max_lems,posns))
 
 val inc_metis_success = map_me_data
- (fn (calls,success,proofs,time,timeout,lemmas,posns)
-  => (calls, success + 1, proofs, time, timeout, lemmas,posns))
+  (fn (calls,success,proofs,time,timeout,lemmas,max_lems,posns)
+    => (calls, success + 1, proofs, time, timeout, lemmas,max_lems,posns))
 
 val inc_metis_proofs = map_me_data
- (fn (calls,success,proofs,time,timeout,lemmas,posns)
-  => (calls, success, proofs + 1, time, timeout, lemmas,posns))
+  (fn (calls,success,proofs,time,timeout,lemmas,max_lems,posns)
+    => (calls, success, proofs + 1, time, timeout, lemmas,max_lems,posns))
 
 fun inc_metis_time t = map_me_data
- (fn (calls,success,proofs,time,timeout,lemmas,posns)
-  => (calls, success, proofs, time + t, timeout, lemmas,posns))
+ (fn (calls,success,proofs,time,timeout,lemmas,max_lems,posns)
+  => (calls, success, proofs, time + t, timeout, lemmas,max_lems,posns))
 
 val inc_metis_timeout = map_me_data
- (fn (calls,success,proofs,time,timeout,lemmas,posns)
-  => (calls, success, proofs, time, timeout + 1, lemmas,posns))
+  (fn (calls,success,proofs,time,timeout,lemmas,max_lems,posns)
+    => (calls, success, proofs, time, timeout + 1, lemmas,max_lems,posns))
 
 fun inc_metis_lemmas n = map_me_data
- (fn (calls,success,proofs,time,timeout,lemmas,posns)
-  => (calls, success, proofs, time, timeout, lemmas + n, posns))
+  (fn (calls,success,proofs,time,timeout,lemmas,max_lems,posns)
+    => (calls, success, proofs, time, timeout, lemmas+n, max_lems, posns))
+
+fun inc_metis_max_lems n = map_me_data
+  (fn (calls,success,proofs,time,timeout,lemmas,max_lems,posns)
+    => (calls,success,proofs,time,timeout,lemmas,Int.max(max_lems,n), posns))
 
 fun inc_metis_posns pos = map_me_data
- (fn (calls,success,proofs,time,timeout,lemmas,posns)
-  => (calls, success, proofs, time, timeout, lemmas, pos::posns))
+  (fn (calls,success,proofs,time,timeout,lemmas,max_lems,posns)
+    => (calls, success, proofs, time, timeout, lemmas,max_lems, pos::posns))
 
 val inc_metis_calls0 = map_me_data0 
-(fn (calls,success,proofs,time,timeout,lemmas,posns)
-  => (calls + 1, success, proofs, time, timeout, lemmas,posns))
+  (fn (calls,success,proofs,time,timeout,lemmas,max_lems,posns)
+    => (calls + 1, success, proofs, time, timeout, lemmas,max_lems,posns))
 
 val inc_metis_success0 = map_me_data0
- (fn (calls,success,proofs,time,timeout,lemmas,posns)
-  => (calls, success + 1, proofs, time, timeout, lemmas,posns))
+  (fn (calls,success,proofs,time,timeout,lemmas,max_lems,posns)
+    => (calls, success + 1, proofs, time, timeout, lemmas,max_lems,posns))
 
 val inc_metis_proofs0 = map_me_data0
- (fn (calls,success,proofs,time,timeout,lemmas,posns)
-  => (calls, success, proofs + 1, time, timeout, lemmas,posns))
+  (fn (calls,success,proofs,time,timeout,lemmas,max_lems,posns)
+    => (calls, success, proofs + 1, time, timeout, lemmas,max_lems,posns))
 
 fun inc_metis_time0 t = map_me_data0
- (fn (calls,success,proofs,time,timeout,lemmas,posns)
-  => (calls, success, proofs, time + t, timeout, lemmas,posns))
+  (fn (calls,success,proofs,time,timeout,lemmas,max_lems,posns)
+    => (calls, success, proofs, time + t, timeout, lemmas,max_lems,posns))
 
 val inc_metis_timeout0 = map_me_data0
- (fn (calls,success,proofs,time,timeout,lemmas,posns)
-  => (calls, success, proofs, time, timeout + 1, lemmas,posns))
+  (fn (calls,success,proofs,time,timeout,lemmas,max_lems,posns)
+    => (calls, success, proofs, time, timeout + 1, lemmas,max_lems,posns))
 
 fun inc_metis_lemmas0 n = map_me_data0
- (fn (calls,success,proofs,time,timeout,lemmas,posns)
-  => (calls, success, proofs, time, timeout, lemmas + n, posns))
+  (fn (calls,success,proofs,time,timeout,lemmas,max_lems,posns)
+    => (calls, success, proofs, time, timeout, lemmas+n, max_lems, posns))
+
+fun inc_metis_max_lems0 n = map_me_data0
+  (fn (calls,success,proofs,time,timeout,lemmas,max_lems,posns)
+    => (calls,success,proofs,time,timeout,lemmas,Int.max(max_lems,n), posns))
 
 fun inc_metis_posns0 pos = map_me_data0
- (fn (calls,success,proofs,time,timeout,lemmas,posns)
-  => (calls, success, proofs, time, timeout, lemmas, pos::posns))
+  (fn (calls,success,proofs,time,timeout,lemmas,max_lems,posns)
+    => (calls, success, proofs, time, timeout, lemmas,max_lems, pos::posns))
 
 local
 
@@ -178,20 +196,21 @@
 fun avg_time t n =
   if n > 0 then (Real.fromInt t / 1000.0) / Real.fromInt n else 0.0
 
-fun log_sh_data log sh_calls sh_success sh_lemmas sh_time_isa sh_time_atp sh_time_atp_fail =
- (log ("Total number of sledgehammer calls: " ^ str sh_calls);
-  log ("Number of successful sledgehammer calls: " ^ str sh_success);
-  log ("Number of sledgehammer lemmas: " ^ str sh_lemmas);
-  log ("Success rate: " ^ percentage sh_success sh_calls ^ "%");
-  log ("Total time for sledgehammer calls (Isabelle): " ^ str3 (time sh_time_isa));
-  log ("Total time for successful sledgehammer calls (ATP): " ^ str3 (time sh_time_atp));
-  log ("Total time for failed sledgehammer calls (ATP): " ^ str3 (time sh_time_atp_fail));
+fun log_sh_data log calls success lemmas max_lems time_isa time_atp time_atp_fail =
+ (log ("Total number of sledgehammer calls: " ^ str calls);
+  log ("Number of successful sledgehammer calls: " ^ str success);
+  log ("Number of sledgehammer lemmas: " ^ str lemmas);
+  log ("Max number of sledgehammer lemmas: " ^ str max_lems);
+  log ("Success rate: " ^ percentage success calls ^ "%");
+  log ("Total time for sledgehammer calls (Isabelle): " ^ str3 (time time_isa));
+  log ("Total time for successful sledgehammer calls (ATP): " ^ str3 (time time_atp));
+  log ("Total time for failed sledgehammer calls (ATP): " ^ str3 (time time_atp_fail));
   log ("Average time for sledgehammer calls (Isabelle): " ^
-    str3 (avg_time sh_time_isa sh_calls));
+    str3 (avg_time time_isa calls));
   log ("Average time for successful sledgehammer calls (ATP): " ^
-    str3 (avg_time sh_time_atp sh_success));
+    str3 (avg_time time_atp success));
   log ("Average time for failed sledgehammer calls (ATP): " ^
-    str3 (avg_time sh_time_atp_fail (sh_calls - sh_success)))
+    str3 (avg_time time_atp_fail (calls - success)))
   )
 
 
@@ -200,13 +219,14 @@
   in str0 (Position.line_of pos) ^ ":" ^ str0 (Position.column_of pos) end
 
 fun log_metis_data log tag sh_calls sh_success metis_calls metis_success metis_proofs metis_time
-    metis_timeout metis_lemmas metis_posns =
+    metis_timeout metis_lemmas metis_max_lems metis_posns =
  (log ("Total number of " ^ tag ^ "metis calls: " ^ str metis_calls);
   log ("Number of successful " ^ tag ^ "metis calls: " ^ str metis_success ^
     " (proof: " ^ str metis_proofs ^ ")");
   log ("Number of " ^ tag ^ "metis timeouts: " ^ str metis_timeout);
   log ("Success rate: " ^ percentage metis_success sh_calls ^ "%");
   log ("Number of successful " ^ tag ^ "metis lemmas: " ^ str metis_lemmas);
+  log ("Max number of successful " ^ tag ^ "metis lemmas: " ^ str metis_max_lems);
   log ("Total time for successful metis calls: " ^ str3 (time metis_time));
   log ("Average time for successful metis calls: " ^
     str3 (avg_time metis_time metis_success));
@@ -224,27 +244,27 @@
 in
 
 fun log_data id log (Data
-   (ShData{calls=sh_calls, lemmas=sh_lemmas, success=sh_success,
+   (ShData{calls=sh_calls, lemmas=sh_lemmas,  max_lems=sh_max_lems, success=sh_success,
       time_isa=sh_time_isa,time_atp=sh_time_atp,time_atp_fail=sh_time_atp_fail},
     MeData{calls=metis_calls0, proofs=metis_proofs0,
       success=metis_success0, time=metis_time0, timeout=metis_timeout0,
-      lemmas=metis_lemmas0,posns=metis_posns0},
+      lemmas=metis_lemmas0,max_lems=metis_max_lems0,posns=metis_posns0},
     MinData{succs=min_succs, ab_ratios=ab_ratios, it_ratios=it_ratios},
     MeData{calls=metis_calls, proofs=metis_proofs,
       success=metis_success, time=metis_time, timeout=metis_timeout,
-      lemmas=metis_lemmas,posns=metis_posns})) =
+      lemmas=metis_lemmas,max_lems=metis_max_lems,posns=metis_posns})) =
   if sh_calls > 0
   then
    (log ("\n\n\nReport #" ^ string_of_int id ^ ":\n");
-    log_sh_data log sh_calls sh_success sh_lemmas sh_time_isa sh_time_atp sh_time_atp_fail;
+    log_sh_data log sh_calls sh_success sh_lemmas sh_max_lems sh_time_isa sh_time_atp sh_time_atp_fail;
     log "";
     if metis_calls > 0 then log_metis_data log "" sh_calls sh_success metis_calls
-      metis_success metis_proofs metis_time metis_timeout metis_lemmas  metis_posns else ();
+      metis_success metis_proofs metis_time metis_timeout metis_lemmas metis_max_lems metis_posns else ();
     log "";
     if metis_calls0 > 0
       then (log_min_data log min_succs ab_ratios it_ratios; log "";
             log_metis_data log "unminimized " sh_calls sh_success metis_calls0
-              metis_success0 metis_proofs0 metis_time0 metis_timeout0 metis_lemmas0 metis_posns0)
+              metis_success0 metis_proofs0 metis_time0 metis_timeout0 metis_lemmas0  metis_max_lems0 metis_posns0)
       else ()
    )
   else ()
@@ -338,15 +358,14 @@
   in
     case result of
       SH_OK (time_isa, time_atp, names) =>
-        let
-          val _ = change_data id inc_sh_success
-          val _ = change_data id (inc_sh_lemmas (length names))
-          val _ = change_data id (inc_sh_time_isa time_isa)
-          val _ = change_data id (inc_sh_time_atp time_atp)
-
-          fun get_thms name = (name, thms_of_name (Proof.context_of st) name)
-          val _ = named_thms := SOME (map get_thms names)
+        let fun get_thms name = (name, thms_of_name (Proof.context_of st) name)
         in
+          change_data id inc_sh_success;
+          change_data id (inc_sh_lemmas (length names));
+          change_data id (inc_sh_max_lems (length names));
+          change_data id (inc_sh_time_isa time_isa);
+          change_data id (inc_sh_time_atp time_atp);
+          named_thms := SOME (map get_thms names);
           log (sh_tag id ^ "succeeded (" ^ string_of_int time_isa ^ "+" ^
             string_of_int time_atp ^ ") [" ^ prover_name ^ "]:\n" ^ msg)
         end
@@ -387,7 +406,7 @@
 
 
 fun run_metis (inc_metis_calls, inc_metis_success, inc_metis_proofs, inc_metis_time, inc_metis_timeout,
-    inc_metis_lemmas, inc_metis_posns) args name named_thms id
+    inc_metis_lemmas, inc_metis_max_lems, inc_metis_posns) args name named_thms id
     ({pre=st, timeout, log, pos, ...}: Mirabelle.run_args) =
   let
     fun metis thms ctxt = MetisTools.metis_tac ctxt thms
@@ -396,6 +415,7 @@
     fun with_time (false, t) = "failed (" ^ string_of_int t ^ ")"
       | with_time (true, t) = (change_data id inc_metis_success;
           change_data id (inc_metis_lemmas (length named_thms));
+          change_data id (inc_metis_max_lems (length named_thms));
           change_data id (inc_metis_time t);
           change_data id (inc_metis_posns pos);
           if name = "proof" then change_data id inc_metis_proofs else ();
@@ -413,13 +433,14 @@
   end
 
 fun sledgehammer_action args id (st as {log, pre, name, ...}: Mirabelle.run_args) =
-  if can Logic.dest_conjunction (Thm.major_prem_of(snd(snd(Proof.get_goal pre))))
+  let val goal = Thm.major_prem_of(snd(snd(Proof.get_goal pre))) in
+  if can Logic.dest_conjunction goal orelse can Logic.dest_equals goal
   then () else
   let
     val metis_fns = (inc_metis_calls, inc_metis_success, inc_metis_proofs, inc_metis_time,
-        inc_metis_timeout, inc_metis_lemmas, inc_metis_posns)
+        inc_metis_timeout, inc_metis_lemmas,  inc_metis_max_lems, inc_metis_posns)
     val metis0_fns = (inc_metis_calls0, inc_metis_success0, inc_metis_proofs0, inc_metis_time0,
-        inc_metis_timeout0, inc_metis_lemmas0, inc_metis_posns0)
+        inc_metis_timeout0, inc_metis_lemmas0, inc_metis_max_lems0, inc_metis_posns0)
     val named_thms = Unsynchronized.ref (NONE : (string * thm list) list option)
     val minimize = AList.defined (op =) args minimizeK
   in 
@@ -435,6 +456,7 @@
        Mirabelle.catch metis_tag (run_metis metis_fns args name (these (!named_thms))) id st)
     else ()
   end
+  end
 
 fun invoke args =
   let
--- a/src/HOL/Record.thy	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/HOL/Record.thy	Thu Oct 01 20:52:18 2009 +0200
@@ -450,10 +450,6 @@
   "Q = R \<Longrightarrow> (P \<and> Q) = (P \<and> R)"
   by simp
 
-lemma meta_all_sameI:
-  "(\<And>x. PROP P x \<equiv> PROP Q x) \<Longrightarrow> (\<And>x. PROP P x) \<equiv> (\<And>x. PROP Q x)"
-  by simp
-
 lemma istuple_UNIV_I: "\<And>x. x\<in>UNIV \<equiv> True"
   by simp
 
--- a/src/HOL/Tools/ATP_Manager/atp_manager.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/HOL/Tools/ATP_Manager/atp_manager.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -41,7 +41,7 @@
 
 local
 
-val atps = Unsynchronized.ref "e remote_vampire";
+val atps = Unsynchronized.ref "e spass remote_vampire";
 val max_atps = Unsynchronized.ref 5;   (* ~1 means infinite number of atps *)
 val timeout = Unsynchronized.ref 60;
 val full_types = Unsynchronized.ref false;
--- a/src/HOL/Tools/record.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/HOL/Tools/record.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -202,22 +202,18 @@
 struct
 
 val eq_reflection = @{thm eq_reflection};
-val Pair_eq = @{thm Product_Type.prod.inject};
 val atomize_all = @{thm HOL.atomize_all};
 val atomize_imp = @{thm HOL.atomize_imp};
 val meta_allE = @{thm Pure.meta_allE};
 val prop_subst = @{thm prop_subst};
-val Pair_sel_convs = [fst_conv, snd_conv];
 val K_record_comp = @{thm K_record_comp};
 val K_comp_convs = [@{thm o_apply}, K_record_comp]
-val transitive_thm = @{thm transitive};
 val o_assoc = @{thm o_assoc};
 val id_apply = @{thm id_apply};
 val id_o_apps = [@{thm id_apply}, @{thm id_o}, @{thm o_id}];
 val Not_eq_iff = @{thm Not_eq_iff};
 
 val refl_conj_eq = @{thm refl_conj_eq};
-val meta_all_sameI = @{thm meta_all_sameI};
 
 val surject_assistI = @{thm "istuple_surjective_proof_assistI"};
 val surject_assist_idE = @{thm "istuple_surjective_proof_assist_idE"};
@@ -250,7 +246,6 @@
 val ext_typeN = "_ext_type";
 val inner_typeN = "_inner_type";
 val extN ="_ext";
-val casesN = "_cases";
 val ext_dest = "_sel";
 val updateN = "_update";
 val updN = "_upd";
@@ -259,10 +254,6 @@
 val extendN = "extend";
 val truncateN = "truncate";
 
-(*see typedef.ML*)
-val RepN = "Rep_";
-val AbsN = "Abs_";
-
 
 
 (*** utilities ***)
@@ -273,24 +264,6 @@
   let fun varify (a, S) = TVar ((a, midx + 1), S);
   in map_type_tfree varify end;
 
-fun domain_type' T =
-  domain_type T handle Match => T;
-
-fun range_type' T =
-  range_type T handle Match => T;
-
-
-(* messages *)  (* FIXME proper context *)
-
-fun trace_thm str thm =
-  tracing (str ^ Pretty.string_of (Display.pretty_thm_without_context thm));
-
-fun trace_thms str thms =
-  (tracing str; map (trace_thm "") thms);
-
-fun trace_term str t =
-  tracing (str ^ Syntax.string_of_term_global Pure.thy t);
-
 
 (* timing *)
 
@@ -302,7 +275,6 @@
 (* syntax *)
 
 fun prune n xs = Library.drop (n, xs);
-fun prefix_base s = Long_Name.map_base_name (fn bname => s ^ bname);
 
 val Trueprop = HOLogic.mk_Trueprop;
 fun All xs t = Term.list_all_free (xs, t);
@@ -311,22 +283,10 @@
 infix 0 :== ===;
 infixr 0 ==>;
 
-val (op $$) = Term.list_comb;
-val (op :==) = PrimitiveDefs.mk_defpair;
-val (op ===) = Trueprop o HOLogic.mk_eq;
-val (op ==>) = Logic.mk_implies;
-
-
-(* morphisms *)
-
-fun mk_RepN name = suffix ext_typeN (prefix_base RepN name);
-fun mk_AbsN name = suffix ext_typeN (prefix_base AbsN name);
-
-fun mk_Rep name repT absT =
-  Const (suffix ext_typeN (prefix_base RepN name), absT --> repT);
-
-fun mk_Abs name repT absT =
-  Const (mk_AbsN name, repT --> absT);
+val op $$ = Term.list_comb;
+val op :== = PrimitiveDefs.mk_defpair;
+val op === = Trueprop o HOLogic.mk_eq;
+val op ==> = Logic.mk_implies;
 
 
 (* constructor *)
@@ -338,15 +298,6 @@
   in list_comb (Const (mk_extC (name, T) Ts), ts) end;
 
 
-(* cases *)
-
-fun mk_casesC (name, T, vT) Ts = (suffix casesN name, (Ts ---> vT) --> T --> vT);
-
-fun mk_cases (name, T, vT) f =
-  let val Ts = binder_types (fastype_of f)
-  in Const (mk_casesC (name, T, vT) Ts) $ f end;
-
-
 (* selector *)
 
 fun mk_selC sT (c, T) = (c, sT --> T);
@@ -369,7 +320,7 @@
 
 (* types *)
 
-fun dest_recT (typ as Type (c_ext_type, Ts as (T :: _))) =
+fun dest_recT (typ as Type (c_ext_type, Ts as (_ :: _))) =
       (case try (unsuffix ext_typeN) c_ext_type of
         NONE => raise TYPE ("Record.dest_recT", [typ], [])
       | SOME c => ((c, Ts), List.last Ts))
@@ -549,8 +500,6 @@
 
 val get_simpset = get_ss_with_context #simpset;
 val get_sel_upd_defs = get_ss_with_context #defset;
-val get_foldcong_ss = get_ss_with_context #foldcong;
-val get_unfoldcong_ss = get_ss_with_context #unfoldcong;
 
 fun get_update_details u thy =
   let val sel_upd = get_sel_upd thy in
@@ -618,8 +567,6 @@
       extfields fieldext;
   in RecordsData.put data thy end;
 
-val get_extsplit = Symtab.lookup o #extsplit o RecordsData.get;
-
 
 (* access 'splits' *)
 
@@ -659,7 +606,7 @@
   let
     val ((name, Ts), moreT) = dest_recT T;
     val recname =
-      let val (nm :: recn :: rst) = rev (Long_Name.explode name)
+      let val (nm :: _ :: rst) = rev (Long_Name.explode name)   (* FIXME !? *)
       in Long_Name.implode (rev (nm :: rst)) end;
     val midx = maxidx_of_typs (moreT :: Ts);
     val varifyT = varifyT midx;
@@ -698,7 +645,7 @@
 
 (* parent records *)
 
-fun add_parents thy NONE parents = parents
+fun add_parents _ NONE parents = parents
   | add_parents thy (SOME (types, name)) parents =
       let
         fun err msg = error (msg ^ " parent record " ^ quote name);
@@ -787,7 +734,7 @@
       | splitargs (_ :: _) [] = raise TERM (msg ^ "expecting more fields", [t])
       | splitargs _ _ = ([], []);
 
-    fun mk_ext (fargs as (name, arg) :: _) =
+    fun mk_ext (fargs as (name, _) :: _) =
           (case get_fieldext thy (Sign.intern_const thy name) of
             SOME (ext, _) =>
               (case get_extfields thy ext of
@@ -816,7 +763,7 @@
       | splitargs (_ :: _) [] = raise TERM (msg ^ "expecting more fields", [t])
       | splitargs _ _ = ([], []);
 
-    fun mk_ext (fargs as (name, arg) :: _) =
+    fun mk_ext (fargs as (name, _) :: _) =
           (case get_fieldext thy (Sign.intern_const thy name) of
             SOME (ext, alphas) =>
               (case get_extfields thy ext of
@@ -838,7 +785,7 @@
                     val more' = mk_ext rest;
                   in
                     list_comb (Syntax.const (suffix sfx ext), alphas' @ [more'])
-                  end handle TYPE_MATCH =>
+                  end handle Type.TYPE_MATCH =>
                     raise TERM (msg ^ "type is no proper record (extension)", [t]))
               | NONE => raise TERM (msg ^ "no fields defined for " ^ ext, [t]))
           | NONE => raise TERM (msg ^ name ^" is no proper field", [t]))
@@ -896,7 +843,7 @@
           (case k of
             Abs (_, _, Abs (_, _, t) $ Bound 0) =>
               if null (loose_bnos t) then t else raise Match
-          | Abs (x, _, t) =>
+          | Abs (_, _, t) =>
               if null (loose_bnos t) then t else raise Match
           | _ => raise Match);
 
@@ -1012,7 +959,7 @@
               if HOLogic.is_unitT (Envir.norm_type subst (varifyT (TFree (zeta, Sign.defaultS thy))))
               then mk_type_abbr subst abbr alphas
               else mk_type_abbr subst (suffix schemeN abbr) (alphas @ [zeta])
-            end handle TYPE_MATCH => default_tr' ctxt tm)
+            end handle Type.TYPE_MATCH => default_tr' ctxt tm)
           else raise Match (*give print translation of specialised record a chance*)
       | _ => raise Match)
     else default_tr' ctxt tm
@@ -1045,8 +992,7 @@
                         val subst = fold2 (curry (Sign.typ_match thy)) alphavars args' Vartab.empty;
                         val flds'' = (map o apsnd) (Envir.norm_type subst o varifyT) flds';
                       in flds'' @ field_lst more end
-                      handle TYPE_MATCH => [("", T)]
-                        | Library.UnequalLengths => [("", T)])
+                      handle Type.TYPE_MATCH => [("", T)] | Library.UnequalLengths => [("", T)])
                   | NONE => [("", T)])
               | NONE => [("", T)])
           | NONE => [("", T)])
@@ -1106,7 +1052,7 @@
   then noopt ()
   else opt ();
 
-fun is_sel_upd_pair thy (Const (s, t)) (Const (u, t')) =
+fun is_sel_upd_pair thy (Const (s, _)) (Const (u, t')) =
   (case get_updates thy u of
     SOME u_name => u_name = s
   | NONE => raise TERM ("is_sel_upd_pair: not update", [Const (u, t')]));
@@ -1130,7 +1076,6 @@
 fun get_accupd_simps thy term defset intros_tac =
   let
     val (acc, [body]) = strip_comb term;
-    val recT = domain_type (fastype_of acc);
     val upd_funs = sort_distinct TermOrd.fast_term_ord (get_upd_funs body);
     fun get_simp upd =
       let
@@ -1140,10 +1085,10 @@
           if is_sel_upd_pair thy acc upd
           then mk_comp (Free ("f", T)) acc
           else mk_comp_id acc;
-        val prop = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
+        val prop = lhs === rhs;
         val othm =
           Goal.prove (ProofContext.init thy) [] [] prop
-            (fn prems =>
+            (fn _ =>
               EVERY
                [simp_tac defset 1,
                 REPEAT_DETERM (intros_tac 1),
@@ -1164,10 +1109,10 @@
       if comp
       then u $ mk_comp f f'
       else mk_comp (u' $ f') (u $ f);
-    val prop = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
+    val prop = lhs === rhs;
     val othm =
       Goal.prove (ProofContext.init thy) [] [] prop
-        (fn prems =>
+        (fn _ =>
           EVERY
            [simp_tac defset 1,
             REPEAT_DETERM (intros_tac 1),
@@ -1177,11 +1122,10 @@
 
 fun get_updupd_simps thy term defset intros_tac =
   let
-    val recT = fastype_of term;
     val upd_funs = get_upd_funs term;
     val cname = fst o dest_Const;
     fun getswap u u' = get_updupd_simp thy defset intros_tac u u' (cname u = cname u');
-    fun build_swaps_to_eq upd [] swaps = swaps
+    fun build_swaps_to_eq _ [] swaps = swaps
       | build_swaps_to_eq upd (u :: us) swaps =
           let
             val key = (cname u, cname upd);
@@ -1192,7 +1136,7 @@
             if cname u = cname upd then newswaps
             else build_swaps_to_eq upd us newswaps
           end;
-    fun swaps_needed [] prev seen swaps = map snd (Symreltab.dest swaps)
+    fun swaps_needed [] _ _ swaps = map snd (Symreltab.dest swaps)
       | swaps_needed (u :: us) prev seen swaps =
           if Symtab.defined seen (cname u)
           then swaps_needed us prev seen (build_swaps_to_eq u prev swaps)
@@ -1201,20 +1145,20 @@
 
 val named_cterm_instantiate = IsTupleSupport.named_cterm_instantiate;
 
-fun prove_unfold_defs thy ss T ex_simps ex_simprs prop =
+fun prove_unfold_defs thy ex_simps ex_simprs prop =
   let
     val defset = get_sel_upd_defs thy;
     val in_tac = IsTupleSupport.istuple_intros_tac thy;
     val prop' = Envir.beta_eta_contract prop;
-    val (lhs, rhs) = Logic.dest_equals (Logic.strip_assums_concl prop');
-    val (head, args) = strip_comb lhs;
+    val (lhs, _) = Logic.dest_equals (Logic.strip_assums_concl prop');
+    val (_, args) = strip_comb lhs;
     val simps =
       if length args = 1
       then get_accupd_simps thy lhs defset in_tac
       else get_updupd_simps thy lhs defset in_tac;
   in
     Goal.prove (ProofContext.init thy) [] [] prop'
-      (fn prems =>
+      (fn _ =>
         simp_tac (HOL_basic_ss addsimps (simps @ [K_record_comp])) 1 THEN
         TRY (simp_tac (HOL_basic_ss addsimps ex_simps addsimprocs ex_simprs) 1))
   end;
@@ -1251,55 +1195,52 @@
 *)
 val record_simproc =
   Simplifier.simproc @{theory HOL} "record_simp" ["x"]
-    (fn thy => fn ss => fn t =>
+    (fn thy => fn _ => fn t =>
       (case t of
-        (sel as Const (s, Type (_, [domS, rangeS]))) $
+        (sel as Const (s, Type (_, [_, rangeS]))) $
             ((upd as Const (u, Type (_, [_, Type (_, [rT, _])]))) $ k $ r) =>
-          if is_selector thy s then
-            (case get_updates thy u of
-              SOME u_name =>
-                let
-                  val {sel_upd = {updates, ...}, extfields, ...} = RecordsData.get thy;
-
-                  fun mk_eq_terms ((upd as Const (u, Type(_, [kT, _]))) $ k $ r) =
-                        (case Symtab.lookup updates u of
-                          NONE => NONE
-                        | SOME u_name =>
-                            if u_name = s then
-                              (case mk_eq_terms r of
-                                NONE =>
-                                  let
-                                    val rv = ("r", rT);
-                                    val rb = Bound 0;
-                                    val (kv, kb) = K_skeleton "k" kT (Bound 1) k;
-                                  in SOME (upd $ kb $ rb, kb $ (sel $ rb), [kv, rv]) end
-                              | SOME (trm, trm', vars) =>
-                                  let
-                                    val (kv, kb) = K_skeleton "k" kT (Bound (length vars)) k;
-                                  in SOME (upd $ kb $ trm, kb $ trm', kv :: vars) end)
-                            else if has_field extfields u_name rangeS orelse
-                              has_field extfields s (domain_type kT) then NONE
-                            else
-                              (case mk_eq_terms r of
-                                SOME (trm, trm', vars) =>
-                                  let val (kv, kb) = K_skeleton "k" kT (Bound (length vars)) k
-                                  in SOME (upd $ kb $ trm, trm', kv :: vars) end
-                              | NONE =>
-                                  let
-                                    val rv = ("r", rT);
-                                    val rb = Bound 0;
-                                    val (kv, kb) = K_skeleton "k" kT (Bound 1) k;
-                                  in SOME (upd $ kb $ rb, sel $ rb, [kv, rv]) end))
-                    | mk_eq_terms r = NONE;
-                in
-                  (case mk_eq_terms (upd $ k $ r) of
-                    SOME (trm, trm', vars) =>
-                      SOME
-                        (prove_unfold_defs thy ss domS [] []
-                          (list_all (vars, Logic.mk_equals (sel $ trm, trm'))))
-                  | NONE => NONE)
-                end
-            | NONE => NONE)
+          if is_selector thy s andalso is_some (get_updates thy u) then
+            let
+              val {sel_upd = {updates, ...}, extfields, ...} = RecordsData.get thy;
+
+              fun mk_eq_terms ((upd as Const (u, Type(_, [kT, _]))) $ k $ r) =
+                    (case Symtab.lookup updates u of
+                      NONE => NONE
+                    | SOME u_name =>
+                        if u_name = s then
+                          (case mk_eq_terms r of
+                            NONE =>
+                              let
+                                val rv = ("r", rT);
+                                val rb = Bound 0;
+                                val (kv, kb) = K_skeleton "k" kT (Bound 1) k;
+                              in SOME (upd $ kb $ rb, kb $ (sel $ rb), [kv, rv]) end
+                          | SOME (trm, trm', vars) =>
+                              let
+                                val (kv, kb) = K_skeleton "k" kT (Bound (length vars)) k;
+                              in SOME (upd $ kb $ trm, kb $ trm', kv :: vars) end)
+                        else if has_field extfields u_name rangeS orelse
+                          has_field extfields s (domain_type kT) then NONE
+                        else
+                          (case mk_eq_terms r of
+                            SOME (trm, trm', vars) =>
+                              let val (kv, kb) = K_skeleton "k" kT (Bound (length vars)) k
+                              in SOME (upd $ kb $ trm, trm', kv :: vars) end
+                          | NONE =>
+                              let
+                                val rv = ("r", rT);
+                                val rb = Bound 0;
+                                val (kv, kb) = K_skeleton "k" kT (Bound 1) k;
+                              in SOME (upd $ kb $ rb, sel $ rb, [kv, rv]) end))
+                | mk_eq_terms _ = NONE;
+            in
+              (case mk_eq_terms (upd $ k $ r) of
+                SOME (trm, trm', vars) =>
+                  SOME
+                    (prove_unfold_defs thy [] []
+                      (list_all (vars, Logic.mk_equals (sel $ trm, trm'))))
+              | NONE => NONE)
+            end
           else NONE
       | _ => NONE));
 
@@ -1311,7 +1252,7 @@
     val prop = concl_of (named_cterm_instantiate insts updacc_cong_triv);
   in
     Goal.prove (ProofContext.init thy) [] [] prop
-      (fn prems =>
+      (fn _ =>
         EVERY
          [simp_tac simpset 1,
           REPEAT_DETERM (in_tac 1),
@@ -1331,7 +1272,7 @@
   both a more update and an update to a field within it.*)
 val record_upd_simproc =
   Simplifier.simproc @{theory HOL} "record_upd_simp" ["x"]
-    (fn thy => fn ss => fn t =>
+    (fn thy => fn _ => fn t =>
       let
         (*We can use more-updators with other updators as long
           as none of the other updators go deeper than any more
@@ -1346,7 +1287,7 @@
               then SOME (if min <= dep then dep else min, max)
               else NONE;
 
-        fun getupdseq (term as (upd as Const (u, T)) $ f $ tm) min max =
+        fun getupdseq (term as (upd as Const (u, _)) $ f $ tm) min max =
               (case get_update_details u thy of
                 SOME (s, dep, ismore) =>
                   (case include_depth (dep, ismore) (min, max) of
@@ -1359,15 +1300,14 @@
 
         val (upds, base, baseT) = getupdseq t 0 ~1;
 
-        fun is_upd_noop s (f as Abs (n, T, Const (s', T') $ tm')) tm =
+        fun is_upd_noop s (Abs (n, T, Const (s', T') $ tm')) tm =
               if s = s' andalso null (loose_bnos tm')
                 andalso subst_bound (HOLogic.unit, tm') = tm
               then (true, Abs (n, T, Const (s', T') $ Bound 1))
               else (false, HOLogic.unit)
-          | is_upd_noop s f tm = (false, HOLogic.unit);
-
-        fun get_noop_simps (upd as Const (u, T))
-            (f as Abs (n, T', (acc as Const (s, T'')) $ _)) =
+          | is_upd_noop _ _ _ = (false, HOLogic.unit);
+
+        fun get_noop_simps (upd as Const _) (Abs (_, _, (acc as Const _) $ _)) =
           let
             val ss = get_sel_upd_defs thy;
             val uathm = get_upd_acc_cong_thm upd acc thy ss;
@@ -1417,17 +1357,16 @@
                       fvar :: vars, dups, true, noops)
                   | NONE => (upd $ skelf $ lhs, upd $ skelf $ rhs, fvar :: vars, dups, simp, noops))
               end
-          | mk_updterm [] above term =
+          | mk_updterm [] _ _ =
               (Bound 0, Bound 0, [("r", baseT)], Symtab.empty, false, Symtab.empty)
-          | mk_updterm us above term =
-              raise TERM ("mk_updterm match", map (fn (x, y, z) => x) us);
-
-        val (lhs, rhs, vars, dups, simp, noops) = mk_updterm upds Symtab.empty base;
+          | mk_updterm us _ _ = raise TERM ("mk_updterm match", map (fn (x, _, _) => x) us);
+
+        val (lhs, rhs, vars, _, simp, noops) = mk_updterm upds Symtab.empty base;
         val noops' = flat (map snd (Symtab.dest noops));
       in
         if simp then
           SOME
-            (prove_unfold_defs thy ss baseT noops' [record_simproc]
+            (prove_unfold_defs thy noops' [record_simproc]
               (list_all (vars, Logic.mk_equals (lhs, rhs))))
         else NONE
       end);
@@ -1473,11 +1412,11 @@
   Simplifier.simproc @{theory HOL} "record_split_simp" ["x"]
     (fn thy => fn _ => fn t =>
       (case t of
-        Const (quantifier, Type (_, [Type (_, [T, _]), _])) $ trm =>
+        Const (quantifier, Type (_, [Type (_, [T, _]), _])) $ _ =>
           if quantifier = "All" orelse quantifier = "all" orelse quantifier = "Ex" then
             (case rec_id ~1 T of
               "" => NONE
-            | name =>
+            | _ =>
                 let val split = P t in
                   if split <> 0 then
                     (case get_splits thy (rec_id split T) of
@@ -1568,10 +1507,10 @@
           simp_tac (HOL_basic_ss addsimps @{thms induct_rulify}) i]
       end;
 
-    fun split_free_tac P i (free as Free (n, T)) =
+    fun split_free_tac P i (free as Free (_, T)) =
           (case rec_id ~1 T of
             "" => NONE
-          | name =>
+          | _ =>
               let val split = P free in
                 if split <> 0 then
                   (case get_splits thy (rec_id split T) of
@@ -1598,8 +1537,6 @@
 (*Split all records in the goal, which are quantified by ! or !!.*)
 fun record_split_tac i st =
   let
-    val thy = Thm.theory_of_thm st;
-
     val has_rec = exists_Const
       (fn (s, Type (_, [Type (_, [T, _]), _])) =>
           (s = "all" orelse s = "All") andalso is_recT T
@@ -1695,40 +1632,16 @@
   in compose_tac (false, rule'', nprems_of rule) i st end;
 
 
-(*!!x1 ... xn. ... ==> EX x1 ... xn. P x1 ... xn;
-  instantiates x1 ... xn with parameters x1 ... xn*)
-fun ex_inst_tac i st =
-  let
-    val thy = Thm.theory_of_thm st;
-    val g = nth (prems_of st) (i - 1);   (* FIXME SUBGOAL *)
-    val params = Logic.strip_params g;
-    val exI' = Thm.lift_rule (Thm.cprem_of st i) exI;
-    val _ $ (_ $ x) = Logic.strip_assums_concl (hd (prems_of exI'));
-    val cx = cterm_of thy (fst (strip_comb x));
-  in
-    Seq.single (Library.foldl (fn (st, v) =>
-      Seq.hd
-        (compose_tac
-          (false,
-            cterm_instantiate [(cx, cterm_of thy (list_abs (params, Bound v)))] exI', 1) i st))
-        (st, (length params - 1) downto 0))
-  end;
-
-fun extension_definition full name fields names alphas zeta moreT more vars thy =
+fun extension_definition name fields alphas zeta moreT more vars thy =
   let
     val base = Long_Name.base_name;
     val fieldTs = (map snd fields);
     val alphas_zeta = alphas @ [zeta];
     val alphas_zetaTs = map (fn n => TFree (n, HOLogic.typeS)) alphas_zeta;
-    val vT = TFree (Name.variant alphas_zeta "'v", HOLogic.typeS);
     val extT_name = suffix ext_typeN name
     val extT = Type (extT_name, alphas_zetaTs);
-    val fields_more = fields @ [(full moreN, moreT)];
     val fields_moreTs = fieldTs @ [moreT];
-    val bfields_more = map (apfst base) fields_more;
-    val r = Free (rN, extT);
-    val len = length fields;
-    val idxms = 0 upto len;
+
 
     (*before doing anything else, create the tree of new types
       that will back the record extension*)
@@ -1739,7 +1652,7 @@
       let
         val suff = if i = 0 then ext_typeN else inner_typeN ^ string_of_int i;
         val nm = suffix suff (Long_Name.base_name name);
-        val (isom, cons, thy') =
+        val (_, cons, thy') =
           IsTupleSupport.add_istuple_type
             (nm, alphas_zeta) (fastype_of left, fastype_of right) thy;
       in
@@ -1763,7 +1676,7 @@
             build_meta_tree_type i' thy' composites more
           end
         else
-          let val (term, (thy', i')) = mk_istuple (mktreeV vars, more) (thy, 0)
+          let val (term, (thy', _)) = mk_istuple (mktreeV vars, more) (thy, 0)
           in (term, thy') end
       end;
 
@@ -1795,16 +1708,15 @@
     val ([ext_def], defs_thy) =
       timeit_msg "record extension constructor def:" mk_defs;
 
+
     (* prepare propositions *)
+
     val _ = timing_msg "record extension preparing propositions";
     val vars_more = vars @ [more];
-    val named_vars_more = (names @ [full moreN]) ~~ vars_more;
     val variants = map (fn Free (x, _) => x) vars_more;
     val ext = mk_ext vars_more;
     val s = Free (rN, extT);
-    val w = Free (wN, extT);
     val P = Free (Name.variant variants "P", extT --> HOLogic.boolT);
-    val C = Free (Name.variant variants "C", HOLogic.boolT);
     val intros_tac = IsTupleSupport.istuple_intros_tac defs_thy;
 
     val inject_prop =
@@ -1819,27 +1731,18 @@
     val induct_prop =
       (All (map dest_Free vars_more) (Trueprop (P $ ext)), Trueprop (P $ s));
 
-    val cases_prop =
-      All (map dest_Free vars_more)
-        (Trueprop (HOLogic.mk_eq (s, ext)) ==> Trueprop C)
-      ==> Trueprop C;
-
     val split_meta_prop =
-      let val P = Free (Name.variant variants "P", extT-->Term.propT) in
+      let val P = Free (Name.variant variants "P", extT --> Term.propT) in
         Logic.mk_equals
          (All [dest_Free s] (P $ s), All (map dest_Free vars_more) (P $ ext))
       end;
 
-    fun prove stndrd = quick_and_dirty_prove stndrd defs_thy;
     val prove_standard = quick_and_dirty_prove true defs_thy;
-    fun prove_simp stndrd simps =
-      let val tac = simp_all_tac HOL_ss simps
-      in fn prop => prove stndrd [] prop (K tac) end;
 
     fun inject_prf () =
       simplify HOL_ss
         (prove_standard [] inject_prop
-          (fn prems =>
+          (fn _ =>
             EVERY
              [simp_tac (HOL_basic_ss addsimps [ext_def]) 1,
               REPEAT_DETERM (resolve_tac [refl_conj_eq] 1 ORELSE
@@ -1872,7 +1775,7 @@
 
     fun split_meta_prf () =
       prove_standard [] split_meta_prop
-        (fn prems =>
+        (fn _ =>
           EVERY
            [rtac equal_intr_rule 1, Goal.norm_hhf_tac 1,
             etac meta_allE 1, atac 1,
@@ -1909,8 +1812,8 @@
   | chop_last [x] = ([], x)
   | chop_last (x :: xs) = let val (tl, l) = chop_last xs in (x :: tl, l) end;
 
-fun subst_last s [] = error "subst_last: list should not be empty"
-  | subst_last s [x] = [s]
+fun subst_last _ [] = error "subst_last: list should not be empty"
+  | subst_last s [_] = [s]
   | subst_last s (x :: xs) = x :: subst_last s xs;
 
 
@@ -1965,7 +1868,6 @@
     val parent_variants = Name.variant_list [moreN, rN, rN ^ "'", wN] (map base parent_names);
     val parent_vars = ListPair.map Free (parent_variants, parent_types);
     val parent_len = length parents;
-    val parents_idx = (map #name parents) ~~ (0 upto (parent_len - 1));
 
     val fields = map (apfst full) bfields;
     val names = map fst fields;
@@ -1979,13 +1881,10 @@
         (map fst bfields);
     val vars = ListPair.map Free (variants, types);
     val named_vars = names ~~ vars;
-    val idxs = 0 upto (len - 1);
     val idxms = 0 upto len;
 
     val all_fields = parent_fields @ fields;
-    val all_names = parent_names @ names;
     val all_types = parent_types @ types;
-    val all_len = parent_fields_len + len;
     val all_variants = parent_variants @ variants;
     val all_vars = parent_vars @ vars;
     val all_named_vars = (parent_names ~~ parent_vars) @ named_vars;
@@ -1997,7 +1896,6 @@
     val full_moreN = full moreN;
     val bfields_more = bfields @ [(moreN, moreT)];
     val fields_more = fields @ [(full_moreN, moreT)];
-    val vars_more = vars @ [more];
     val named_vars_more = named_vars @ [(full_moreN, more)];
     val all_vars_more = all_vars @ [more];
     val all_named_vars_more = all_named_vars @ [(full_moreN, more)];
@@ -2008,7 +1906,7 @@
     val (extension_thy, extT, ext_induct, ext_inject, ext_split, ext_def) =
       thy
       |> Sign.add_path bname
-      |> extension_definition full extN fields names alphas_ext zeta moreT more vars;
+      |> extension_definition extN fields alphas_ext zeta moreT more vars;
 
     val _ = timing_msg "record preparing definitions";
     val Type extension_scheme = extT;
@@ -2080,16 +1978,6 @@
 
     (* prepare definitions *)
 
-    fun parent_more s =
-      if null parents then s
-      else mk_sel s (Long_Name.qualify (#name (List.last parents)) moreN, extT);
-
-    fun parent_more_upd v s =
-      if null parents then v $ s
-      else
-        let val mp = Long_Name.qualify (#name (List.last parents)) moreN;
-        in mk_upd updateN mp v s end;
-
     (*record (scheme) type abbreviation*)
     val recordT_specs =
       [(Binding.name (suffix schemeN bname), alphas @ [zeta], rec_schemeT0, Syntax.NoSyn),
@@ -2233,14 +2121,12 @@
 
     (*cases*)
     val cases_scheme_prop =
-      (All (map dest_Free all_vars_more)
-        (Trueprop (HOLogic.mk_eq (r0, r_rec0)) ==> Trueprop C))
-      ==> Trueprop C;
+      (All (map dest_Free all_vars_more) ((r0 === r_rec0) ==> Trueprop C))
+        ==> Trueprop C;
 
     val cases_prop =
-      (All (map dest_Free all_vars)
-        (Trueprop (HOLogic.mk_eq (r_unit0, r_rec_unit0)) ==> Trueprop C))
-       ==> Trueprop C;
+      (All (map dest_Free all_vars) ((r_unit0 === r_rec_unit0) ==> Trueprop C))
+         ==> Trueprop C;
 
     (*split*)
     val split_meta_prop =
@@ -2359,7 +2245,7 @@
         val init_ss = HOL_basic_ss addsimps ext_defs;
       in
         prove_standard [] surjective_prop
-          (fn prems =>
+          (fn _ =>
             EVERY
              [rtac surject_assist_idE 1,
               simp_tac init_ss 1,
@@ -2369,7 +2255,7 @@
 
     fun split_meta_prf () =
       prove false [] split_meta_prop
-        (fn prems =>
+        (fn _ =>
           EVERY
            [rtac equal_intr_rule 1, Goal.norm_hhf_tac 1,
             etac meta_allE 1, atac 1,
@@ -2423,7 +2309,7 @@
         val so' = named_cterm_instantiate ([(P_nm, not_P)]) split_object;
         val so'' = simplify ss so';
       in
-        prove_standard [] split_ex_prop (fn prems => resolve_tac [so''] 1)
+        prove_standard [] split_ex_prop (fn _ => resolve_tac [so''] 1)
       end;
     val split_ex = timeit_msg "record split_ex proof:" split_ex_prf;
 
--- a/src/HOL/Tools/transfer.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/HOL/Tools/transfer.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -1,5 +1,5 @@
 (*  Author:     Amine Chaieb, University of Cambridge, 2009
-                Jeremy Avigad, Carnegie Mellon University
+    Author:     Jeremy Avigad, Carnegie Mellon University
 *)
 
 signature TRANSFER =
@@ -14,10 +14,8 @@
 structure Transfer : TRANSFER =
 struct
 
-val eq_thm = Thm.eq_thm;
-
 type entry = {inj : thm list, emb : thm list, ret : thm list, cong : thm list,
-  guess : bool, hints : string list}; 
+  guess : bool, hints : string list};
 type data = simpset * (thm * entry) list;
 
 structure Data = GenericDataFun
@@ -26,36 +24,34 @@
   val empty = (HOL_ss, []);
   val extend  = I;
   fun merge _ ((ss1, e1), (ss2, e2)) =
-    (merge_ss (ss1, ss2), AList.merge eq_thm (K true) (e1, e2));
+    (merge_ss (ss1, ss2), AList.merge Thm.eq_thm (K true) (e1, e2));
 );
 
 val get = Data.get o Context.Proof;
 
-fun del_data key = apsnd (remove (eq_fst eq_thm) (key, []));
+fun del_data key = apsnd (remove (eq_fst Thm.eq_thm) (key, []));
 
 val del = Thm.declaration_attribute (Data.map o del_data);
-val add_ss = Thm.declaration_attribute 
+val add_ss = Thm.declaration_attribute
    (fn th => Data.map (fn (ss,data) => (ss addsimps [th], data)));
 
-val del_ss = Thm.declaration_attribute 
+val del_ss = Thm.declaration_attribute
    (fn th => Data.map (fn (ss,data) => (ss delsimps [th], data)));
 
 val transM_pat = (Thm.dest_arg1 o Thm.dest_arg o cprop_of) @{thm TransferMorphism_def};
 
 fun merge_update eq m (k,v) [] = [(k,v)]
-  | merge_update eq m (k,v) ((k',v')::al) = 
+  | merge_update eq m (k,v) ((k',v')::al) =
            if eq (k,k') then (k',m (v,v')):: al else (k',v') :: merge_update eq m (k,v) al
 
-fun C f x y = f y x
-
-fun simpset_of_entry injonly {inj = inj, emb = emb, ret = ret, cong = cg, guess = g, hints = hints} = 
+fun simpset_of_entry injonly {inj = inj, emb = emb, ret = ret, cong = cg, guess = g, hints = hints} =
  HOL_ss addsimps inj addsimps (if injonly then [] else emb@ret) addcongs cg;
 
-fun basic_transfer_rule injonly a0 D0 e leave ctxt0 th = 
- let 
+fun basic_transfer_rule injonly a0 D0 e leave ctxt0 th =
+ let
   val ([a,D], ctxt) = apfst (map Drule.dest_term o snd) (Variable.import true (map Drule.mk_term [a0, D0]) ctxt0)
-  val (aT,bT) = 
-     let val T = typ_of (ctyp_of_term a) 
+  val (aT,bT) =
+     let val T = typ_of (ctyp_of_term a)
      in (Term.range_type T, Term.domain_type T)
      end
   val ctxt' = (Variable.declare_term (term_of a) o Variable.declare_term (term_of D) o Variable.declare_thm th) ctxt
@@ -65,60 +61,64 @@
   val cfis = map ((cterm_of o ProofContext.theory_of) ctxt'' o (fn n => Free (n, bT))) ins
   val cis = map (Thm.capply a) cfis
   val (hs,ctxt''') = Assumption.add_assumes (map (fn ct => Thm.capply @{cterm "Trueprop"} (Thm.capply D ct)) cfis) ctxt''
-  val th1 = Drule.cterm_instantiate (cns~~ cis) th
-  val th2 = fold (C implies_elim) hs (fold_rev implies_intr (map cprop_of hs) th1)
-  val th3 = Simplifier.asm_full_simplify (Simplifier.context ctxt''' (simpset_of_entry injonly e)) 
+  val th1 = Drule.cterm_instantiate (cns ~~ cis) th
+  val th2 = fold Thm.elim_implies hs (fold_rev implies_intr (map cprop_of hs) th1)
+  val th3 = Simplifier.asm_full_simplify (Simplifier.context ctxt''' (simpset_of_entry injonly e))
                                          (fold_rev implies_intr (map cprop_of hs) th2)
 in hd (Variable.export ctxt''' ctxt0 [th3]) end;
 
 local
-fun transfer_ruleh a D leave ctxt th = 
+fun transfer_ruleh a D leave ctxt th =
  let val (ss,al) = get ctxt
      val a0 = cterm_of (ProofContext.theory_of ctxt) a
      val D0 = cterm_of (ProofContext.theory_of ctxt) D
-     fun h (th', e) = let val (a',D') = (Thm.dest_binop o Thm.dest_arg o cprop_of) th' 
+     fun h (th', e) = let val (a',D') = (Thm.dest_binop o Thm.dest_arg o cprop_of) th'
                  in if a0 aconvc a' andalso D0 aconvc D' then SOME e else NONE
                  end
  in case get_first h al of
       SOME e => basic_transfer_rule false a0 D0 e leave ctxt th
     | NONE => error "Transfer: corresponding instance not found in context-data"
  end
-in fun transfer_rule (a,D) leave (gctxt,th) = 
+in fun transfer_rule (a,D) leave (gctxt,th) =
    (gctxt, transfer_ruleh a D leave (Context.proof_of gctxt) th)
 end;
 
 fun  splits P [] = []
-   | splits P (xxs as (x::xs)) = 
+   | splits P (xxs as (x::xs)) =
     let val pss = filter (P x) xxs
         val qss = filter_out (P x) xxs
     in if null pss then [qss] else if null qss then [pss] else pss:: splits P qss
     end
 
-fun all_transfers leave (gctxt,th) = 
- let 
+fun all_transfers leave (gctxt,th) =
+ let
   val ctxt = Context.proof_of gctxt
   val tys = map snd (Term.add_vars (prop_of th) [])
   val _ = if null tys then error "transfer: Unable to guess instance" else ()
-  val tyss = splits (curry Type.could_unify) tys 
+  val tyss = splits (curry Type.could_unify) tys
   val get_ty = typ_of o ctyp_of_term o fst o Thm.dest_binop o Thm.dest_arg o cprop_of
   val get_aD = Thm.dest_binop o Thm.dest_arg o cprop_of
-  val insts = 
-    map_filter (fn tys => 
-          get_first (fn (k,ss) => if Type.could_unify (hd tys, range_type (get_ty k)) 
-                                  then SOME (get_aD k, ss) 
-                                  else NONE) (snd (get ctxt))) tyss
-  val _ = if null insts then error "Transfer guesser: there were no possible instances, use direction: in order to provide a direction" else ()
+  val insts =
+    map_filter (fn tys =>
+      get_first (fn (k,ss) =>
+        if Type.could_unify (hd tys, range_type (get_ty k))
+        then SOME (get_aD k, ss)
+        else NONE) (snd (get ctxt))) tyss
+  val _ =
+    if null insts then
+      error "Transfer guesser: there were no possible instances, use direction: in order to provide a direction"
+    else ()
   val ths = map  (fn ((a,D),e) => basic_transfer_rule false a D e leave ctxt th) insts
   val cth = Conjunction.intr_balanced ths
  in (gctxt, cth)
  end;
 
-fun transfer_rule_by_hint ls leave (gctxt,th) = 
- let 
+fun transfer_rule_by_hint ls leave (gctxt,th) =
+ let
   val ctxt = Context.proof_of gctxt
   val get_aD = Thm.dest_binop o Thm.dest_arg o cprop_of
-  val insts = 
-    map_filter (fn (k,e) => if exists (fn l => l mem_string (#hints e)) ls 
+  val insts =
+    map_filter (fn (k,e) => if exists (member (op =) (#hints e)) ls
           then SOME (get_aD k, e) else NONE)
         (snd (get ctxt))
   val _ = if null insts then error "Transfer: No labels provided are stored in the context" else ()
@@ -128,53 +128,58 @@
  end;
 
 
-fun transferred_attribute ls NONE leave = 
+fun transferred_attribute ls NONE leave =
          if null ls then all_transfers leave else transfer_rule_by_hint ls leave
   | transferred_attribute _ (SOME (a,D)) leave = transfer_rule (a,D) leave
 
-                                                    (* Add data to the context *)
+
+(* Add data to the context *)
+
 fun gen_merge_entries {inj = inj0, emb = emb0, ret = ret0, cong = cg0, guess = g0, hints = hints0}
-                      ({inj = inj1, emb = emb1, ret = ret1, cong = cg1, guess = g1, hints = hints1}, 
+                      ({inj = inj1, emb = emb1, ret = ret1, cong = cg1, guess = g1, hints = hints1},
                        {inj = inj2, emb = emb2, ret = ret2, cong = cg2, guess = g2, hints = hints2})
- = 
+ =
  let fun h xs0 xs ys = subtract Thm.eq_thm xs0 (merge Thm.eq_thm (xs,ys)) in
- {inj = h inj0 inj1 inj2, emb = h emb0 emb1 emb2, 
+ {inj = h inj0 inj1 inj2, emb = h emb0 emb1 emb2,
   ret = h ret0 ret1 ret2, cong = h cg0 cg1 cg2, guess = g1 andalso g2,
-  hints = subtract (op = : string*string -> bool) hints0 
+  hints = subtract (op = : string*string -> bool) hints0
             (hints1 union_string hints2)}
  end;
 
 local
  val h = curry (merge Thm.eq_thm)
 in
-fun merge_entries ({inj = inj1, emb = emb1, ret = ret1, cong = cg1, guess = g1, hints = hints1}, 
-                   {inj = inj2, emb = emb2, ret = ret2, cong = cg2, guess = g2, hints = hints2}) = 
+fun merge_entries ({inj = inj1, emb = emb1, ret = ret1, cong = cg1, guess = g1, hints = hints1},
+                   {inj = inj2, emb = emb2, ret = ret2, cong = cg2, guess = g2, hints = hints2}) =
     {inj = h inj1 inj2, emb = h emb1 emb2, ret = h ret1 ret2, cong = h cg1 cg2, guess = g1 andalso g2, hints = hints1 union_string hints2}
-end; 
+end;
 
 fun add ((inja,injd), (emba,embd), (reta,retd), (cga,cgd), g, (hintsa, hintsd)) =
   Thm.declaration_attribute (fn key => fn context => context |> Data.map
-   (fn (ss, al) => 
+   (fn (ss, al) =>
      let
-      val _ = ((let val _ = Thm.match (transM_pat, (Thm.dest_arg o cprop_of) key) 
-                in 0 end) 
-                handle MATCH => error "Attribute expected Theorem of the form : TransferMorphism A a B b")
+      val _ = Thm.match (transM_pat, Thm.dest_arg (Thm.cprop_of key))
+        handle Pattern.MATCH =>
+          error "Attribute expected Theorem of the form : TransferMorphism A a B b"
       val e0 = {inj = inja, emb = emba, ret = reta, cong = cga, guess = g, hints = hintsa}
       val ed = {inj = injd, emb = embd, ret = retd, cong = cgd, guess = g, hints = hintsd}
-      val entry = 
-        if g then 
+      val entry =
+        if g then
          let val (a0,D0) = (Thm.dest_binop o Thm.dest_arg o cprop_of) key
              val ctxt0 = ProofContext.init (Thm.theory_of_thm key)
-             val inj' = if null inja then #inj (case AList.lookup eq_thm al key of SOME e => e | NONE => error "Transfer: can not generate return rules on the fly, either add injectivity axiom or force manual mode with mode: manual") 
-                        else inja
+             val inj' =
+               if null inja then
+                #inj
+                  (case AList.lookup Thm.eq_thm al key of SOME e => e
+                  | NONE => error "Transfer: can not generate return rules on the fly, either add injectivity axiom or force manual mode with mode: manual")
+               else inja
              val ret' = merge Thm.eq_thm (reta,  map (fn th => basic_transfer_rule true a0 D0 {inj = inj', emb = [], ret = [], cong = cga, guess = g, hints = hintsa} [] ctxt0 th RS sym) emba)
-         in {inj = inja, emb = emba, ret = ret', cong = cga, guess = g, hints = hintsa} end 
+         in {inj = inja, emb = emba, ret = ret', cong = cga, guess = g, hints = hintsa} end
         else e0
-    in (ss, merge_update eq_thm (gen_merge_entries ed) (key, entry) al)
+    in (ss, merge_update Thm.eq_thm (gen_merge_entries ed) (key, entry) al)
     end));
 
 
-
 (* concrete syntax *)
 
 local
@@ -199,7 +204,7 @@
 
 val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat
 val terms = thms >> map Drule.dest_term
-val types = thms >> (Logic.dest_type o HOLogic.dest_Trueprop o prop_of o hd) 
+val types = thms >> (Logic.dest_type o HOLogic.dest_Trueprop o prop_of o hd)
 val name = Scan.lift Args.name
 val names = Scan.repeat (Scan.unless any_keyword name)
 fun optional scan = Scan.optional scan []
--- a/src/Pure/Concurrent/future.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/Concurrent/future.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -30,6 +30,7 @@
   type task = Task_Queue.task
   type group = Task_Queue.group
   val is_worker: unit -> bool
+  val worker_task: unit -> Task_Queue.task option
   val worker_group: unit -> Task_Queue.group option
   type 'a future
   val task_of: 'a future -> task
@@ -71,6 +72,7 @@
 end;
 
 val is_worker = is_some o thread_data;
+val worker_task = Option.map #2 o thread_data;
 val worker_group = Option.map #3 o thread_data;
 
 
@@ -347,7 +349,8 @@
   | SOME res => res);
 
 fun join_wait x =
-  Synchronized.guarded_access (result_of x) (fn NONE => NONE | some => SOME ((), some));
+  Synchronized.guarded_access (result_of x)
+    (fn NONE => NONE | some => SOME ((), some));
 
 fun join_next deps = (*requires SYNCHRONIZED*)
   if null deps then NONE
@@ -357,10 +360,14 @@
     | (NONE, deps') => (worker_wait work_finished; join_next deps')
     | (SOME work, deps') => SOME (work, deps'));
 
-fun join_work deps =
-  (case SYNCHRONIZED "join" (fn () => join_next deps) of
-    NONE => ()
-  | SOME (work, deps') => (execute "join" work; join_work deps'));
+fun execute_work NONE = ()
+  | execute_work (SOME (work, deps')) = (execute "join" work; join_work deps')
+and join_work deps =
+  execute_work (SYNCHRONIZED "join" (fn () => join_next deps));
+
+fun join_depend task deps =
+  execute_work (SYNCHRONIZED "join" (fn () =>
+    (Unsynchronized.change queue (Task_Queue.depend task deps); join_next deps)));
 
 in
 
@@ -368,11 +375,11 @@
   if forall is_finished xs then map get_result xs
   else if Multithreading.self_critical () then
     error "Cannot join future values within critical section"
-  else uninterruptible (fn _ => fn () =>
-     (if is_worker ()
-      then join_work (map task_of xs)
-      else List.app join_wait xs;
-      map get_result xs)) ();
+  else
+    (case worker_task () of
+      SOME task => join_depend task (map task_of xs)
+    | NONE => List.app join_wait xs;
+    map get_result xs);
 
 end;
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Concurrent/lazy.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -0,0 +1,58 @@
+(*  Title:      Pure/Concurrent/lazy.ML
+    Author:     Makarius
+
+Lazy evaluation based on futures.
+*)
+
+signature LAZY =
+sig
+  type 'a lazy
+  val peek: 'a lazy -> 'a Exn.result option
+  val lazy: (unit -> 'a) -> 'a lazy
+  val value: 'a -> 'a lazy
+  val force_result: 'a lazy -> 'a Exn.result
+  val force: 'a lazy -> 'a
+  val map_force: ('a -> 'b) -> 'a lazy -> 'b lazy
+end;
+
+structure Lazy: LAZY =
+struct
+
+(* datatype *)
+
+datatype 'a expr =
+  Expr of unit -> 'a |
+  Future of 'a future;
+
+abstype 'a lazy = Lazy of 'a expr Synchronized.var
+with
+
+fun peek (Lazy var) =
+  (case Synchronized.value var of
+    Expr _ => NONE
+  | Future x => Future.peek x);
+
+fun lazy e = Lazy (Synchronized.var "lazy" (Expr e));
+fun value a = Lazy (Synchronized.var "lazy" (Future (Future.value a)));
+
+
+(* force result *)
+
+fun force_result (Lazy var) =
+  (case peek (Lazy var) of
+    SOME res => res
+  | NONE =>
+      Synchronized.guarded_access var
+        (fn Expr e => let val x = Future.fork e in SOME (x, Future x) end
+          | Future x => SOME (x, Future x))
+      |> Future.join_result);
+
+fun force r = Exn.release (force_result r);
+
+fun map_force f = value o f o force;
+
+end;
+end;
+
+type 'a lazy = 'a Lazy.lazy;
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Concurrent/lazy_sequential.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -0,0 +1,43 @@
+(*  Title:      Pure/Concurrent/lazy_sequential.ML
+    Author:     Florian Haftmann and Makarius, TU Muenchen
+
+Lazy evaluation with memoing (sequential version).
+*)
+
+structure Lazy: LAZY =
+struct
+
+(* datatype *)
+
+datatype 'a expr =
+  Expr of unit -> 'a |
+  Result of 'a Exn.result;
+
+abstype 'a lazy = Lazy of 'a expr Unsynchronized.ref
+with
+
+fun peek (Lazy r) =
+  (case ! r of
+    Expr _ => NONE
+  | Result x => SOME x);
+
+fun lazy e = Lazy (Unsynchronized.ref (Expr e));
+fun value a = Lazy (Unsynchronized.ref (Result (Exn.Result a)));
+
+
+(* force result *)
+
+fun force_result (Lazy r) =
+  (case ! r of
+    Expr e => Exn.capture e ()
+  | Result res => res);
+
+fun force r = Exn.release (force_result r);
+
+fun map_force f = value o f o force;
+
+end;
+end;
+
+type 'a lazy = 'a Lazy.lazy;
+
--- a/src/Pure/Concurrent/par_list_dummy.ML	Thu Oct 01 20:49:46 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,16 +0,0 @@
-(*  Title:      Pure/Concurrent/par_list_dummy.ML
-    Author:     Makarius
-
-Dummy version of parallel list combinators -- plain sequential evaluation.
-*)
-
-structure Par_List: PAR_LIST =
-struct
-
-val map = map;
-val get_some = get_first;
-val find_some = find_first;
-val exists = exists;
-val forall = forall;
-
-end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Concurrent/par_list_sequential.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -0,0 +1,16 @@
+(*  Title:      Pure/Concurrent/par_list_sequential.ML
+    Author:     Makarius
+
+Dummy version of parallel list combinators -- plain sequential evaluation.
+*)
+
+structure Par_List: PAR_LIST =
+struct
+
+val map = map;
+val get_some = get_first;
+val find_some = find_first;
+val exists = exists;
+val forall = forall;
+
+end;
--- a/src/Pure/Concurrent/synchronized_dummy.ML	Thu Oct 01 20:49:46 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,25 +0,0 @@
-(*  Title:      Pure/Concurrent/synchronized_dummy.ML
-    Author:     Makarius
-
-Dummy version of state variables -- plain refs for sequential access.
-*)
-
-structure Synchronized: SYNCHRONIZED =
-struct
-
-datatype 'a var = Var of 'a Unsynchronized.ref;
-
-fun var _ x = Var (Unsynchronized.ref x);
-fun value (Var var) = ! var;
-
-fun timed_access (Var var) _ f =
-  (case f (! var) of
-    SOME (y, x') => (var := x'; SOME y)
-  | NONE => Thread.unavailable ());
-
-fun guarded_access var f = the (timed_access var (K NONE) f);
-
-fun change_result var f = guarded_access var (SOME o f);
-fun change var f = change_result var (fn x => ((), f x));
-
-end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Concurrent/synchronized_sequential.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -0,0 +1,27 @@
+(*  Title:      Pure/Concurrent/synchronized_sequential.ML
+    Author:     Makarius
+
+Sequential version of state variables -- plain refs.
+*)
+
+structure Synchronized: SYNCHRONIZED =
+struct
+
+abstype 'a var = Var of 'a Unsynchronized.ref
+with
+
+fun var _ x = Var (Unsynchronized.ref x);
+fun value (Var var) = ! var;
+
+fun timed_access (Var var) _ f =
+  (case f (! var) of
+    SOME (y, x') => (var := x'; SOME y)
+  | NONE => Thread.unavailable ());
+
+fun guarded_access var f = the (timed_access var (K NONE) f);
+
+fun change_result var f = guarded_access var (SOME o f);
+fun change var f = change_result var (fn x => ((), f x));
+
+end;
+end;
--- a/src/Pure/Concurrent/task_queue.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/Concurrent/task_queue.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -27,6 +27,7 @@
   val enqueue: group -> task list -> int -> (bool -> bool) -> queue -> (task * bool) * queue
   val extend: task -> (bool -> bool) -> queue -> queue option
   val dequeue: Thread.thread -> queue -> (task * group * (bool -> bool) list) option * queue
+  val depend: task -> task list -> queue -> queue
   val dequeue_towards: Thread.thread -> task list -> queue ->
     (((task * group * (bool -> bool) list) option * task list) * queue)
   val finish: task -> queue -> bool * queue
@@ -101,6 +102,11 @@
 fun add_job task dep (jobs: jobs) =
   Task_Graph.add_edge (dep, task) jobs handle Task_Graph.UNDEF _ => jobs;
 
+fun add_dep task dep (jobs: jobs) =
+  if Task_Graph.is_edge jobs (task, dep) then
+    raise Fail "Cyclic dependency of future tasks"
+  else add_job task dep jobs;
+
 fun get_deps (jobs: jobs) task =
   Task_Graph.imm_preds jobs task handle Task_Graph.UNDEF _ => [];
 
@@ -125,7 +131,7 @@
 fun status (Queue {jobs, ...}) =
   let
     val (x, y, z) =
-      Task_Graph.fold (fn (task, ((_, job), (deps, _))) => fn (x, y, z) =>
+      Task_Graph.fold (fn (_, ((_, job), (deps, _))) => fn (x, y, z) =>
           (case job of
             Job _ => if null deps then (x + 1, y, z) else (x, y + 1, z)
           | Running _ => (x, y, z + 1)))
@@ -205,6 +211,9 @@
 
 (* dequeue_towards -- adhoc dependencies *)
 
+fun depend task deps (Queue {groups, jobs, ...}) =
+  make_queue groups (fold (add_dep task) deps jobs) Unknown;
+
 fun dequeue_towards thread deps (queue as Queue {groups, jobs, ...}) =
   let
     fun ready task =
--- a/src/Pure/General/heap.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/General/heap.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -78,8 +78,8 @@
 
 nonfix upto;
 
-fun upto _ (h as Empty) = ([], h)
-  | upto limit (h as Heap (_, x, a, b)) =
+fun upto _ Empty = ([], Empty)
+  | upto limit (h as Heap (_, x, _, _)) =
       (case ord (x, limit) of
         GREATER => ([], h)
       | _ => upto limit (delete_min h) |>> cons x);
--- a/src/Pure/General/lazy.ML	Thu Oct 01 20:49:46 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,65 +0,0 @@
-(*  Title:      Pure/General/lazy.ML
-    Author:     Florian Haftmann and Makarius, TU Muenchen
-
-Lazy evaluation with memoing.  Concurrency may lead to multiple
-attempts of evaluation -- the first result persists.
-*)
-
-signature LAZY =
-sig
-  type 'a lazy
-  val same: 'a lazy * 'a lazy -> bool
-  val lazy: (unit -> 'a) -> 'a lazy
-  val value: 'a -> 'a lazy
-  val peek: 'a lazy -> 'a Exn.result option
-  val force_result: 'a lazy -> 'a Exn.result
-  val force: 'a lazy -> 'a
-  val map_force: ('a -> 'a) -> 'a lazy -> 'a lazy
-end
-
-structure Lazy :> LAZY =
-struct
-
-(* datatype *)
-
-datatype 'a T =
-  Lazy of unit -> 'a |
-  Result of 'a Exn.result;
-
-type 'a lazy = 'a T Unsynchronized.ref;
-
-fun same (r1: 'a lazy, r2) = r1 = r2;
-
-fun lazy e = Unsynchronized.ref (Lazy e);
-fun value x = Unsynchronized.ref (Result (Exn.Result x));
-
-fun peek r =
-  (case ! r of
-    Lazy _ => NONE
-  | Result res => SOME res);
-
-
-(* force result *)
-
-fun force_result r =
-  let
-    (*potentially concurrent evaluation*)
-    val res =
-      (case ! r of
-        Lazy e => Exn.capture e ()
-      | Result res => res);
-    (*assign result -- first one persists*)
-    val res' = NAMED_CRITICAL "lazy" (fn () =>
-      (case ! r of
-        Lazy _ => (r := Result res; res)
-      | Result res' => res'));
-  in res' end;
-
-fun force r = Exn.release (force_result r);
-
-fun map_force f = value o f o force;
-
-end;
-
-type 'a lazy = 'a Lazy.lazy;
-
--- a/src/Pure/General/pretty.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/General/pretty.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -197,11 +197,11 @@
 fun setdepth dp = (depth := dp);
 
 local
-  fun pruning dp (Block (m, bes, indent, wd)) =
+  fun pruning dp (Block (m, bes, indent, _)) =
         if dp > 0
         then block_markup m (indent, map (pruning (dp - 1)) bes)
         else str "..."
-    | pruning dp e = e
+    | pruning _ e = e;
 in
   fun prune e = if ! depth > 0 then pruning (! depth) e else e;
 end;
@@ -219,7 +219,7 @@
   pos = 0,
   nl = 0};
 
-fun newline {tx, ind, pos, nl} : text =
+fun newline {tx, ind = _, pos = _, nl} : text =
  {tx = Buffer.add (Output.output "\n") tx,
   ind = Buffer.empty,
   pos = 0,
@@ -250,22 +250,22 @@
 (*Add the lengths of the expressions until the next Break; if no Break then
   include "after", to account for text following this block.*)
 fun breakdist (Block (_, _, _, len) :: es, after) = len + breakdist (es, after)
-  | breakdist (String (s, len) :: es, after) = len + breakdist (es, after)
-  | breakdist (Break _ :: es, after) = 0
+  | breakdist (String (_, len) :: es, after) = len + breakdist (es, after)
+  | breakdist (Break _ :: _, _) = 0
   | breakdist ([], after) = after;
 
 (*Search for the next break (at this or higher levels) and force it to occur.*)
 fun forcenext [] = []
-  | forcenext (Break (_, wd) :: es) = Break (true, 0) :: es
+  | forcenext (Break _ :: es) = Break (true, 0) :: es
   | forcenext (e :: es) = e :: forcenext es;
 
 (*es is list of expressions to print;
   blockin is the indentation of the current block;
   after is the width of the following context until next break.*)
 fun format ([], _, _) text = text
-  | format (e :: es, block as (blockind, blockin), after) (text as {ind, pos, nl, ...}) =
+  | format (e :: es, block as (_, blockin), after) (text as {ind, pos, nl, ...}) =
       (case e of
-        Block ((bg, en), bes, indent, wd) =>
+        Block ((bg, en), bes, indent, _) =>
           let
             val {emergencypos, ...} = ! margin_info;
             val pos' = pos + indent;
--- a/src/Pure/General/scan.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/General/scan.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -273,7 +273,7 @@
 val empty_lexicon = Lexicon Symtab.empty;
 
 fun is_literal _ [] = false
-  | is_literal (Lexicon tab) (chs as c :: cs) =
+  | is_literal (Lexicon tab) (c :: cs) =
       (case Symtab.lookup tab c of
         SOME (tip, lex) => tip andalso null cs orelse is_literal lex cs
       | NONE => false);
@@ -286,7 +286,7 @@
     fun finish (SOME (res, rest)) = (rev res, rest)
       | finish NONE = raise FAIL NONE;
     fun scan _ res (Lexicon tab) [] = if Symtab.is_empty tab then finish res else raise MORE NONE
-      | scan path res (Lexicon tab) (chs as c :: cs) =
+      | scan path res (Lexicon tab) (c :: cs) =
           (case Symtab.lookup tab (fst c) of
             SOME (tip, lex) =>
               let val path' = c :: path
@@ -300,7 +300,7 @@
 fun extend_lexicon chrs lexicon =
   let
     fun ext [] lex = lex
-      | ext (chs as c :: cs) (Lexicon tab) =
+      | ext (c :: cs) (Lexicon tab) =
           (case Symtab.lookup tab c of
             SOME (tip, lex) => Lexicon (Symtab.update (c, (tip orelse null cs, ext cs lex)) tab)
           | NONE => Lexicon (Symtab.update (c, (null cs, ext cs empty_lexicon)) tab));
--- a/src/Pure/General/symbol_pos.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/General/symbol_pos.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -161,7 +161,7 @@
 
 fun pad [] = []
   | pad [(s, _)] = [s]
-  | pad ((s1, pos1) :: (rest as (s2, pos2) :: _)) =
+  | pad ((s1, pos1) :: (rest as (_, pos2) :: _)) =
       let
         val end_pos1 = Position.advance s1 pos1;
         val d = Int.max (0, Position.distance_of end_pos1 pos2);
--- a/src/Pure/IsaMakefile	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/IsaMakefile	Thu Oct 01 20:52:18 2009 +0200
@@ -43,18 +43,19 @@
 Pure: $(OUT)/Pure
 
 $(OUT)/Pure: $(BOOTSTRAP_FILES) Concurrent/future.ML			\
+  Concurrent/lazy.ML Concurrent/lazy_sequential.ML			\
   Concurrent/mailbox.ML Concurrent/par_list.ML				\
-  Concurrent/par_list_dummy.ML Concurrent/simple_thread.ML		\
-  Concurrent/synchronized.ML Concurrent/synchronized_dummy.ML		\
+  Concurrent/par_list_sequential.ML Concurrent/simple_thread.ML		\
+  Concurrent/synchronized.ML Concurrent/synchronized_sequential.ML	\
   Concurrent/task_queue.ML General/alist.ML General/antiquote.ML	\
   General/balanced_tree.ML General/basics.ML General/binding.ML		\
   General/buffer.ML General/file.ML General/graph.ML General/heap.ML	\
-  General/integer.ML General/lazy.ML General/long_name.ML		\
-  General/markup.ML General/name_space.ML General/ord_list.ML		\
-  General/output.ML General/path.ML General/position.ML			\
-  General/pretty.ML General/print_mode.ML General/properties.ML		\
-  General/queue.ML General/same.ML General/scan.ML General/secure.ML	\
-  General/seq.ML General/source.ML General/stack.ML General/symbol.ML	\
+  General/integer.ML General/long_name.ML General/markup.ML		\
+  General/name_space.ML General/ord_list.ML General/output.ML		\
+  General/path.ML General/position.ML General/pretty.ML			\
+  General/print_mode.ML General/properties.ML General/queue.ML		\
+  General/same.ML General/scan.ML General/secure.ML General/seq.ML	\
+  General/source.ML General/stack.ML General/symbol.ML			\
   General/symbol_pos.ML General/table.ML General/url.ML General/xml.ML	\
   General/yxml.ML Isar/args.ML Isar/attrib.ML Isar/auto_bind.ML		\
   Isar/calculation.ML Isar/class.ML Isar/class_target.ML Isar/code.ML	\
--- a/src/Pure/Isar/args.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/Isar/args.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -283,7 +283,7 @@
 
 (** syntax wrapper **)
 
-fun syntax kind scan (src as Src ((s, args), pos)) st =
+fun syntax kind scan (Src ((s, args), pos)) st =
   (case Scan.error (Scan.finite' T.stopper (Scan.option scan)) (st, args) of
     (SOME x, (st', [])) => (x, st')
   | (_, (_, args')) =>
--- a/src/Pure/Isar/context_rules.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/Isar/context_rules.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -131,8 +131,8 @@
 (* retrieving rules *)
 
 fun untaglist [] = []
-  | untaglist [(k : int * int, x)] = [x]
-  | untaglist ((k, x) :: (rest as (k', x') :: _)) =
+  | untaglist [(_ : int * int, x)] = [x]
+  | untaglist ((k, x) :: (rest as (k', _) :: _)) =
       if k = k' then untaglist rest
       else x :: untaglist rest;
 
@@ -160,7 +160,7 @@
 (* wrappers *)
 
 fun gen_add_wrapper upd w =
-  Context.theory_map (Rules.map (fn (rs as Rules {next, rules, netpairs, wrappers}) =>
+  Context.theory_map (Rules.map (fn Rules {next, rules, netpairs, wrappers} =>
     make_rules next rules netpairs (upd (fn ws => (w, stamp ()) :: ws) wrappers)));
 
 val addSWrapper = gen_add_wrapper Library.apfst;
--- a/src/Pure/Isar/expression.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/Isar/expression.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -311,7 +311,7 @@
   | finish_primitive _ close (Defines defs) = close (Defines defs)
   | finish_primitive _ _ (Notes facts) = Notes facts;
 
-fun finish_inst ctxt parms do_close (loc, (prfx, inst)) =
+fun finish_inst ctxt (loc, (prfx, inst)) =
   let
     val thy = ProofContext.theory_of ctxt;
     val (parm_names, parm_types) = Locale.params_of thy loc |> map #1 |> split_list;
@@ -323,7 +323,7 @@
 
 fun finish ctxt parms do_close insts elems =
   let
-    val deps = map (finish_inst ctxt parms do_close) insts;
+    val deps = map (finish_inst ctxt) insts;
     val elems' = map (finish_elem ctxt parms do_close) elems;
   in (deps, elems') end;
 
--- a/src/Pure/Isar/isar_document.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/Isar/isar_document.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -13,6 +13,7 @@
   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 init: unit -> unit
 end;
 
 structure Isar_Document: ISAR_DOCUMENT =
--- a/src/Pure/Isar/obtain.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/Isar/obtain.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -215,7 +215,6 @@
     val thy = ProofContext.theory_of ctxt;
     val certT = Thm.ctyp_of thy;
     val cert = Thm.cterm_of thy;
-    val string_of_typ = Syntax.string_of_typ ctxt;
     val string_of_term = setmp show_types true (Syntax.string_of_term ctxt);
 
     fun err msg th = error (msg ^ ":\n" ^ Display.string_of_thm ctxt th);
--- a/src/Pure/Isar/proof.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/Isar/proof.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -581,7 +581,6 @@
   let
     val _ = assert_forward state;
     val thy = theory_of state;
-    val ctxt = context_of state;
 
     val (raw_name_atts, (raw_vars, raw_rhss)) = args |> split_list ||> split_list;
     val name_atts = map (apsnd (map (prep_att thy))) raw_name_atts;
@@ -1008,7 +1007,7 @@
   let
     val _ = assert_backward state;
     val (goal_ctxt, (_, goal)) = find_goal state;
-    val {statement as (kind, propss, prop), messages, using, goal, before_qed, after_qed} = goal;
+    val {statement as (kind, _, prop), messages, using, goal, before_qed, after_qed} = goal;
     val goal_txt = prop :: map Thm.term_of (Assumption.all_assms_of goal_ctxt);
 
     val _ = is_relevant state andalso error "Cannot fork relevant proof";
--- a/src/Pure/Isar/proof_context.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/Isar/proof_context.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -1033,7 +1033,7 @@
 
 local
 
-fun const_syntax ctxt (Free (x, T), mx) = SOME (true, (x, T, mx))
+fun const_syntax _ (Free (x, T), mx) = SOME (true, (x, T, mx))
   | const_syntax ctxt (Const (c, _), mx) =
       Option.map (pair false) (try (Consts.syntax (consts_of ctxt)) (c, mx))
   | const_syntax _ _ = NONE;
@@ -1106,7 +1106,7 @@
 
 (* fixes vs. frees *)
 
-fun auto_fixes (arg as (ctxt, (propss, x))) =
+fun auto_fixes (ctxt, (propss, x)) =
   ((fold o fold) Variable.auto_fixes propss ctxt, (propss, x));
 
 fun bind_fixes xs ctxt =
--- a/src/Pure/Isar/proof_node.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/Isar/proof_node.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -41,7 +41,7 @@
 
 (* apply transformer *)
 
-fun applys f (ProofNode (node as (st, _), n)) =
+fun applys f (ProofNode ((st, _), n)) =
   (case Seq.pull (f st) of
     NONE => error "empty result sequence -- proof command failed"
   | SOME res => ProofNode (res, n + 1));
--- a/src/Pure/Isar/rule_insts.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/Isar/rule_insts.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -266,8 +266,8 @@
   let
     val thy = ProofContext.theory_of ctxt;
     (* Separate type and term insts *)
-    fun has_type_var ((x, _), _) = (case Symbol.explode x of
-          "'"::cs => true | cs => false);
+    fun has_type_var ((x, _), _) =
+      (case Symbol.explode x of "'" :: _ => true | _ => false);
     val Tinsts = List.filter has_type_var insts;
     val tinsts = filter_out has_type_var insts;
 
--- a/src/Pure/Isar/specification.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/Isar/specification.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -136,9 +136,6 @@
   prepare prep_vars parse_prop prep_att do_close
     vars (map single_spec specs) #>> apsnd (map the_spec);
 
-fun prep_specification prep_vars parse_prop prep_att vars specs =
-  prepare prep_vars parse_prop prep_att true [specs];
-
 in
 
 fun check_spec x = prep_spec ProofContext.cert_vars (K I) (K I) true x;
--- a/src/Pure/Isar/theory_target.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/Isar/theory_target.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -45,7 +45,7 @@
 
 (* pretty *)
 
-fun pretty_thy ctxt target is_locale is_class =
+fun pretty_thy ctxt target is_class =
   let
     val thy = ProofContext.theory_of ctxt;
     val target_name = (if is_class then "class " else "locale ") ^ Locale.extern thy target;
@@ -63,12 +63,12 @@
       (map (Pretty.chunks o Element.pretty_ctxt ctxt) elems)]
   end;
 
-fun pretty (Target {target, is_locale, is_class, instantiation, overloading, ...}) ctxt =
+fun pretty (Target {target, is_class, instantiation, overloading, ...}) ctxt =
   Pretty.block [Pretty.str "theory", Pretty.brk 1,
       Pretty.str (Context.theory_name (ProofContext.theory_of ctxt))] ::
     (if not (null overloading) then [Overloading.pretty ctxt]
      else if not (null (#1 instantiation)) then [Class_Target.pretty_instantiation ctxt]
-     else pretty_thy ctxt target is_locale is_class);
+     else pretty_thy ctxt target is_class);
 
 
 (* target declarations *)
@@ -260,8 +260,7 @@
 
 (* define *)
 
-fun define (ta as Target {target, is_locale, is_class, ...})
-    kind ((b, mx), ((name, atts), rhs)) lthy =
+fun define ta kind ((b, mx), ((name, atts), rhs)) lthy =
   let
     val thy = ProofContext.theory_of lthy;
     val thy_ctxt = ProofContext.init thy;
--- a/src/Pure/Isar/toplevel.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/Isar/toplevel.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -126,7 +126,6 @@
   SkipProof of int * (generic_theory * generic_theory);
     (*proof depth, resulting theory, original theory*)
 
-val the_global_theory = fn Theory (Context.Theory thy, _) => thy | _ => raise UNDEF;
 val theory_node = fn Theory (gthy, _) => SOME gthy | _ => NONE;
 val proof_node = fn Proof (prf, _) => SOME prf | _ => NONE;
 
@@ -256,7 +255,7 @@
 
 in
 
-fun apply_transaction pos f g (node, exit) =
+fun apply_transaction f g (node, exit) =
   let
     val _ = is_draft_theory node andalso error "Illegal draft theory in toplevel state";
     val cont_node = reset_presentation node;
@@ -293,29 +292,29 @@
 
 local
 
-fun apply_tr int _ (Init (_, f, exit)) (State (NONE, _)) =
+fun apply_tr int (Init (_, f, exit)) (State (NONE, _)) =
       State (SOME (Theory (Context.Theory (Theory.checkpoint (f int)), NONE), exit), NONE)
-  | apply_tr _ _ Exit (State (prev as SOME (Theory (Context.Theory _, _), _), _)) =
+  | apply_tr _ Exit (State (prev as SOME (Theory (Context.Theory _, _), _), _)) =
       State (NONE, prev)
-  | apply_tr _ _ CommitExit (State (NONE, SOME (Theory (Context.Theory thy, _), exit))) =
+  | apply_tr _ CommitExit (State (NONE, SOME (Theory (Context.Theory thy, _), exit))) =
       (Runtime.controlled_execution exit thy; toplevel)
-  | apply_tr int _ (Keep f) state =
+  | apply_tr int (Keep f) state =
       Runtime.controlled_execution (fn x => tap (f int) x) state
-  | apply_tr int pos (Transaction (f, g)) (State (SOME state, _)) =
-      apply_transaction pos (fn x => f int x) g state
-  | apply_tr _ _ _ _ = raise UNDEF;
+  | apply_tr int (Transaction (f, g)) (State (SOME state, _)) =
+      apply_transaction (fn x => f int x) g state
+  | apply_tr _ _ _ = raise UNDEF;
 
-fun apply_union _ _ [] state = raise FAILURE (state, UNDEF)
-  | apply_union int pos (tr :: trs) state =
-      apply_union int pos trs state
-        handle Runtime.UNDEF => apply_tr int pos tr state
-          | FAILURE (alt_state, UNDEF) => apply_tr int pos tr alt_state
+fun apply_union _ [] state = raise FAILURE (state, UNDEF)
+  | apply_union int (tr :: trs) state =
+      apply_union int trs state
+        handle Runtime.UNDEF => apply_tr int tr state
+          | FAILURE (alt_state, UNDEF) => apply_tr int tr alt_state
           | exn as FAILURE _ => raise exn
           | exn => raise FAILURE (state, exn);
 
 in
 
-fun apply_trans int pos trs state = (apply_union int pos trs state, NONE)
+fun apply_trans int trs state = (apply_union int trs state, NONE)
   handle FAILURE (alt_state, exn) => (alt_state, SOME exn) | exn => (state, SOME exn);
 
 end;
@@ -562,14 +561,14 @@
 
 local
 
-fun app int (tr as Transition {trans, pos, print, no_timing, ...}) =
+fun app int (tr as Transition {trans, print, no_timing, ...}) =
   setmp_thread_position tr (fn state =>
     let
       fun do_timing f x = (warning (command_msg "" tr); timeap f x);
       fun do_profiling f x = profile (! profiling) f x;
 
       val (result, status) =
-         state |> (apply_trans int pos trans
+         state |> (apply_trans int trans
           |> (if ! profiling > 0 andalso not no_timing then do_profiling else I)
           |> (if ! profiling > 0 orelse ! timing andalso not no_timing then do_timing else I));
 
--- a/src/Pure/Isar/value_parse.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/Isar/value_parse.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -9,6 +9,7 @@
   val comma: 'a parser -> 'a parser
   val equal: 'a parser -> 'a parser
   val parens: 'a parser -> 'a parser
+  val unit: unit parser
   val pair: 'a parser -> 'b parser -> ('a * 'b) parser
   val triple: 'a parser -> 'b parser -> 'c parser -> ('a * 'b * 'c) parser
   val list: 'a parser -> 'a list parser
--- a/src/Pure/ML/ml_antiquote.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/ML/ml_antiquote.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -41,7 +41,7 @@
     (Context.Proof ctxt, fn {background, ...} => (K ("", ""), background)))));
 
 fun inline name scan = ML_Context.add_antiq name
-  (fn _ => scan >> (fn s => fn {struct_name, background} => (K ("", s), background)));
+  (fn _ => scan >> (fn s => fn {background, ...} => (K ("", s), background)));
 
 fun declaration kind name scan = ML_Context.add_antiq name
   (fn _ => scan >> (fn s => fn {struct_name, background} =>
--- a/src/Pure/Proof/extraction.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/Proof/extraction.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -81,8 +81,7 @@
   {next = next - 1, rs = r :: rs, net = Net.insert_term (K false)
      (Envir.eta_contract lhs, (next, r)) net};
 
-fun merge_rules
-  ({next, rs = rs1, net} : rules) ({next = next2, rs = rs2, ...} : rules) =
+fun merge_rules ({next, rs = rs1, net} : rules) ({rs = rs2, ...} : rules) =
   List.foldr add_rule {next = next, rs = rs1, net = net} (subtract (op =) rs1 rs2);
 
 fun condrew thy rules procs =
@@ -141,7 +140,7 @@
   map_proof_terms (subst_TVars tye) (typ_subst_TVars tye);
 
 fun relevant_vars types prop = List.foldr (fn
-      (Var ((a, i), T), vs) => (case strip_type T of
+      (Var ((a, _), T), vs) => (case strip_type T of
         (_, Type (s, _)) => if member (op =) types s then a :: vs else vs
       | _ => vs)
     | (_, vs) => vs) [] (vars_of prop);
--- a/src/Pure/ROOT.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/ROOT.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -45,7 +45,6 @@
 use "General/long_name.ML";
 use "General/binding.ML";
 use "General/name_space.ML";
-use "General/lazy.ML";
 use "General/path.ML";
 use "General/url.ML";
 use "General/buffer.ML";
@@ -58,12 +57,17 @@
 
 use "Concurrent/simple_thread.ML";
 use "Concurrent/synchronized.ML";
-if Multithreading.available then () else use "Concurrent/synchronized_dummy.ML";
+if Multithreading.available then ()
+else use "Concurrent/synchronized_sequential.ML";
 use "Concurrent/mailbox.ML";
 use "Concurrent/task_queue.ML";
 use "Concurrent/future.ML";
+use "Concurrent/lazy.ML";
+if Multithreading.available then ()
+else use "Concurrent/lazy_sequential.ML";
 use "Concurrent/par_list.ML";
-if Multithreading.available then () else use "Concurrent/par_list_dummy.ML";
+if Multithreading.available then ()
+else use "Concurrent/par_list_sequential.ML";
 
 
 (* fundamental structures *)
--- a/src/Pure/Syntax/ast.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/Syntax/ast.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -94,7 +94,7 @@
 
 (** check translation rules **)
 
-fun rule_error (rule as (lhs, rhs)) =
+fun rule_error (lhs, rhs) =
   let
     fun add_vars (Constant _) = I
       | add_vars (Variable x) = cons x
--- a/src/Pure/Syntax/mixfix.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/Syntax/mixfix.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -128,16 +128,6 @@
 
 fun const_mixfix c mx = (const_name mx c, fix_mixfix c mx);
 
-fun map_mixfix _ NoSyn = NoSyn
-  | map_mixfix f (Mixfix (sy, ps, p)) = Mixfix (f sy, ps, p)
-  | map_mixfix f (Delimfix sy) = Delimfix (f sy)
-  | map_mixfix f (InfixName (sy, p)) = InfixName (f sy, p)
-  | map_mixfix f (InfixlName (sy, p)) = InfixlName (f sy, p)
-  | map_mixfix f (InfixrName (sy, p)) = InfixrName (f sy, p)
-  | map_mixfix f (Binder (sy, p, q)) = Binder (f sy, p, q)
-  | map_mixfix _ Structure = Structure
-  | map_mixfix _ _ = raise Fail ("map_mixfix: illegal occurrence of unnamed infix");
-
 fun mixfix_args NoSyn = 0
   | mixfix_args (Mixfix (sy, _, _)) = SynExt.mfix_args sy
   | mixfix_args (Delimfix sy) = SynExt.mfix_args sy
--- a/src/Pure/Syntax/syn_ext.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/Syntax/syn_ext.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -270,9 +270,9 @@
     fun rem_pri (Argument (s, _)) = Argument (s, chain_pri)
       | rem_pri sym = sym;
 
-    fun logify_types copy_prod (a as (Argument (s, p))) =
+    fun logify_types (a as (Argument (s, p))) =
           if s <> "prop" andalso is_logtype s then Argument (logic, p) else a
-      | logify_types _ a = a;
+      | logify_types a = a;
 
 
     val raw_symbs = read_mfix sy handle ERROR msg => err_in_mfix msg mfix;
@@ -305,7 +305,7 @@
       if convert andalso not copy_prod then
        (if lhs = "prop" then sprop else if is_logtype lhs then logic else lhs)
       else lhs;
-    val symbs' = map (logify_types copy_prod) symbs;
+    val symbs' = map logify_types symbs;
     val xprod = XProd (lhs', symbs', const', pri);
 
     val _ = (List.app check_pri pris; check_pri pri; check_blocks symbs');
--- a/src/Pure/Syntax/syn_trans.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/Syntax/syn_trans.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -155,20 +155,11 @@
   | bigimpl_ast_tr (*"_bigimpl"*) asts = raise Ast.AST ("bigimpl_ast_tr", asts);
 
 
-(* meta conjunction *)
-
-fun conjunction_tr [t, u] = Lexicon.const "Pure.conjunction" $ t $ u
-  | conjunction_tr ts = raise TERM ("conjunction_tr", ts);
-
-
 (* type/term reflection *)
 
 fun type_tr (*"_TYPE"*) [ty] = mk_type ty
   | type_tr (*"_TYPE"*) ts = raise TERM ("type_tr", ts);
 
-fun term_tr [t] = Lexicon.const "Pure.term" $ t
-  | term_tr ts = raise TERM ("term_tr", ts);
-
 
 (* dddot *)
 
--- a/src/Pure/Syntax/syntax.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/Syntax/syntax.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -404,7 +404,7 @@
 
 fun pretty_gram (Syntax (tabs, _)) =
   let
-    val {lexicon, prmodes, gram, prtabs, ...} = tabs;
+    val {lexicon, prmodes, gram, ...} = tabs;
     val prmodes' = sort_strings (filter_out (fn s => s = "") prmodes);
   in
     [pretty_strs_qs "lexicon:" (Scan.dest_lexicon lexicon),
@@ -639,7 +639,7 @@
 
 local
 
-fun unparse_t t_to_ast prt_t markup ctxt (syn as Syntax (tabs, _)) curried t =
+fun unparse_t t_to_ast prt_t markup ctxt (Syntax (tabs, _)) curried t =
   let
     val {print_trtab, print_ruletab, print_ast_trtab, tokentrtab, prtabs, ...} = tabs;
     val ast = t_to_ast ctxt (lookup_tr' print_trtab) t;
--- a/src/Pure/Syntax/type_ext.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/Syntax/type_ext.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -82,8 +82,8 @@
           if Lexicon.is_tid x then TFree (x, get_sort (x, ~1))
           else Type (x, [])
       | typ_of (Var (xi, _)) = TVar (xi, get_sort xi)
-      | typ_of (Const ("_tfree",_) $ (t as Free x)) = typ_of t
-      | typ_of (Const ("_tvar",_) $ (t as Var x)) = typ_of t
+      | typ_of (Const ("_tfree",_) $ (t as Free _)) = typ_of t
+      | typ_of (Const ("_tvar",_) $ (t as Var _)) = typ_of t
       | typ_of (Const ("_ofsort", _) $ Free (x, _) $ _) = TFree (x, get_sort (x, ~1))
       | typ_of (Const ("_ofsort", _) $ (Const ("_tfree",_) $ Free (x, _)) $ _) =
           TFree (x, get_sort (x, ~1))
--- a/src/Pure/System/isabelle_process.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/System/isabelle_process.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -133,6 +133,7 @@
  (Unsynchronized.change print_mode (update (op =) isabelle_processN);
   setup_channels out |> init_message;
   OuterKeyword.report ();
+  Isar_Document.init ();
   Output.status (Markup.markup Markup.ready "");
   Isar.toplevel_loop {init = true, welcome = false, sync = true, secure = true});
 
--- a/src/Pure/System/isar.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/System/isar.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -48,7 +48,6 @@
   in edit count (! global_state, ! global_history) end);
 
 fun state () = NAMED_CRITICAL "Isar" (fn () => ! global_state);
-fun set_state state = NAMED_CRITICAL "Isar" (fn () => global_state := state);
 
 fun exn () = NAMED_CRITICAL "Isar" (fn () => ! global_exn);
 fun set_exn exn =  NAMED_CRITICAL "Isar" (fn () => global_exn := exn);
--- a/src/Pure/Thy/html.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/Thy/html.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -222,7 +222,6 @@
     in (implode syms, width) end;
 
 val output = #1 o output_width;
-val output_symbols = map #2 o output_syms;
 
 val _ = Output.add_mode htmlN output_width Symbol.encode_raw;
 
--- a/src/Pure/Thy/thm_deps.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/Thy/thm_deps.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -40,7 +40,7 @@
                path = "",
                parents = parents};
           in cons entry end;
-    val deps = Proofterm.fold_body_thms (add_dep o #2) (map Thm.proof_body_of thms) [];
+    val deps = Proofterm.fold_body_thms add_dep (map Thm.proof_body_of thms) [];
   in Present.display_graph (sort_wrt #ID deps) end;
 
 
@@ -55,9 +55,10 @@
       fold (Facts.fold_static add_fact o PureThy.facts_of) thys []
       |> sort_distinct (string_ord o pairself #1);
 
-    val tab = Proofterm.fold_body_thms
-      (fn (_, (name, prop, _)) => name <> "" ? Symtab.insert_list (op =) (name, prop))
-      (map (Proofterm.strip_thm o Thm.proof_body_of o snd) thms) Symtab.empty;
+    val tab =
+      Proofterm.fold_body_thms
+        (fn (name, prop, _) => name <> "" ? Symtab.insert_list (op =) (name, prop))
+        (map (Proofterm.strip_thm o Thm.proof_body_of o snd) thms) Symtab.empty;
     fun is_unused (name, th) =
       not (member (op aconv) (Symtab.lookup_list tab name) (Thm.prop_of th));
 
--- a/src/Pure/Thy/thy_info.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/Thy/thy_info.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -283,7 +283,7 @@
 
 local
 
-fun provide path name info (deps as SOME {update_time, master, text, parents, files}) =
+fun provide path name info (SOME {update_time, master, text, parents, files}) =
      (if AList.defined (op =) files path then ()
       else warning (loader_msg "undeclared dependency of theory" [name] ^
         " on file: " ^ quote (Path.implode path));
@@ -338,7 +338,7 @@
 fun load_thy time upd_time initiators name =
   let
     val _ = priority ("Loading theory " ^ quote name ^ required_by " " initiators);
-    val (pos, text, files) =
+    val (pos, text, _) =
       (case get_deps name of
         SOME {master = SOME (master_path, _), text as _ :: _, files, ...} =>
           (Path.position master_path, text, files)
@@ -364,7 +364,7 @@
 
 local
 
-fun schedule_futures task_graph =
+fun schedule_futures task_graph = uninterruptible (fn _ => fn () =>
   let
     val tasks = Graph.topological_order task_graph |> map_filter (fn name =>
       (case Graph.get_node task_graph name of Task body => SOME (name, body) | _ => NONE));
@@ -397,7 +397,7 @@
         val _ = after_load ();
       in [] end handle exn => (kill_thy name; [exn]));
 
-  in ignore (Exn.release_all (map Exn.Exn (rev exns))) end;
+  in ignore (Exn.release_all (map Exn.Exn (rev exns))) end) ();
 
 fun schedule_seq tasks =
   Graph.topological_order tasks
@@ -410,7 +410,7 @@
 
 in
 
-fun schedule_tasks tasks n =
+fun schedule_tasks tasks =
   if not (Multithreading.enabled ()) then schedule_seq tasks
   else if Multithreading.self_critical () then
      (warning (loader_msg "no multithreading within critical section" []);
@@ -438,7 +438,7 @@
   | NONE =>
       let val {master, text, imports = parents, uses = files} = ThyLoad.deps_thy dir name
       in (false, init_deps (SOME master) text parents files, parents) end
-  | SOME (deps as SOME {update_time, master, text, parents, files}) =>
+  | SOME (SOME {update_time, master, text, parents, files}) =>
       let
         val (thy_path, thy_id) = ThyLoad.check_thy dir name;
         val master' = SOME (thy_path, thy_id);
@@ -471,7 +471,7 @@
     val dir' = Path.append dir (Path.dir path);
     val _ = member (op =) initiators name andalso error (cycle_msg initiators);
   in
-    (case try (Graph.get_node (fst tasks)) name of
+    (case try (Graph.get_node tasks) name of
       SOME task => (task_finished task, tasks)
     | NONE =>
         let
@@ -481,7 +481,7 @@
                 required_by "\n" initiators);
           val parent_names = map base_name parents;
 
-          val (parents_current, (tasks_graph', tasks_len')) =
+          val (parents_current, tasks_graph') =
             require_thys time (name :: initiators)
               (Path.append dir (master_dir' deps)) parents tasks;
 
@@ -496,8 +496,7 @@
           val tasks_graph'' = tasks_graph' |> new_deps name parent_names
            (if all_current then Finished
             else Task (fn () => load_thy time upd_time initiators name));
-          val tasks_len'' = if all_current then tasks_len' else tasks_len' + 1;
-        in (all_current, (tasks_graph'', tasks_len'')) end)
+        in (all_current, tasks_graph'') end)
   end;
 
 end;
@@ -508,8 +507,7 @@
 local
 
 fun gen_use_thy' req dir arg =
-  let val (_, (tasks, n)) = req [] dir arg (Graph.empty, 0)
-  in schedule_tasks tasks n end;
+  schedule_tasks (snd (req [] dir arg Graph.empty));
 
 fun gen_use_thy req str =
   let val name = base_name str in
--- a/src/Pure/Thy/thy_load.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/Thy/thy_load.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -73,7 +73,7 @@
 
 (* check files *)
 
-fun check_ext exts paths dir src_path =
+fun check_ext exts paths src_path =
   let
     val path = Path.expand src_path;
     val _ = Path.is_current path andalso error "Bad file specification";
@@ -84,15 +84,15 @@
     fun try_prfx prfx = get_first (try_ext (Path.append prfx path)) ("" :: exts);
   in get_first try_prfx paths end;
 
-fun check_file dir path = check_ext [] (search_path dir path) dir path;
-fun check_ml dir path = check_ext ml_exts (search_path dir path) dir path;
+fun check_file dir path = check_ext [] (search_path dir path) path;
+fun check_ml dir path = check_ext ml_exts (search_path dir path) path;
 
 fun check_thy dir name =
   let
     val thy_file = thy_path name;
     val paths = search_path dir thy_file;
   in
-    (case check_ext [] paths dir thy_file of
+    (case check_ext [] paths thy_file of
       NONE => error ("Could not find theory file " ^ quote (Path.implode thy_file) ^
         " in " ^ commas_quote (map Path.implode paths))
     | SOME thy_id => thy_id)
--- a/src/Pure/Tools/find_consts.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/Tools/find_consts.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -107,7 +107,7 @@
                 val tye = Sign.typ_match thy (qty, ty) Vartab.empty;
                 val sub_size =
                   Vartab.fold (fn (_, (_, t)) => fn n => Term.size_of_typ t + n) tye 0;
-              in SOME sub_size end handle MATCH => NONE
+              in SOME sub_size end handle Type.TYPE_MATCH => NONE
           end
       | make_match (Loose arg) =
           check_const (matches_subtype thy (make_pattern arg) o snd)
--- a/src/Pure/Tools/find_theorems.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/Tools/find_theorems.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -76,18 +76,9 @@
 
 fun is_nontrivial thy = Term.is_Const o Term.head_of o ObjectLogic.drop_judgment thy;
 
-(* Note: ("op =" : "bool --> bool --> bool") does not exist in Pure. *)
-fun is_Iff c =
-  (case dest_Const c of
-     ("op =", ty) =>
-       (ty
-        |> strip_type
-        |> swap
-        |> (op ::)
-        |> map (fst o dest_Type)
-        |> forall (curry (op =) "bool")
-        handle TYPE _ => false)
-   | _ => false);
+(*educated guesses on HOL*)  (* FIXME broken *)
+val boolT = Type ("bool", []);
+val iff_const = Const ("op =", boolT --> boolT --> boolT);
 
 (*extract terms from term_src, refine them to the parts that concern us,
   if po try match them against obj else vice versa.
@@ -97,19 +88,20 @@
   let
     val thy = ProofContext.theory_of ctxt;
 
-    val chkmatch = obj |> (if po then rpair else pair) #> Pattern.matches thy;
+    fun check_match pat = Pattern.matches thy (if po then (pat, obj) else (obj, pat));
     fun matches pat =
       let
         val jpat = ObjectLogic.drop_judgment thy pat;
         val c = Term.head_of jpat;
         val pats =
           if Term.is_Const c
-          then if doiff andalso is_Iff c
-               then pat :: map (ObjectLogic.ensure_propT thy) ((snd o strip_comb) jpat)
-                    |> filter (is_nontrivial thy)
-               else [pat]
+          then
+            if doiff andalso c = iff_const then
+              (pat :: map (ObjectLogic.ensure_propT thy) (snd (strip_comb jpat)))
+                |> filter (is_nontrivial thy)
+            else [pat]
           else [];
-      in filter chkmatch pats end;
+      in filter check_match pats end;
 
     fun substsize pat =
       let val (_, subst) =
@@ -117,12 +109,11 @@
       in Vartab.fold (fn (_, (_, t)) => fn n => size_of_term t + n) subst 0 end;
 
     fun bestmatch [] = NONE
-     |  bestmatch xs = SOME (foldr1 Int.min xs);
+      | bestmatch xs = SOME (foldr1 Int.min xs);
 
     val match_thm = matches o refine_term;
   in
-    map match_thm (extract_terms term_src)
-    |> flat
+    maps match_thm (extract_terms term_src)
     |> map substsize
     |> bestmatch
   end;
@@ -178,8 +169,8 @@
         is_matching_thm false (single, I) ctxt true (goal_tree prem) rule_tree;
       val successful = prems |> map_filter try_subst;
     in
-    (*elim rules always have assumptions, so an elim with one
-      assumption is as good as an intro rule with none*)
+      (*elim rules always have assumptions, so an elim with one
+        assumption is as good as an intro rule with none*)
       if is_nontrivial (ProofContext.theory_of ctxt) (Thm.major_prem_of thm)
         andalso not (null successful)
       then SOME (Thm.nprems_of thm - 1, foldr1 Int.min successful) else NONE
@@ -190,15 +181,13 @@
 
 fun filter_solves ctxt goal =
   let
-    val baregoal = Logic.get_goal (Thm.prop_of goal) 1;
-
     fun etacn thm i = Seq.take (! tac_limit) o etac thm i;
     fun try_thm thm =
       if Thm.no_prems thm then rtac thm 1 goal
       else (etacn thm THEN_ALL_NEW (Goal.norm_hhf_tac THEN' Method.assm_tac ctxt)) 1 goal;
   in
     fn (_, thm) =>
-      if (is_some o Seq.pull o try_thm) thm
+      if is_some (Seq.pull (try_thm thm))
       then SOME (Thm.nprems_of thm, 0) else NONE
   end;
 
@@ -218,7 +207,7 @@
 
 (* filter_pattern *)
 
-fun get_names t = (Term.add_const_names t []) union (Term.add_free_names t []);
+fun get_names t = Term.add_const_names t (Term.add_free_names t []);
 fun get_thm_names (_, thm) = get_names (Thm.full_prop_of thm);
 
 (*Including all constants and frees is only sound because
@@ -238,10 +227,9 @@
 
     fun check (t, NONE) = check (t, SOME (get_thm_names t))
       | check ((_, thm), c as SOME thm_consts) =
-          (if pat_consts subset_string thm_consts
-              andalso (Pattern.matches_subterm (ProofContext.theory_of ctxt)
-                                               (pat, Thm.full_prop_of thm))
-           then SOME (0, 0) else NONE, c);
+         (if pat_consts subset_string thm_consts andalso
+            Pattern.matches_subterm (ProofContext.theory_of ctxt) (pat, Thm.full_prop_of thm)
+          then SOME (0, 0) else NONE, c);
   in check end;
 
 
@@ -253,7 +241,6 @@
   error ("Current goal required for " ^ c ^ " search criterion");
 
 val fix_goal = Thm.prop_of;
-val fix_goalo = Option.map fix_goal;
 
 fun filter_crit _ _ (Name name) = apfst (filter_name name)
   | filter_crit _ NONE Intro = err_no_goal "intro"
@@ -276,7 +263,7 @@
 fun app_filters thm =
   let
     fun app (NONE, _, _) = NONE
-      | app (SOME v, consts, []) = SOME (v, thm)
+      | app (SOME v, _, []) = SOME (v, thm)
       | app (r, consts, f :: fs) =
           let val (r', consts') = f (thm, consts)
           in app (opt_add r r', consts', fs) end;
@@ -439,6 +426,7 @@
   end;
 
 
+
 (** command syntax **)
 
 fun find_theorems_cmd ((opt_lim, rem_dups), spec) =
--- a/src/Pure/axclass.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/axclass.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -150,7 +150,6 @@
   let val params = #2 (get_axclasses thy);
   in fold (fn (x, c) => if pred c then cons x else I) params [] end;
 
-fun params_of thy c = get_params thy (fn c' => c' = c);
 fun all_params_of thy S = get_params thy (fn c => Sign.subsort thy (S, [c]));
 
 fun class_of_param thy = AList.lookup (op =) (#2 (get_axclasses thy));
@@ -263,7 +262,8 @@
 
 fun unoverload_const thy (c_ty as (c, _)) =
   case class_of_param thy c
-   of SOME class => (case get_inst_tyco (Sign.consts_of thy) c_ty
+   of SOME class (* FIXME unused? *) =>
+     (case get_inst_tyco (Sign.consts_of thy) c_ty
        of SOME tyco => try (param_of_inst thy) (c, tyco) |> the_default c
         | NONE => c)
     | NONE => c;
@@ -585,7 +585,7 @@
         val hyps = vars
           |> map (fn (T, S) => (T, Thm.of_sort (certT T, S) ~~ S));
         val ths = (typ, sort)
-          |> Sorts.of_sort_derivation (Syntax.pp_global thy) (Sign.classes_of thy)
+          |> Sorts.of_sort_derivation (Sign.classes_of thy)
            {class_relation =
               fn (th, c1) => fn c2 => th RS the_classrel thy (c1, c2),
             type_constructor =
--- a/src/Pure/consts.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/consts.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -199,7 +199,7 @@
 
 fun subscript (Type (_, Ts)) (i :: is) = subscript (nth Ts i) is
   | subscript T [] = T
-  | subscript T _ = raise Subscript;
+  | subscript _ _ = raise Subscript;
 
 in
 
--- a/src/Pure/context.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/context.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -455,7 +455,7 @@
 
 fun init_proof thy = Prf (init_data thy, check_thy thy);
 
-fun transfer_proof thy' (prf as Prf (data, thy_ref)) =
+fun transfer_proof thy' (Prf (data, thy_ref)) =
   let
     val thy = deref thy_ref;
     val _ = subthy (thy, thy') orelse error "transfer proof context: not a super theory";
--- a/src/Pure/defs.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/defs.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -123,7 +123,7 @@
 fun contained (U as TVar _) (Type (_, Ts)) = exists (fn T => T = U orelse contained U T) Ts
   | contained _ _ = false;
 
-fun acyclic pp defs (c, args) (d, Us) =
+fun acyclic pp (c, args) (d, Us) =
   c <> d orelse
   exists (fn U => exists (contained U) args) Us orelse
   is_none (match_args (args, Us)) orelse
@@ -150,7 +150,7 @@
       if forall (is_none o #1) reds then NONE
       else SOME (fold_rev
         (fn (NONE, dp) => insert (op =) dp | (SOME dps, _) => fold (insert (op =)) dps) reds []);
-    val _ = forall (acyclic pp defs const) (the_default deps deps');
+    val _ = forall (acyclic pp const) (the_default deps deps');
   in deps' end;
 
 in
@@ -163,8 +163,7 @@
           (args, perhaps (reduction pp defs (c, args)) deps));
       in
         if reducts = reducts' then (changed, defs)
-        else (true, defs |> map_def c (fn (specs, restricts, reducts) =>
-          (specs, restricts, reducts')))
+        else (true, defs |> map_def c (fn (specs, restricts, _) => (specs, restricts, reducts')))
       end;
     fun norm_all defs =
       (case Symtab.fold norm_update defs (false, defs) of
@@ -206,7 +205,7 @@
   let
     val restr =
       if plain_args args orelse
-        (case args of [Type (a, rec_args)] => plain_args rec_args | _ => false)
+        (case args of [Type (_, rec_args)] => plain_args rec_args | _ => false)
       then [] else [(args, name)];
     val spec =
       (serial (), {is_def = is_def, name = name, lhs = args, rhs = deps});
--- a/src/Pure/display.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/display.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -177,7 +177,7 @@
     val axioms = (Theory.axiom_space thy, Theory.axiom_table thy);
     val defs = Theory.defs_of thy;
     val {restricts, reducts} = Defs.dest defs;
-    val {naming, syn = _, tsig, consts} = Sign.rep_sg thy;
+    val {naming = _, syn = _, tsig, consts} = Sign.rep_sg thy;
     val {constants, constraints} = Consts.dest consts;
     val extern_const = NameSpace.extern (#1 constants);
     val {classes, default, types, ...} = Type.rep_tsig tsig;
--- a/src/Pure/envir.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/envir.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -62,8 +62,8 @@
   tenv: tenv,           (*assignments to Vars*)
   tyenv: Type.tyenv};   (*assignments to TVars*)
 
-fun make_env (maxidx, tenv, tyenv) = Envir {maxidx = maxidx, tenv = tenv, tyenv = tyenv};
-fun map_env f (Envir {maxidx, tenv, tyenv}) = make_env (f (maxidx, tenv, tyenv));
+fun make_env (maxidx, tenv, tyenv) =
+  Envir {maxidx = maxidx, tenv = tenv, tyenv = tyenv};
 
 fun maxidx_of (Envir {maxidx, ...}) = maxidx;
 fun term_env (Envir {tenv, ...}) = tenv;
--- a/src/Pure/goal.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/goal.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -300,22 +300,22 @@
   | SOME st' => Seq.single st');
 
 (*parallel refinement of non-schematic goal by single results*)
+exception FAILED of unit;
 fun PARALLEL_GOALS tac st =
   let
     val st0 = Thm.adjust_maxidx_thm ~1 st;
     val _ = Thm.maxidx_of st0 >= 0 andalso
       raise THM ("PARALLEL_GOALS: schematic goal state", 0, [st]);
 
-    exception FAILED;
     fun try_tac g =
       (case SINGLE tac g of
-        NONE => raise FAILED
+        NONE => raise FAILED ()
       | SOME g' => g');
 
     val goals = Drule.strip_imp_prems (Thm.cprop_of st0);
     val results = Par_List.map (try_tac o init) goals;
   in ALLGOALS (fn i => retrofit i 1 (nth results (i - 1))) st0 end
-  handle FAILED => Seq.empty;
+  handle FAILED () => Seq.empty;
 
 end;
 
--- a/src/Pure/logic.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/logic.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -222,7 +222,7 @@
 fun mk_of_class (ty, c) =
   Const (const_of_class c, Term.itselfT ty --> propT) $ mk_type ty;
 
-fun dest_of_class (t as Const (c_class, _) $ ty) = (dest_type ty, class_of_const c_class)
+fun dest_of_class (Const (c_class, _) $ ty) = (dest_type ty, class_of_const c_class)
   | dest_of_class t = raise TERM ("dest_of_class", [t]);
 
 
--- a/src/Pure/meta_simplifier.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/meta_simplifier.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -229,11 +229,6 @@
 
 fun make_simpset (args1, args2) = Simpset (make_ss1 args1, make_ss2 args2);
 
-fun map_simpset f (Simpset ({rules, prems, bounds, depth, context},
-    {congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers})) =
-  make_simpset (f ((rules, prems, bounds, depth, context),
-    (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers)));
-
 fun map_simpset1 f (Simpset (r1, r2)) = Simpset (map_ss1 f r1, r2);
 fun map_simpset2 f (Simpset (r1, r2)) = Simpset (r1, map_ss2 f r2);
 
@@ -388,7 +383,7 @@
     (Net.delete_term eq_rrule (term_of elhs, rrule) rules, prems, bounds, depth, context))
   handle Net.DELETE => (cond_warn_thm "Rewrite rule not in simpset:" ss thm; ss);
 
-fun insert_rrule (rrule as {thm, name, elhs, ...}) ss =
+fun insert_rrule (rrule as {thm, name, ...}) ss =
  (trace_named_thm (fn () => "Adding rewrite rule") ss (thm, name);
   ss |> map_simpset1 (fn (rules, prems, bounds, depth, context) =>
     let
@@ -455,7 +450,6 @@
   | SOME eq_True =>
       let
         val (_, _, lhs, elhs, _, _) = decomp_simp eq_True;
-        val extra = rrule_extra_vars elhs eq_True;
       in [{thm = eq_True, name = name, lhs = lhs, elhs = elhs, perm = false}] end);
 
 (*create the rewrite rule and possibly also the eq_True variant,
@@ -869,7 +863,7 @@
   Otherwise those vars may become instantiated with unnormalized terms
   while the premises are solved.*)
 
-fun cond_skel (args as (congs, (lhs, rhs))) =
+fun cond_skel (args as (_, (lhs, rhs))) =
   if Term.add_vars rhs [] subset Term.add_vars lhs [] then uncond_skel args
   else skel0;
 
@@ -892,8 +886,7 @@
     val eta_t = term_of eta_t';
     fun rew {thm, name, lhs, elhs, extra, fo, perm} =
       let
-        val thy = Thm.theory_of_thm thm;
-        val {prop, maxidx, ...} = rep_thm thm;
+        val prop = Thm.prop_of thm;
         val (rthm, elhs') =
           if maxt = ~1 orelse not extra then (thm, elhs)
           else (Thm.incr_indexes (maxt + 1) thm, Thm.incr_indexes_cterm (maxt + 1) elhs);
@@ -979,7 +972,7 @@
       (* Thm.match can raise Pattern.MATCH;
          is handled when congc is called *)
       val thm' = Thm.instantiate insts (Thm.rename_boundvars (term_of rlhs) (term_of t) rthm);
-      val unit = trace_thm (fn () => "Applying congruence rule:") ss thm';
+      val _ = trace_thm (fn () => "Applying congruence rule:") ss thm';
       fun err (msg, thm) = (trace_thm (fn () => msg) ss thm; NONE)
   in case prover thm' of
        NONE => err ("Congruence proof failed.  Could not prove", thm')
@@ -1025,7 +1018,7 @@
 
     and subc skel (ss as Simpset ({bounds, ...}, {congs, ...})) t0 =
        (case term_of t0 of
-           Abs (a, T, t) =>
+           Abs (a, T, _) =>
              let
                  val b = Name.bound (#1 bounds);
                  val (v, t') = Thm.dest_abs (SOME b) t0;
@@ -1124,7 +1117,7 @@
       end
 
     and rebuild [] _ _ _ _ eq = eq
-      | rebuild (prem :: prems) concl (rrs :: rrss) (asm :: asms) ss eq =
+      | rebuild (prem :: prems) concl (_ :: rrss) (_ :: asms) ss eq =
           let
             val ss' = add_rrules (rev rrss, rev asms) ss;
             val concl' =
@@ -1231,7 +1224,7 @@
   let
     val thy = Thm.theory_of_cterm raw_ct;
     val ct = Thm.adjust_maxidx_cterm ~1 raw_ct;
-    val {t, maxidx, ...} = Thm.rep_cterm ct;
+    val {maxidx, ...} = Thm.rep_cterm ct;
     val ss = inc_simp_depth (activate_context thy raw_ss);
     val depth = simp_depth ss;
     val _ =
@@ -1297,8 +1290,8 @@
 (* for folding definitions, handling critical pairs *)
 
 (*The depth of nesting in a term*)
-fun term_depth (Abs(a,T,t)) = 1 + term_depth t
-  | term_depth (f$t) = 1 + Int.max(term_depth f, term_depth t)
+fun term_depth (Abs (_, _, t)) = 1 + term_depth t
+  | term_depth (f $ t) = 1 + Int.max (term_depth f, term_depth t)
   | term_depth _ = 0;
 
 val lhs_of_thm = #1 o Logic.dest_equals o prop_of;
--- a/src/Pure/proofterm.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/proofterm.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -40,8 +40,7 @@
   val proof_of: proof_body -> proof
   val join_proof: proof_body future -> proof
   val fold_proof_atoms: bool -> (proof -> 'a -> 'a) -> proof list -> 'a -> 'a
-  val fold_body_thms: (serial * (string * term * proof_body) -> 'a -> 'a) ->
-    proof_body list -> 'a -> 'a
+  val fold_body_thms: (string * term * proof_body -> 'a -> 'a) -> proof_body list -> 'a -> 'a
   val join_bodies: proof_body list -> unit
   val status_of: proof_body list -> {failed: bool, oracle: bool, unfinished: bool}
 
@@ -110,7 +109,7 @@
   val axm_proof: string -> term -> proof
   val oracle_proof: string -> term -> oracle * proof
   val promise_proof: theory -> serial -> term -> proof
-  val fulfill_proof: theory -> serial -> (serial * proof_body) list -> proof_body -> proof_body
+  val fulfill_proof: theory -> (serial * proof_body) list -> proof_body -> proof_body
   val thm_proof: theory -> string -> term list -> term ->
     (serial * proof_body future) list -> proof_body -> pthm * proof
   val get_name: term list -> term -> proof -> string
@@ -182,7 +181,7 @@
           let
             val body' = Future.join body;
             val (x', seen') = app body' (x, Inttab.update (i, ()) seen);
-          in (f (i, (name, prop, body')) x', seen') end));
+          in (f (name, prop, body') x', seen') end));
   in fn bodies => fn x => #1 (fold app bodies (x, Inttab.empty)) end;
 
 fun join_bodies bodies = fold_body_thms (fn _ => fn () => ()) bodies ();
@@ -959,7 +958,7 @@
   if ! proofs = 0 then ((name, dummy), Oracle (name, dummy, NONE))
   else ((name, prop), gen_axm_proof Oracle name prop);
 
-fun shrink_proof thy =
+val shrink_proof =
   let
     fun shrink ls lev (prf as Abst (a, T, body)) =
           let val (b, is, ch, body') = shrink ls (lev+1) body
@@ -1279,16 +1278,12 @@
         | _ => false));
   in Promise (i, prop, map TVar (Term.add_tvars prop [])) end;
 
-fun fulfill_proof _ _ [] body0 = body0
-  | fulfill_proof thy id ps body0 =
+fun fulfill_proof _ [] body0 = body0
+  | fulfill_proof thy ps body0 =
       let
         val PBody {oracles = oracles0, thms = thms0, proof = proof0} = body0;
-        val bodies = map snd ps;
-        val _ = fold_body_thms (fn (i, (name, _, _)) => fn () =>
-          if i = id then error ("Cyclic reference to theorem " ^ quote name)
-          else ()) bodies ();
-        val oracles = fold (fn PBody {oracles, ...} => merge_oracles oracles) bodies oracles0;
-        val thms = fold (fn PBody {thms, ...} => merge_thms thms) bodies thms0;
+        val oracles = fold (fn (_, PBody {oracles, ...}) => merge_oracles oracles) ps oracles0;
+        val thms = fold (fn (_, PBody {thms, ...}) => merge_thms thms) ps thms0;
         val proofs = fold (fn (i, PBody {proof, ...}) => Inttab.update (i, proof)) ps Inttab.empty;
 
         fun fill (Promise (i, prop, Ts)) =
@@ -1300,18 +1295,18 @@
         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 id promises body =
+fun fulfill_proof_future _ [] body = Future.value body
+  | fulfill_proof_future thy promises body =
       Future.fork_deps (map snd promises) (fn () =>
-        fulfill_proof thy id (map (apsnd Future.join) promises) body);
+        fulfill_proof thy (map (apsnd Future.join) promises) body);
 
 
 (***** theorems *****)
 
-fun thm_proof thy name hyps prop promises body =
+fun thm_proof thy name hyps concl promises body =
   let
     val PBody {oracles = oracles0, thms = thms0, proof = prf} = body;
-    val prop = Logic.list_implies (hyps, prop);
+    val prop = Logic.list_implies (hyps, concl);
     val nvs = needed_vars prop;
     val args = map (fn (v as Var (ixn, _)) =>
         if member (op =) nvs ixn then SOME v else NONE) (vars_of prop) @
@@ -1319,13 +1314,11 @@
 
     val proof0 =
       if ! proofs = 2 then
-        #4 (shrink_proof thy [] 0 (rew_proof thy (fold_rev implies_intr_proof hyps prf)))
+        #4 (shrink_proof [] 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 () =
-      let val id = serial ()
-      in (id, name, prop, fulfill_proof_future thy id promises body0) end;
+    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_thy.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/pure_thy.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -239,7 +239,6 @@
 (*** Pure theory syntax and logical content ***)
 
 val typ = SimpleSyntax.read_typ;
-val term = SimpleSyntax.read_term;
 val prop = SimpleSyntax.read_prop;
 
 val typeT = Syntax.typeT;
--- a/src/Pure/sign.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/sign.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -68,7 +68,6 @@
   val certify_typ_mode: Type.mode -> theory -> typ -> typ
   val certify': bool -> Pretty.pp -> bool -> Consts.T -> theory -> term -> term * typ * int
   val certify_term: theory -> term -> term * typ * int
-  val certify_prop: theory -> term -> term * typ * int
   val cert_term: theory -> term -> term
   val cert_prop: theory -> term -> term
   val no_frees: Pretty.pp -> term -> term
@@ -369,11 +368,9 @@
   in (if tm = tm'' then tm else tm'', T, Term.maxidx_of_term tm'') end;
 
 fun certify_term thy = certify' false (Syntax.pp_global thy) true (consts_of thy) thy;
-fun certify_prop thy = certify' true (Syntax.pp_global thy) true (consts_of thy) thy;
-
 fun cert_term_abbrev thy = #1 o certify' false (Syntax.pp_global thy) false (consts_of thy) thy;
 val cert_term = #1 oo certify_term;
-val cert_prop = #1 oo certify_prop;
+fun cert_prop thy = #1 o certify' true (Syntax.pp_global thy) true (consts_of thy) thy;
 
 end;
 
--- a/src/Pure/simplifier.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/simplifier.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -287,8 +287,6 @@
 
 val simpN = "simp";
 val congN = "cong";
-val addN = "add";
-val delN = "del";
 val onlyN = "only";
 val no_asmN = "no_asm";
 val no_asm_useN = "no_asm_use";
--- a/src/Pure/sorts.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/sorts.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -57,7 +57,7 @@
   val meet_sort_typ: algebra -> typ * sort -> typ -> typ   (*exception CLASS_ERROR*)
   val of_sort: algebra -> typ * sort -> bool
   val weaken: algebra -> ('a * class -> class -> 'a) -> 'a * class -> class -> 'a
-  val of_sort_derivation: Pretty.pp -> algebra ->
+  val of_sort_derivation: algebra ->
     {class_relation: 'a * class -> class -> 'a,
      type_constructor: string -> ('a * class) list list -> class -> 'a,
      type_variable: typ -> ('a * class) list} ->
@@ -401,7 +401,7 @@
     | cs :: _ => path (x, cs))
   end;
 
-fun of_sort_derivation pp algebra {class_relation, type_constructor, type_variable} =
+fun of_sort_derivation algebra {class_relation, type_constructor, type_variable} =
   let
     val weaken = weaken algebra class_relation;
     val arities = arities_of algebra;
--- a/src/Pure/term.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/term.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -796,7 +796,7 @@
       let
         fun subst (Const (a, T)) = Const (a, typ_subst_TVars instT T)
           | subst (Free (a, T)) = Free (a, typ_subst_TVars instT T)
-          | subst (t as Var (xi, T)) =
+          | subst (Var (xi, T)) =
               (case AList.lookup (op =) inst xi of
                 NONE => Var (xi, typ_subst_TVars instT T)
               | SOME t => t)
--- a/src/Pure/theory.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/theory.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -94,7 +94,8 @@
   val empty = make_thy (NameSpace.empty_table, Defs.empty, ([], []));
   val copy = I;
 
-  fun extend (Thy {axioms, defs, wrappers}) = make_thy (NameSpace.empty_table, defs, wrappers);
+  fun extend (Thy {axioms = _, defs, wrappers}) =
+    make_thy (NameSpace.empty_table, defs, wrappers);
 
   fun merge pp (thy1, thy2) =
     let
@@ -155,7 +156,7 @@
 
 fun cert_axm thy (b, raw_tm) =
   let
-    val (t, T, _) = Sign.certify_prop thy raw_tm
+    val t = Sign.cert_prop thy raw_tm
       handle TYPE (msg, _, _) => error msg
         | TERM (msg, _) => error msg;
   in
--- a/src/Pure/thm.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/thm.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -181,7 +181,7 @@
     val sorts = Sorts.insert_typ T [];
   in Ctyp {thy_ref = Theory.check_thy thy, T = T, maxidx = maxidx, sorts = sorts} end;
 
-fun dest_ctyp (Ctyp {thy_ref, T = Type (s, Ts), maxidx, sorts}) =
+fun dest_ctyp (Ctyp {thy_ref, T = Type (_, Ts), maxidx, sorts}) =
       map (fn T => Ctyp {thy_ref = thy_ref, T = T, maxidx = maxidx, sorts = sorts}) Ts
   | dest_ctyp cT = raise TYPE ("dest_ctyp", [typ_of cT], []);
 
@@ -218,31 +218,31 @@
     val sorts = Sorts.insert_term t [];
   in Cterm {thy_ref = Theory.check_thy thy, t = t, T = T, maxidx = maxidx, sorts = sorts} end;
 
-fun merge_thys0 (Cterm {thy_ref = r1, t = t1, ...}) (Cterm {thy_ref = r2, t = t2, ...}) =
+fun merge_thys0 (Cterm {thy_ref = r1, ...}) (Cterm {thy_ref = r2, ...}) =
   Theory.merge_refs (r1, r2);
 
 
 (* destructors *)
 
-fun dest_comb (ct as Cterm {t = c $ a, T, thy_ref, maxidx, sorts}) =
+fun dest_comb (Cterm {t = c $ a, T, thy_ref, maxidx, sorts}) =
       let val A = Term.argument_type_of c 0 in
         (Cterm {t = c, T = A --> T, thy_ref = thy_ref, maxidx = maxidx, sorts = sorts},
          Cterm {t = a, T = A, thy_ref = thy_ref, maxidx = maxidx, sorts = sorts})
       end
   | dest_comb ct = raise CTERM ("dest_comb", [ct]);
 
-fun dest_fun (ct as Cterm {t = c $ _, T, thy_ref, maxidx, sorts}) =
+fun dest_fun (Cterm {t = c $ _, T, thy_ref, maxidx, sorts}) =
       let val A = Term.argument_type_of c 0
       in Cterm {t = c, T = A --> T, thy_ref = thy_ref, maxidx = maxidx, sorts = sorts} end
   | dest_fun ct = raise CTERM ("dest_fun", [ct]);
 
-fun dest_arg (ct as Cterm {t = c $ a, T = _, thy_ref, maxidx, sorts}) =
+fun dest_arg (Cterm {t = c $ a, T = _, thy_ref, maxidx, sorts}) =
       let val A = Term.argument_type_of c 0
       in Cterm {t = a, T = A, thy_ref = thy_ref, maxidx = maxidx, sorts = sorts} end
   | dest_arg ct = raise CTERM ("dest_arg", [ct]);
 
 
-fun dest_fun2 (Cterm {t = c $ a $ b, T, thy_ref, maxidx, sorts}) =
+fun dest_fun2 (Cterm {t = c $ _ $ _, T, thy_ref, maxidx, sorts}) =
       let
         val A = Term.argument_type_of c 0;
         val B = Term.argument_type_of c 1;
@@ -254,8 +254,7 @@
       in Cterm {t = a, T = A, thy_ref = thy_ref, maxidx = maxidx, sorts = sorts} end
   | dest_arg1 ct = raise CTERM ("dest_arg1", [ct]);
 
-fun dest_abs a (ct as
-        Cterm {t = Abs (x, T, t), T = Type ("fun", [_, U]), thy_ref, maxidx, sorts}) =
+fun dest_abs a (Cterm {t = Abs (x, T, t), T = Type ("fun", [_, U]), thy_ref, maxidx, sorts}) =
       let val (y', t') = Term.dest_abs (the_default x a, T, t) in
         (Cterm {t = Free (y', T), T = T, thy_ref = thy_ref, maxidx = maxidx, sorts = sorts},
           Cterm {t = t', T = U, thy_ref = thy_ref, maxidx = maxidx, sorts = sorts})
@@ -392,10 +391,10 @@
 
 (* merge theories of cterms/thms -- trivial absorption only *)
 
-fun merge_thys1 (Cterm {thy_ref = r1, ...}) (th as Thm (_, {thy_ref = r2, ...})) =
+fun merge_thys1 (Cterm {thy_ref = r1, ...}) (Thm (_, {thy_ref = r2, ...})) =
   Theory.merge_refs (r1, r2);
 
-fun merge_thys2 (th1 as Thm (_, {thy_ref = r1, ...})) (th2 as Thm (_, {thy_ref = r2, ...})) =
+fun merge_thys2 (Thm (_, {thy_ref = r1, ...})) (Thm (_, {thy_ref = r2, ...})) =
   Theory.merge_refs (r1, r2);
 
 
@@ -541,7 +540,7 @@
 fun raw_body (Thm (Deriv {body, ...}, _)) = body;
 
 fun fulfill_body (Thm (Deriv {promises, body}, {thy_ref, ...})) =
-  Pt.fulfill_proof (Theory.deref thy_ref) ~1
+  Pt.fulfill_proof (Theory.deref thy_ref)
     (map #1 promises ~~ fulfill_bodies (map #2 promises)) body
 and fulfill_bodies futures = map fulfill_body (Exn.release_all (Future.join_results futures));
 
@@ -808,7 +807,7 @@
 (*Reflexivity
   t == t
 *)
-fun reflexive (ct as Cterm {thy_ref, t, T, maxidx, sorts}) =
+fun reflexive (Cterm {thy_ref, t, T = _, maxidx, sorts}) =
   Thm (deriv_rule0 Pt.reflexive,
    {thy_ref = thy_ref,
     tags = [],
@@ -825,7 +824,7 @@
 *)
 fun symmetric (th as Thm (der, {thy_ref, maxidx, shyps, hyps, tpairs, prop, ...})) =
   (case prop of
-    (eq as Const ("==", Type (_, [T, _]))) $ t $ u =>
+    (eq as Const ("==", _)) $ t $ u =>
       Thm (deriv_rule1 Pt.symmetric der,
        {thy_ref = thy_ref,
         tags = [],
@@ -868,7 +867,7 @@
   (%x. t)(u) == t[u/x]
   fully beta-reduces the term if full = true
 *)
-fun beta_conversion full (Cterm {thy_ref, t, T, maxidx, sorts}) =
+fun beta_conversion full (Cterm {thy_ref, t, T = _, maxidx, sorts}) =
   let val t' =
     if full then Envir.beta_norm t
     else
@@ -885,7 +884,7 @@
       prop = Logic.mk_equals (t, t')})
   end;
 
-fun eta_conversion (Cterm {thy_ref, t, T, maxidx, sorts}) =
+fun eta_conversion (Cterm {thy_ref, t, T = _, maxidx, sorts}) =
   Thm (deriv_rule0 Pt.reflexive,
    {thy_ref = thy_ref,
     tags = [],
@@ -895,7 +894,7 @@
     tpairs = [],
     prop = Logic.mk_equals (t, Envir.eta_contract t)});
 
-fun eta_long_conversion (Cterm {thy_ref, t, T, maxidx, sorts}) =
+fun eta_long_conversion (Cterm {thy_ref, t, T = _, maxidx, sorts}) =
   Thm (deriv_rule0 Pt.reflexive,
    {thy_ref = thy_ref,
     tags = [],
@@ -951,7 +950,7 @@
       prop = prop2, ...}) = th2;
     fun chktypes fT tT =
       (case fT of
-        Type ("fun", [T1, T2]) =>
+        Type ("fun", [T1, _]) =>
           if T1 <> tT then
             raise THM ("combination: types", 0, [th1, th2])
           else ()
@@ -1264,7 +1263,7 @@
 val varifyT = #2 o varifyT' [];
 
 (* Replace all TVars by new TFrees *)
-fun freezeT (Thm (der, {thy_ref, maxidx, shyps, hyps, tpairs, prop, ...})) =
+fun freezeT (Thm (der, {thy_ref, shyps, hyps, tpairs, prop, ...})) =
   let
     val prop1 = attach_tpairs tpairs prop;
     val prop2 = Type.freeze prop1;
@@ -1329,7 +1328,7 @@
 (*Solve subgoal Bi of proof state B1...Bn/C by assumption. *)
 fun assumption i state =
   let
-    val Thm (der, {thy_ref, maxidx, shyps, hyps, prop, ...}) = state;
+    val Thm (der, {thy_ref, maxidx, shyps, hyps, ...}) = state;
     val thy = Theory.deref thy_ref;
     val (tpairs, Bs, Bi, C) = dest_state (state, i);
     fun newth n (env, tpairs) =
@@ -1365,7 +1364,7 @@
   Checks if Bi's conclusion is alpha-convertible to one of its assumptions*)
 fun eq_assumption i state =
   let
-    val Thm (der, {thy_ref, maxidx, shyps, hyps, prop, ...}) = state;
+    val Thm (der, {thy_ref, maxidx, shyps, hyps, ...}) = state;
     val (tpairs, Bs, Bi, C) = dest_state (state, i);
     val (_, asms, concl) = Logic.assum_problems (~1, Bi);
   in
@@ -1386,7 +1385,7 @@
 (*For rotate_tac: fast rotation of assumptions of subgoal i*)
 fun rotate_rule k i state =
   let
-    val Thm (der, {thy_ref, maxidx, shyps, hyps, prop, ...}) = state;
+    val Thm (der, {thy_ref, maxidx, shyps, hyps, ...}) = state;
     val (tpairs, Bs, Bi, C) = dest_state (state, i);
     val params = Term.strip_all_vars Bi
     and rest   = Term.strip_all_body Bi;
@@ -1558,7 +1557,7 @@
       in Term.all T' $ Abs (a, T', norm_term_skip env n t) end
   | norm_term_skip env n (Const ("==>", _) $ A $ B) =
       Logic.mk_implies (A, norm_term_skip env (n - 1) B)
-  | norm_term_skip env n t = error "norm_term_skip: too few assumptions??";
+  | norm_term_skip _ _ _ = error "norm_term_skip: too few assumptions??";
 
 
 (*Composition of object rule r=(A1...Am/B) with proof state s=(B1...Bn/C)
--- a/src/Pure/type.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/type.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -140,7 +140,7 @@
 fun cert_class (TSig {classes, ...}) = Sorts.certify_class (#2 classes);
 fun cert_sort (TSig {classes, ...}) = Sorts.certify_sort (#2 classes);
 
-fun witness_sorts (tsig as TSig {classes, log_types, ...}) =
+fun witness_sorts (TSig {classes, log_types, ...}) =
   Sorts.witness_sorts (#2 classes) log_types;
 
 
@@ -280,7 +280,7 @@
     val used = Name.context
       |> fold_types (fold_atyps (fn TVar ((a, _), _) => Name.declare a | _ => I)) t;
     val fmap = fs ~~ map (rpair 0) (#1 (Name.variants (map fst fs) used));
-    fun thaw (f as (a, S)) =
+    fun thaw (f as (_, S)) =
       (case AList.lookup (op =) fmap f of
         NONE => TFree f
       | SOME xi => TVar (xi, S));
@@ -412,10 +412,10 @@
       (case lookup tye v of
         SOME U => devar tye U
       | NONE => T)
-  | devar tye T = T;
+  | devar _ T = T;
 
 (*order-sorted unification*)
-fun unify (tsig as TSig {classes = (_, classes), ...}) TU (tyenv, maxidx) =
+fun unify (TSig {classes = (_, classes), ...}) TU (tyenv, maxidx) =
   let
     val tyvar_count = Unsynchronized.ref maxidx;
     fun gen_tyvar S = TVar ((Name.aT, Unsynchronized.inc tyvar_count), S);
@@ -465,7 +465,7 @@
 (*purely structural unification*)
 fun raw_unify (ty1, ty2) tye =
   (case (devar tye ty1, devar tye ty2) of
-    (T as TVar (v, S1), U as TVar (w, S2)) =>
+    (T as TVar (v, S1), TVar (w, S2)) =>
       if Term.eq_ix (v, w) then
         if S1 = S2 then tye else tvar_clash v S1 S2
       else Vartab.update_new (w, (S2, T)) tye
@@ -545,7 +545,7 @@
     let
       val rel' = pairself (cert_class tsig) rel
         handle TYPE (msg, _, _) => error msg;
-      val classes' = classes |> Sorts.add_classrel pp rel;
+      val classes' = classes |> Sorts.add_classrel pp rel';
     in ((space, classes'), default, types) end);
 
 
--- a/src/Pure/variable.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Pure/variable.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -89,7 +89,7 @@
 structure Data = ProofDataFun
 (
   type T = data;
-  fun init thy =
+  fun init _ =
     make_data (false, Name.context, Symtab.empty, [], Vartab.empty, Symtab.empty,
       ~1, [], (Vartab.empty, Vartab.empty));
 );
--- a/src/Tools/Code/code_preproc.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Tools/Code/code_preproc.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -403,7 +403,7 @@
         @ (maps o maps) fst xs;
     fun type_variable (TFree (_, sort)) = map (pair []) (proj_sort sort);
   in
-    flat (Sorts.of_sort_derivation (Syntax.pp_global thy) algebra
+    flat (Sorts.of_sort_derivation algebra
       { class_relation = class_relation, type_constructor = type_constructor,
         type_variable = type_variable } (T, proj_sort sort)
        handle Sorts.CLASS_ERROR _ => [] (*permissive!*))
--- a/src/Tools/Code/code_thingol.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Tools/Code/code_thingol.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -750,7 +750,6 @@
   #>> (fn sort => (unprefix "'" v, sort))
 and translate_dicts thy (algbr as (proj_sort, algebra)) funcgr thm (ty, sort) =
   let
-    val pp = Syntax.pp_global thy;
     datatype typarg =
         Global of (class * string) * typarg list list
       | Local of (class * class) list * (string * (int * sort));
@@ -764,7 +763,7 @@
       let
         val sort' = proj_sort sort;
       in map_index (fn (n, class) => (Local ([], (v, (n, sort'))), class)) sort' end;
-    val typargs = Sorts.of_sort_derivation pp algebra
+    val typargs = Sorts.of_sort_derivation algebra
       {class_relation = class_relation, type_constructor = type_constructor,
        type_variable = type_variable} (ty, proj_sort sort)
       handle Sorts.CLASS_ERROR e => not_wellsorted thy thm ty sort e;
--- a/src/Tools/Code/lib/Tools/codegen	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Tools/Code/lib/Tools/codegen	Thu Oct 01 20:52:18 2009 +0200
@@ -60,6 +60,6 @@
 
 CTXT_CMD="ML_Context.eval_in (SOME (ProofContext.init (theory \"HOL\"))) false Position.none \"Code_Target.shell_command thyname cmd\";"
 
-FULL_CMD="$QND_CMD quick_and_dirty; val thyname = \"$THY\"; val cmd = \"$CODE_CMD\"; $CTXT_CMD"
+FULL_CMD="Unsynchronized.$QND_CMD quick_and_dirty; val thyname = \"$THY\"; val cmd = \"$CODE_CMD\"; $CTXT_CMD"
 
 "$ISABELLE" -q -e "$FULL_CMD" "$IMAGE" || exit 1
--- a/src/Tools/more_conv.ML	Thu Oct 01 20:49:46 2009 +0200
+++ b/src/Tools/more_conv.ML	Thu Oct 01 20:52:18 2009 +0200
@@ -46,16 +46,18 @@
   Conv.arg_conv (Conv.abs_conv (fn (_, cx) => cv cx) ctxt)
 
 
-fun cache_conv conv =   (* FIXME not thread-safe *)
-  let 
-    val tab = Unsynchronized.ref Termtab.empty
-    fun add_result t thm =
-      let val _ = Unsynchronized.change tab (Termtab.insert Thm.eq_thm (t, thm))
-      in thm end
-    fun cconv ct =  
-      (case Termtab.lookup (!tab) (Thm.term_of ct) of
+fun cache_conv conv =
+  let
+    val cache = Synchronized.var "cache_conv" Termtab.empty
+    fun lookup t =
+      Synchronized.change_result cache (fn tab => (Termtab.lookup tab t, tab))
+    val keep = Synchronized.change cache o Termtab.insert Thm.eq_thm
+    fun keep_result t thm = (keep (t, thm); thm)
+
+    fun cconv ct =
+      (case lookup (Thm.term_of ct) of
         SOME thm => thm
-      | NONE => add_result (Thm.term_of ct) (conv ct))
+      | NONE => keep_result (Thm.term_of ct) (conv ct))
   in cconv end
 
 end