merged
authorhaftmann
Mon, 10 Aug 2009 10:25:00 +0200
changeset 32351 96f9e6402403
parent 32350 5ef633275b15 (current diff)
parent 32336 e88b295aae35 (diff)
child 32352 4839a704939a
merged
lib/scripts/SystemOnTPTP
lib/scripts/neos/NeosCSDPClient.py
lib/scripts/neos/config.py
src/HOL/Library/sos_wrapper.ML
src/HOL/Library/sum_of_squares.ML
src/HOL/Tools/atp_manager.ML
src/HOL/Tools/atp_minimal.ML
src/HOL/Tools/atp_wrapper.ML
src/HOL/Tools/inductive_set.ML
src/HOL/ex/Predicate_Compile.thy
src/HOL/ex/Predicate_Compile_ex.thy
src/HOL/ex/predicate_compile.ML
--- a/Admin/isatest/isatest-makeall	Mon Aug 10 08:37:37 2009 +0200
+++ b/Admin/isatest/isatest-makeall	Mon Aug 10 10:25:00 2009 +0200
@@ -80,6 +80,11 @@
         NICE=""
         ;;
 
+    macbroy21)
+        MFLAGS="-k"
+        NICE=""
+        ;;
+
     macbroy23)
         MFLAGS="-k -j 2"
         NICE="nice"
--- a/Admin/isatest/isatest-makedist	Mon Aug 10 08:37:37 2009 +0200
+++ b/Admin/isatest/isatest-makedist	Mon Aug 10 10:25:00 2009 +0200
@@ -106,7 +106,7 @@
 sleep 15
 $SSH atbroy101 "$MAKEALL $HOME/settings/at64-poly"
 sleep 15
-$SSH macbroy2 "$MAKEALL $HOME/settings/mac-poly-M4; $MAKEALL $HOME/settings/mac-poly-M8; $MAKEALL $HOME/settings/mac-poly64-M4"
+$SSH macbroy2 "$MAKEALL $HOME/settings/mac-poly-M4; $MAKEALL $HOME/settings/mac-poly-M8; $MAKEALL $HOME/settings/mac-poly64-M4; $MAKEALL $HOME/settings/mac-poly64-M8"
 sleep 15
 $SSH macbroy5 "$MAKEALL $HOME/settings/mac-poly"
 sleep 15
--- a/Admin/isatest/isatest-stats	Mon Aug 10 08:37:37 2009 +0200
+++ b/Admin/isatest/isatest-stats	Mon Aug 10 10:25:00 2009 +0200
@@ -6,7 +6,7 @@
 
 THIS=$(cd "$(dirname "$0")"; pwd -P)
 
-PLATFORMS="at-poly at64-poly mac-poly-M4 mac-poly-M8 at-poly-5.1-para-e at64-poly-5.1-para at-mac-poly-5.1-para afp at-sml-dev"
+PLATFORMS="at-poly at64-poly mac-poly-M4 mac-poly64-M4 mac-poly-M8 at-poly-5.1-para-e at64-poly-5.1-para at-mac-poly-5.1-para afp at-sml-dev"
 
 ISABELLE_SESSIONS="\
   HOL-Plain \
--- a/Admin/isatest/settings/at-poly	Mon Aug 10 08:37:37 2009 +0200
+++ b/Admin/isatest/settings/at-poly	Mon Aug 10 10:25:00 2009 +0200
@@ -22,6 +22,6 @@
 ISABELLE_OUTPUT="$ISABELLE_HOME_USER/heaps"
 ISABELLE_BROWSER_INFO="$ISABELLE_HOME_USER/browser_info"
 
-ISABELLE_USEDIR_OPTIONS="-i true -d pdf -v true"
+ISABELLE_USEDIR_OPTIONS="-i true -d pdf -v true -t true"
 
 HOL_USEDIR_OPTIONS="-p 2"
--- a/Admin/isatest/settings/mac-poly64-M4	Mon Aug 10 08:37:37 2009 +0200
+++ b/Admin/isatest/settings/mac-poly64-M4	Mon Aug 10 10:25:00 2009 +0200
@@ -4,7 +4,7 @@
   ML_SYSTEM="polyml-experimental"
   ML_PLATFORM="x86_64-darwin"
   ML_HOME="$POLYML_HOME/$ML_PLATFORM"
-  ML_OPTIONS="--mutable 2000 --immutable 2000"
+  ML_OPTIONS="--mutable 5000 --immutable 2000"
 
 
 ISABELLE_HOME_USER=~/isabelle-mac-poly64-M4
@@ -23,6 +23,6 @@
 ISABELLE_OUTPUT="$ISABELLE_HOME_USER/heaps"
 ISABELLE_BROWSER_INFO="$ISABELLE_HOME_USER/browser_info"
 
-ISABELLE_USEDIR_OPTIONS="-i false -d false -M 4"
+ISABELLE_USEDIR_OPTIONS="-i false -d false -M 4 -q 2 -t true"
 
-HOL_USEDIR_OPTIONS="-p 2 -q 1"
+HOL_USEDIR_OPTIONS="-p 2 -q 2"
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Admin/isatest/settings/mac-poly64-M8	Mon Aug 10 10:25:00 2009 +0200
@@ -0,0 +1,28 @@
+# -*- shell-script -*- :mode=shellscript:
+
+  POLYML_HOME="/home/polyml/polyml-svn"
+  ML_SYSTEM="polyml-experimental"
+  ML_PLATFORM="x86_64-darwin"
+  ML_HOME="$POLYML_HOME/$ML_PLATFORM"
+  ML_OPTIONS="--mutable 5000 --immutable 2000"
+
+
+ISABELLE_HOME_USER=~/isabelle-mac-poly64-M8
+
+# Where to look for isabelle tools (multiple dirs separated by ':').
+ISABELLE_TOOLS="$ISABELLE_HOME/lib/Tools"
+
+# Location for temporary files (should be on a local file system).
+ISABELLE_TMP_PREFIX="/tmp/isabelle-$USER"
+
+
+# Heap input locations. ML system identifier is included in lookup.
+ISABELLE_PATH="$ISABELLE_HOME_USER/heaps:$ISABELLE_HOME/heaps"
+
+# Heap output location. ML system identifier is appended automatically later on.
+ISABELLE_OUTPUT="$ISABELLE_HOME_USER/heaps"
+ISABELLE_BROWSER_INFO="$ISABELLE_HOME_USER/browser_info"
+
+ISABELLE_USEDIR_OPTIONS="-i false -d false -M 8 -q 2 -t true"
+
+HOL_USEDIR_OPTIONS="-p 2 -q 2"
--- a/NEWS	Mon Aug 10 08:37:37 2009 +0200
+++ b/NEWS	Mon Aug 10 10:25:00 2009 +0200
@@ -175,6 +175,12 @@
 
 *** System ***
 
+* Support for additional "Isabelle components" via etc/components, see
+also the system manual.
+
+* The isabelle makeall tool now operates on all components with
+IsaMakefile, not just hardwired "logics".
+
 * Discontinued support for Poly/ML 4.x versions.
 
 * Removed "compress" option from isabelle-process and isabelle usedir;
--- a/bin/isabelle	Mon Aug 10 08:37:37 2009 +0200
+++ b/bin/isabelle	Mon Aug 10 10:25:00 2009 +0200
@@ -17,6 +17,8 @@
 ISABELLE_HOME="$(cd "$(dirname "$0")"; cd "$(pwd -P)"; cd ..; pwd)"
 source "$ISABELLE_HOME/lib/scripts/getsettings" || exit 2
 
+ORIG_IFS="$IFS"; IFS=":"; declare -a TOOLS=($ISABELLE_TOOLS); IFS="$ORIG_IFS"
+
 
 ## diagnostics
 
@@ -28,24 +30,19 @@
   echo "  Start Isabelle tool NAME with ARGS; pass \"-?\" for tool specific help."
   echo
   echo "  Available tools are:"
-  (
-    ORIG_IFS="$IFS"
-    IFS=":"
-    for DIR in $ISABELLE_TOOLS
-    do
-      if [ -d "$DIR" ]; then
-        cd "$DIR"
-        for T in *
-        do
-          if [ -f "$T" -a -x "$T" ]; then
-            DESCRLINE=$(fgrep DESCRIPTION: "$T" | sed -e 's/^.*DESCRIPTION: *//')
-            echo "    $T - $DESCRLINE"
-          fi
-        done
-      fi
-    done
-    IFS="$ORIG_IFS"
-  )
+  for DIR in ${TOOLS[@]}
+  do
+    if [ -d "$DIR" ]; then
+      for TOOL in "$DIR"/*
+      do
+        if [ -f "$TOOL" -a -x "$TOOL" ]; then
+          NAME="$(basename "$TOOL")"
+          DESCRLINE="$(fgrep DESCRIPTION: "$TOOL" | sed -e 's/^.*DESCRIPTION: *//')"
+          echo "    $NAME - $DESCRLINE"
+        fi
+      done
+    fi
+  done
   exit 1
 }
 
@@ -66,13 +63,10 @@
 
 ## main
 
-ORIG_IFS="$IFS"
-IFS=":"
-for DIR in $ISABELLE_TOOLS
+for DIR in "${TOOLS[@]}"
 do
   TOOL="$DIR/$TOOLNAME"
   [ -f "$TOOL" -a -x "$TOOL" ] && exec "$TOOL" "$@"
 done
-IFS="$ORIG_IFS"
 
 fail "Unknown Isabelle tool: $TOOLNAME"
--- a/bin/isabelle-process	Mon Aug 10 08:37:37 2009 +0200
+++ b/bin/isabelle-process	Mon Aug 10 10:25:00 2009 +0200
@@ -160,15 +160,13 @@
     INFILE=""
     ISA_PATH=""
 
-    ORIG_IFS="$IFS"
-    IFS=":"
-    for DIR in $ISABELLE_PATH
+    ORIG_IFS="$IFS"; IFS=":"; declare -a PATHS=($ISABELLE_PATH); IFS="$ORIG_IFS"
+    for DIR in "${PATHS[@]}"
     do
       DIR="$DIR/$ML_IDENTIFIER"
       ISA_PATH="$ISA_PATH  $DIR\n"
       [ -z "$INFILE" -a -f "$DIR/$INPUT" ] && INFILE="$DIR/$INPUT"
     done
-    IFS="$ORIG_IFS"
 
     if [ -z "$INFILE" ]; then
       echo "Unknown logic \"$INPUT\" -- no heap file found in:" >&2
--- a/doc-src/IsarImplementation/Thy/Proof.thy	Mon Aug 10 08:37:37 2009 +0200
+++ b/doc-src/IsarImplementation/Thy/Proof.thy	Mon Aug 10 10:25:00 2009 +0200
@@ -94,7 +94,7 @@
   @{index_ML Variable.export: "Proof.context -> Proof.context -> thm list -> thm list"} \\
   @{index_ML Variable.polymorphic: "Proof.context -> term list -> term list"} \\
   @{index_ML Variable.import: "bool -> thm list -> Proof.context ->
-  ((ctyp list * cterm list) * thm list) * Proof.context"} \\
+  (((ctyp * ctyp) list * (cterm * cterm) list) * thm list) * Proof.context"} \\
   @{index_ML Variable.focus: "cterm -> Proof.context -> ((string * cterm) list * cterm) * Proof.context"} \\
   \end{mldecls}
 
--- a/doc-src/IsarImplementation/Thy/document/Proof.tex	Mon Aug 10 08:37:37 2009 +0200
+++ b/doc-src/IsarImplementation/Thy/document/Proof.tex	Mon Aug 10 10:25:00 2009 +0200
@@ -112,7 +112,7 @@
   \indexdef{}{ML}{Variable.export}\verb|Variable.export: Proof.context -> Proof.context -> thm list -> thm list| \\
   \indexdef{}{ML}{Variable.polymorphic}\verb|Variable.polymorphic: Proof.context -> term list -> term list| \\
   \indexdef{}{ML}{Variable.import}\verb|Variable.import: bool -> thm list -> Proof.context ->|\isasep\isanewline%
-\verb|  ((ctyp list * cterm list) * thm list) * Proof.context| \\
+\verb|  (((ctyp * ctyp) list * (cterm * cterm) list) * thm list) * Proof.context| \\
   \indexdef{}{ML}{Variable.focus}\verb|Variable.focus: cterm -> Proof.context -> ((string * cterm) list * cterm) * Proof.context| \\
   \end{mldecls}
 
--- a/doc-src/System/Thy/Basics.thy	Mon Aug 10 08:37:37 2009 +0200
+++ b/doc-src/System/Thy/Basics.thy	Mon Aug 10 10:25:00 2009 +0200
@@ -59,11 +59,14 @@
 *}
 
 
-subsection {* Building the environment *}
+subsection {* Bootstrapping the environment \label{sec:boot} *}
 
-text {*
-  Whenever any of the Isabelle executables is run, their settings
-  environment is put together as follows.
+text {* Isabelle executables need to be run within a proper settings
+  environment.  This is bootstrapped as described below, on the first
+  invocation of one of the outer wrapper scripts (such as
+  @{executable_ref isabelle}).  This happens only once for each
+  process tree, i.e.\ the environment is passed to subprocesses
+  according to regular Unix conventions.
 
   \begin{enumerate}
 
@@ -78,7 +81,7 @@
   links are admissible, but a plain copy of the @{"file"
   "$ISABELLE_HOME/bin"} files will not work!
 
-  \item The file @{"file" "$ISABELLE_HOME/etc/settings"} ist run as a
+  \item The file @{"file" "$ISABELLE_HOME/etc/settings"} is run as a
   @{executable_ref bash} shell script with the auto-export option for
   variables enabled.
   
@@ -252,6 +255,52 @@
 *}
 
 
+subsection {* Additional components \label{sec:components} *}
+
+text {* Any directory may be registered as an explicit \emph{Isabelle
+  component}.  The general layout conventions are that of the main
+  Isabelle distribution itself, and the following two files (both
+  optional) have a special meaning:
+
+  \begin{itemize}
+
+  \item @{verbatim "etc/settings"} holds additional settings that are
+  initialized when bootstrapping the overall Isabelle environment,
+  cf.\ \secref{sec:boot}.  As usual, the content is interpreted as a
+  @{verbatim bash} script.  It may refer to the component's enclosing
+  directory via the @{verbatim "COMPONENT"} shell variable.
+
+  For example, the following setting allows to refer to files within
+  the component later on, without having to hardwire absolute paths:
+
+  \begin{ttbox}
+  MY_COMPONENT_HOME="$COMPONENT"
+  \end{ttbox}
+
+  Components can also add to existing Isabelle settings such as
+  @{setting_def ISABELLE_TOOLS}, in order to provide
+  component-specific tools that can be invoked by end-users.  For
+  example:
+
+  \begin{ttbox}
+  ISABELLE_TOOLS="$ISABELLE_TOOLS:$COMPONENT/lib/Tools"
+  \end{ttbox}
+
+  \item @{verbatim "etc/components"} holds a list of further
+  sub-components of the same structure.  The directory specifications
+  given here can be either absolute (with leading @{verbatim "/"}) or
+  relative to the component's main directory.
+
+  \end{itemize}
+
+  The root of component initialization is @{setting ISABELLE_HOME}
+  itself.  After initializing all of its sub-components recursively,
+  @{setting ISABELLE_HOME_USER} is included in the same manner (if
+  that directory exists).  Thus users can easily add private
+  components to @{verbatim "$ISABELLE_HOME_USER/etc/components"}.
+*}
+
+
 section {* The raw Isabelle process *}
 
 text {*
--- a/doc-src/System/Thy/Misc.thy	Mon Aug 10 08:37:37 2009 +0200
+++ b/doc-src/System/Thy/Misc.thy	Mon Aug 10 10:25:00 2009 +0200
@@ -225,13 +225,13 @@
 
 section {* Make all logics *}
 
-text {*
-  The @{tool_def makeall} utility applies Isabelle make to all logic
-  directories of the distribution:
+text {* The @{tool_def makeall} utility applies Isabelle make to any
+  Isabelle component (cf.\ \secref{sec:components}) that contains an
+  @{verbatim IsaMakefile}:
 \begin{ttbox}
 Usage: makeall [ARGS ...]
 
-  Apply isabelle make to all logics (passing ARGS).
+  Apply isabelle make to all components with IsaMakefile (passing ARGS).
 \end{ttbox}
 
   The arguments @{verbatim ARGS} are just passed verbatim to each
--- a/doc-src/System/Thy/document/Basics.tex	Mon Aug 10 08:37:37 2009 +0200
+++ b/doc-src/System/Thy/document/Basics.tex	Mon Aug 10 10:25:00 2009 +0200
@@ -76,13 +76,17 @@
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
-\isamarkupsubsection{Building the environment%
+\isamarkupsubsection{Bootstrapping the environment \label{sec:boot}%
 }
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-Whenever any of the Isabelle executables is run, their settings
-  environment is put together as follows.
+Isabelle executables need to be run within a proper settings
+  environment.  This is bootstrapped as described below, on the first
+  invocation of one of the outer wrapper scripts (such as
+  \indexref{}{executable}{isabelle}\hyperlink{executable.isabelle}{\mbox{\isa{\isatt{isabelle}}}}).  This happens only once for each
+  process tree, i.e.\ the environment is passed to subprocesses
+  according to regular Unix conventions.
 
   \begin{enumerate}
 
@@ -96,7 +100,7 @@
   executable objects created by the \hyperlink{tool.install}{\mbox{\isa{\isatt{install}}}} utility.  Symbolic
   links are admissible, but a plain copy of the \hyperlink{file.$ISABELLE-HOME/bin}{\mbox{\isa{\isatt{{\isachardollar}ISABELLE{\isacharunderscore}HOME{\isacharslash}bin}}}} files will not work!
 
-  \item The file \hyperlink{file.$ISABELLE-HOME/etc/settings}{\mbox{\isa{\isatt{{\isachardollar}ISABELLE{\isacharunderscore}HOME{\isacharslash}etc{\isacharslash}settings}}}} ist run as a
+  \item The file \hyperlink{file.$ISABELLE-HOME/etc/settings}{\mbox{\isa{\isatt{{\isachardollar}ISABELLE{\isacharunderscore}HOME{\isacharslash}etc{\isacharslash}settings}}}} is run as a
   \indexref{}{executable}{bash}\hyperlink{executable.bash}{\mbox{\isa{\isatt{bash}}}} shell script with the auto-export option for
   variables enabled.
   
@@ -259,6 +263,55 @@
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
+\isamarkupsubsection{Additional components \label{sec:components}%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+Any directory may be registered as an explicit \emph{Isabelle
+  component}.  The general layout conventions are that of the main
+  Isabelle distribution itself, and the following two files (both
+  optional) have a special meaning:
+
+  \begin{itemize}
+
+  \item \verb|etc/settings| holds additional settings that are
+  initialized when bootstrapping the overall Isabelle environment,
+  cf.\ \secref{sec:boot}.  As usual, the content is interpreted as a
+  \verb|bash| script.  It may refer to the component's enclosing
+  directory via the \verb|COMPONENT| shell variable.
+
+  For example, the following setting allows to refer to files within
+  the component later on, without having to hardwire absolute paths:
+
+  \begin{ttbox}
+  MY_COMPONENT_HOME="$COMPONENT"
+  \end{ttbox}
+
+  Components can also add to existing Isabelle settings such as
+  \indexdef{}{setting}{ISABELLE\_TOOLS}\hypertarget{setting.ISABELLE-TOOLS}{\hyperlink{setting.ISABELLE-TOOLS}{\mbox{\isa{\isatt{ISABELLE{\isacharunderscore}TOOLS}}}}}, in order to provide
+  component-specific tools that can be invoked by end-users.  For
+  example:
+
+  \begin{ttbox}
+  ISABELLE_TOOLS="$ISABELLE_TOOLS:$COMPONENT/lib/Tools"
+  \end{ttbox}
+
+  \item \verb|etc/components| holds a list of further
+  sub-components of the same structure.  The directory specifications
+  given here can be either absolute (with leading \verb|/|) or
+  relative to the component's main directory.
+
+  \end{itemize}
+
+  The root of component initialization is \hyperlink{setting.ISABELLE-HOME}{\mbox{\isa{\isatt{ISABELLE{\isacharunderscore}HOME}}}}
+  itself.  After initializing all of its sub-components recursively,
+  \hyperlink{setting.ISABELLE-HOME-USER}{\mbox{\isa{\isatt{ISABELLE{\isacharunderscore}HOME{\isacharunderscore}USER}}}} is included in the same manner (if
+  that directory exists).  Thus users can easily add private
+  components to \verb|$ISABELLE_HOME_USER/etc/components|.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
 \isamarkupsection{The raw Isabelle process%
 }
 \isamarkuptrue%
--- a/doc-src/System/Thy/document/Misc.tex	Mon Aug 10 08:37:37 2009 +0200
+++ b/doc-src/System/Thy/document/Misc.tex	Mon Aug 10 10:25:00 2009 +0200
@@ -259,12 +259,13 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-The \indexdef{}{tool}{makeall}\hypertarget{tool.makeall}{\hyperlink{tool.makeall}{\mbox{\isa{\isatt{makeall}}}}} utility applies Isabelle make to all logic
-  directories of the distribution:
+The \indexdef{}{tool}{makeall}\hypertarget{tool.makeall}{\hyperlink{tool.makeall}{\mbox{\isa{\isatt{makeall}}}}} utility applies Isabelle make to any
+  Isabelle component (cf.\ \secref{sec:components}) that contains an
+  \verb|IsaMakefile|:
 \begin{ttbox}
 Usage: makeall [ARGS ...]
 
-  Apply isabelle make to all logics (passing ARGS).
+  Apply isabelle make to all components with IsaMakefile (passing ARGS).
 \end{ttbox}
 
   The arguments \verb|ARGS| are just passed verbatim to each
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/etc/components	Mon Aug 10 10:25:00 2009 +0200
@@ -0,0 +1,16 @@
+#main object logics
+src/Pure
+src/FOL
+src/HOL
+src/ZF
+src/CCL
+src/CTT
+src/Cube
+src/FOLP
+src/HOLCF
+src/LCF
+src/Sequents
+#misc components
+src/HOL/Tools/ATP_Manager
+src/HOL/Library/Sum_Of_Squares
+
--- a/etc/settings	Mon Aug 10 08:37:37 2009 +0200
+++ b/etc/settings	Mon Aug 10 10:25:00 2009 +0200
@@ -144,6 +144,7 @@
 # Site settings check -- just to make it a little bit harder to copy this file verbatim!
 [ -n "$ISABELLE_SITE_SETTINGS_PRESENT" ] && \
   { echo >&2 "### Isabelle site settings already present!  Maybe copied etc/settings in full?"; }
+ISABELLE_SITE_SETTINGS_PRESENT=true
 
 
 ###
@@ -221,7 +222,6 @@
 #JEDIT_JAVA_OPTIONS="-server -Xms128m -Xmx512m"
 JEDIT_OPTIONS="-reuseview -noserver -nobackground"
 
-
 ###
 ### External reasoning tools
 ###
@@ -273,6 +273,9 @@
 # Jerusat 1.3 (SAT Solver, cf. Isabelle/src/HOL/Tools/sat_solver.ML)
 #JERUSAT_HOME=/usr/local/bin
 
+# CSDP (SDP Solver, cf. Isabelle/src/HOL/Library/Sum_of_Squares/sos_wrapper.ML)
+#CSDP_EXE=csdp
+
 # For configuring HOL/Matrix/cplex
 # LP_SOLVER is the default solver. It can be changed during runtime via Cplex.set_solver.
 # First option: use the commercial cplex solver
--- a/lib/Tools/doc	Mon Aug 10 08:37:37 2009 +0200
+++ b/lib/Tools/doc	Mon Aug 10 10:25:00 2009 +0200
@@ -34,28 +34,23 @@
 
 ## main
 
+ORIG_IFS="$IFS"; IFS=":"; declare -a DOCS=($ISABELLE_DOCS); IFS="$ORIG_IFS"
+
 if [ -z "$DOC" ]; then
-  ORIG_IFS="$IFS"
-  IFS=":"
-  for DIR in $ISABELLE_DOCS
+  for DIR in "${DOCS[@]}"
   do
     [ -d "$DIR" ] || fail "Bad document directory: $DIR"
     [ -f "$DIR/Contents" ] && grep -v "^>>" "$DIR/Contents"
   done
-  IFS="$ORIG_IFS"
 else
-  ORIG_IFS="$IFS"
-  IFS=":"
-  for DIR in $ISABELLE_DOCS
+  for DIR in "${DOCS[@]}"
   do
-    IFS="$ORIG_IFS"
     [ -d "$DIR" ] || fail "Bad document directory: $DIR"
     for FMT in "$ISABELLE_DOC_FORMAT" dvi
     do
       [ -f "$DIR/$DOC.$FMT" ] && { cd "$DIR"; exec "$ISABELLE_TOOL" display "$DOC.$FMT"; }
     done
   done
-  IFS="$ORIG_IFS"
   fail "Unknown Isabelle document: $DOC"  
 fi
 
--- a/lib/Tools/document	Mon Aug 10 08:37:37 2009 +0200
+++ b/lib/Tools/document	Mon Aug 10 10:25:00 2009 +0200
@@ -38,7 +38,7 @@
 CLEAN=""
 NAME="document"
 OUTFORMAT=dvi
-TAGS=""
+declare -a TAGS=()
 
 while getopts "cn:o:t:" OPT
 do
@@ -53,7 +53,7 @@
       OUTFORMAT="$OPTARG"
       ;;
     t)
-      TAGS="$OPTARG"
+      ORIG_IFS="$IFS"; IFS=","; TAGS=($OPTARG); IFS="$ORIG_IFS"
       ;;
     \?)
       usage
@@ -90,21 +90,20 @@
 function prep_tags ()
 {
   (
-    IFS=","
-    for TAG in $TAGS
+    for TAG in "${TAGS[@]}"
     do
       case "$TAG" in
         /*)
-  	  echo "\\isafoldtag{${TAG:1}}"
+          echo "\\isafoldtag{${TAG:1}}"
           ;;
         -*)
-  	  echo "\\isadroptag{${TAG:1}}"
+          echo "\\isadroptag{${TAG:1}}"
           ;;
         +*)
-  	  echo "\\isakeeptag{${TAG:1}}"
+          echo "\\isakeeptag{${TAG:1}}"
           ;;
         *)
-  	  echo "\\isakeeptag{${TAG}}"
+          echo "\\isakeeptag{${TAG}}"
           ;;
       esac
     done
--- a/lib/Tools/findlogics	Mon Aug 10 08:37:37 2009 +0200
+++ b/lib/Tools/findlogics	Mon Aug 10 10:25:00 2009 +0200
@@ -22,22 +22,21 @@
 
 [ "$#" -ne 0 ] && usage
 
-
-LOGICS=""
+declare -a LOGICS=()
+declare -a ISABELLE_PATHS=()
 
-ORIG_IFS="$IFS"
-IFS=":"
-for DIR in $ISABELLE_PATH
+ORIG_IFS="$IFS"; IFS=":"; ISABELLE_PATHS=($ISABELLE_PATH); IFS=$ORIG_IFS
+
+for DIR in "${ISABELLE_PATHS[@]}"
 do
   DIR="$DIR/$ML_IDENTIFIER"
   for FILE in "$DIR"/*
   do
     if [ -f "$FILE" ]; then
       NAME=$(basename "$FILE")
-      LOGICS="$LOGICS $NAME"
+      LOGICS+=("$NAME")
     fi
   done
 done
-IFS="$ORIG_IFS"
 
-echo $({ for L in $LOGICS; do echo "$L"; done; } | sort | uniq)
+echo $({ for L in ${LOGICS[@]}; do echo "$L"; done; } | sort | uniq)
--- a/lib/Tools/makeall	Mon Aug 10 08:37:37 2009 +0200
+++ b/lib/Tools/makeall	Mon Aug 10 10:25:00 2009 +0200
@@ -4,11 +4,6 @@
 #
 # DESCRIPTION: apply make utility to all logics
 
-## global settings
-
-ALL_LOGICS="Pure FOL HOL ZF CCL CTT Cube FOLP HOLCF LCF Sequents"
-
-
 ## diagnostics
 
 PRG="$(basename "$0")"
@@ -18,7 +13,7 @@
   echo
   echo "Usage: isabelle $PRG [ARGS ...]"
   echo
-  echo "  Apply isabelle make to all logics (passing ARGS)."
+  echo "  Apply isabelle make to all components with IsaMakefile (passing ARGS)."
   echo
   exit 1
 }
@@ -29,6 +24,7 @@
   exit 2
 }
 
+
 ## main
 
 [ "$1" = "-?" ] && usage
@@ -38,9 +34,14 @@
 echo "Started at $(date) ($ML_IDENTIFIER on $(hostname))"
 . "$ISABELLE_HOME/lib/scripts/timestart.bash"
 
-for L in $ALL_LOGICS
+ORIG_IFS="$IFS"; IFS=":"; declare -a COMPONENTS=($ISABELLE_COMPONENTS); IFS="$ORIG_IFS"
+
+for DIR in "${COMPONENTS[@]}"
 do
-  ( cd "$ISABELLE_HOME/src/$L"; "$ISABELLE_TOOL" make "$@" ) || FAIL="$FAIL$L "
+  if [ -f "$DIR/IsaMakefile" ]; then
+    NAME="$(basename "$DIR")"
+    ( cd "$DIR"; "$ISABELLE_TOOL" make "$@" ) || FAIL="$FAIL$NAME "
+  fi
 done
 
 echo -n "Finished at "; date
--- a/lib/Tools/mkdir	Mon Aug 10 08:37:37 2009 +0200
+++ b/lib/Tools/mkdir	Mon Aug 10 10:25:00 2009 +0200
@@ -187,8 +187,8 @@
   [ -z "$QUIET" ] && echo "creating $PREFIX/ROOT.ML" >&2
   cat >ROOT.ML <<EOF
 (*
-  no_document use_thy "ThisTheory";
-  use_thy "ThatTheory";
+  no_document use_thys ["This_Theory1", "This_Theory2"];
+  use_thys ["That_Theory1", "That_Theory2", "That_Theory3"];
 *)
 EOF
 fi
--- a/lib/scripts/SystemOnTPTP	Mon Aug 10 08:37:37 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,141 +0,0 @@
-#!/usr/bin/env perl
-#
-# Wrapper for custom remote provers on SystemOnTPTP
-# Author: Fabian Immler, TU Muenchen
-#
-
-use warnings;
-use strict;
-use Getopt::Std;
-use HTTP::Request::Common;
-use LWP;
-
-my $SystemOnTPTPFormReplyURL =
-  "http://www.cs.miami.edu/~tptp/cgi-bin/SystemOnTPTPFormReply";
-
-# default parameters
-my %URLParameters = (
-    "NoHTML" => 1,
-    "QuietFlag" => "-q01",
-    "SubmitButton" => "RunSelectedSystems",
-    "ProblemSource" => "UPLOAD",
-    );
-
-#----Get format and transform options if specified
-my %Options;
-getopts("hwxs:t:c:",\%Options);
-
-#----Usage
-sub usage() {
-  print("Usage: remote [<options>] <File name>\n");
-  print("    <options> are ...\n");
-  print("    -h            - print this help\n");
-  print("    -w            - list available ATP systems\n");
-  print("    -x            - use X2TPTP to convert output of prover\n");
-  print("    -s<system>    - specified system to use\n");
-  print("    -t<timelimit> - CPU time limit for system\n");
-  print("    -c<command>   - custom command for system\n");
-  print("    <File name>   - TPTP problem file\n");
-  exit(0);
-}
-if (exists($Options{'h'})) {
-  usage();
-}
-
-#----What systems flag
-if (exists($Options{'w'})) {
-    $URLParameters{"SubmitButton"} = "ListSystems";
-    delete($URLParameters{"ProblemSource"});
-}
-
-#----X2TPTP
-if (exists($Options{'x'})) {
-    $URLParameters{"X2TPTP"} = "-S";
-}
-
-#----Selected system
-my $System;
-if (exists($Options{'s'})) {
-    $System = $Options{'s'};
-} else {
-    # use Vampire as default
-    $System = "Vampire---9.0";
-}
-$URLParameters{"System___$System"} = $System;
-
-#----Time limit
-if (exists($Options{'t'})) {
-    $URLParameters{"TimeLimit___$System"} = $Options{'t'};
-}
-#----Custom command
-if (exists($Options{'c'})) {
-    $URLParameters{"Command___$System"} = $Options{'c'};
-}
-
-#----Get single file name
-if (exists($URLParameters{"ProblemSource"})) {
-    if (scalar(@ARGV) >= 1) {
-        $URLParameters{"UPLOADProblem"} = [shift(@ARGV)];
-    } else {
-      print("Missing problem file\n");
-      usage();
-      die;
-    }
-}
-
-# Query Server
-my $Agent = LWP::UserAgent->new;
-if (exists($Options{'t'})) {
-  # give server more time to respond
-  $Agent->timeout($Options{'t'} + 10);
-}
-my $Request = POST($SystemOnTPTPFormReplyURL,
-	Content_Type => 'form-data',Content => \%URLParameters);
-my $Response = $Agent->request($Request);
-
-#catch errors / failure
-if(!$Response->is_success) {
-  print "HTTP-Error: " . $Response->message . "\n";
-  exit(-1);
-} elsif (exists($Options{'w'})) {
-  print $Response->content;
-  exit (0);
-} elsif ($Response->content =~ /WARNING: (\S*) does not exist/) {
-  print "Specified System $1 does not exist\n";
-  exit(-1);
-} elsif (exists($Options{'x'}) &&
-  $Response->content =~
-    /%\s*Result\s*:\s*Unsatisfiable.*\n%\s*Output\s*:\s*(CNF)?Refutation.*\n%/ &&
-  $Response->content !~ /ERROR: Could not form TPTP format derivation/ )
-{
-  # converted output: extract proof
-  my @lines = split( /\n/, $Response->content);
-  my $extract = "";
-  foreach my $line (@lines){
-      #ignore comments
-      if ($line !~ /^%/ && !($line eq "")) {
-          $extract .= "$line";
-      }
-  }
-  # insert newlines after ').'
-  $extract =~ s/\s//g;
-  $extract =~ s/\)\.cnf/\)\.\ncnf/g;
-
-  print "========== ~~/lib/scripts/SystemOnTPTP extracted proof: ==========\n";
-  # orientation for res_reconstruct.ML
-  print "# SZS output start CNFRefutation.\n";
-  print "$extract\n";
-  print "# SZS output end CNFRefutation.\n";
-  # can be useful for debugging; Isabelle ignores this
-  print "============== original response from SystemOnTPTP: ==============\n";
-  print $Response->content;
-  exit(0);
-} elsif (!exists($Options{'x'})) {
-  # pass output directly to Isabelle
-  print $Response->content;
-  exit(0);
-}else {
-  print "Remote-script could not extract proof:\n".$Response->content;
-  exit(-1);
-}
-
--- a/lib/scripts/getsettings	Mon Aug 10 08:37:37 2009 +0200
+++ b/lib/scripts/getsettings	Mon Aug 10 10:25:00 2009 +0200
@@ -68,14 +68,42 @@
   done
 }
 
-#get actual settings
-source "$ISABELLE_HOME/etc/settings" || exit 2
-ISABELLE_SITE_SETTINGS_PRESENT=true
+#nested components
+ISABELLE_COMPONENTS=""
+function init_component ()
+{
+  local COMPONENT="$1"
+
+  if [ ! -d "$COMPONENT" ]; then
+    echo >&2 "Bad Isabelle component: $COMPONENT"
+    exit 2
+  elif [ -z "$ISABELLE_COMPONENTS" ]; then
+    ISABELLE_COMPONENTS="$COMPONENT"
+  else
+    ISABELLE_COMPONENTS="$ISABELLE_COMPONENTS:$COMPONENT"
+  fi
+  if [ -f "$COMPONENT/etc/settings" ]; then
+    source "$COMPONENT/etc/settings" || exit 2
+  fi
+  if [ -f "$COMPONENT/etc/components" ]; then
+    {
+      while read; do
+        case "$REPLY" in
+          \#* | "") ;;
+          /*) init_component "$REPLY" ;;
+          *) init_component "$COMPONENT/$REPLY" ;;
+        esac
+      done
+    } < "$COMPONENT/etc/components"
+  fi
+}
+
+#main components
+init_component "$ISABELLE_HOME"
 
 [ "$ISABELLE_HOME" -ef "$ISABELLE_HOME_USER" ] && \
-  { echo >&2 "### ISABELLE_HOME and ISABELLE_HOME_USER should not be the same directory!"; }
-[ -z "$ISABELLE_IGNORE_USER_SETTINGS" -a -f "$ISABELLE_HOME_USER/etc/settings" ] && \
-  { source "$ISABELLE_HOME_USER/etc/settings" || exit 2; }
+  { echo >&2 "### ISABELLE_HOME and ISABELLE_HOME_USER must not be the same directory!"; }
+[ -d "$ISABELLE_HOME_USER" ] && init_component "$ISABELLE_HOME_USER"
 
 #ML system identifier
 if [ -z "$ML_PLATFORM" ]; then
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/lib/scripts/mirabelle	Mon Aug 10 10:25:00 2009 +0200
@@ -0,0 +1,129 @@
+#!/usr/bin/perl -w
+
+use strict;
+use File::Basename;
+
+# Taken from http://www.skywayradio.com/tech/perl/trim_blanks.html
+sub trim {
+    my @out = @_;
+    for (@out) {
+        s/^\s+//;
+        s/\s+$//;
+    }
+    return wantarray ? @out : $out[0];
+}
+
+sub quote {
+    my $str = pop;
+    return "\"" . $str . "\"";
+}
+
+sub print_usage_and_quit {
+    print STDERR "Usage: mirabelle actions file1.thy...\n" .
+                 "  actions: action1:...:actionN\n" .
+                 "  action: name or name[key1=value1,...,keyM=valueM]\n";
+    exit 1;
+}
+
+my $num_args = $#ARGV + 1;
+if ($num_args < 2) {
+    print_usage_and_quit();
+}
+
+my @action_names;
+my @action_settings;
+
+foreach (split(/:/, $ARGV[0])) {
+    my %settings;
+
+    $_ =~ /([^[]*)(?:\[(.*)\])?/;
+    my ($name, $settings_str) = ($1, $2 || "");
+    my @setting_strs = split(/,/, $settings_str);
+    foreach (@setting_strs) {
+        $_ =~ /(.*)=(.*)/;
+	    my $key = $1;
+	    my $value = $2;
+	    $settings{trim($key)} = trim($value);
+    }
+
+    push @action_names, trim($name);
+    push @action_settings, \%settings;
+}
+
+my $output_path = "/tmp/mirabelle"; # FIXME: generate path
+my $mirabellesetup_thy_name = $output_path . "/MirabelleSetup";
+my $mirabellesetup_file = $mirabellesetup_thy_name . ".thy";
+my $mirabelle_log_file = $output_path . "/mirabelle.log";
+
+mkdir $output_path, 0755;
+
+open(FILE, ">$mirabellesetup_file")
+    || die "Could not create file '$mirabellesetup_file'";
+
+my $invoke_lines;
+
+for my $i (0 .. $#action_names) { 
+    my $settings_str = "";
+    my $settings = $action_settings[$i];
+    my $key;
+    my $value;
+
+    while (($key, $value) = each(%$settings)) {
+        $settings_str .= "(" . quote ($key) . ", " . quote ($value) . "), ";
+    }
+    $settings_str =~ s/, $//;
+
+    $invoke_lines .= "setup {* Mirabelle.invoke \"$action_names[$i]\" ";
+    $invoke_lines .= "[$settings_str] *}\n"
+}
+
+print FILE <<EOF;
+theory MirabelleSetup
+imports Mirabelle
+begin
+
+setup {* Mirabelle.set_logfile "$mirabelle_log_file" *}
+
+$invoke_lines
+
+end
+EOF
+
+my $root_text = "";
+my @new_thy_files;
+
+for my $i (1 .. $num_args - 1) {
+    my $old_thy_file = $ARGV[$i];
+    my ($base, $dir, $ext) = fileparse($old_thy_file, "\.thy");
+    my $new_thy_name = $base . "Mirabelle";
+    my $new_thy_file = $dir . $new_thy_name . $ext;
+
+    open(OLD_FILE, "<$old_thy_file")
+        || die "Cannot open file $old_thy_file";
+    my @lines = <OLD_FILE>;
+    close(OLD_FILE);
+
+    my $thy_text = join("", @lines);
+    my $old_len = length($thy_text);
+    $thy_text =~ s/\btheory\b[^\n]*\s*\bimports\s/theory $new_thy_name\nimports "$mirabellesetup_thy_name" /gm;
+    die "No 'imports' found" if length($thy_text) == $old_len;
+
+    open(NEW_FILE, ">$new_thy_file");
+    print NEW_FILE $thy_text;
+    close(NEW_FILE);
+
+    $root_text .= "use_thy \"" . $dir . $new_thy_name . "\";\n";
+
+    push @new_thy_files, $new_thy_file;
+}
+
+my $root_file = "ROOT_mirabelle.ML";
+open(ROOT_FILE, ">$root_file") || die "Cannot open file $root_file";
+print ROOT_FILE $root_text;
+close(ROOT_FILE);
+
+system "isabelle-process -e 'use \"ROOT_mirabelle.ML\";' -f -q HOL";
+
+# unlink $mirabellesetup_file;
+unlink $root_file;
+unlink @new_thy_files;
--- a/lib/scripts/neos/NeosCSDPClient.py	Mon Aug 10 08:37:37 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,56 +0,0 @@
-#!/usr/bin/env python
-import sys
-import xmlrpclib
-import time
-import re
-
-from config import Variables
-
-if len(sys.argv) < 3 or len(sys.argv) > 3:
-  sys.stderr.write("Usage: NeosCSDPClient <input_filename> <output_filename>\n")
-  sys.exit(1)
-
-neos=xmlrpclib.Server("http://%s:%d" % (Variables.NEOS_HOST, Variables.NEOS_PORT))
-
-xmlfile = open(sys.argv[1],"r")
-xml_pre = "<document>\n<category>sdp</category>\n<solver>csdp</solver>\n<inputMethod>SPARSE_SDPA</inputMethod>\n<dat><![CDATA["
-xml_post = "]]></dat>\n</document>\n"
-xml = xml_pre
-buffer = 1
-while buffer:
-  buffer = xmlfile.read()
-  xml += buffer
-xmlfile.close()
-xml += xml_post
-
-(jobNumber,password) = neos.submitJob(xml)
-
-if jobNumber == 0:
-  sys.stdout.write("error submitting job: %s" % password)
-  sys.exit(-1)
-else:
-  sys.stdout.write("jobNumber = %d\tpassword = %s\n" % (jobNumber,password))
-
-offset=0
-status="Waiting"
-while status == "Running" or status=="Waiting":
-  time.sleep(1)
-  (msg,offset) = neos.getIntermediateResults(jobNumber,password,offset)
-  sys.stdout.write(msg.data)
-  status = neos.getJobStatus(jobNumber, password)
-
-msg = neos.getFinalResults(jobNumber, password).data
-result = msg.split("Solution:")
-
-sys.stdout.write(result[0])
-if len(result) > 1:
-  plain_msg = result[1].strip()
-  if plain_msg != "":
-    output = open(sys.argv[2],"w")
-    output.write(plain_msg)
-    output.close()
-    sys.exit(0)
-
-sys.exit(2)
-
-
--- a/lib/scripts/neos/config.py	Mon Aug 10 08:37:37 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,3 +0,0 @@
-class Variables:
-  NEOS_HOST="neos.mcs.anl.gov"
-  NEOS_PORT=3332
--- a/src/HOL/ATP_Linkup.thy	Mon Aug 10 08:37:37 2009 +0200
+++ b/src/HOL/ATP_Linkup.thy	Mon Aug 10 10:25:00 2009 +0200
@@ -15,9 +15,9 @@
   ("Tools/res_hol_clause.ML")
   ("Tools/res_reconstruct.ML")
   ("Tools/res_atp.ML")
-  ("Tools/atp_manager.ML")
-  ("Tools/atp_wrapper.ML")
-  ("Tools/atp_minimal.ML")
+  ("Tools/ATP_Manager/atp_manager.ML")
+  ("Tools/ATP_Manager/atp_wrapper.ML")
+  ("Tools/ATP_Manager/atp_minimal.ML")
   "~~/src/Tools/Metis/metis.ML"
   ("Tools/metis_tools.ML")
 begin
@@ -96,10 +96,9 @@
 use "Tools/res_reconstruct.ML" setup ResReconstruct.setup
 use "Tools/res_atp.ML"
 
-use "Tools/atp_manager.ML"
-use "Tools/atp_wrapper.ML"
-
-use "Tools/atp_minimal.ML"
+use "Tools/ATP_Manager/atp_manager.ML"
+use "Tools/ATP_Manager/atp_wrapper.ML"
+use "Tools/ATP_Manager/atp_minimal.ML"
 
 text {* basic provers *}
 setup {* AtpManager.add_prover "spass" AtpWrapper.spass *}
--- a/src/HOL/IsaMakefile	Mon Aug 10 08:37:37 2009 +0200
+++ b/src/HOL/IsaMakefile	Mon Aug 10 10:25:00 2009 +0200
@@ -231,12 +231,13 @@
   $(SRC)/Provers/Arith/combine_numerals.ML \
   $(SRC)/Provers/Arith/extract_common_term.ML \
   $(SRC)/Tools/Metis/metis.ML \
+  Tools/ATP_Manager/atp_manager.ML \
+  Tools/ATP_Manager/atp_minimal.ML \
+  Tools/ATP_Manager/atp_wrapper.ML \
   Tools/Groebner_Basis/groebner.ML \
   Tools/Groebner_Basis/misc.ML \
+  Tools/Groebner_Basis/normalizer.ML \
   Tools/Groebner_Basis/normalizer_data.ML \
-  Tools/Groebner_Basis/normalizer.ML \
-  Tools/atp_manager.ML \
-  Tools/atp_wrapper.ML \
   Tools/int_arith.ML \
   Tools/list_code.ML \
   Tools/meson.ML \
@@ -315,49 +316,45 @@
 
 HOL-Library: HOL $(LOG)/HOL-Library.gz
 
-
 $(LOG)/HOL-Library.gz: $(OUT)/HOL Library/SetsAndFunctions.thy		\
-  Library/Abstract_Rat.thy \
-  Library/BigO.thy Library/ContNotDenum.thy Library/Efficient_Nat.thy	\
-  Library/Euclidean_Space.thy Library/Sum_Of_Squares.thy Library/positivstellensatz.ML	\
-  Library/Fset.thy  Library/Convex_Euclidean_Space.thy \
-  Library/sum_of_squares.ML Library/Glbs.thy Library/normarith.ML \
-  Library/Executable_Set.thy Library/Infinite_Set.thy			\
-  Library/FuncSet.thy Library/Permutations.thy Library/Determinants.thy\
-  Library/Bit.thy Library/Topology_Euclidean_Space.thy \
-  Library/Finite_Cartesian_Product.thy \
-  Library/FrechetDeriv.thy Library/Fraction_Field.thy\
-  Library/Fundamental_Theorem_Algebra.thy \
-  Library/Inner_Product.thy Library/Kleene_Algebra.thy Library/Lattice_Syntax.thy \
-  Library/Legacy_GCD.thy \
-  Library/Library.thy Library/List_Prefix.thy Library/List_Set.thy Library/State_Monad.thy	\
-  Library/Nat_Int_Bij.thy Library/Multiset.thy Library/Permutation.thy	\
-  Library/Primes.thy Library/Pocklington.thy Library/Quotient.thy	\
-  Library/Quicksort.thy Library/Nat_Infinity.thy Library/Word.thy	\
-  Library/README.html Library/Continuity.thy Library/Order_Relation.thy \
-  Library/Nested_Environment.thy Library/Ramsey.thy Library/Zorn.thy	\
-  Library/Library/ROOT.ML Library/Library/document/root.tex		\
-  Library/Library/document/root.bib Library/While_Combinator.thy	\
-  Library/Product_ord.thy Library/Char_nat.thy Library/Char_ord.thy	\
-  Library/Option_ord.thy Library/Sublist_Order.thy			\
-  Library/List_lexord.thy Library/Commutative_Ring.thy			\
-  Library/comm_ring.ML Library/Coinductive_List.thy			\
-  Library/AssocList.thy	Library/Formal_Power_Series.thy	\
-  Library/Binomial.thy Library/Eval_Witness.thy				\
-  Library/Code_Char.thy				\
+  Library/Abstract_Rat.thy Library/BigO.thy Library/ContNotDenum.thy	\
+  Library/Efficient_Nat.thy Library/Euclidean_Space.thy			\
+  Library/Sum_Of_Squares.thy Library/Sum_Of_Squares/sos_wrapper.ML	\
+  Library/Sum_Of_Squares/sum_of_squares.ML Library/Fset.thy		\
+  Library/Convex_Euclidean_Space.thy Library/Glbs.thy			\
+  Library/normarith.ML Library/Executable_Set.thy			\
+  Library/Infinite_Set.thy Library/FuncSet.thy				\
+  Library/Permutations.thy Library/Determinants.thy Library/Bit.thy	\
+  Library/Topology_Euclidean_Space.thy					\
+  Library/Finite_Cartesian_Product.thy Library/FrechetDeriv.thy		\
+  Library/Fraction_Field.thy Library/Fundamental_Theorem_Algebra.thy	\
+  Library/Inner_Product.thy Library/Kleene_Algebra.thy			\
+  Library/Lattice_Syntax.thy Library/Legacy_GCD.thy			\
+  Library/Library.thy Library/List_Prefix.thy Library/List_Set.thy	\
+  Library/State_Monad.thy Library/Nat_Int_Bij.thy Library/Multiset.thy	\
+  Library/Permutation.thy Library/Primes.thy Library/Pocklington.thy	\
+  Library/Quotient.thy Library/Quicksort.thy Library/Nat_Infinity.thy	\
+  Library/Word.thy Library/README.html Library/Continuity.thy		\
+  Library/Order_Relation.thy Library/Nested_Environment.thy		\
+  Library/Ramsey.thy Library/Zorn.thy Library/Library/ROOT.ML		\
+  Library/Library/document/root.tex Library/Library/document/root.bib	\
+  Library/While_Combinator.thy Library/Product_ord.thy			\
+  Library/Char_nat.thy Library/Char_ord.thy Library/Option_ord.thy	\
+  Library/Sublist_Order.thy Library/List_lexord.thy			\
+  Library/Commutative_Ring.thy Library/comm_ring.ML			\
+  Library/Coinductive_List.thy Library/AssocList.thy			\
+  Library/Formal_Power_Series.thy Library/Binomial.thy			\
+  Library/Eval_Witness.thy Library/Code_Char.thy			\
   Library/Code_Char_chr.thy Library/Code_Integer.thy			\
   Library/Mapping.thy Library/Numeral_Type.thy Library/Reflection.thy	\
   Library/Boolean_Algebra.thy Library/Countable.thy			\
   Library/Diagonalize.thy Library/RBT.thy Library/Univ_Poly.thy		\
-  Library/Poly_Deriv.thy \
-  Library/Polynomial.thy \
-  Library/Preorder.thy \
-  Library/Product_plus.thy \
-  Library/Product_Vector.thy \
-  Library/Tree.thy \
-  Library/Enum.thy Library/Float.thy $(SRC)/Tools/float.ML $(SRC)/HOL/Tools/float_arith.ML \
-  Library/reify_data.ML Library/reflection.ML \
-  Library/LaTeXsugar.thy Library/OptionalSugar.thy
+  Library/Poly_Deriv.thy Library/Polynomial.thy Library/Preorder.thy	\
+  Library/Product_plus.thy Library/Product_Vector.thy Library/Tree.thy	\
+  Library/Enum.thy Library/Float.thy $(SRC)/Tools/float.ML		\
+  $(SRC)/HOL/Tools/float_arith.ML Library/positivstellensatz.ML		\
+  Library/reify_data.ML Library/reflection.ML Library/LaTeXsugar.thy	\
+  Library/OptionalSugar.thy
 	@cd Library; $(ISABELLE_TOOL) usedir $(OUT)/HOL Library
 
 
@@ -903,7 +900,8 @@
   ex/Sudoku.thy ex/Tarski.thy \
   ex/Termination.thy ex/Unification.thy ex/document/root.bib		\
   ex/document/root.tex ex/set.thy ex/svc_funcs.ML ex/svc_test.thy \
-  ex/Predicate_Compile.thy ex/predicate_compile.ML ex/Predicate_Compile_ex.thy
+  ex/Predicate_Compile.thy ex/predicate_compile.ML ex/Predicate_Compile_ex.thy \
+  ex/Mirabelle.thy ex/mirabelle.ML
 	@$(ISABELLE_TOOL) usedir $(OUT)/HOL ex
 
 
--- a/src/HOL/Library/Sum_Of_Squares.thy	Mon Aug 10 08:37:37 2009 +0200
+++ b/src/HOL/Library/Sum_Of_Squares.thy	Mon Aug 10 10:25:00 2009 +0200
@@ -1,26 +1,39 @@
-(* Title:      Library/Sum_Of_Squares
+(* Title:      HOL/Library/Sum_Of_Squares.thy
    Author:     Amine Chaieb, University of Cambridge
-
-In order to use the method sos, call it with (sos remote_csdp) to use the remote solver
-or install CSDP (https://projects.coin-or.org/Csdp/), put the executable csdp on your path,
-and call it with (sos csdp). By default, sos calls remote_csdp. This can take of the order
-of a minute for one sos call, because sos calls CSDP repeatedly.
-If you install CSDP locally, sos calls typically takes only a few seconds.
-
 *)
 
 header {* A decision method for universal multivariate real arithmetic with addition, 
-          multiplication and ordering using semidefinite programming*}
+  multiplication and ordering using semidefinite programming *}
 
 theory Sum_Of_Squares
-  imports Complex_Main (* "~~/src/HOL/Decision_Procs/Dense_Linear_Order" *)
-  uses "positivstellensatz.ML" "sum_of_squares.ML" "sos_wrapper.ML"
-  begin
+imports Complex_Main (* "~~/src/HOL/Decision_Procs/Dense_Linear_Order" *)
+uses
+  ("positivstellensatz.ML")  (* duplicate use!? -- cf. Euclidian_Space.thy *)
+  ("Sum_Of_Squares/sum_of_squares.ML")
+  ("Sum_Of_Squares/sos_wrapper.ML")
+begin
 
-(* setup sos tactic *)
+text {*
+  In order to use the method sos, call it with @{text "(sos
+  remote_csdp)"} to use the remote solver.  Or install CSDP
+  (https://projects.coin-or.org/Csdp), configure the Isabelle setting
+  @{text CSDP_EXE}, and call it with @{text "(sos csdp)"}.  By
+  default, sos calls @{text remote_csdp}.  This can take of the order
+  of a minute for one sos call, because sos calls CSDP repeatedly.  If
+  you install CSDP locally, sos calls typically takes only a few
+  seconds.
+*}
+
+text {* setup sos tactic *}
+
+use "positivstellensatz.ML"
+use "Sum_Of_Squares/sum_of_squares.ML"
+use "Sum_Of_Squares/sos_wrapper.ML"
+
 setup SosWrapper.setup
 
-text{* Tests -- commented since they work only when csdp is installed  or take too long with remote csdps *}
+text {* Tests -- commented since they work only when csdp is installed
+  or take too long with remote csdps *}
 
 (*
 lemma "(3::real) * x + 7 * a < 4 & 3 < 2 * x \<Longrightarrow> a < 0" by sos
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Sum_Of_Squares/etc/settings	Mon Aug 10 10:25:00 2009 +0200
@@ -0,0 +1,1 @@
+ISABELLE_SUM_OF_SQUARES="$COMPONENT"
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Sum_Of_Squares/neos_csdp_client	Mon Aug 10 10:25:00 2009 +0200
@@ -0,0 +1,65 @@
+#!/usr/bin/env python
+import sys
+import xmlrpclib
+import time
+import re
+
+# Neos server config
+NEOS_HOST="neos.mcs.anl.gov"
+NEOS_PORT=3332
+
+if len(sys.argv) < 3 or len(sys.argv) > 3:
+  sys.stderr.write("Usage: NeosCSDPClient <input_filename> <output_filename>\n")
+  sys.exit(1)
+
+neos=xmlrpclib.Server("http://%s:%d" % (NEOS_HOST, NEOS_PORT))
+
+inputfile = open(sys.argv[1],"r")
+xml_pre = "<document>\n<category>sdp</category>\n<solver>csdp</solver>\n<inputMethod>SPARSE_SDPA</inputMethod>\n<dat><![CDATA["
+xml_post = "]]></dat>\n</document>\n"
+xml = xml_pre
+buffer = 1
+while buffer:
+  buffer = inputfile.read()
+  xml += buffer
+inputfile.close()
+xml += xml_post
+
+(jobNumber,password) = neos.submitJob(xml)
+
+if jobNumber == 0:
+  sys.stdout.write("error submitting job: %s" % password)
+  sys.exit(20)
+else:
+  sys.stdout.write("jobNumber = %d\tpassword = %s\n" % (jobNumber,password))
+
+offset=0
+messages = ""
+status="Waiting"
+while status == "Running" or status=="Waiting":
+  time.sleep(1)
+  (msg,offset) = neos.getIntermediateResults(jobNumber,password,offset)
+  messages += msg.data
+  sys.stdout.write(msg.data)
+  status = neos.getJobStatus(jobNumber, password)
+
+msg = neos.getFinalResults(jobNumber, password).data
+sys.stdout.write("---------- Begin CSDP Output -------------\n");
+sys.stdout.write(msg)
+
+# extract solution
+result = msg.split("Solution:")
+if len(result) > 1:
+  output = open(sys.argv[2],"w")
+  output.write(result[1].strip())
+  output.close()
+
+# extract return code
+p = re.compile(r"^Error: Command exited with non-zero status (\d+)$", re.MULTILINE)
+m = p.search(messages)
+if m:
+  sys.exit(int(m.group(1)))
+else:
+  sys.exit(0)
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Sum_Of_Squares/sos_wrapper.ML	Mon Aug 10 10:25:00 2009 +0200
@@ -0,0 +1,140 @@
+(* Title:      sos_wrapper.ML
+   Author:     Philipp Meyer, TU Muenchen
+
+Added functionality for sums of squares, e.g. calling a remote prover
+*)
+
+signature SOS_WRAPPER =
+sig
+
+  datatype prover_result = Success | Failure | Error
+
+  val setup: theory -> theory
+  val destdir: string ref
+end
+
+structure SosWrapper : SOS_WRAPPER=
+struct
+
+datatype prover_result = Success | Failure | Error
+fun str_of_result Success = "Success"
+  | str_of_result Failure = "Failure"
+  | str_of_result Error = "Error"
+
+(*** output control ***)
+
+fun debug s = if ! Sos.debugging then Output.writeln s else ()
+val write = Output.priority
+
+(*** calling provers ***)
+
+val destdir = ref ""
+
+fun filename dir name =
+  let
+    val probfile = Path.basic (name ^ serial_string ())
+    val dir_path = Path.explode dir
+  in
+    if dir = "" then
+      File.tmp_path probfile
+    else
+      if File.exists dir_path then
+        Path.append dir_path probfile
+      else
+        error ("No such directory: " ^ dir)
+  end
+
+fun run_solver name cmd find_failure input =
+  let
+    val _ = write ("Calling solver: " ^ name)
+
+    (* create input file *)
+    val dir = ! destdir
+    val input_file = filename dir "sos_in" 
+    val _ = File.write input_file input
+
+    (* call solver *)
+    val output_file = filename dir "sos_out"
+    val (output, rv) = system_out (
+      if File.exists cmd then space_implode " "
+        [File.shell_path cmd, File.platform_path input_file, File.platform_path output_file]
+      else error ("Bad executable: " ^ File.shell_path cmd))
+ 
+    (* read and analysize output *)
+    val (res, res_msg) = find_failure rv
+    val result = if File.exists output_file then File.read output_file else ""
+
+    (* remove temporary files *)
+    val _ = if dir = "" then
+        (File.rm input_file ; if File.exists output_file then File.rm output_file else ())
+        else ()
+
+    val _ = debug ("Solver output:\n" ^ output)
+
+    val _ = write (str_of_result res ^ ": " ^ res_msg)
+  in
+    case res of
+      Success => result
+    | Failure => raise Sos.Failure res_msg
+    | Error => error ("Prover failed: " ^ res_msg)
+  end
+
+(*** various provers ***)
+
+(* local csdp client *)
+
+fun find_csdp_failure rv =
+  case rv of
+    0 => (Success, "SDP solved")
+  | 1 => (Failure, "SDP is primal infeasible")
+  | 2 => (Failure, "SDP is dual infeasible")
+  | 3 => (Success, "SDP solved with reduced accuracy")
+  | 4 => (Failure, "Maximum iterations reached")
+  | 5 => (Failure, "Stuck at edge of primal feasibility")
+  | 6 => (Failure, "Stuck at edge of dual infeasibility")
+  | 7 => (Failure, "Lack of progress")
+  | 8 => (Failure, "X, Z, or O was singular")
+  | 9 => (Failure, "Detected NaN or Inf values")
+  | _ => (Error, "return code is " ^ string_of_int rv)
+
+val csdp = ("$CSDP_EXE", find_csdp_failure)
+
+(* remote neos server *)
+
+fun find_neos_failure rv =
+  case rv of
+    20 => (Error, "error submitting job")
+  | 21 => (Error, "no solution")
+  |  _ => find_csdp_failure rv
+
+val neos_csdp = ("$ISABELLE_SUM_OF_SQUARES/neos_csdp_client", find_neos_failure)
+
+(* save provers in table *)
+
+val provers =
+     Symtab.make [("remote_csdp", neos_csdp),("csdp", csdp)]
+
+fun get_prover name =
+  case Symtab.lookup provers name of
+    SOME prover => prover
+  | NONE => error ("unknown prover: " ^ name)
+
+fun call_solver name =
+  let
+    val (cmd, find_failure) = get_prover name
+  in
+    run_solver name (Path.explode cmd) find_failure
+  end
+
+(* setup tactic *)
+
+val def_solver = "remote_csdp"
+
+fun sos_solver name = (SIMPLE_METHOD' o (Sos.sos_tac (call_solver name))) 
+
+val sos_method = Scan.optional (Scan.lift OuterParse.xname) def_solver >> sos_solver
+
+val setup = Method.setup @{binding sos} sos_method
+  "Prove universal problems over the reals using sums of squares"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML	Mon Aug 10 10:25:00 2009 +0200
@@ -0,0 +1,1512 @@
+(* Title:      sum_of_squares.ML
+   Authors:    Amine Chaieb, University of Cambridge
+               Philipp Meyer, TU Muenchen
+
+A tactic for proving nonlinear inequalities
+*)
+
+signature SOS =
+sig
+
+  val sos_tac : (string -> string) -> Proof.context -> int -> Tactical.tactic
+
+  val debugging : bool ref;
+  
+  exception Failure of string;
+end
+
+structure Sos : SOS = 
+struct
+
+val rat_0 = Rat.zero;
+val rat_1 = Rat.one;
+val rat_2 = Rat.two;
+val rat_10 = Rat.rat_of_int 10;
+val rat_1_2 = rat_1 // rat_2;
+val max = curry IntInf.max;
+val min = curry IntInf.min;
+
+val denominator_rat = Rat.quotient_of_rat #> snd #> Rat.rat_of_int;
+val numerator_rat = Rat.quotient_of_rat #> fst #> Rat.rat_of_int;
+fun int_of_rat a = 
+    case Rat.quotient_of_rat a of (i,1) => i | _ => error "int_of_rat: not an int";
+fun lcm_rat x y = Rat.rat_of_int (Integer.lcm (int_of_rat x) (int_of_rat y));
+
+fun rat_pow r i = 
+ let fun pow r i = 
+   if i = 0 then rat_1 else 
+   let val d = pow r (i div 2)
+   in d */ d */ (if i mod 2 = 0 then rat_1 else r)
+   end
+ in if i < 0 then pow (Rat.inv r) (~ i) else pow r i end;
+
+fun round_rat r = 
+ let val (a,b) = Rat.quotient_of_rat (Rat.abs r)
+     val d = a div b
+     val s = if r </ rat_0 then (Rat.neg o Rat.rat_of_int) else Rat.rat_of_int
+     val x2 = 2 * (a - (b * d))
+ in s (if x2 >= b then d + 1 else d) end
+
+val abs_rat = Rat.abs;
+val pow2 = rat_pow rat_2;
+val pow10 = rat_pow rat_10;
+
+val debugging = ref false;
+
+exception Sanity;
+
+exception Unsolvable;
+
+exception Failure of string;
+
+(* Turn a rational into a decimal string with d sig digits.                  *)
+
+local
+fun normalize y =
+  if abs_rat y </ (rat_1 // rat_10) then normalize (rat_10 */ y) - 1
+  else if abs_rat y >=/ rat_1 then normalize (y // rat_10) + 1
+  else 0 
+ in
+fun decimalize d x =
+  if x =/ rat_0 then "0.0" else
+  let 
+   val y = Rat.abs x
+   val e = normalize y
+   val z = pow10(~ e) */ y +/ rat_1
+   val k = int_of_rat (round_rat(pow10 d */ z))
+  in (if x </ rat_0 then "-0." else "0.") ^
+     implode(tl(explode(string_of_int k))) ^
+     (if e = 0 then "" else "e"^string_of_int e)
+  end
+end;
+
+(* Iterations over numbers, and lists indexed by numbers.                    *)
+
+fun itern k l f a =
+  case l of
+    [] => a
+  | h::t => itern (k + 1) t f (f h k a);
+
+fun iter (m,n) f a =
+  if n < m then a
+  else iter (m+1,n) f (f m a);
+
+(* The main types.                                                           *)
+
+fun strict_ord ord (x,y) = case ord (x,y) of LESS => LESS | _ => GREATER
+
+structure Intpairfunc = FuncFun(type key = int*int val ord = prod_ord int_ord int_ord);
+
+type vector = int* Rat.rat Intfunc.T;
+
+type matrix = (int*int)*(Rat.rat Intpairfunc.T);
+
+type monomial = int Ctermfunc.T;
+
+val cterm_ord = (fn (s,t) => TermOrd.fast_term_ord(term_of s, term_of t))
+ fun monomial_ord (m1,m2) = list_ord (prod_ord cterm_ord int_ord) (Ctermfunc.graph m1, Ctermfunc.graph m2)
+structure Monomialfunc = FuncFun(type key = monomial val ord = monomial_ord)
+
+type poly = Rat.rat Monomialfunc.T;
+
+ fun iszero (k,r) = r =/ rat_0;
+
+fun fold_rev2 f l1 l2 b =
+  case (l1,l2) of
+    ([],[]) => b
+  | (h1::t1,h2::t2) => f h1 h2 (fold_rev2 f t1 t2 b)
+  | _ => error "fold_rev2";
+ 
+(* Vectors. Conventionally indexed 1..n.                                     *)
+
+fun vector_0 n = (n,Intfunc.undefined):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;
+
+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))
+ end;
+
+fun vector_neg (v:vector) = (fst v,Intfunc.mapf 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 
+ end;
+
+fun vector_sub v1 v2 = vector_add v1 (vector_neg v2);
+
+fun vector_dot (v1:vector) (v2:vector) =
+ 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
+ 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
+ end;
+
+(* Matrices; again rows and columns indexed from 1.                          *)
+
+fun matrix_0 (m,n) = ((m,n),Intpairfunc.undefined):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;;
+
+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))
+ end;
+
+fun matrix_neg (m:matrix) = 
+  (dimensions m, Intpairfunc.mapf 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
+ end;;
+
+fun matrix_sub m1 m2 = matrix_add m1 (matrix_neg m2);
+
+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
+ 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)
+   : 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
+ 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
+ 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)
+   end
+ end;
+
+(* Monomials.                                                                *)
+
+fun monomial_eval assig (m:monomial) =
+  Ctermfunc.fold (fn (x, k) => fn a => a */ rat_pow (Ctermfunc.apply assig x) k)
+        m rat_1;
+val monomial_1 = (Ctermfunc.undefined:monomial);
+
+fun monomial_var x = Ctermfunc.onefunc (x, 1) :monomial;
+
+val (monomial_mul:monomial->monomial->monomial) =
+  Ctermfunc.combine (curry op +) (K false);
+
+fun monomial_pow (m:monomial) k =
+  if k = 0 then monomial_1
+  else Ctermfunc.mapf (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_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
+  else error "monomial_div: non-divisible"
+ end;
+
+fun monomial_degree x (m:monomial) = 
+  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_multidegree (m:monomial) = 
+ Ctermfunc.fold (fn (x, k) => fn a => k + a) m 0;;
+
+fun monomial_variables m = 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;
+
+val poly_0 = (Monomialfunc.undefined:poly);
+
+fun poly_isconst (p:poly) = 
+  Monomialfunc.fold (fn (m, c) => fn a => Ctermfunc.is_undefined m andalso a) p true;
+
+fun poly_var x = Monomialfunc.onefunc (monomial_var x,rat_1) :poly;
+
+fun poly_const c =
+  if c =/ rat_0 then poly_0 else Monomialfunc.onefunc(monomial_1, c);
+
+fun poly_cmul c (p:poly) =
+  if c =/ rat_0 then poly_0
+  else Monomialfunc.mapf (fn x => c */ x) p;
+
+fun poly_neg (p:poly) = (Monomialfunc.mapf Rat.neg p :poly);;
+
+fun poly_add (p1:poly) (p2:poly) =
+  (Monomialfunc.combine (curry op +/) (fn x => x =/ rat_0) p1 p2 :poly);
+
+fun poly_sub p1 p2 = poly_add p1 (poly_neg p2);
+
+fun poly_cmmul (c,m) (p:poly) =
+ 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;
+
+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_div (p1:poly) (p2:poly) =
+ if not(poly_isconst p2) 
+ then error "poly_div: non-constant" else
+ let val c = eval Ctermfunc.undefined p2 
+ in if c =/ rat_0 then error "poly_div: division by zero"
+    else poly_cmul (Rat.inv c) p1
+ end;
+
+fun poly_square p = poly_mul p p;
+
+fun poly_pow p k =
+ if k = 0 then poly_const rat_1
+ else if k = 1 then p
+ else let val q = poly_square(poly_pow p (k div 2)) in
+      if k mod 2 = 1 then poly_mul p q else q end;
+
+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));
+
+fun degree x (p:poly) = 
+ 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 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 []);;
+
+(* 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);
+
+local
+ fun ord (l1,l2) = case (l1,l2) of
+  (_,[]) => LESS 
+ | ([],_) => GREATER
+ | (h1::t1,h2::t2) => 
+   (case humanorder_varpow (h1, h2) of 
+     LESS => LESS
+   | 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))
+end;
+
+fun fold1 f l =  case l of
+   []     => error "fold1"
+ | [x]    => x
+ | (h::t) => f h (fold1 f t);
+
+(* Conversions to strings.                                                   *)
+
+fun string_of_vector min_size max_size (v:vector) =
+ let val n_raw = dim v 
+ 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 "[" ^ fold1 (fn s => fn t => s ^ ", " ^ t) xs ^
+  (if n_raw > max_size then ", ...]" else "]")
+  end
+ end;
+
+fun string_of_matrix max_size (m:matrix) =
+ let 
+  val (i_raw,j_raw) = dimensions m
+  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 "["^ fold1 (fn s => fn t => s^";\n "^t) rstr ^
+  (if j > max_size then "\n ...]" else "]")
+ end;
+
+fun string_of_term t = 
+ case t of
+   a$b => "("^(string_of_term a)^" "^(string_of_term b)^")"
+ | Abs x => 
+    let val (xn, b) = Term.dest_abs x
+    in "(\\"^xn^"."^(string_of_term b)^")"
+    end
+ | Const(s,_) => s
+ | Free (s,_) => s
+ | Var((s,_),_) => s
+ | _ => error "string_of_term";
+
+val string_of_cterm = string_of_term o term_of;
+
+fun string_of_varpow x k =
+  if k = 1 then string_of_cterm x 
+  else string_of_cterm x^"^"^string_of_int k;
+
+fun string_of_monomial m =
+ if Ctermfunc.is_undefined 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 fold1 (fn s => fn t => s^"*"^t) vps
+ end;
+
+fun string_of_cmonomial (c,m) =
+ if Ctermfunc.is_undefined 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:poly) =
+ if Monomialfunc.is_undefined p then "<<0>>" else
+ let 
+  val cms = sort (fn ((m1,_),(m2,_)) => humanorder_monomial m1  m2) (Monomialfunc.graph 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))
+          cms ""
+  val s1 = String.substring (s, 0, 3)
+  val s2 = String.substring (s, 3, String.size s - 3) 
+ in "<<" ^(if s1 = " + " then s2 else "-"^s2)^">>"
+ end;
+
+(* Conversion from HOL term.                                                 *)
+
+local
+ val neg_tm = @{cterm "uminus :: real => _"}
+ val add_tm = @{cterm "op + :: real => _"}
+ val sub_tm = @{cterm "op - :: real => _"}
+ val mul_tm = @{cterm "op * :: real => _"}
+ val inv_tm = @{cterm "inverse :: real => _"}
+ val div_tm = @{cterm "op / :: real => _"}
+ val pow_tm = @{cterm "op ^ :: real => _"}
+ val zero_tm = @{cterm "0:: real"}
+ val is_numeral = can (HOLogic.dest_number o term_of)
+ fun is_comb t = case t of _$_ => true | _ => false
+ fun poly_of_term tm =
+  if tm aconvc zero_tm then poly_0
+  else if RealArith.is_ratconst tm 
+       then poly_const(RealArith.dest_ratconst tm)
+  else 
+  (let val (lop,r) = Thm.dest_comb tm
+   in if lop aconvc neg_tm then poly_neg(poly_of_term r)
+      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))
+          else error "poly_of_term: inverse of non-constant polyomial"
+       end
+   else (let val (opr,l) = Thm.dest_comb lop
+         in 
+          if opr aconvc pow_tm andalso is_numeral r 
+          then poly_pow (poly_of_term l) ((snd o HOLogic.dest_number o term_of) r)
+          else if opr aconvc add_tm 
+           then poly_add (poly_of_term l) (poly_of_term r)
+          else if opr aconvc sub_tm 
+           then poly_sub (poly_of_term l) (poly_of_term r)
+          else if opr aconvc mul_tm 
+           then poly_mul (poly_of_term l) (poly_of_term r)
+          else if opr aconvc div_tm 
+           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
+                   else error "poly_of_term: division by non-constant polynomial"
+                end
+          else poly_var tm
+ 
+         end
+         handle CTERM ("dest_comb",_) => poly_var tm)
+   end
+   handle CTERM ("dest_comb",_) => poly_var tm)
+in
+val poly_of_term = fn tm =>
+ if type_of (term_of tm) = @{typ real} then poly_of_term tm
+ else error "poly_of_term: term does not have real type"
+end;
+
+(* String of vector (just a list of space-separated numbers).                *)
+
+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 fold1 (fn x => fn y => x ^ " " ^ y) strs ^ "\n"
+ end;
+
+fun increasing f ord (x,y) = ord (f x, f y);
+fun triple_int_ord ((a,b,c),(a',b',c')) = 
+ prod_ord int_ord (prod_ord int_ord int_ord) 
+    ((a,(b,c)),(a',(b',c')));
+structure Inttriplefunc = FuncFun(type key = int*int*int val ord = triple_int_ord);
+
+(* String for block diagonal matrix numbered k.                              *)
+
+fun sdpa_of_blockdiagonal k m =
+ let 
+  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
+ 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 ""
+ end;
+
+(* String for a matrix numbered k, in SDPA sparse format.                    *)
+
+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 
+ in fold_rev (fn ((i,j),c) => fn a =>
+     pfx ^ string_of_int i ^ " " ^ string_of_int j ^
+     " " ^ decimalize 20 c ^ "\n" ^ a) mss ""
+ end;;
+
+(* ------------------------------------------------------------------------- *)
+(* String in SDPA sparse format for standard SDP problem:                    *)
+(*                                                                           *)
+(*    X = v_1 * [M_1] + ... + v_m * [M_m] - [M_0] must be PSD                *)
+(*    Minimize obj_1 * v_1 + ... obj_m * v_m                                 *)
+(* ------------------------------------------------------------------------- *)
+
+fun sdpa_of_problem obj mats =
+ let 
+  val m = length mats - 1
+  val (n,_) = dimensions (hd mats) 
+ in
+  string_of_int m ^ "\n" ^
+  "1\n" ^
+  string_of_int n ^ "\n" ^
+  sdpa_of_vector obj ^
+  fold_rev2 (fn k => fn m => fn a => sdpa_of_matrix (k - 1) m ^ a) (1 upto length mats) mats ""
+ end;
+
+fun index_char str chr pos =
+  if pos >= String.size str then ~1
+  else if String.sub(str,pos) = chr then pos
+  else index_char str chr (pos + 1);
+fun rat_of_quotient (a,b) = if b = 0 then rat_0 else Rat.rat_of_quotient (a,b);
+fun rat_of_string s = 
+ let val n = index_char s #"/" 0 in
+  if n = ~1 then s |> IntInf.fromString |> valOf |> Rat.rat_of_int
+  else 
+   let val SOME numer = IntInf.fromString(String.substring(s,0,n))
+       val SOME den = IntInf.fromString (String.substring(s,n+1,String.size s - n - 1))
+   in rat_of_quotient(numer, den)
+   end
+ end;
+
+fun isspace x = x = " " ;
+fun isnum x = x mem_string ["0","1","2","3","4","5","6","7","8","9"]
+
+(* More parser basics.                                                       *)
+
+local
+ open Scan
+in 
+ val word = 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
+    >> (fn s => rat_of_string(implode s) // pow10 (length s))
+ val decimalsig =
+    decimalint -- option (Scan.$$ "." |-- decimalfrac)
+    >> (fn (h,NONE) => h | (h,SOME x) => h +/ x)
+ fun signed prs =
+       $$ "-" |-- prs >> Rat.neg 
+    || $$ "+" |-- prs
+    || prs;
+
+fun emptyin def xs = if null xs then (def,xs) else Scan.fail xs
+
+ val exponent = ($$ "e" || $$ "E") |-- signed decimalint;
+
+ 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) 
+  in if null rst then x 
+     else error "mkparser: unparsed input"
+  end;;
+
+(* Parse back csdp output.                                                      *)
+
+ fun ignore inp = ((),[])
+ fun csdpoutput inp =  ((decimal -- Scan.bulk (Scan.$$ " " |-- Scan.option decimal) >> (fn (h,to) => map_filter I ((SOME h)::to))) --| ignore >> vector_of_list) inp
+ val parse_csdpoutput = mkparser csdpoutput
+
+(* Run prover on a problem in linear form.                       *)
+
+fun run_problem prover obj mats =
+  parse_csdpoutput (prover (sdpa_of_problem obj mats))
+
+(* Try some apparently sensible scaling first. Note that this is purely to   *)
+(* get a cleaner translation to floating-point, and doesn't affect any of    *)
+(* the results, in principle. In practice it seems a lot better when there   *)
+(* are extreme numbers in the original problem.                              *)
+
+  (* Version for (int*int) keys *)
+local
+  fun max_rat x y = if x </ y then y else x
+  fun common_denominator fld amat acc =
+      fld (fn (m,c) => fn a => lcm_rat (denominator_rat c) a) amat acc
+  fun maximal_element fld amat acc =
+    fld (fn (m,c) => fn maxa => max_rat maxa (abs_rat c)) amat acc 
+fun float_of_rat x = let val (a,b) = Rat.quotient_of_rat x
+                     in Real.fromLargeInt a / Real.fromLargeInt b end;
+in
+
+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 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 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 obj'' = vector_cmul scal2 obj' 
+ in solver obj'' mats''
+  end
+end;
+
+(* Try some apparently sensible scaling first. Note that this is purely to   *)
+(* get a cleaner translation to floating-point, and doesn't affect any of    *)
+(* the results, in principle. In practice it seems a lot better when there   *)
+(* are extreme numbers in the original problem.                              *)
+
+  (* Version for (int*int*int) keys *)
+local
+  fun max_rat x y = if x </ y then y else x
+  fun common_denominator fld amat acc =
+      fld (fn (m,c) => fn a => lcm_rat (denominator_rat c) a) amat acc
+  fun maximal_element fld amat acc =
+    fld (fn (m,c) => fn maxa => max_rat maxa (abs_rat c)) amat acc 
+fun float_of_rat x = let val (a,b) = Rat.quotient_of_rat x
+                     in Real.fromLargeInt a / Real.fromLargeInt b end;
+fun int_of_float x = (trunc x handle Overflow => 0 | Domain => 0)
+in
+
+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 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 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 obj'' = vector_cmul scal2 obj' 
+ in solver obj'' mats''
+  end
+end;
+
+(* Round a vector to "nice" rationals.                                       *)
+
+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 => 
+   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
+
+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;
+
+fun tri_equation_add eq1 eq2 = Inttriplefunc.combine (curry op +/) (fn x => x =/ rat_0) eq1 eq2;
+
+fun tri_equation_eval assig eq =
+ let fun value v = Inttriplefunc.apply assig v 
+ in Inttriplefunc.fold (fn (v, c) => fn a => a +/ value v */ c) eq rat_0
+ end;
+
+(* Eliminate among linear equations: return unconstrained variables and      *)
+(* assignments for the others in terms of them. We give one pseudo-variable  *)
+(* "one" that's used for a constant term.                                    *)
+
+local
+  fun extract_first p l = case l of  (* FIXME : use find_first instead *)
+   [] => error "extract_first"
+ | 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
+        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)
+    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)
+   end)
+  handle Failure _ => eliminate vs dun eqs)
+in
+fun tri_eliminate_equations one vars eqs =
+ let 
+  val assig = eliminate vars Inttriplefunc.undefined 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
+end;
+
+(* Eliminate all variables, in an essentially arbitrary order.               *)
+
+fun tri_eliminate_all_equations one =
+ let 
+  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" 
+         else fst (Inttriplefunc.choose eq')
+      end
+    else v 
+   end
+  fun eliminate dun eqs = case eqs of 
+    [] => dun
+  | eq::oeqs =>
+    if Inttriplefunc.is_undefined 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)
+        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)) 
+                 (map elim oeqs) 
+    end
+in fn eqs =>
+ let 
+  val assig = eliminate Inttriplefunc.undefined 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
+end;
+ 
+(* Solve equations by assigning arbitrary numbers.                           *)
+
+fun tri_solve_equations one eqs =
+ let 
+  val (vars,assigs) = tri_eliminate_all_equations one eqs
+  val vfn = fold_rev (fn v => Inttriplefunc.update(v,rat_0)) vars 
+            (Inttriplefunc.onefunc(one, Rat.rat_of_int ~1))
+  val ass =
+    Inttriplefunc.combine (curry op +/) (K false) 
+    (Inttriplefunc.mapf (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
+ 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 =>
+   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 
+   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;;
+
+val tri_epoly_neg = tri_epoly_cmul (Rat.rat_of_int ~1);
+
+val tri_epoly_add = Inttriplefunc.combine tri_equation_add Inttriplefunc.is_undefined;
+
+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;
+
+fun pi_equation_add eq1 eq2 = Inttriplefunc.combine (curry op +/) (fn x => x =/ rat_0) eq1 eq2;
+
+fun pi_equation_eval assig eq =
+ let fun value v = Inttriplefunc.apply assig v 
+ in Inttriplefunc.fold (fn (v, c) => fn a => a +/ value v */ c) eq rat_0
+ end;
+
+(* Eliminate among linear equations: return unconstrained variables and      *)
+(* assignments for the others in terms of them. We give one pseudo-variable  *)
+(* "one" that's used for a constant term.                                    *)
+
+local
+fun extract_first p l = case l of 
+   [] => error "extract_first"
+ | 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
+        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)
+    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)
+   end
+  handle Failure _ => eliminate vs dun eqs
+in
+fun pi_eliminate_equations one vars eqs =
+ let 
+  val assig = eliminate vars Inttriplefunc.undefined 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
+end;
+
+(* Eliminate all variables, in an essentially arbitrary order.               *)
+
+fun pi_eliminate_all_equations one =
+ let 
+  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" 
+         else fst (Inttriplefunc.choose eq')
+      end
+    else v 
+   end
+  fun eliminate dun eqs = case eqs of 
+    [] => dun
+  | eq::oeqs =>
+    if Inttriplefunc.is_undefined 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)
+        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)) 
+                 (map elim oeqs) 
+    end
+in fn eqs =>
+ let 
+  val assig = eliminate Inttriplefunc.undefined 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
+end;
+ 
+(* Solve equations by assigning arbitrary numbers.                           *)
+
+fun pi_solve_equations one eqs =
+ let 
+  val (vars,assigs) = pi_eliminate_all_equations one eqs
+  val vfn = fold_rev (fn v => Inttriplefunc.update(v,rat_0)) vars 
+            (Inttriplefunc.onefunc(one, Rat.rat_of_int ~1))
+  val ass =
+    Inttriplefunc.combine (curry op +/) (K false) 
+    (Inttriplefunc.mapf (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
+ 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 =>
+   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 
+   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;;
+
+val pi_epoly_neg = pi_epoly_cmul (Rat.rat_of_int ~1);
+
+val pi_epoly_add = Inttriplefunc.combine pi_equation_add Inttriplefunc.is_undefined;
+
+fun pi_epoly_sub p q = pi_epoly_add p (pi_epoly_neg q);;
+
+fun allpairs f l1 l2 =  fold_rev (fn x => (curry (op @)) (map (f x) l2)) l1 [];
+
+(* Hence produce the "relevant" monomials: those whose squares lie in the    *)
+(* Newton polytope of the monomials in the input. (This is enough according  *)
+(* to Reznik: "Extremal PSD forms with few terms", Duke Math. Journal,       *)
+(* vol 45, pp. 363--374, 1978.                                               *)
+(*                                                                           *)
+(* These are ordered in sort of decreasing degree. In particular the         *)
+(* constant monomial is last; this gives an order in diagonalization of the  *)
+(* quadratic form that will tend to display constants.                       *)
+
+(* Diagonalize (Cholesky/LDU) the matrix corresponding to a quadratic form.  *)
+
+local
+fun diagonalize n i m =
+ if Intpairfunc.is_undefined (snd m) then [] 
+ else
+  let val a11 = 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
+          else raise Failure "diagonalize: not PSD ___ "
+    else
+     let 
+      val v = row i m
+      val v' = (fst v, 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
+      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)
+     in (a11,v')::diagonalize n (i + 1) m' 
+     end
+  end
+in
+fun diag m =
+ let 
+   val nn = dimensions m 
+   val n = fst nn 
+ in if snd nn <> n then error "diagonalize: non-square matrix" 
+    else diagonalize n 1 m
+ end
+end;
+
+fun gcd_rat a b = Rat.rat_of_int (Integer.gcd (int_of_rat a) (int_of_rat b));
+
+(* 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 mapa f (d,v) = 
+  (d, Intfunc.fold (fn (i,c) => fn a => upd0 i (f c) a) v Intfunc.undefined)
+ fun adj (c,l) =
+ let val a = 
+  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)) 
+    (snd l) rat_0
+  in ((c // (a */ a)),mapa (fn x => a */ x) l)
+  end
+in
+fun deration d = if null d then (rat_0,d) else
+ let val d' = map adj d
+     val a = fold (lcm_rat o denominator_rat o fst) d' rat_1 //
+          fold (gcd_rat o numerator_rat o fst) d' rat_0 
+ in ((rat_1 // a),map (fn (c,l) => (a */ c,l)) d')
+ end
+end;
+ 
+(* Enumeration of monomials with given multidegree bound.                    *)
+
+fun enumerate_monomials d vars = 
+ if d < 0 then []
+ else if d = 0 then [Ctermfunc.undefined]
+ 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 fold1 (curry op @) 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)] 
+else if d < 0 then [] else
+case pols of 
+   [] => [(poly_const rat_1,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)))
+         (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;
+
+(* String for block diagonal matrix numbered k.                              *)
+
+fun sdpa_of_blockdiagonal k m =
+ let 
+  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 
+ 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 ""
+ end;
+
+(* SDPA for problem using block diagonal (i.e. multiple SDPs)                *)
+
+fun sdpa_of_blockproblem nblocks blocksizes obj mats =
+ let val m = length mats - 1 
+ in
+  string_of_int m ^ "\n" ^
+  string_of_int nblocks ^ "\n" ^
+  (fold1 (fn s => fn t => s^" "^t) (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)
+    (1 upto length mats) mats ""
+ end;
+
+(* Run prover on a problem in block diagonal form.                       *)
+
+fun run_blockproblem prover nblocks blocksizes obj mats=
+  parse_csdpoutput (prover (sdpa_of_blockproblem nblocks blocksizes obj mats))
+
+(* 3D versions of matrix operations to consider blocks separately.           *)
+
+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;
+
+val bmatrix_neg = bmatrix_cmul (Rat.rat_of_int ~1);
+fun bmatrix_sub m1 m2 = bmatrix_add m1 (bmatrix_neg m2);;
+
+(* Smash a block matrix into components.                                     *)
+
+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 
+      in (((bs,bs),m):matrix) end)
+ (blocksizes ~~ (1 upto length blocksizes));;
+
+(* FIXME : Get rid of this !!!*)
+local
+  fun tryfind_with msg f [] = raise Failure msg
+    | tryfind_with msg f (x::xs) = (f x handle Failure s => tryfind_with s f xs);
+in 
+  fun tryfind f = tryfind_with "tryfind" f
+end
+
+(*
+fun tryfind f [] = error "tryfind"
+  | tryfind f (x::xs) = (f x handle ERROR _ => tryfind f xs);
+*)
+
+(* 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)::
+      (filter (fn (p,c) => multidegree p <= d) leqs)
+    else enumerate_products d leqs
+ val nblocks = length monoid
+ fun mk_idmultiplier k p =
+  let 
+   val e = d - multidegree p
+   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)
+  end
+
+ fun mk_sqmultiplier k (p,c) =
+  let 
+   val e = (d - multidegree p) div 2
+   val mons = enumerate_monomials e vars
+   val nons = mons ~~ (1 upto length mons) 
+  in (mons, 
+      fold_rev (fn (m1,n1) =>
+       fold_rev (fn (m2,n2) => fn  a =>
+        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
+          end
+        end)  nons)
+       nons Monomialfunc.undefined)
+  end
+
+  val (sqmonlist,sqs) = split_list (map2 mk_sqmultiplier (1 upto length monoid) monoid)
+  val (idmonlist,ids) =  split_list(map2 mk_idmultiplier (1 upto length eqs) eqs)
+  val blocksizes = map length sqmonlist
+  val bigsum =
+    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 (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
+  fun mk_matrix v =
+    Inttriplefunc.fold (fn ((b,i,j), ass) => fn m => 
+        if b < 0 then m else
+         let val c = Inttriplefunc.tryapplyd ass v rat_0
+         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
+  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
+
+  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)
+  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 find_rounding d =
+   let 
+    val _ = if !debugging 
+          then writeln ("Trying rounding with limit "^Rat.string_of_rat d ^ "\n") 
+          else ()
+    val vec = nice_vector d raw_vec
+    val blockmat = iter (1,dim vec)
+     (fn i => fn a => bmatrix_add (bmatrix_cmul (int_element vec i) (nth mats i)) a)
+     (bmatrix_neg (nth mats 0))
+    val allmats = blocks blocksizes blockmat 
+   in (vec,map diag allmats)
+   end
+  val (vec,ratdias) =
+    if null pvs then find_rounding rat_1
+    else tryfind find_rounding (map Rat.rat_of_int (1 upto 31) @
+                                map pow2 (5 upto 66))
+  val newassigs =
+    fold_rev (fn k => Inttriplefunc.update (nth pvs (k - 1), int_element vec k))
+           (1 upto dim vec) (Inttriplefunc.onefunc ((0,0,0), Rat.rat_of_int ~1))
+  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
+  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)
+   in map mk_sq
+   end
+  val sqs = map2 mk_sos sqmonlist ratdias
+  val cfs = map poly_of_epoly ids
+  val msq = filter (fn (a,b) => not (null b)) (map2 pair monoid sqs)
+  fun eval_sq sqs = fold_rev (fn (c,q) => poly_add (poly_cmul c (poly_mul q q))) sqs poly_0
+  val sanity =
+    fold_rev (fn ((p,c),s) => poly_add (poly_mul p (eval_sq s))) msq
+           (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
+  (cfs,map (fn (a,b) => (snd a,b)) msq)
+ end
+
+
+end;
+
+(* Iterative deepening.                                                      *)
+
+fun deepen f n = 
+  (writeln ("Searching with depth limit " ^ string_of_int n) ; (f n handle Failure s => (writeln ("failed with message: " ^ s) ; deepen f (n+1))))
+
+(* The ordering so we can create canonical HOL polynomials.                  *)
+
+fun dest_monomial mon = sort (increasing fst cterm_ord) (Ctermfunc.graph mon);
+
+fun monomial_order (m1,m2) =
+ if Ctermfunc.is_undefined m2 then LESS 
+ else if Ctermfunc.is_undefined m1 then GREATER 
+ else
+  let val mon1 = dest_monomial m1 
+      val mon2 = dest_monomial m2
+      val deg1 = fold (curry op + o snd) mon1 0
+      val deg2 = fold (curry op + o snd) mon2 0 
+  in if deg1 < deg2 then GREATER else if deg1 > deg2 then LESS
+     else list_ord (prod_ord cterm_ord int_ord) (mon1,mon2)
+  end;
+
+fun dest_poly p =
+  map (fn (m,c) => (c,dest_monomial m))
+      (sort (prod_ord monomial_order (K EQUAL)) (Monomialfunc.graph p));
+
+(* Map back polynomials and their composites to HOL.                         *)
+
+local
+ open Thm Numeral RealArith
+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_monomial m = 
+ if Ctermfunc.is_undefined m then @{cterm "1::real"} 
+ else 
+  let 
+   val m' = dest_monomial m
+   val vps = fold_rev (fn (x,k) => cons (cterm_of_varpow x k)) m' [] 
+  in fold1 (fn s => fn t => capply (capply @{cterm "op * :: real => _"} s) t) vps
+  end
+
+fun cterm_of_cmonomial (m,c) = if Ctermfunc.is_undefined 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);
+
+fun cterm_of_poly p = 
+ if Monomialfunc.is_undefined p then @{cterm "0::real"} 
+ else
+  let 
+   val cms = map cterm_of_cmonomial
+     (sort (prod_ord monomial_order (K EQUAL)) (Monomialfunc.graph p))
+  in fold1 (fn t1 => fn t2 => capply(capply @{cterm "op + :: real => _"} t1) t2) cms
+  end;
+
+fun cterm_of_sqterm (c,p) = Product(Rational_lt c,Square(cterm_of_poly p));
+
+fun cterm_of_sos (pr,sqs) = if null sqs then pr
+  else Product(pr,fold1 (fn a => fn b => Sum(a,b)) (map cterm_of_sqterm sqs));
+
+end
+
+(* Interface to HOL.                                                         *)
+local
+  open Thm Conv RealArith
+  val concl = 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 !! *)
+fun real_nonlinear_prover prover ctxt =
+ let 
+  val {add,mul,neg,pow,sub,main} =  Normalizer.semiring_normalizers_ord_wrapper ctxt
+      (valOf (NormalizerData.match ctxt @{cterm "(0::real) + 1"})) 
+     simple_cterm_ord
+  val (real_poly_add_conv,real_poly_mul_conv,real_poly_neg_conv,
+       real_poly_pow_conv,real_poly_sub_conv,real_poly_conv) = (add,mul,neg,pow,sub,main)
+  fun mainf  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 (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 
+                     else raise Failure "trivial_axiom: Not a trivial axiom"
+     | Axiom_le n => if eval Ctermfunc.undefined 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 
+                     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 end)
+   handle Failure _ => (
+    let 
+     val pol = fold_rev poly_mul (map fst ltp) (poly_const Rat.one)
+     val leq = lep @ ltp
+     fun tryall d =
+      let val e = multidegree pol
+          val k = if e = 0 then 0 else d div e
+          val eq' = map fst eq 
+      in tryfind (fn i => (d,i,real_positivnullstellensatz_general prover false d eq' leq
+                            (poly_neg(poly_pow pol i))))
+              (0 upto k)
+      end
+    val (d,i,(cert_ideal,cert_cone)) = deepen tryall 0
+    val proofs_ideal =
+      map2 (fn q => fn (p,ax) => Eqmul(cterm_of_poly 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 = fold1 (fn s => fn t => Product(s,t)) (map snd ltp) 
+      in  funpow i (fn q => Product(p,q)) (Rational_lt Rat.one)
+      end
+    val proof = fold1 (fn s => fn t => Sum(s,t))
+                           (proof_ne :: proofs_ideal @ proofs_cone) 
+    in writeln "Translating proof certificate to HOL";
+       translator (eqs,les,lts) proof
+    end))
+   end
+ in mainf end
+end
+
+fun C f x y = f y x;
+  (* FIXME : This is very bad!!!*)
+fun subst_conv eqs t = 
+ let 
+  val t' = fold (Thm.cabs o Thm.lhs_of) eqs t
+ in Conv.fconv_rule (Thm.beta_conversion true) (fold (C combination) eqs (reflexive t'))
+ end
+
+(* A wrapper that tries to substitute away variables first.                  *)
+
+local
+ open Thm Conv RealArith
+  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 shuffle1 = 
+   fconv_rule (rewr_conv @{lemma "(a + x == y) == (x == y - (a::real))" by (atomize (full)) (simp add: ring_simps) })
+ val shuffle2 =
+    fconv_rule (rewr_conv @{lemma "(x + a == y) ==  (x == y - (a::real))" by (atomize (full)) (simp add: ring_simps)})
+ fun substitutable_monomial fvs tm = case term_of tm of
+    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"
+  | @{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))
+  | _ => raise Failure "substitutable_monomial"
+
+  fun isolate_variable v th = 
+   let val w = 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 
+              else isolate_variable v (shuffle1 th)
+          | _ => error "isolate variable : This should not happen?"
+   end 
+in
+
+fun real_nonlinear_subst_prover prover ctxt =
+ let 
+  val {add,mul,neg,pow,sub,main} =  Normalizer.semiring_normalizers_ord_wrapper ctxt
+      (valOf (NormalizerData.match ctxt @{cterm "(0::real) + 1"})) 
+     simple_cterm_ord
+
+  val (real_poly_add_conv,real_poly_mul_conv,real_poly_neg_conv,
+       real_poly_pow_conv,real_poly_sub_conv,real_poly_conv) = (add,mul,neg,pow,sub,main)
+
+  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 th2 = fconv_rule (binop_conv real_poly_mul_conv) th1
+   in fconv_rule (arg_conv real_poly_conv) (isolate_variable v th2)
+   end
+   fun oprconv cv ct = 
+    let val g = Thm.dest_fun2 ct
+    in if g aconvc @{cterm "op <= :: real => _"} 
+         orelse g aconvc @{cterm "op < :: real => _"} 
+       then arg_conv cv ct else arg1_conv cv ct
+    end
+  fun mainf translator =
+   let 
+    fun substfirst(eqs,les,lts) =
+      ((let 
+           val eth = tryfind make_substitution eqs
+           val modify = fconv_rule (arg_conv (oprconv(subst_conv [eth] then_conv real_poly_conv)))
+       in  substfirst
+             (filter_out (fn t => (Thm.dest_arg1 o Thm.dest_arg o cprop_of) t 
+                                   aconvc @{cterm "0::real"}) (map modify eqs),
+                                   map modify les,map modify lts)
+       end)
+       handle Failure  _ => real_nonlinear_prover prover ctxt translator (rev eqs, rev les, rev lts))
+    in substfirst
+   end
+
+
+ in mainf
+ end
+
+(* Overall function. *)
+
+fun real_sos prover ctxt t = gen_prover_real_arith ctxt (real_nonlinear_subst_prover prover ctxt) t;
+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)
+
+fun core_sos_conv prover ctxt t = Drule.arg_cong_rule @{cterm Trueprop} (real_sos prover ctxt (Thm.dest_arg t) RS @{thm Eq_TrueI})
+
+val known_sos_constants = 
+  [@{term "op ==>"}, @{term "Trueprop"}, 
+   @{term "op -->"}, @{term "op &"}, @{term "op |"}, 
+   @{term "Not"}, @{term "op = :: bool => _"}, 
+   @{term "All :: (real => _) => _"}, @{term "Ex :: (real => _) => _"}, 
+   @{term "op = :: real => _"}, @{term "op < :: real => _"}, 
+   @{term "op <= :: real => _"}, 
+   @{term "op + :: real => _"}, @{term "op - :: real => _"}, 
+   @{term "op * :: real => _"}, @{term "uminus :: real => _"}, 
+   @{term "op / :: real => _"}, @{term "inverse :: real => _"},
+   @{term "op ^ :: real => _"}, @{term "abs :: real => _"}, 
+   @{term "min :: real => _"}, @{term "max :: real => _"},
+   @{term "0::real"}, @{term "1::real"}, @{term "number_of :: int => real"},
+   @{term "number_of :: int => nat"},
+   @{term "Int.Bit0"}, @{term "Int.Bit1"}, 
+   @{term "Int.Pls"}, @{term "Int.Min"}];
+
+fun check_sos kcts ct = 
+ let
+  val t = term_of ct
+  val _ = if not (null (Term.add_tfrees t []) 
+                  andalso null (Term.add_tvars t [])) 
+          then error "SOS: not sos. Additional type varables" else ()
+  val fs = Term.add_frees t []
+  val _ = if exists (fn ((_,T)) => not (T = @{typ "real"})) fs 
+          then error "SOS: not sos. Variables with type not real" else ()
+  val vs = Term.add_vars t []
+  val _ = if exists (fn ((_,T)) => not (T = @{typ "real"})) fs 
+          then error "SOS: not sos. Variables with type not real" else ()
+  val ukcs = subtract (fn (t,p) => Const p aconv t) kcts (Term.add_consts t [])
+  val _ = if  null ukcs then () 
+              else error ("SOSO: Unknown constants in Subgoal:" ^ commas (map fst ukcs))
+in () end
+
+fun core_sos_tac prover ctxt = CSUBGOAL (fn (ct, i) => 
+  let val _ = check_sos known_sos_constants ct
+      val (avs, p) = strip_all ct
+      val th = standard (fold_rev forall_intr avs (real_sos prover ctxt (Thm.dest_arg p)))
+  in rtac th i end);
+
+fun default_SOME f NONE v = SOME v
+  | default_SOME f (SOME v) _ = SOME v;
+
+fun lift_SOME f NONE a = f a
+  | lift_SOME f (SOME a) _ = SOME a;
+
+
+local
+ val is_numeral = can (HOLogic.dest_number o term_of)
+in
+fun get_denom b ct = case term_of ct of
+  @{term "op / :: real => _"} $ _ $ _ => 
+     if is_numeral (Thm.dest_arg ct) then get_denom b (Thm.dest_arg1 ct)
+     else default_SOME (get_denom b) (get_denom b (Thm.dest_arg ct))   (Thm.dest_arg ct, b)
+ | @{term "op < :: real => _"} $ _ $ _ => lift_SOME (get_denom true) (get_denom true (Thm.dest_arg ct)) (Thm.dest_arg1 ct)
+ | @{term "op <= :: real => _"} $ _ $ _ => lift_SOME (get_denom true) (get_denom true (Thm.dest_arg ct)) (Thm.dest_arg1 ct)
+ | _ $ _ => lift_SOME (get_denom b) (get_denom b (Thm.dest_fun ct)) (Thm.dest_arg ct)
+ | _ => NONE
+end;
+
+fun elim_one_denom_tac ctxt = 
+CSUBGOAL (fn (P,i) => 
+ case get_denom false P of 
+   NONE => no_tac
+ | SOME (d,ord) => 
+     let 
+      val ss = simpset_of ctxt addsimps @{thms field_simps} 
+               addsimps [@{thm nonzero_power_divide}, @{thm power_divide}]
+      val th = instantiate' [] [SOME d, SOME (Thm.dest_arg P)] 
+         (if ord then @{lemma "(d=0 --> P) & (d>0 --> P) & (d<(0::real) --> P) ==> P" by auto}
+          else @{lemma "(d=0 --> P) & (d ~= (0::real) --> P) ==> P" by blast})
+     in (rtac th i THEN Simplifier.asm_full_simp_tac ss i) end);
+
+fun elim_denom_tac ctxt i = REPEAT (elim_one_denom_tac ctxt i);
+
+fun sos_tac prover ctxt = ObjectLogic.full_atomize_tac THEN' elim_denom_tac ctxt THEN' core_sos_tac prover ctxt
+
+
+end;
--- a/src/HOL/Library/sos_wrapper.ML	Mon Aug 10 08:37:37 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,158 +0,0 @@
-(* Title:      sos_wrapper.ML
-   Author:     Philipp Meyer, TU Muenchen
-
-Added functionality for sums of squares, e.g. calling a remote prover
-*)
-
-signature SOS_WRAPPER =
-sig
-
-  datatype prover_result = Success | PartialSuccess | Failure | Error
-  type prover = string * (int -> string -> prover_result * string)
-
-  val setup: theory -> theory
-end
-
-structure SosWrapper : SOS_WRAPPER=
-struct
-
-datatype prover_result = Success | PartialSuccess | Failure | Error
-type prover =
-  string *                           (* command name *)
-  (int -> string ->prover_result * string)   (* function to find failure from return value and output *)
-
-
-(*** output control ***)
-
-fun debug s = Output.debug (fn () => s)
-val answer = Output.priority
-val write = Output.writeln
-
-(*** calling provers ***)
-
-val destdir = ref ""
-
-fun filename dir name =
-  let
-    val probfile = Path.basic (name ^ serial_string ())
-  in
-    if dir = "" then
-      File.tmp_path probfile
-    else
-      if File.exists (Path.explode dir) then
-        Path.append  (Path.explode dir) probfile
-      else
-        error ("No such directory: " ^ dir)
-  end
-
-fun is_success Success = true
-  | is_success PartialSuccess = true
-  | is_success _ = false
-fun str_of_status Success = "Success"
-  | str_of_status PartialSuccess = "Partial Success"
-  | str_of_status Failure= "Failure"
-  | str_of_status Error= "Error"
-
-fun run_solver name (cmd, find_failure) input =
-  let
-    val _ = answer ("Calling solver: " ^ name)
-
-    (* create input file *)
-    val dir = ! destdir
-    val input_file = filename dir "sos_in" 
-    val _ = File.write input_file input
-
-    val _ = debug "Solver input:"
-    val _ = debug input
-
-    (* call solver *)
-    val output_file = filename dir "sos_out"
-    val (output, rv) = system_out (cmd ^ " " ^ (Path.implode input_file) ^
-                         " " ^ (Path.implode output_file))
- 
-    (* read and analysize output *)
-    val (res, res_msg) = find_failure rv output
-    val result = if is_success res then File.read output_file else ""
-
-    (* remove temporary files *)
-    val _ = if dir = "" then (File.rm input_file ; if File.exists output_file then File.rm output_file else ()) else ()
-
-    val _ = debug "Solver output:"
-    val _ = debug output
-    val _ = debug "Solver result:"
-    val _ = debug result
-
-    val _ = answer (str_of_status res ^ ": " ^ res_msg)
-
-  in
-    if is_success res then
-      result
-    else
-      error ("Prover failed: " ^ res_msg)
-  end
-
-(*** various provers ***)
-
-(* local csdp client *)
-
-fun find_csdp_run_failure rv _ =
-  case rv of
-    0 => (Success, "SDP solved")
-  | 1 => (Failure, "SDP is primal infeasible")
-  | 2 => (Failure, "SDP is dual infeasible")
-  | 3 => (PartialSuccess, "SDP solved with reduced accuracy")
-  | _ => (Failure, "return code is " ^ string_of_int rv)
-
-val csdp = ("csdp", find_csdp_run_failure)
-
-(* remote neos server *)
-
-fun find_neos_failure rv output =
-  if rv = 2 then (Failure, "no solution") else
-  if rv <> 0 then (Error, "return code is " ^ string_of_int rv) else
-  let
-    fun find_success str =
-      if String.isPrefix "Success: " str then
-        SOME (Success, unprefix "Success: " str)
-      else if String.isPrefix "Partial Success: " str then
-        SOME (PartialSuccess, unprefix "Partial Success: " str)
-      else if String.isPrefix "Failure: " str then
-        SOME (Failure, unprefix "Failure: " str)
-      else
-        NONE 
-    val exit_line = get_first find_success (split_lines output)
-  in
-    case exit_line of
-      SOME (status, msg) =>
-        if String.isPrefix "SDP solved" msg then
-          (status, msg)
-        else (Failure, msg)
-    | NONE => (Failure, "no success")
-  end
-
-val neos_csdp = ("$ISABELLE_HOME/lib/scripts/neos/NeosCSDPClient.py", find_neos_failure)
-
-(* save provers in table *)
-
-val provers =
-     Symtab.make [("remote_csdp", neos_csdp),("csdp", csdp)]
-
-fun get_prover name =
-  case Symtab.lookup provers name of
-    SOME prover => prover
-  | NONE => error ("unknown prover: " ^ name)
-
-fun call_solver name =
-    run_solver name (get_prover name)
-
-(* setup tactic *)
-
-val def_solver = "remote_csdp"
-
-fun sos_solver name = (SIMPLE_METHOD' o (Sos.sos_tac (call_solver name))) 
-
-val sos_method = Scan.optional (Scan.lift OuterParse.xname) def_solver >> sos_solver
-
-val setup = Method.setup @{binding sos} sos_method "Prove universal problems over the reals using sums of squares"
-
-end
--- a/src/HOL/Library/sum_of_squares.ML	Mon Aug 10 08:37:37 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1754 +0,0 @@
-(* Title:      sum_of_squares.ML
-   Authors:    Amine Chaieb, University of Cambridge
-               Philipp Meyer, TU Muenchen
-
-A tactic for proving nonlinear inequalities
-*)
-
-signature SOS =
-sig
-
-  val sos_tac : (string -> string) -> Proof.context -> int -> Tactical.tactic
-
-end
-
-structure Sos : SOS = 
-struct
-
-
-val rat_0 = Rat.zero;
-val rat_1 = Rat.one;
-val rat_2 = Rat.two;
-val rat_10 = Rat.rat_of_int 10;
-val rat_1_2 = rat_1 // rat_2;
-val max = curry IntInf.max;
-val min = curry IntInf.min;
-
-val denominator_rat = Rat.quotient_of_rat #> snd #> Rat.rat_of_int;
-val numerator_rat = Rat.quotient_of_rat #> fst #> Rat.rat_of_int;
-fun int_of_rat a = 
-    case Rat.quotient_of_rat a of (i,1) => i | _ => error "int_of_rat: not an int";
-fun lcm_rat x y = Rat.rat_of_int (Integer.lcm (int_of_rat x) (int_of_rat y));
-
-fun rat_pow r i = 
- let fun pow r i = 
-   if i = 0 then rat_1 else 
-   let val d = pow r (i div 2)
-   in d */ d */ (if i mod 2 = 0 then rat_1 else r)
-   end
- in if i < 0 then pow (Rat.inv r) (~ i) else pow r i end;
-
-fun round_rat r = 
- let val (a,b) = Rat.quotient_of_rat (Rat.abs r)
-     val d = a div b
-     val s = if r </ rat_0 then (Rat.neg o Rat.rat_of_int) else Rat.rat_of_int
-     val x2 = 2 * (a - (b * d))
- in s (if x2 >= b then d + 1 else d) end
-
-val abs_rat = Rat.abs;
-val pow2 = rat_pow rat_2;
-val pow10 = rat_pow rat_10;
-
-val debugging = ref false;
-
-exception Sanity;
-
-exception Unsolvable;
-
-(* Turn a rational into a decimal string with d sig digits.                  *)
-
-local
-fun normalize y =
-  if abs_rat y </ (rat_1 // rat_10) then normalize (rat_10 */ y) - 1
-  else if abs_rat y >=/ rat_1 then normalize (y // rat_10) + 1
-  else 0 
- in
-fun decimalize d x =
-  if x =/ rat_0 then "0.0" else
-  let 
-   val y = Rat.abs x
-   val e = normalize y
-   val z = pow10(~ e) */ y +/ rat_1
-   val k = int_of_rat (round_rat(pow10 d */ z))
-  in (if x </ rat_0 then "-0." else "0.") ^
-     implode(tl(explode(string_of_int k))) ^
-     (if e = 0 then "" else "e"^string_of_int e)
-  end
-end;
-
-(* Iterations over numbers, and lists indexed by numbers.                    *)
-
-fun itern k l f a =
-  case l of
-    [] => a
-  | h::t => itern (k + 1) t f (f h k a);
-
-fun iter (m,n) f a =
-  if n < m then a
-  else iter (m+1,n) f (f m a);
-
-(* The main types.                                                           *)
-
-fun strict_ord ord (x,y) = case ord (x,y) of LESS => LESS | _ => GREATER
-
-structure Intpairfunc = FuncFun(type key = int*int val ord = prod_ord int_ord int_ord);
-
-type vector = int* Rat.rat Intfunc.T;
-
-type matrix = (int*int)*(Rat.rat Intpairfunc.T);
-
-type monomial = int Ctermfunc.T;
-
-val cterm_ord = (fn (s,t) => TermOrd.fast_term_ord(term_of s, term_of t))
- fun monomial_ord (m1,m2) = list_ord (prod_ord cterm_ord int_ord) (Ctermfunc.graph m1, Ctermfunc.graph m2)
-structure Monomialfunc = FuncFun(type key = monomial val ord = monomial_ord)
-
-type poly = Rat.rat Monomialfunc.T;
-
- fun iszero (k,r) = r =/ rat_0;
-
-fun fold_rev2 f l1 l2 b =
-  case (l1,l2) of
-    ([],[]) => b
-  | (h1::t1,h2::t2) => f h1 h2 (fold_rev2 f t1 t2 b)
-  | _ => error "fold_rev2";
- 
-(* Vectors. Conventionally indexed 1..n.                                     *)
-
-fun vector_0 n = (n,Intfunc.undefined):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;
-
-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))
- end;
-
-fun vector_neg (v:vector) = (fst v,Intfunc.mapf 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 
- end;
-
-fun vector_sub v1 v2 = vector_add v1 (vector_neg v2);
-
-fun vector_dot (v1:vector) (v2:vector) =
- 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
- 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
- end;
-
-(* Matrices; again rows and columns indexed from 1.                          *)
-
-fun matrix_0 (m,n) = ((m,n),Intpairfunc.undefined):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;;
-
-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))
- end;
-
-fun matrix_neg (m:matrix) = 
-  (dimensions m, Intpairfunc.mapf 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
- end;;
-
-fun matrix_sub m1 m2 = matrix_add m1 (matrix_neg m2);
-
-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
- 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)
-   : 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
- 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
- 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)
-   end
- end;
-
-(* Monomials.                                                                *)
-
-fun monomial_eval assig (m:monomial) =
-  Ctermfunc.fold (fn (x, k) => fn a => a */ rat_pow (Ctermfunc.apply assig x) k)
-        m rat_1;
-val monomial_1 = (Ctermfunc.undefined:monomial);
-
-fun monomial_var x = Ctermfunc.onefunc (x, 1) :monomial;
-
-val (monomial_mul:monomial->monomial->monomial) =
-  Ctermfunc.combine (curry op +) (K false);
-
-fun monomial_pow (m:monomial) k =
-  if k = 0 then monomial_1
-  else Ctermfunc.mapf (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_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
-  else error "monomial_div: non-divisible"
- end;
-
-fun monomial_degree x (m:monomial) = 
-  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_multidegree (m:monomial) = 
- Ctermfunc.fold (fn (x, k) => fn a => k + a) m 0;;
-
-fun monomial_variables m = 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;
-
-val poly_0 = (Monomialfunc.undefined:poly);
-
-fun poly_isconst (p:poly) = 
-  Monomialfunc.fold (fn (m, c) => fn a => Ctermfunc.is_undefined m andalso a) p true;
-
-fun poly_var x = Monomialfunc.onefunc (monomial_var x,rat_1) :poly;
-
-fun poly_const c =
-  if c =/ rat_0 then poly_0 else Monomialfunc.onefunc(monomial_1, c);
-
-fun poly_cmul c (p:poly) =
-  if c =/ rat_0 then poly_0
-  else Monomialfunc.mapf (fn x => c */ x) p;
-
-fun poly_neg (p:poly) = (Monomialfunc.mapf Rat.neg p :poly);;
-
-fun poly_add (p1:poly) (p2:poly) =
-  (Monomialfunc.combine (curry op +/) (fn x => x =/ rat_0) p1 p2 :poly);
-
-fun poly_sub p1 p2 = poly_add p1 (poly_neg p2);
-
-fun poly_cmmul (c,m) (p:poly) =
- 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;
-
-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_div (p1:poly) (p2:poly) =
- if not(poly_isconst p2) 
- then error "poly_div: non-constant" else
- let val c = eval Ctermfunc.undefined p2 
- in if c =/ rat_0 then error "poly_div: division by zero"
-    else poly_cmul (Rat.inv c) p1
- end;
-
-fun poly_square p = poly_mul p p;
-
-fun poly_pow p k =
- if k = 0 then poly_const rat_1
- else if k = 1 then p
- else let val q = poly_square(poly_pow p (k div 2)) in
-      if k mod 2 = 1 then poly_mul p q else q end;
-
-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));
-
-fun degree x (p:poly) = 
- 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 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 []);;
-
-(* 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);
-
-local
- fun ord (l1,l2) = case (l1,l2) of
-  (_,[]) => LESS 
- | ([],_) => GREATER
- | (h1::t1,h2::t2) => 
-   (case humanorder_varpow (h1, h2) of 
-     LESS => LESS
-   | 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))
-end;
-
-fun fold1 f l =  case l of
-   []     => error "fold1"
- | [x]    => x
- | (h::t) => f h (fold1 f t);
-
-(* Conversions to strings.                                                   *)
-
-fun string_of_vector min_size max_size (v:vector) =
- let val n_raw = dim v 
- 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 "[" ^ fold1 (fn s => fn t => s ^ ", " ^ t) xs ^
-  (if n_raw > max_size then ", ...]" else "]")
-  end
- end;
-
-fun string_of_matrix max_size (m:matrix) =
- let 
-  val (i_raw,j_raw) = dimensions m
-  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 "["^ fold1 (fn s => fn t => s^";\n "^t) rstr ^
-  (if j > max_size then "\n ...]" else "]")
- end;
-
-fun string_of_term t = 
- case t of
-   a$b => "("^(string_of_term a)^" "^(string_of_term b)^")"
- | Abs x => 
-    let val (xn, b) = Term.dest_abs x
-    in "(\\"^xn^"."^(string_of_term b)^")"
-    end
- | Const(s,_) => s
- | Free (s,_) => s
- | Var((s,_),_) => s
- | _ => error "string_of_term";
-
-val string_of_cterm = string_of_term o term_of;
-
-fun string_of_varpow x k =
-  if k = 1 then string_of_cterm x 
-  else string_of_cterm x^"^"^string_of_int k;
-
-fun string_of_monomial m =
- if Ctermfunc.is_undefined 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 fold1 (fn s => fn t => s^"*"^t) vps
- end;
-
-fun string_of_cmonomial (c,m) =
- if Ctermfunc.is_undefined 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:poly) =
- if Monomialfunc.is_undefined p then "<<0>>" else
- let 
-  val cms = sort (fn ((m1,_),(m2,_)) => humanorder_monomial m1  m2) (Monomialfunc.graph 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))
-          cms ""
-  val s1 = String.substring (s, 0, 3)
-  val s2 = String.substring (s, 3, String.size s - 3) 
- in "<<" ^(if s1 = " + " then s2 else "-"^s2)^">>"
- end;
-
-(* Conversion from HOL term.                                                 *)
-
-local
- val neg_tm = @{cterm "uminus :: real => _"}
- val add_tm = @{cterm "op + :: real => _"}
- val sub_tm = @{cterm "op - :: real => _"}
- val mul_tm = @{cterm "op * :: real => _"}
- val inv_tm = @{cterm "inverse :: real => _"}
- val div_tm = @{cterm "op / :: real => _"}
- val pow_tm = @{cterm "op ^ :: real => _"}
- val zero_tm = @{cterm "0:: real"}
- val is_numeral = can (HOLogic.dest_number o term_of)
- fun is_comb t = case t of _$_ => true | _ => false
- fun poly_of_term tm =
-  if tm aconvc zero_tm then poly_0
-  else if RealArith.is_ratconst tm 
-       then poly_const(RealArith.dest_ratconst tm)
-  else 
-  (let val (lop,r) = Thm.dest_comb tm
-   in if lop aconvc neg_tm then poly_neg(poly_of_term r)
-      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))
-          else error "poly_of_term: inverse of non-constant polyomial"
-       end
-   else (let val (opr,l) = Thm.dest_comb lop
-         in 
-          if opr aconvc pow_tm andalso is_numeral r 
-          then poly_pow (poly_of_term l) ((snd o HOLogic.dest_number o term_of) r)
-          else if opr aconvc add_tm 
-           then poly_add (poly_of_term l) (poly_of_term r)
-          else if opr aconvc sub_tm 
-           then poly_sub (poly_of_term l) (poly_of_term r)
-          else if opr aconvc mul_tm 
-           then poly_mul (poly_of_term l) (poly_of_term r)
-          else if opr aconvc div_tm 
-           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
-                   else error "poly_of_term: division by non-constant polynomial"
-                end
-          else poly_var tm
- 
-         end
-         handle CTERM ("dest_comb",_) => poly_var tm)
-   end
-   handle CTERM ("dest_comb",_) => poly_var tm)
-in
-val poly_of_term = fn tm =>
- if type_of (term_of tm) = @{typ real} then poly_of_term tm
- else error "poly_of_term: term does not have real type"
-end;
-
-(* String of vector (just a list of space-separated numbers).                *)
-
-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 fold1 (fn x => fn y => x ^ " " ^ y) strs ^ "\n"
- end;
-
-fun increasing f ord (x,y) = ord (f x, f y);
-fun triple_int_ord ((a,b,c),(a',b',c')) = 
- prod_ord int_ord (prod_ord int_ord int_ord) 
-    ((a,(b,c)),(a',(b',c')));
-structure Inttriplefunc = FuncFun(type key = int*int*int val ord = triple_int_ord);
-
-(* String for block diagonal matrix numbered k.                              *)
-
-fun sdpa_of_blockdiagonal k m =
- let 
-  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
- 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 ""
- end;
-
-(* String for a matrix numbered k, in SDPA sparse format.                    *)
-
-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 
- in fold_rev (fn ((i,j),c) => fn a =>
-     pfx ^ string_of_int i ^ " " ^ string_of_int j ^
-     " " ^ decimalize 20 c ^ "\n" ^ a) mss ""
- end;;
-
-(* ------------------------------------------------------------------------- *)
-(* String in SDPA sparse format for standard SDP problem:                    *)
-(*                                                                           *)
-(*    X = v_1 * [M_1] + ... + v_m * [M_m] - [M_0] must be PSD                *)
-(*    Minimize obj_1 * v_1 + ... obj_m * v_m                                 *)
-(* ------------------------------------------------------------------------- *)
-
-fun sdpa_of_problem obj mats =
- let 
-  val m = length mats - 1
-  val (n,_) = dimensions (hd mats) 
- in
-  string_of_int m ^ "\n" ^
-  "1\n" ^
-  string_of_int n ^ "\n" ^
-  sdpa_of_vector obj ^
-  fold_rev2 (fn k => fn m => fn a => sdpa_of_matrix (k - 1) m ^ a) (1 upto length mats) mats ""
- end;
-
-fun index_char str chr pos =
-  if pos >= String.size str then ~1
-  else if String.sub(str,pos) = chr then pos
-  else index_char str chr (pos + 1);
-fun rat_of_quotient (a,b) = if b = 0 then rat_0 else Rat.rat_of_quotient (a,b);
-fun rat_of_string s = 
- let val n = index_char s #"/" 0 in
-  if n = ~1 then s |> IntInf.fromString |> valOf |> Rat.rat_of_int
-  else 
-   let val SOME numer = IntInf.fromString(String.substring(s,0,n))
-       val SOME den = IntInf.fromString (String.substring(s,n+1,String.size s - n - 1))
-   in rat_of_quotient(numer, den)
-   end
- end;
-
-fun isspace x = x = " " ;
-fun isnum x = x mem_string ["0","1","2","3","4","5","6","7","8","9"]
-
-(* More parser basics.                                                       *)
-
-local
- open Scan
-in 
- val word = 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
-    >> (fn s => rat_of_string(implode s) // pow10 (length s))
- val decimalsig =
-    decimalint -- option (Scan.$$ "." |-- decimalfrac)
-    >> (fn (h,NONE) => h | (h,SOME x) => h +/ x)
- fun signed prs =
-       $$ "-" |-- prs >> Rat.neg 
-    || $$ "+" |-- prs
-    || prs;
-
-fun emptyin def xs = if null xs then (def,xs) else Scan.fail xs
-
- val exponent = ($$ "e" || $$ "E") |-- signed decimalint;
-
- 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) 
-  in if null rst then x 
-     else error "mkparser: unparsed input"
-  end;;
-val parse_decimal = mkparser decimal;
-
-fun fix err prs = 
-  prs || (fn x=> error err);
-
-fun listof prs sep err =
-  prs -- Scan.bulk (sep |-- fix err prs) >> uncurry cons;
-
-(* Parse back a vector.                                                      *)
-
- val vector = 
-    token "{" |-- listof decimal (token ",") "decimal" --| token "}"
-               >> vector_of_list 
- val parse_vector = mkparser vector
- fun skipupto dscr prs inp =
-   (dscr |-- prs 
-    || Scan.one (K true) |-- skipupto dscr prs) inp 
- fun ignore inp = ((),[])
- fun sdpaoutput inp =  skipupto (word "xVec" -- token "=")
-             (vector --| ignore) inp
- fun csdpoutput inp =  ((decimal -- Scan.bulk (Scan.$$ " " |-- Scan.option decimal) >> (fn (h,to) => map_filter I ((SOME h)::to))) --| ignore >> vector_of_list) inp
- val parse_sdpaoutput = mkparser sdpaoutput
- val parse_csdpoutput = mkparser csdpoutput
-
-(* Run prover on a problem in linear form.                       *)
-
-fun run_problem prover obj mats =
-  parse_csdpoutput (prover (sdpa_of_problem obj mats))
-
-(*
-UNUSED
-
-(* Also parse the SDPA output to test success (CSDP yields a return code).   *)
-
-local
- val prs = 
-  skipupto (word "phase.value" -- token "=")
-   (Scan.option (Scan.$$ "p") -- Scan.option (Scan.$$ "d") 
-    -- (word "OPT" || word "FEAS")) 
-in
- fun sdpa_run_succeeded s = 
-  (prs (explode s); true) handle _ => false
-end;
-
-(* The default parameters. Unfortunately this goes to a fixed file.          *)
-
-val sdpa_default_parameters =
-"100     unsigned int maxIteration; \n1.0E-7  double 0.0 < epsilonStar;\n1.0E2   double 0.0 < lambdaStar;\n2.0     double 1.0 < omegaStar;\n-1.0E5  double lowerBound;\n1.0E5   double upperBound;\n0.1     double 0.0 <= betaStar <  1.0;\n0.2     double 0.0 <= betaBar  <  1.0, betaStar <= betaBar;\n0.9     double 0.0 < gammaStar  <  1.0;\n1.0E-7  double 0.0 < epsilonDash;\n";;
-
-(* These were suggested by Makoto Yamashita for problems where we are        *)
-(* right at the edge of the semidefinite cone, as sometimes happens.         *)
-
-val sdpa_alt_parameters =
-"1000    unsigned int maxIteration;\n1.0E-7  double 0.0 < epsilonStar;\n1.0E4   double 0.0 < lambdaStar;\n2.0     double 1.0 < omegaStar;\n-1.0E5  double lowerBound;\n1.0E5   double upperBound;\n0.1     double 0.0 <= betaStar <  1.0;\n0.2     double 0.0 <= betaBar  <  1.0, betaStar <= betaBar;\n0.9     double 0.0 < gammaStar  <  1.0;\n1.0E-7  double 0.0 < epsilonDash;\n";;
-
-val sdpa_params = sdpa_alt_parameters;;
-
-(* CSDP parameters; so far I'm sticking with the defaults.                   *)
-
-val csdp_default_parameters =
-"axtol=1.0e-8\natytol=1.0e-8\nobjtol=1.0e-8\npinftol=1.0e8\ndinftol=1.0e8\nmaxiter=100\nminstepfrac=0.9\nmaxstepfrac=0.97\nminstepp=1.0e-8\nminstepd=1.0e-8\nusexzgap=1\ntweakgap=0\naffine=0\nprintlevel=1\n";;
-
-val csdp_params = csdp_default_parameters;;
-
-fun tmp_file pre suf =
- let val i = string_of_int (round (random()))
-   val name = Path.append (Path.variable "ISABELLE_TMP") (Path.explode (pre ^ i ^ suf))
- in 
-   if File.exists name then tmp_file pre suf 
-   else name 
- end;
-
-(* Now call SDPA on a problem and parse back the output.                     *)
-
-fun run_sdpa dbg obj mats =
- let 
-  val input_file = tmp_file "sos" ".dat-s"
-  val output_file = tmp_file "sos" ".out"
-  val params_file = tmp_file "param" ".sdpa" 
-  val current_dir = File.pwd()
-  val _ = File.write input_file 
-                         (sdpa_of_problem "" obj mats)
-  val _ = File.write params_file sdpa_params
-  val _ = File.cd (Path.variable "ISABELLE_TMP")
-  val _ = File.system_command ("sdpa "^ (Path.implode input_file) ^ " " ^ 
-                               (Path.implode output_file) ^
-                               (if dbg then "" else "> /dev/null"))
-  val opr = File.read output_file 
- in if not(sdpa_run_succeeded opr) then error "sdpa: call failed" 
-    else
-      let val res = parse_sdpaoutput opr 
-      in ((if dbg then ()
-           else (File.rm input_file; File.rm output_file ; File.cd current_dir));
-          res)
-      end
- end;
-
-fun sdpa obj mats = run_sdpa (!debugging) obj mats;
-
-(* The same thing with CSDP.                                                 *)
-
-fun run_csdp dbg obj mats =
- let 
-  val input_file = tmp_file "sos" ".dat-s"
-  val output_file = tmp_file "sos" ".out"
-  val params_file = tmp_file "param" ".csdp"
-  val current_dir = File.pwd()
-  val _ = File.write input_file (sdpa_of_problem "" obj mats)
-  val _ = File.write params_file csdp_params
-  val _ = File.cd (Path.variable "ISABELLE_TMP")
-  val rv = system ("csdp "^(Path.implode input_file) ^ " " 
-                   ^ (Path.implode output_file) ^
-                   (if dbg then "" else "> /dev/null"))
-  val  opr = File.read output_file 
-  val res = parse_csdpoutput opr 
- in
-    ((if dbg then ()
-      else (File.rm input_file; File.rm output_file ; File.cd current_dir));
-     (rv,res))
- end;
-
-fun csdp obj mats =
- let 
-  val (rv,res) = run_csdp (!debugging) obj mats 
- in
-   ((if rv = 1 orelse rv = 2 then error "csdp: Problem is infeasible"
-    else if rv = 3 then writeln "csdp warning: Reduced accuracy"
-    else if rv <> 0 then error ("csdp: error "^string_of_int rv)
-    else ());
-   res)
- end;
-
-*)
-
-(* Try some apparently sensible scaling first. Note that this is purely to   *)
-(* get a cleaner translation to floating-point, and doesn't affect any of    *)
-(* the results, in principle. In practice it seems a lot better when there   *)
-(* are extreme numbers in the original problem.                              *)
-
-  (* Version for (int*int) keys *)
-local
-  fun max_rat x y = if x </ y then y else x
-  fun common_denominator fld amat acc =
-      fld (fn (m,c) => fn a => lcm_rat (denominator_rat c) a) amat acc
-  fun maximal_element fld amat acc =
-    fld (fn (m,c) => fn maxa => max_rat maxa (abs_rat c)) amat acc 
-fun float_of_rat x = let val (a,b) = Rat.quotient_of_rat x
-                     in Real.fromLargeInt a / Real.fromLargeInt b end;
-in
-
-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 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 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 obj'' = vector_cmul scal2 obj' 
- in solver obj'' mats''
-  end
-end;
-
-(* Try some apparently sensible scaling first. Note that this is purely to   *)
-(* get a cleaner translation to floating-point, and doesn't affect any of    *)
-(* the results, in principle. In practice it seems a lot better when there   *)
-(* are extreme numbers in the original problem.                              *)
-
-  (* Version for (int*int*int) keys *)
-local
-  fun max_rat x y = if x </ y then y else x
-  fun common_denominator fld amat acc =
-      fld (fn (m,c) => fn a => lcm_rat (denominator_rat c) a) amat acc
-  fun maximal_element fld amat acc =
-    fld (fn (m,c) => fn maxa => max_rat maxa (abs_rat c)) amat acc 
-fun float_of_rat x = let val (a,b) = Rat.quotient_of_rat x
-                     in Real.fromLargeInt a / Real.fromLargeInt b end;
-fun int_of_float x = (trunc x handle Overflow => 0 | Domain => 0)
-in
-
-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 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 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 obj'' = vector_cmul scal2 obj' 
- in solver obj'' mats''
-  end
-end;
-
-(* Round a vector to "nice" rationals.                                       *)
-
-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 => 
-   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
-
-(*
-UNUSED
-
-(* Reduce linear program to SDP (diagonal matrices) and test with CSDP. This *)
-(* one tests A [-1;x1;..;xn] >= 0 (i.e. left column is negated constants).   *)
-
-fun linear_program_basic a =
- let 
-  val (m,n) = dimensions a
-  val mats =  map (fn j => diagonal (column j a)) (1 upto n)
-  val obj = vector_const rat_1 m 
-  val (rv,res) = run_csdp false obj mats 
- in if rv = 1 orelse rv = 2 then false
-    else if rv = 0 then true
-    else error "linear_program: An error occurred in the SDP solver"
- end;
-
-(* Alternative interface testing A x >= b for matrix A, vector b.            *)
-
-fun linear_program a b =
- let val (m,n) = dimensions a 
- in if dim b <> m then error "linear_program: incompatible dimensions" 
-    else
-    let 
-     val mats = diagonal b :: map (fn j => diagonal (column j a)) (1 upto n)
-     val obj = vector_const rat_1 m 
-     val (rv,res) = run_csdp false obj mats 
-    in if rv = 1 orelse rv = 2 then false
-       else if rv = 0 then true
-       else error "linear_program: An error occurred in the SDP solver"
-    end
- end;
-
-(* Test whether a point is in the convex hull of others. Rather than use     *)
-(* computational geometry, express as linear inequalities and call CSDP.     *)
-(* This is a bit lazy of me, but it's easy and not such a bottleneck so far. *)
-
-fun in_convex_hull pts pt =
- let 
-  val pts1 = (1::pt) :: map (fn x => 1::x) pts 
-  val pts2 = map (fn p => map (fn x => ~x) p @ p) pts1
-  val n = length pts + 1
-  val v = 2 * (length pt + 1)
-  val m = v + n - 1 
-  val mat = ((m,n),
-  itern 1 pts2 (fn pts => fn j => itern 1 pts 
-               (fn x => fn i => Intpairfunc.update ((i,j), Rat.rat_of_int x)))
-  (iter (1,n) (fn i => Intpairfunc.update((v + i,i+1), rat_1)) 
-      Intpairfunc.undefined))
- in linear_program_basic mat
- end;
-
-(* Filter down a set of points to a minimal set with the same convex hull.   *)
-
-local
- fun augment1 (m::ms) = if in_convex_hull ms m then ms else ms@[m]
- fun augment m ms = funpow 3 augment1 (m::ms)
-in
-fun minimal_convex_hull mons =
- let val mons' = fold_rev augment (tl mons) [hd mons] 
- in funpow (length mons') augment1 mons'
- end
-end;
-
-*)
-
-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;
-
-fun tri_equation_add eq1 eq2 = Inttriplefunc.combine (curry op +/) (fn x => x =/ rat_0) eq1 eq2;
-
-fun tri_equation_eval assig eq =
- let fun value v = Inttriplefunc.apply assig v 
- in Inttriplefunc.fold (fn (v, c) => fn a => a +/ value v */ c) eq rat_0
- end;
-
-(* Eliminate among linear equations: return unconstrained variables and      *)
-(* assignments for the others in terms of them. We give one pseudo-variable  *)
-(* "one" that's used for a constant term.                                    *)
-
-local
-  fun extract_first p l = case l of  (* FIXME : use find_first instead *)
-   [] => error "extract_first"
- | 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
-        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)
-    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)
-   end)
-  handle ERROR _ => eliminate vs dun eqs)
-in
-fun tri_eliminate_equations one vars eqs =
- let 
-  val assig = eliminate vars Inttriplefunc.undefined 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
-end;
-
-(* Eliminate all variables, in an essentially arbitrary order.               *)
-
-fun tri_eliminate_all_equations one =
- let 
-  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" 
-         else fst (Inttriplefunc.choose eq')
-      end
-    else v 
-   end
-  fun eliminate dun eqs = case eqs of 
-    [] => dun
-  | eq::oeqs =>
-    if Inttriplefunc.is_undefined 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)
-        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)) 
-                 (map elim oeqs) 
-    end
-in fn eqs =>
- let 
-  val assig = eliminate Inttriplefunc.undefined 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
-end;
- 
-(* Solve equations by assigning arbitrary numbers.                           *)
-
-fun tri_solve_equations one eqs =
- let 
-  val (vars,assigs) = tri_eliminate_all_equations one eqs
-  val vfn = fold_rev (fn v => Inttriplefunc.update(v,rat_0)) vars 
-            (Inttriplefunc.onefunc(one, Rat.rat_of_int ~1))
-  val ass =
-    Inttriplefunc.combine (curry op +/) (K false) 
-    (Inttriplefunc.mapf (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
- 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 =>
-   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 
-   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;;
-
-val tri_epoly_neg = tri_epoly_cmul (Rat.rat_of_int ~1);
-
-val tri_epoly_add = Inttriplefunc.combine tri_equation_add Inttriplefunc.is_undefined;
-
-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;
-
-fun pi_equation_add eq1 eq2 = Inttriplefunc.combine (curry op +/) (fn x => x =/ rat_0) eq1 eq2;
-
-fun pi_equation_eval assig eq =
- let fun value v = Inttriplefunc.apply assig v 
- in Inttriplefunc.fold (fn (v, c) => fn a => a +/ value v */ c) eq rat_0
- end;
-
-(* Eliminate among linear equations: return unconstrained variables and      *)
-(* assignments for the others in terms of them. We give one pseudo-variable  *)
-(* "one" that's used for a constant term.                                    *)
-
-local
-fun extract_first p l = case l of 
-   [] => error "extract_first"
- | 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
-        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)
-    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)
-   end
-  handle ERROR _ => eliminate vs dun eqs
-in
-fun pi_eliminate_equations one vars eqs =
- let 
-  val assig = eliminate vars Inttriplefunc.undefined 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
-end;
-
-(* Eliminate all variables, in an essentially arbitrary order.               *)
-
-fun pi_eliminate_all_equations one =
- let 
-  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" 
-         else fst (Inttriplefunc.choose eq')
-      end
-    else v 
-   end
-  fun eliminate dun eqs = case eqs of 
-    [] => dun
-  | eq::oeqs =>
-    if Inttriplefunc.is_undefined 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)
-        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)) 
-                 (map elim oeqs) 
-    end
-in fn eqs =>
- let 
-  val assig = eliminate Inttriplefunc.undefined 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
-end;
- 
-(* Solve equations by assigning arbitrary numbers.                           *)
-
-fun pi_solve_equations one eqs =
- let 
-  val (vars,assigs) = pi_eliminate_all_equations one eqs
-  val vfn = fold_rev (fn v => Inttriplefunc.update(v,rat_0)) vars 
-            (Inttriplefunc.onefunc(one, Rat.rat_of_int ~1))
-  val ass =
-    Inttriplefunc.combine (curry op +/) (K false) 
-    (Inttriplefunc.mapf (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
- 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 =>
-   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 
-   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;;
-
-val pi_epoly_neg = pi_epoly_cmul (Rat.rat_of_int ~1);
-
-val pi_epoly_add = Inttriplefunc.combine pi_equation_add Inttriplefunc.is_undefined;
-
-fun pi_epoly_sub p q = pi_epoly_add p (pi_epoly_neg q);;
-
-fun allpairs f l1 l2 =  fold_rev (fn x => (curry (op @)) (map (f x) l2)) l1 [];
-
-(* Hence produce the "relevant" monomials: those whose squares lie in the    *)
-(* Newton polytope of the monomials in the input. (This is enough according  *)
-(* to Reznik: "Extremal PSD forms with few terms", Duke Math. Journal,       *)
-(* vol 45, pp. 363--374, 1978.                                               *)
-(*                                                                           *)
-(* These are ordered in sort of decreasing degree. In particular the         *)
-(* constant monomial is last; this gives an order in diagonalization of the  *)
-(* quadratic form that will tend to display constants.                       *)
-
-(*
-UNUSED
-
-fun newton_polytope pol =
- let 
-  val vars = poly_variables pol
-  val mons = map (fn m => map (fn x => monomial_degree x m) vars) 
-             (Monomialfunc.dom pol)
-  val ds = map (fn x => (degree x pol + 1) div 2) vars
-  val all = fold_rev (fn n => allpairs cons (0 upto n)) ds [[]]
-  val mons' = minimal_convex_hull mons
-  val all' =
-    filter (fn m => in_convex_hull mons' (map (fn x => 2 * x) m)) all 
- in map (fn m => fold_rev2 (fn v => fn i => fn a => if i = 0 then a else Ctermfunc.update (v,i) a)
-                        vars m monomial_1) (rev all')
- end;
-
-*)
-
-(* Diagonalize (Cholesky/LDU) the matrix corresponding to a quadratic form.  *)
-
-local
-fun diagonalize n i m =
- if Intpairfunc.is_undefined (snd m) then [] 
- else
-  let val a11 = Intpairfunc.tryapplyd (snd m) (i,i) rat_0 
-  in if a11 </ rat_0 then error "diagonalize: not PSD"
-    else if a11 =/ rat_0 then
-          if Intfunc.is_undefined (snd (row i m)) then diagonalize n (i + 1) m
-          else error "diagonalize: not PSD ___ "
-    else
-     let 
-      val v = row i m
-      val v' = (fst v, 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
-      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)
-     in (a11,v')::diagonalize n (i + 1) m' 
-     end
-  end
-in
-fun diag m =
- let 
-   val nn = dimensions m 
-   val n = fst nn 
- in if snd nn <> n then error "diagonalize: non-square matrix" 
-    else diagonalize n 1 m
- end
-end;
-
-fun gcd_rat a b = Rat.rat_of_int (Integer.gcd (int_of_rat a) (int_of_rat b));
-
-(* 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 mapa f (d,v) = 
-  (d, Intfunc.fold (fn (i,c) => fn a => upd0 i (f c) a) v Intfunc.undefined)
- fun adj (c,l) =
- let val a = 
-  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)) 
-    (snd l) rat_0
-  in ((c // (a */ a)),mapa (fn x => a */ x) l)
-  end
-in
-fun deration d = if null d then (rat_0,d) else
- let val d' = map adj d
-     val a = fold (lcm_rat o denominator_rat o fst) d' rat_1 //
-          fold (gcd_rat o numerator_rat o fst) d' rat_0 
- in ((rat_1 // a),map (fn (c,l) => (a */ c,l)) d')
- end
-end;
- 
-(* Enumeration of monomials with given multidegree bound.                    *)
-
-fun enumerate_monomials d vars = 
- if d < 0 then []
- else if d = 0 then [Ctermfunc.undefined]
- 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 fold1 (curry op @) 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)] 
-else if d < 0 then [] else
-case pols of 
-   [] => [(poly_const rat_1,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)))
-         (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;
-
-(* String for block diagonal matrix numbered k.                              *)
-
-fun sdpa_of_blockdiagonal k m =
- let 
-  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 
- 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 ""
- end;
-
-(* SDPA for problem using block diagonal (i.e. multiple SDPs)                *)
-
-fun sdpa_of_blockproblem nblocks blocksizes obj mats =
- let val m = length mats - 1 
- in
-  string_of_int m ^ "\n" ^
-  string_of_int nblocks ^ "\n" ^
-  (fold1 (fn s => fn t => s^" "^t) (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)
-    (1 upto length mats) mats ""
- end;
-
-(* Run prover on a problem in block diagonal form.                       *)
-
-fun run_blockproblem prover nblocks blocksizes obj mats=
-  parse_csdpoutput (prover (sdpa_of_blockproblem nblocks blocksizes obj mats))
-
-(*
-UNUSED
-
-(* Hence run CSDP on a problem in block diagonal form.                       *)
-
-fun run_csdp dbg nblocks blocksizes obj mats =
- let 
-  val input_file = tmp_file "sos" ".dat-s" 
-  val output_file = tmp_file "sos" ".out"
-  val params_file = tmp_file "param" ".csdp" 
-  val _ = File.write input_file
-   (sdpa_of_blockproblem "" nblocks blocksizes obj mats)
-  val _ = File.write params_file csdp_params
-  val current_dir = File.pwd()
-  val _ = File.cd (Path.variable "ISABELLE_TMP")
-  val rv = system ("csdp "^(Path.implode input_file) ^ " " 
-                   ^ (Path.implode output_file) ^
-                   (if dbg then "" else "> /dev/null"))
-  val  opr = File.read output_file 
-  val res = parse_csdpoutput opr 
- in
-   ((if dbg then ()
-     else (File.rm input_file ; File.rm output_file ; File.cd current_dir));
-    (rv,res))
- end;
-
-fun csdp nblocks blocksizes obj mats =
- let 
-  val (rv,res) = run_csdp (!debugging) nblocks blocksizes obj mats 
- in ((if rv = 1 orelse rv = 2 then error "csdp: Problem is infeasible"
-     else if rv = 3 then writeln "csdp warning: Reduced accuracy"
-     else if rv <> 0 then error ("csdp: error "^string_of_int rv)
-     else ());
-     res)
- end;
-*)
-
-(* 3D versions of matrix operations to consider blocks separately.           *)
-
-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;
-
-val bmatrix_neg = bmatrix_cmul (Rat.rat_of_int ~1);
-fun bmatrix_sub m1 m2 = bmatrix_add m1 (bmatrix_neg m2);;
-
-(* Smash a block matrix into components.                                     *)
-
-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 
-      in (((bs,bs),m):matrix) end)
- (blocksizes ~~ (1 upto length blocksizes));;
-
-(* FIXME : Get rid of this !!!*)
-local
-  fun tryfind_with msg f [] = error msg
-    | tryfind_with msg f (x::xs) = (f x handle ERROR s => tryfind_with s f xs);
-in 
-  fun tryfind f = tryfind_with "tryfind" f
-end
-
-(*
-fun tryfind f [] = error "tryfind"
-  | tryfind f (x::xs) = (f x handle ERROR _ => tryfind f xs);
-*)
-
-(* 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)::
-      (filter (fn (p,c) => multidegree p <= d) leqs)
-    else enumerate_products d leqs
- val nblocks = length monoid
- fun mk_idmultiplier k p =
-  let 
-   val e = d - multidegree p
-   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)
-  end
-
- fun mk_sqmultiplier k (p,c) =
-  let 
-   val e = (d - multidegree p) div 2
-   val mons = enumerate_monomials e vars
-   val nons = mons ~~ (1 upto length mons) 
-  in (mons, 
-      fold_rev (fn (m1,n1) =>
-       fold_rev (fn (m2,n2) => fn  a =>
-        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
-          end
-        end)  nons)
-       nons Monomialfunc.undefined)
-  end
-
-  val (sqmonlist,sqs) = split_list (map2 mk_sqmultiplier (1 upto length monoid) monoid)
-  val (idmonlist,ids) =  split_list(map2 mk_idmultiplier (1 upto length eqs) eqs)
-  val blocksizes = map length sqmonlist
-  val bigsum =
-    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 (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
-  fun mk_matrix v =
-    Inttriplefunc.fold (fn ((b,i,j), ass) => fn m => 
-        if b < 0 then m else
-         let val c = Inttriplefunc.tryapplyd ass v rat_0
-         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
-  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
-
-  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)
-  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 find_rounding d =
-   let 
-    val _ = if !debugging 
-          then writeln ("Trying rounding with limit "^Rat.string_of_rat d ^ "\n") 
-          else ()
-    val vec = nice_vector d raw_vec
-    val blockmat = iter (1,dim vec)
-     (fn i => fn a => bmatrix_add (bmatrix_cmul (int_element vec i) (nth mats i)) a)
-     (bmatrix_neg (nth mats 0))
-    val allmats = blocks blocksizes blockmat 
-   in (vec,map diag allmats)
-   end
-  val (vec,ratdias) =
-    if null pvs then find_rounding rat_1
-    else tryfind find_rounding (map Rat.rat_of_int (1 upto 31) @
-                                map pow2 (5 upto 66))
-  val newassigs =
-    fold_rev (fn k => Inttriplefunc.update (nth pvs (k - 1), int_element vec k))
-           (1 upto dim vec) (Inttriplefunc.onefunc ((0,0,0), Rat.rat_of_int ~1))
-  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
-  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)
-   in map mk_sq
-   end
-  val sqs = map2 mk_sos sqmonlist ratdias
-  val cfs = map poly_of_epoly ids
-  val msq = filter (fn (a,b) => not (null b)) (map2 pair monoid sqs)
-  fun eval_sq sqs = fold_rev (fn (c,q) => poly_add (poly_cmul c (poly_mul q q))) sqs poly_0
-  val sanity =
-    fold_rev (fn ((p,c),s) => poly_add (poly_mul p (eval_sq s))) msq
-           (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
-  (cfs,map (fn (a,b) => (snd a,b)) msq)
- end
-
-
-end;
-
-(* Iterative deepening.                                                      *)
-
-fun deepen f n = 
-  (writeln ("Searching with depth limit " ^ string_of_int n) ; (f n handle ERROR s => (writeln ("failed with message: " ^ s) ; deepen f (n+1))))
-
-(* The ordering so we can create canonical HOL polynomials.                  *)
-
-fun dest_monomial mon = sort (increasing fst cterm_ord) (Ctermfunc.graph mon);
-
-fun monomial_order (m1,m2) =
- if Ctermfunc.is_undefined m2 then LESS 
- else if Ctermfunc.is_undefined m1 then GREATER 
- else
-  let val mon1 = dest_monomial m1 
-      val mon2 = dest_monomial m2
-      val deg1 = fold (curry op + o snd) mon1 0
-      val deg2 = fold (curry op + o snd) mon2 0 
-  in if deg1 < deg2 then GREATER else if deg1 > deg2 then LESS
-     else list_ord (prod_ord cterm_ord int_ord) (mon1,mon2)
-  end;
-
-fun dest_poly p =
-  map (fn (m,c) => (c,dest_monomial m))
-      (sort (prod_ord monomial_order (K EQUAL)) (Monomialfunc.graph p));
-
-(* Map back polynomials and their composites to HOL.                         *)
-
-local
- open Thm Numeral RealArith
-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_monomial m = 
- if Ctermfunc.is_undefined m then @{cterm "1::real"} 
- else 
-  let 
-   val m' = dest_monomial m
-   val vps = fold_rev (fn (x,k) => cons (cterm_of_varpow x k)) m' [] 
-  in fold1 (fn s => fn t => capply (capply @{cterm "op * :: real => _"} s) t) vps
-  end
-
-fun cterm_of_cmonomial (m,c) = if Ctermfunc.is_undefined 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);
-
-fun cterm_of_poly p = 
- if Monomialfunc.is_undefined p then @{cterm "0::real"} 
- else
-  let 
-   val cms = map cterm_of_cmonomial
-     (sort (prod_ord monomial_order (K EQUAL)) (Monomialfunc.graph p))
-  in fold1 (fn t1 => fn t2 => capply(capply @{cterm "op + :: real => _"} t1) t2) cms
-  end;
-
-fun cterm_of_sqterm (c,p) = Product(Rational_lt c,Square(cterm_of_poly p));
-
-fun cterm_of_sos (pr,sqs) = if null sqs then pr
-  else Product(pr,fold1 (fn a => fn b => Sum(a,b)) (map cterm_of_sqterm sqs));
-
-end
-
-(* Interface to HOL.                                                         *)
-local
-  open Thm Conv RealArith
-  val concl = 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 !! *)
-fun real_nonlinear_prover prover ctxt =
- let 
-  val {add,mul,neg,pow,sub,main} =  Normalizer.semiring_normalizers_ord_wrapper ctxt
-      (valOf (NormalizerData.match ctxt @{cterm "(0::real) + 1"})) 
-     simple_cterm_ord
-  val (real_poly_add_conv,real_poly_mul_conv,real_poly_neg_conv,
-       real_poly_pow_conv,real_poly_sub_conv,real_poly_conv) = (add,mul,neg,pow,sub,main)
-  fun mainf  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 (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 
-                     else error "trivial_axiom: Not a trivial axiom"
-     | Axiom_le n => if eval Ctermfunc.undefined p </ Rat.zero then nth les n 
-                     else error "trivial_axiom: Not a trivial axiom"
-     | Axiom_lt n => if eval Ctermfunc.undefined p <=/ Rat.zero then nth lts n 
-                     else error "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 end)
-   handle ERROR _ => (
-    let 
-     val pol = fold_rev poly_mul (map fst ltp) (poly_const Rat.one)
-     val leq = lep @ ltp
-     fun tryall d =
-      let val e = multidegree pol
-          val k = if e = 0 then 0 else d div e
-          val eq' = map fst eq 
-      in tryfind (fn i => (d,i,real_positivnullstellensatz_general prover false d eq' leq
-                            (poly_neg(poly_pow pol i))))
-              (0 upto k)
-      end
-    val (d,i,(cert_ideal,cert_cone)) = deepen tryall 0
-    val proofs_ideal =
-      map2 (fn q => fn (p,ax) => Eqmul(cterm_of_poly 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 = fold1 (fn s => fn t => Product(s,t)) (map snd ltp) 
-      in  funpow i (fn q => Product(p,q)) (Rational_lt Rat.one)
-      end
-    val proof = fold1 (fn s => fn t => Sum(s,t))
-                           (proof_ne :: proofs_ideal @ proofs_cone) 
-    in writeln "Translating proof certificate to HOL";
-       translator (eqs,les,lts) proof
-    end))
-   end
- in mainf end
-end
-
-fun C f x y = f y x;
-  (* FIXME : This is very bad!!!*)
-fun subst_conv eqs t = 
- let 
-  val t' = fold (Thm.cabs o Thm.lhs_of) eqs t
- in Conv.fconv_rule (Thm.beta_conversion true) (fold (C combination) eqs (reflexive t'))
- end
-
-(* A wrapper that tries to substitute away variables first.                  *)
-
-local
- open Thm Conv RealArith
-  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 shuffle1 = 
-   fconv_rule (rewr_conv @{lemma "(a + x == y) == (x == y - (a::real))" by (atomize (full)) (simp add: ring_simps) })
- val shuffle2 =
-    fconv_rule (rewr_conv @{lemma "(x + a == y) ==  (x == y - (a::real))" by (atomize (full)) (simp add: ring_simps)})
- fun substitutable_monomial fvs tm = case term_of tm of
-    Free(_,@{typ real}) => if not (member (op aconvc) fvs tm) then (Rat.one,tm) 
-                           else error "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 error "substitutable_monomial"
-  | @{term "op + :: real => _"}$s$t => 
-       (substitutable_monomial (add_cterm_frees (dest_arg tm) fvs) (dest_arg1 tm)
-        handle ERROR _ => substitutable_monomial (add_cterm_frees (dest_arg1 tm) fvs) (dest_arg tm))
-  | _ => error "substitutable_monomial"
-
-  fun isolate_variable v th = 
-   let val w = 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 
-              else isolate_variable v (shuffle1 th)
-          | _ => error "isolate variable : This should not happen?"
-   end 
-in
-
-fun real_nonlinear_subst_prover prover ctxt =
- let 
-  val {add,mul,neg,pow,sub,main} =  Normalizer.semiring_normalizers_ord_wrapper ctxt
-      (valOf (NormalizerData.match ctxt @{cterm "(0::real) + 1"})) 
-     simple_cterm_ord
-
-  val (real_poly_add_conv,real_poly_mul_conv,real_poly_neg_conv,
-       real_poly_pow_conv,real_poly_sub_conv,real_poly_conv) = (add,mul,neg,pow,sub,main)
-
-  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 th2 = fconv_rule (binop_conv real_poly_mul_conv) th1
-   in fconv_rule (arg_conv real_poly_conv) (isolate_variable v th2)
-   end
-   fun oprconv cv ct = 
-    let val g = Thm.dest_fun2 ct
-    in if g aconvc @{cterm "op <= :: real => _"} 
-         orelse g aconvc @{cterm "op < :: real => _"} 
-       then arg_conv cv ct else arg1_conv cv ct
-    end
-  fun mainf translator =
-   let 
-    fun substfirst(eqs,les,lts) =
-      ((let 
-           val eth = tryfind make_substitution eqs
-           val modify = fconv_rule (arg_conv (oprconv(subst_conv [eth] then_conv real_poly_conv)))
-       in  substfirst
-             (filter_out (fn t => (Thm.dest_arg1 o Thm.dest_arg o cprop_of) t 
-                                   aconvc @{cterm "0::real"}) (map modify eqs),
-                                   map modify les,map modify lts)
-       end)
-       handle ERROR  _ => real_nonlinear_prover prover ctxt translator (rev eqs, rev les, rev lts))
-    in substfirst
-   end
-
-
- in mainf
- end
-
-(* Overall function. *)
-
-fun real_sos prover ctxt t = gen_prover_real_arith ctxt (real_nonlinear_subst_prover prover ctxt) t;
-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)
-
-fun core_sos_conv prover ctxt t = Drule.arg_cong_rule @{cterm Trueprop} (real_sos prover ctxt (Thm.dest_arg t) RS @{thm Eq_TrueI})
-
-val known_sos_constants = 
-  [@{term "op ==>"}, @{term "Trueprop"}, 
-   @{term "op -->"}, @{term "op &"}, @{term "op |"}, 
-   @{term "Not"}, @{term "op = :: bool => _"}, 
-   @{term "All :: (real => _) => _"}, @{term "Ex :: (real => _) => _"}, 
-   @{term "op = :: real => _"}, @{term "op < :: real => _"}, 
-   @{term "op <= :: real => _"}, 
-   @{term "op + :: real => _"}, @{term "op - :: real => _"}, 
-   @{term "op * :: real => _"}, @{term "uminus :: real => _"}, 
-   @{term "op / :: real => _"}, @{term "inverse :: real => _"},
-   @{term "op ^ :: real => _"}, @{term "abs :: real => _"}, 
-   @{term "min :: real => _"}, @{term "max :: real => _"},
-   @{term "0::real"}, @{term "1::real"}, @{term "number_of :: int => real"},
-   @{term "number_of :: int => nat"},
-   @{term "Int.Bit0"}, @{term "Int.Bit1"}, 
-   @{term "Int.Pls"}, @{term "Int.Min"}];
-
-fun check_sos kcts ct = 
- let
-  val t = term_of ct
-  val _ = if not (null (Term.add_tfrees t []) 
-                  andalso null (Term.add_tvars t [])) 
-          then error "SOS: not sos. Additional type varables" else ()
-  val fs = Term.add_frees t []
-  val _ = if exists (fn ((_,T)) => not (T = @{typ "real"})) fs 
-          then error "SOS: not sos. Variables with type not real" else ()
-  val vs = Term.add_vars t []
-  val _ = if exists (fn ((_,T)) => not (T = @{typ "real"})) fs 
-          then error "SOS: not sos. Variables with type not real" else ()
-  val ukcs = subtract (fn (t,p) => Const p aconv t) kcts (Term.add_consts t [])
-  val _ = if  null ukcs then () 
-              else error ("SOSO: Unknown constants in Subgoal:" ^ commas (map fst ukcs))
-in () end
-
-fun core_sos_tac prover ctxt = CSUBGOAL (fn (ct, i) => 
-  let val _ = check_sos known_sos_constants ct
-      val (avs, p) = strip_all ct
-      val th = standard (fold_rev forall_intr avs (real_sos prover ctxt (Thm.dest_arg p)))
-  in rtac th i end);
-
-fun default_SOME f NONE v = SOME v
-  | default_SOME f (SOME v) _ = SOME v;
-
-fun lift_SOME f NONE a = f a
-  | lift_SOME f (SOME a) _ = SOME a;
-
-
-local
- val is_numeral = can (HOLogic.dest_number o term_of)
-in
-fun get_denom b ct = case term_of ct of
-  @{term "op / :: real => _"} $ _ $ _ => 
-     if is_numeral (Thm.dest_arg ct) then get_denom b (Thm.dest_arg1 ct)
-     else default_SOME (get_denom b) (get_denom b (Thm.dest_arg ct))   (Thm.dest_arg ct, b)
- | @{term "op < :: real => _"} $ _ $ _ => lift_SOME (get_denom true) (get_denom true (Thm.dest_arg ct)) (Thm.dest_arg1 ct)
- | @{term "op <= :: real => _"} $ _ $ _ => lift_SOME (get_denom true) (get_denom true (Thm.dest_arg ct)) (Thm.dest_arg1 ct)
- | _ $ _ => lift_SOME (get_denom b) (get_denom b (Thm.dest_fun ct)) (Thm.dest_arg ct)
- | _ => NONE
-end;
-
-fun elim_one_denom_tac ctxt = 
-CSUBGOAL (fn (P,i) => 
- case get_denom false P of 
-   NONE => no_tac
- | SOME (d,ord) => 
-     let 
-      val ss = simpset_of ctxt addsimps @{thms field_simps} 
-               addsimps [@{thm nonzero_power_divide}, @{thm power_divide}]
-      val th = instantiate' [] [SOME d, SOME (Thm.dest_arg P)] 
-         (if ord then @{lemma "(d=0 --> P) & (d>0 --> P) & (d<(0::real) --> P) ==> P" by auto}
-          else @{lemma "(d=0 --> P) & (d ~= (0::real) --> P) ==> P" by blast})
-     in (rtac th i THEN Simplifier.asm_full_simp_tac ss i) end);
-
-fun elim_denom_tac ctxt i = REPEAT (elim_one_denom_tac ctxt i);
-
-fun sos_tac prover ctxt = ObjectLogic.full_atomize_tac THEN' elim_denom_tac ctxt THEN' core_sos_tac prover ctxt
-
-
-end;
--- a/src/HOL/Nominal/nominal_inductive2.ML	Mon Aug 10 08:37:37 2009 +0200
+++ b/src/HOL/Nominal/nominal_inductive2.ML	Mon Aug 10 10:25:00 2009 +0200
@@ -8,7 +8,7 @@
 
 signature NOMINAL_INDUCTIVE2 =
 sig
-  val prove_strong_ind: string -> (string * string list) list -> local_theory -> Proof.state
+  val prove_strong_ind: string -> string option -> (string * string list) list -> local_theory -> Proof.state
 end
 
 structure NominalInductive2 : NOMINAL_INDUCTIVE2 =
@@ -150,7 +150,7 @@
     map (Envir.subst_term env #> cterm_of thy) vs ~~ cts) th
   end;
 
-fun prove_strong_ind s avoids ctxt =
+fun prove_strong_ind s alt_name avoids ctxt =
   let
     val thy = ProofContext.theory_of ctxt;
     val ({names, ...}, {raw_induct, intrs, elims, ...}) =
@@ -461,8 +461,13 @@
             (strong_raw_induct, [ind_case_names, RuleCases.consumes 0])
           else (strong_raw_induct RSN (2, rev_mp),
             [ind_case_names, RuleCases.consumes 1]);
+        val (induct_name, inducts_name) =
+          case alt_name of
+            NONE => (rec_qualified (Binding.name "strong_induct"),
+                     rec_qualified (Binding.name "strong_inducts"))
+          | SOME s => (Binding.name s, Binding.name (s ^ "s"));
         val ((_, [strong_induct']), ctxt') = LocalTheory.note Thm.generatedK
-          ((rec_qualified (Binding.name "strong_induct"),
+          ((induct_name,
             map (Attrib.internal o K) (#2 strong_induct)), [#1 strong_induct])
           ctxt;
         val strong_inducts =
@@ -470,7 +475,7 @@
       in
         ctxt' |>
         LocalTheory.note Thm.generatedK
-          ((rec_qualified (Binding.name "strong_inducts"),
+          ((inducts_name,
             [Attrib.internal (K ind_case_names),
              Attrib.internal (K (RuleCases.consumes 1))]),
            strong_inducts) |> snd
@@ -486,9 +491,11 @@
 val _ =
   OuterSyntax.local_theory_to_proof "nominal_inductive2"
     "prove strong induction theorem for inductive predicate involving nominal datatypes" K.thy_goal
-    (P.xname -- Scan.optional (P.$$$ "avoids" |-- P.enum1 "|" (P.name --
-      (P.$$$ ":" |-- P.and_list1 P.term))) [] >> (fn (name, avoids) =>
-        prove_strong_ind name avoids));
+    (P.xname -- 
+     Scan.option (P.$$$ "(" |-- P.!!! (P.name --| P.$$$ ")")) --
+     (Scan.optional (P.$$$ "avoids" |-- P.enum1 "|" (P.name --
+      (P.$$$ ":" |-- P.and_list1 P.term))) []) >> (fn ((name, rule_name), avoids) =>
+        prove_strong_ind name rule_name avoids));
 
 end;
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/ATP_Manager/SystemOnTPTP	Mon Aug 10 10:25:00 2009 +0200
@@ -0,0 +1,141 @@
+#!/usr/bin/env perl
+#
+# Wrapper for custom remote provers on SystemOnTPTP
+# Author: Fabian Immler, TU Muenchen
+#
+
+use warnings;
+use strict;
+use Getopt::Std;
+use HTTP::Request::Common;
+use LWP;
+
+my $SystemOnTPTPFormReplyURL =
+  "http://www.cs.miami.edu/~tptp/cgi-bin/SystemOnTPTPFormReply";
+
+# default parameters
+my %URLParameters = (
+    "NoHTML" => 1,
+    "QuietFlag" => "-q01",
+    "SubmitButton" => "RunSelectedSystems",
+    "ProblemSource" => "UPLOAD",
+    );
+
+#----Get format and transform options if specified
+my %Options;
+getopts("hwxs:t:c:",\%Options);
+
+#----Usage
+sub usage() {
+  print("Usage: remote [<options>] <File name>\n");
+  print("    <options> are ...\n");
+  print("    -h            - print this help\n");
+  print("    -w            - list available ATP systems\n");
+  print("    -x            - use X2TPTP to convert output of prover\n");
+  print("    -s<system>    - specified system to use\n");
+  print("    -t<timelimit> - CPU time limit for system\n");
+  print("    -c<command>   - custom command for system\n");
+  print("    <File name>   - TPTP problem file\n");
+  exit(0);
+}
+if (exists($Options{'h'})) {
+  usage();
+}
+
+#----What systems flag
+if (exists($Options{'w'})) {
+    $URLParameters{"SubmitButton"} = "ListSystems";
+    delete($URLParameters{"ProblemSource"});
+}
+
+#----X2TPTP
+if (exists($Options{'x'})) {
+    $URLParameters{"X2TPTP"} = "-S";
+}
+
+#----Selected system
+my $System;
+if (exists($Options{'s'})) {
+    $System = $Options{'s'};
+} else {
+    # use Vampire as default
+    $System = "Vampire---9.0";
+}
+$URLParameters{"System___$System"} = $System;
+
+#----Time limit
+if (exists($Options{'t'})) {
+    $URLParameters{"TimeLimit___$System"} = $Options{'t'};
+}
+#----Custom command
+if (exists($Options{'c'})) {
+    $URLParameters{"Command___$System"} = $Options{'c'};
+}
+
+#----Get single file name
+if (exists($URLParameters{"ProblemSource"})) {
+    if (scalar(@ARGV) >= 1) {
+        $URLParameters{"UPLOADProblem"} = [shift(@ARGV)];
+    } else {
+      print("Missing problem file\n");
+      usage();
+      die;
+    }
+}
+
+# Query Server
+my $Agent = LWP::UserAgent->new;
+if (exists($Options{'t'})) {
+  # give server more time to respond
+  $Agent->timeout($Options{'t'} + 10);
+}
+my $Request = POST($SystemOnTPTPFormReplyURL,
+	Content_Type => 'form-data',Content => \%URLParameters);
+my $Response = $Agent->request($Request);
+
+#catch errors / failure
+if(!$Response->is_success) {
+  print "HTTP-Error: " . $Response->message . "\n";
+  exit(-1);
+} elsif (exists($Options{'w'})) {
+  print $Response->content;
+  exit (0);
+} elsif ($Response->content =~ /WARNING: (\S*) does not exist/) {
+  print "Specified System $1 does not exist\n";
+  exit(-1);
+} elsif (exists($Options{'x'}) &&
+  $Response->content =~
+    /%\s*Result\s*:\s*Unsatisfiable.*\n%\s*Output\s*:\s*(CNF)?Refutation.*\n%/ &&
+  $Response->content !~ /ERROR: Could not form TPTP format derivation/ )
+{
+  # converted output: extract proof
+  my @lines = split( /\n/, $Response->content);
+  my $extract = "";
+  foreach my $line (@lines){
+      #ignore comments
+      if ($line !~ /^%/ && !($line eq "")) {
+          $extract .= "$line";
+      }
+  }
+  # insert newlines after ').'
+  $extract =~ s/\s//g;
+  $extract =~ s/\)\.cnf/\)\.\ncnf/g;
+
+  print "========== ~~/lib/scripts/SystemOnTPTP extracted proof: ==========\n";
+  # orientation for res_reconstruct.ML
+  print "# SZS output start CNFRefutation.\n";
+  print "$extract\n";
+  print "# SZS output end CNFRefutation.\n";
+  # can be useful for debugging; Isabelle ignores this
+  print "============== original response from SystemOnTPTP: ==============\n";
+  print $Response->content;
+  exit(0);
+} elsif (!exists($Options{'x'})) {
+  # pass output directly to Isabelle
+  print $Response->content;
+  exit(0);
+}else {
+  print "Remote-script could not extract proof:\n".$Response->content;
+  exit(-1);
+}
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/ATP_Manager/atp_manager.ML	Mon Aug 10 10:25:00 2009 +0200
@@ -0,0 +1,403 @@
+(*  Title:      HOL/Tools/ATP_Manager/atp_manager.ML
+    Author:     Fabian Immler, TU Muenchen
+
+ATP threads are registered here.
+Threads with the same birth-time are seen as one group.
+All threads of a group are killed when one thread of it has been successful,
+or after a certain time,
+or when the maximum number of threads exceeds; then the oldest thread is killed.
+*)
+
+signature ATP_MANAGER =
+sig
+  val get_atps: unit -> string
+  val set_atps: string -> unit
+  val get_max_atps: unit -> int
+  val set_max_atps: int -> unit
+  val get_timeout: unit -> int
+  val set_timeout: int -> unit
+  val get_full_types: unit -> bool
+  val set_full_types: bool -> unit
+  val kill: unit -> unit
+  val info: unit -> unit
+  val messages: int option -> unit
+  type prover = int -> (thm * (string * int)) list option ->
+    (thm * (string * int)) list option -> string -> int ->
+    Proof.context * (thm list * thm) ->
+    bool * string * string * string vector * (thm * (string * int)) list
+  val add_prover: string -> prover -> theory -> theory
+  val print_provers: theory -> unit
+  val get_prover: string -> theory -> prover option
+  val sledgehammer: string list -> Proof.state -> unit
+end;
+
+structure AtpManager: ATP_MANAGER =
+struct
+
+(** preferences **)
+
+val message_store_limit = 20;
+val message_display_limit = 5;
+
+local
+
+val atps = ref "e remote_vampire";
+val max_atps = ref 5;   (* ~1 means infinite number of atps *)
+val timeout = ref 60;
+val full_types = ref false;
+
+in
+
+fun get_atps () = CRITICAL (fn () => ! atps);
+fun set_atps str = CRITICAL (fn () => atps := str);
+
+fun get_max_atps () = CRITICAL (fn () => ! max_atps);
+fun set_max_atps number = CRITICAL (fn () => max_atps := number);
+
+fun get_timeout () = CRITICAL (fn () => ! timeout);
+fun set_timeout time = CRITICAL (fn () => timeout := time);
+
+fun get_full_types () = CRITICAL (fn () => ! full_types);
+fun set_full_types bool = CRITICAL (fn () => full_types := bool);
+
+val _ =
+  ProofGeneralPgip.add_preference Preferences.category_proof
+    (Preferences.string_pref atps
+      "ATP: provers" "Default automatic provers (separated by whitespace)");
+
+val _ =
+  ProofGeneralPgip.add_preference Preferences.category_proof
+    (Preferences.int_pref max_atps
+      "ATP: maximum number" "How many provers may run in parallel");
+
+val _ =
+  ProofGeneralPgip.add_preference Preferences.category_proof
+    (Preferences.int_pref timeout
+      "ATP: timeout" "ATPs will be interrupted after this time (in seconds)");
+
+val _ =
+  ProofGeneralPgip.add_preference Preferences.category_proof
+    (Preferences.bool_pref full_types
+      "ATP: full types" "ATPs will use full type information");
+
+end;
+
+
+
+(** thread management **)
+
+(* data structures over threads *)
+
+structure ThreadHeap = HeapFun
+(
+  type elem = Time.time * Thread.thread;
+  fun ord ((a, _), (b, _)) = Time.compare (a, b);
+);
+
+fun lookup_thread xs = AList.lookup Thread.equal xs;
+fun delete_thread xs = AList.delete Thread.equal xs;
+fun update_thread xs = AList.update Thread.equal xs;
+
+
+(* state of thread manager *)
+
+datatype T = State of
+ {managing_thread: Thread.thread option,
+  timeout_heap: ThreadHeap.T,
+  oldest_heap: ThreadHeap.T,
+  active: (Thread.thread * (Time.time * Time.time * string)) list,
+  cancelling: (Thread.thread * (Time.time * Time.time * string)) list,
+  messages: string list,
+  store: string list};
+
+fun make_state managing_thread timeout_heap oldest_heap active cancelling messages store =
+  State {managing_thread = managing_thread, timeout_heap = timeout_heap, oldest_heap = oldest_heap,
+    active = active, cancelling = cancelling, messages = messages, store = store};
+
+val state = Synchronized.var "atp_manager"
+  (make_state NONE ThreadHeap.empty ThreadHeap.empty [] [] [] []);
+
+
+(* unregister thread *)
+
+fun unregister (success, message) thread = Synchronized.change state
+  (fn state as State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
+    (case lookup_thread active thread of
+      SOME (birthtime, _, description) =>
+        let
+          val (group, active') =
+            if success then List.partition (fn (_, (tb, _, _)) => tb = birthtime) active
+            else List.partition (fn (th, _) => Thread.equal (th, thread)) active
+
+          val now = Time.now ()
+          val cancelling' =
+            fold (fn (th, (tb, _, desc)) => update_thread (th, (tb, now, desc))) group cancelling
+
+          val message' = description ^ "\n" ^ message ^
+            (if length group <= 1 then ""
+             else "\nInterrupted " ^ string_of_int (length group - 1) ^ " other group members")
+          val store' = message' ::
+            (if length store <= message_store_limit then store
+             else #1 (chop message_store_limit store))
+        in make_state
+          managing_thread timeout_heap oldest_heap active' cancelling' (message' :: messages) store'
+        end
+    | NONE => state));
+
+
+(* kill excessive atp threads *)
+
+fun excessive_atps active =
+  let val max = get_max_atps ()
+  in length active > max andalso max > ~1 end;
+
+local
+
+fun kill_oldest () =
+  let exception Unchanged in
+    Synchronized.change_result state
+      (fn State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
+        if ThreadHeap.is_empty oldest_heap orelse not (excessive_atps active)
+        then raise Unchanged
+        else
+          let val ((_, oldest_thread), oldest_heap') = ThreadHeap.min_elem oldest_heap
+          in (oldest_thread,
+          make_state managing_thread timeout_heap oldest_heap' active cancelling messages store) end)
+      |> unregister (false, "Interrupted (maximum number of ATPs exceeded)")
+    handle Unchanged => ()
+  end;
+
+in
+
+fun kill_excessive () =
+  let val State {active, ...} = Synchronized.value state
+  in if excessive_atps active then (kill_oldest (); kill_excessive ()) else () end;
+
+end;
+
+fun print_new_messages () =
+  let val to_print = Synchronized.change_result state
+    (fn State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
+      (messages, make_state managing_thread timeout_heap oldest_heap active cancelling [] store))
+  in
+    if null to_print then ()
+    else priority ("Sledgehammer: " ^ space_implode "\n\n" to_print)
+  end;
+
+
+(* start a watching thread -- only one may exist *)
+
+fun check_thread_manager () = Synchronized.change state
+  (fn State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
+    if (case managing_thread of SOME thread => Thread.isActive thread | NONE => false)
+    then make_state managing_thread timeout_heap oldest_heap active cancelling messages store
+    else let val managing_thread = SOME (SimpleThread.fork false (fn () =>
+      let
+        val min_wait_time = Time.fromMilliseconds 300
+        val max_wait_time = Time.fromSeconds 10
+
+        (* wait for next thread to cancel, or maximum*)
+        fun time_limit (State {timeout_heap, ...}) =
+          (case try ThreadHeap.min timeout_heap of
+            NONE => SOME (Time.+ (Time.now (), max_wait_time))
+          | SOME (time, _) => SOME time)
+
+        (* action: find threads whose timeout is reached, and interrupt cancelling threads *)
+        fun action (State {managing_thread, timeout_heap, oldest_heap, active, cancelling,
+                           messages, store}) =
+          let val (timeout_threads, timeout_heap') =
+            ThreadHeap.upto (Time.now (), Thread.self ()) timeout_heap
+          in
+            if null timeout_threads andalso null cancelling andalso not (excessive_atps active)
+            then NONE
+            else
+              let
+                val _ = List.app (SimpleThread.interrupt o #1) cancelling
+                val cancelling' = filter (Thread.isActive o #1) cancelling
+                val state' = make_state
+                  managing_thread timeout_heap' oldest_heap active cancelling' messages store
+              in SOME (map #2 timeout_threads, state') end
+          end
+      in
+        while Synchronized.change_result state
+          (fn st as
+            State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
+            if (null active) andalso (null cancelling) andalso (null messages)
+            then (false, make_state NONE timeout_heap oldest_heap active cancelling messages store)
+            else (true, st))
+        do
+          (Synchronized.timed_access state time_limit action
+            |> these
+            |> List.app (unregister (false, "Interrupted (reached timeout)"));
+            kill_excessive ();
+            print_new_messages ();
+            (*give threads time to respond to interrupt*)
+            OS.Process.sleep min_wait_time)
+      end))
+    in make_state managing_thread timeout_heap oldest_heap active cancelling messages store end);
+
+
+(* thread is registered here by sledgehammer *)
+
+fun register birthtime deadtime (thread, desc) =
+ (Synchronized.change state
+    (fn State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
+      let
+        val timeout_heap' = ThreadHeap.insert (deadtime, thread) timeout_heap
+        val oldest_heap' = ThreadHeap.insert (birthtime, thread) oldest_heap
+        val active' = update_thread (thread, (birthtime, deadtime, desc)) active
+      in make_state managing_thread timeout_heap' oldest_heap' active' cancelling messages store end);
+  check_thread_manager ());
+
+
+
+(** user commands **)
+
+(* kill: move all threads to cancelling *)
+
+fun kill () = Synchronized.change state
+  (fn State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
+    let val formerly_active = map (fn (th, (tb, _, desc)) => (th, (tb, Time.now (), desc))) active
+    in make_state
+      managing_thread timeout_heap oldest_heap [] (formerly_active @ cancelling) messages store
+    end);
+
+
+(* ATP info *)
+
+fun info () =
+  let
+    val State {active, cancelling, ...} = Synchronized.value state
+
+    fun running_info (_, (birth_time, dead_time, desc)) = "Running: "
+        ^ (string_of_int o Time.toSeconds) (Time.- (Time.now (), birth_time))
+        ^ " s  --  "
+        ^ (string_of_int o Time.toSeconds) (Time.- (dead_time, Time.now ()))
+        ^ " s to live:\n" ^ desc
+    fun cancelling_info (_, (_, dead_time, desc)) = "Trying to interrupt thread since "
+        ^ (string_of_int o Time.toSeconds) (Time.- (Time.now (), dead_time))
+        ^ " s:\n" ^ desc
+
+    val running =
+      if null active then "No ATPs running."
+      else space_implode "\n\n" ("Running ATPs:" :: map running_info active)
+    val interrupting =
+      if null cancelling then ""
+      else space_implode "\n\n"
+        ("Trying to interrupt the following ATPs:" :: map cancelling_info cancelling)
+
+  in writeln (running ^ "\n" ^ interrupting) end;
+
+fun messages opt_limit =
+  let
+    val limit = the_default message_display_limit opt_limit;
+    val State {store = msgs, ...} = Synchronized.value state
+    val header = "Recent ATP messages" ^
+      (if length msgs <= limit then ":" else " (" ^ string_of_int limit ^ " displayed):");
+  in writeln (space_implode "\n\n" (header :: #1 (chop limit msgs))) end;
+
+
+
+(** The Sledgehammer **)
+
+(* named provers *)
+
+type prover = int -> (thm * (string * int)) list option ->
+  (thm * (string * int)) list option -> string -> int ->
+  Proof.context * (thm list * thm) ->
+  bool * string * string * string vector * (thm * (string * int)) list
+
+fun err_dup_prover name = error ("Duplicate prover: " ^ quote name);
+
+structure Provers = TheoryDataFun
+(
+  type T = (prover * stamp) Symtab.table
+  val empty = Symtab.empty
+  val copy = I
+  val extend = I
+  fun merge _ tabs : T = Symtab.merge (eq_snd op =) tabs
+    handle Symtab.DUP dup => err_dup_prover dup
+);
+
+fun add_prover name prover thy =
+  Provers.map (Symtab.update_new (name, (prover, stamp ()))) thy
+    handle Symtab.DUP dup => err_dup_prover dup;
+
+fun print_provers thy = Pretty.writeln
+  (Pretty.strs ("external provers:" :: sort_strings (Symtab.keys (Provers.get thy))));
+
+fun get_prover name thy = case Symtab.lookup (Provers.get thy) name of
+  NONE => NONE
+| SOME (prover, _) => SOME prover;
+
+(* start prover thread *)
+
+fun start_prover name birthtime deadtime i proof_state =
+  (case get_prover name (Proof.theory_of proof_state) of
+    NONE => warning ("Unknown external prover: " ^ quote name)
+  | SOME prover =>
+      let
+        val (ctxt, (_, goal)) = Proof.get_goal proof_state
+        val desc =
+          "external prover " ^ quote name ^ " for subgoal " ^ string_of_int i ^ ":\n" ^
+            Syntax.string_of_term ctxt (Thm.term_of (Thm.cprem_of goal i))
+        val _ = SimpleThread.fork true (fn () =>
+          let
+            val _ = register birthtime deadtime (Thread.self (), desc)
+            val result =
+              let val (success, message, _, _, _) =
+                prover (get_timeout ()) NONE NONE name i (Proof.get_goal proof_state)
+              in (success, message) end
+              handle ResHolClause.TOO_TRIVIAL
+                => (true, "Empty clause: Try this command: " ^ Markup.markup Markup.sendback "apply metis")
+              | ERROR msg
+                => (false, "Error: " ^ msg)
+            val _ = unregister result (Thread.self ())
+          in () end handle Interrupt => ())
+      in () end);
+
+
+(* sledghammer for first subgoal *)
+
+fun sledgehammer names proof_state =
+  let
+    val provers =
+      if null names then String.tokens (Symbol.is_ascii_blank o String.str) (get_atps ())
+      else names
+    val birthtime = Time.now ()
+    val deadtime = Time.+ (birthtime, Time.fromSeconds (get_timeout ()))
+  in List.app (fn name => start_prover name birthtime deadtime 1 proof_state) provers end;
+
+
+
+(** Isar command syntax **)
+
+local structure K = OuterKeyword and P = OuterParse in
+
+val _ =
+  OuterSyntax.improper_command "atp_kill" "kill all managed provers" K.diag
+    (Scan.succeed (Toplevel.no_timing o Toplevel.imperative kill));
+
+val _ =
+  OuterSyntax.improper_command "atp_info" "print information about managed provers" K.diag
+    (Scan.succeed (Toplevel.no_timing o Toplevel.imperative info));
+
+val _ =
+  OuterSyntax.improper_command "atp_messages" "print recent messages issued by managed provers" K.diag
+    (Scan.option (P.$$$ "(" |-- P.nat --| P.$$$ ")") >>
+      (fn limit => Toplevel.no_timing o Toplevel.imperative (fn () => messages limit)));
+
+val _ =
+  OuterSyntax.improper_command "print_atps" "print external provers" K.diag
+    (Scan.succeed (Toplevel.no_timing o Toplevel.unknown_theory o
+      Toplevel.keep (print_provers o Toplevel.theory_of)));
+
+val _ =
+  OuterSyntax.command "sledgehammer" "call all automatic theorem provers" K.diag
+    (Scan.repeat P.xname >> (fn names => Toplevel.no_timing o Toplevel.unknown_proof o
+      Toplevel.keep (sledgehammer names o Toplevel.proof_of)));
+
+end;
+
+end;
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/ATP_Manager/atp_minimal.ML	Mon Aug 10 10:25:00 2009 +0200
@@ -0,0 +1,223 @@
+(*  Title:      HOL/Tools/ATP_Manager/atp_minimal.ML
+    Author:     Philipp Meyer, TU Muenchen
+
+Minimalization of theorem list for metis by using an external automated theorem prover
+*)
+
+structure AtpMinimal: sig end =
+struct
+
+(* output control *)
+
+fun debug str = Output.debug (fn () => str)
+fun debug_fn f = if ! Output.debugging then f () else ()
+fun answer str = Output.writeln str
+fun println str = Output.priority str
+
+fun order_unique name_list = OrdList.make (String.collate Char.compare) name_list
+fun length_string namelist = Int.toString (length namelist)
+
+fun print_names name_thms_pairs =
+  let
+    val names = map fst name_thms_pairs
+    val ordered = order_unique names
+  in
+    app (fn name => (debug ("  " ^ name))) ordered
+  end
+
+
+(* minimalization algorithm *)
+
+local
+  fun isplit (l,r) [] = (l,r)
+    | isplit (l,r) (h::[]) = (h::l, r)
+    | isplit (l,r) (h1::h2::t) = isplit (h1::l, h2::r) t
+in
+  fun split lst = isplit ([],[]) lst
+end
+
+local
+  fun min p sup [] = raise Empty
+    | min p sup [s0] = [s0]
+    | min p sup s =
+      let
+        val (l0, r0) = split s
+      in
+        if p (sup @ l0)
+        then min p sup l0
+        else
+          if p (sup @ r0)
+          then min p sup r0
+          else
+            let
+              val l = min p (sup @ r0) l0
+              val r = min p (sup @ l) r0
+            in
+              l @ r
+            end
+      end
+in
+  (* return a minimal subset v of s that satisfies p
+   @pre p(s) & ~p([]) & monotone(p)
+   @post v subset s & p(v) &
+         forall e in v. ~p(v \ e)
+   *)
+  fun minimal p s = min p [] s
+end
+
+
+(* failure check and producing answer *)
+
+datatype 'a prove_result = Success of 'a | Failure | Timeout | Error
+
+val string_of_result =
+  fn Success _ => "Success"
+   | Failure => "Failure"
+   | Timeout => "Timeout"
+   | Error => "Error"
+
+val failure_strings =
+  [("SPASS beiseite: Ran out of time.", Timeout),
+   ("Timeout", Timeout),
+   ("time limit exceeded", Timeout),
+   ("# Cannot determine problem status within resource limit", Timeout),
+   ("Error", Error)]
+
+fun produce_answer (success, message, result_string, thm_name_vec, filtered) =
+  if success then
+    (Success (Vector.foldr op:: [] thm_name_vec, filtered), result_string)
+  else
+    let
+      val failure = failure_strings |> get_first (fn (s, t) =>
+          if String.isSubstring s result_string then SOME (t, result_string) else NONE)
+    in
+      if is_some failure then
+        the failure
+      else
+        (Failure, result_string)
+    end
+
+
+(* wrapper for calling external prover *)
+
+fun sh_test_thms prover prover_name time_limit subgoalno state filtered name_thms_pairs =
+  let
+    val _ = println ("Testing " ^ (length_string name_thms_pairs) ^ " theorems... ")
+    val name_thm_pairs =
+      flat (map (fn (n, ths) => map_index (fn (i, th) => (n, th)) ths) name_thms_pairs)
+    val _ = debug_fn (fn () => print_names name_thm_pairs)
+    val axclauses = ResAxioms.cnf_rules_pairs (Proof.theory_of state) name_thm_pairs
+    val (result, proof) =
+      produce_answer
+        (prover time_limit (SOME axclauses) filtered prover_name subgoalno (Proof.get_goal state))
+    val _ = println (string_of_result result)
+    val _ = debug proof
+  in
+    (result, proof)
+  end
+
+
+(* minimalization of thms *)
+
+fun minimalize prover prover_name time_limit state name_thms_pairs =
+  let
+    val _ =
+      println ("Minimize called with " ^ length_string name_thms_pairs ^ " theorems, prover: "
+        ^ prover_name ^ ", time limit: " ^ Int.toString time_limit ^ " seconds")
+    val _ = debug_fn (fn () => app (fn (n, tl) =>
+        (debug n; app (fn t =>
+          debug ("  " ^ Display.string_of_thm (Proof.context_of state) t)) tl)) name_thms_pairs)
+    val test_thms_fun = sh_test_thms prover prover_name time_limit 1 state
+    fun test_thms filtered thms =
+      case test_thms_fun filtered thms of (Success _, _) => true | _ => false
+  in
+    (* try prove first to check result and get used theorems *)
+    (case test_thms_fun NONE name_thms_pairs of
+      (Success (used, filtered), _) =>
+        let
+          val ordered_used = order_unique used
+          val to_use =
+            if length ordered_used < length name_thms_pairs then
+              filter (fn (name1, _) => List.exists (equal name1) ordered_used) name_thms_pairs
+            else
+              name_thms_pairs
+          val min_thms = (minimal (test_thms (SOME filtered)) to_use)
+          val min_names = order_unique (map fst min_thms)
+          val _ = println ("Minimal " ^ (length_string min_thms) ^ " theorems")
+          val _ = debug_fn (fn () => print_names min_thms)
+        in
+          answer ("Try this command: " ^
+            Markup.markup Markup.sendback ("apply (metis " ^ space_implode " " min_names ^ ")"))
+        end
+    | (Timeout, _) =>
+        answer ("Timeout: You may need to increase the time limit of " ^
+          Int.toString time_limit ^ " seconds. Call atp_minimize [time=...] ")
+    | (Error, msg) =>
+        answer ("Error in prover: " ^ msg)
+    | (Failure, _) =>
+        answer "Failure: No proof with the theorems supplied")
+    handle ResHolClause.TOO_TRIVIAL =>
+        answer ("Trivial: Try this command: " ^ Markup.markup Markup.sendback "apply metis")
+    | ERROR msg =>
+        answer ("Error: " ^ msg)
+  end
+
+
+(* Isar command and parsing input *)
+
+local structure K = OuterKeyword and P = OuterParse and T = OuterLex in
+
+fun get_thms context =
+  map (fn (name, interval) =>
+    let
+      val thmref = Facts.Named ((name, Position.none), interval)
+      val ths = ProofContext.get_fact context thmref
+      val name' = Facts.string_of_ref thmref
+    in
+      (name', ths)
+    end)
+
+val default_prover = "remote_vampire"
+val default_time_limit = 5
+
+fun get_time_limit_arg time_string =
+  (case Int.fromString time_string of
+    SOME t => t
+  | NONE => error ("Invalid time limit: " ^ quote time_string))
+
+val get_options =
+  let
+    val def = (default_prover, default_time_limit)
+  in
+    foldl (fn ((name, a), (p, t)) =>
+      (case name of
+        "time" => (p, (get_time_limit_arg a))
+      | "atp" => (a, t)
+      | n => error ("Invalid argument: " ^ n))) def
+  end
+
+fun sh_min_command args thm_names state =
+  let
+    val (prover_name, time_limit) = get_options args
+    val prover =
+      case AtpManager.get_prover prover_name (Proof.theory_of state) of
+        SOME prover => prover
+      | NONE => error ("Unknown prover: " ^ quote prover_name)
+    val name_thms_pairs = get_thms (Proof.context_of state) thm_names
+  in
+    minimalize prover prover_name time_limit state name_thms_pairs
+  end
+
+val parse_args = Scan.optional (Args.bracks (P.list (P.xname --| P.$$$ "=" -- P.xname))) []
+val parse_thm_names = Scan.repeat (P.xname -- Scan.option Attrib.thm_sel)
+
+val _ =
+  OuterSyntax.command "atp_minimize" "minimize theorem list with external prover" K.diag
+    (parse_args -- parse_thm_names >> (fn (args, thm_names) =>
+      Toplevel.no_timing o Toplevel.unknown_proof o
+        Toplevel.keep (sh_min_command args thm_names o Toplevel.proof_of)))
+
+end
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/ATP_Manager/atp_wrapper.ML	Mon Aug 10 10:25:00 2009 +0200
@@ -0,0 +1,217 @@
+(*  Title:      HOL/Tools/ATP_Manager/atp_wrapper.ML
+    Author:     Fabian Immler, TU Muenchen
+
+Wrapper functions for external ATPs.
+*)
+
+signature ATP_WRAPPER =
+sig
+  val destdir: string ref
+  val problem_name: string ref
+  val tptp_prover_opts_full: int -> bool -> bool -> Path.T * string -> AtpManager.prover
+  val tptp_prover_opts: int -> bool -> Path.T * string -> AtpManager.prover
+  val tptp_prover: Path.T * string -> AtpManager.prover
+  val full_prover_opts: int -> bool -> Path.T * string -> AtpManager.prover
+  val full_prover: Path.T * string  -> AtpManager.prover
+  val vampire_opts: int -> bool -> AtpManager.prover
+  val vampire: AtpManager.prover
+  val vampire_opts_full: int -> bool -> AtpManager.prover
+  val vampire_full: AtpManager.prover
+  val eprover_opts: int -> bool  -> AtpManager.prover
+  val eprover: AtpManager.prover
+  val eprover_opts_full: int -> bool -> AtpManager.prover
+  val eprover_full: AtpManager.prover
+  val spass_opts: int -> bool  -> AtpManager.prover
+  val spass: AtpManager.prover
+  val remote_prover_opts: int -> bool -> string -> string -> AtpManager.prover
+  val remote_prover: string -> string -> AtpManager.prover
+  val refresh_systems: unit -> unit
+end;
+
+structure AtpWrapper: ATP_WRAPPER =
+struct
+
+(** generic ATP wrapper **)
+
+(* global hooks for writing problemfiles *)
+
+val destdir = ref "";   (*Empty means write files to /tmp*)
+val problem_name = ref "prob";
+
+
+(* basic template *)
+
+fun external_prover relevance_filter preparer writer (cmd, args) find_failure produce_answer
+  timeout axiom_clauses filtered_clauses name subgoalno goal =
+  let
+    (* path to unique problem file *)
+    val destdir' = ! destdir
+    val problem_name' = ! problem_name
+    fun prob_pathname nr =
+      let val probfile = Path.basic (problem_name' ^ serial_string () ^ "_" ^ string_of_int nr)
+      in if destdir' = "" then File.tmp_path probfile
+        else if File.exists (Path.explode (destdir'))
+        then Path.append  (Path.explode (destdir')) probfile
+        else error ("No such directory: " ^ destdir')
+      end
+
+    (* get clauses and prepare them for writing *)
+    val (ctxt, (chain_ths, th)) = goal
+    val thy = ProofContext.theory_of ctxt
+    val chain_ths = map (Thm.put_name_hint ResReconstruct.chained_hint) chain_ths
+    val goal_cls = #1 (ResAxioms.neg_conjecture_clauses ctxt th subgoalno)
+    val _ = app (fn th => Output.debug (fn _ => Display.string_of_thm ctxt th)) goal_cls
+    val the_filtered_clauses =
+      case filtered_clauses of
+          NONE => relevance_filter goal goal_cls
+        | SOME fcls => fcls
+    val the_axiom_clauses =
+      case axiom_clauses of
+          NONE => the_filtered_clauses
+        | SOME axcls => axcls
+    val (thm_names, clauses) =
+      preparer goal_cls chain_ths the_axiom_clauses the_filtered_clauses thy
+
+    (* write out problem file and call prover *)
+    val probfile = prob_pathname subgoalno
+    val conj_pos = writer probfile clauses
+    val (proof, rc) = system_out (
+      if File.exists cmd then
+        space_implode " " ["exec", File.shell_path cmd, args, File.platform_path probfile]
+      else error ("Bad executable: " ^ Path.implode cmd))
+
+    (* if problemfile has not been exported, delete problemfile; otherwise export proof, too *)
+    val _ =
+      if destdir' = "" then File.rm probfile
+      else File.write (Path.explode (Path.implode probfile ^ "_proof")) proof
+
+    (* check for success and print out some information on failure *)
+    val failure = find_failure proof
+    val success = rc = 0 andalso is_none failure
+    val message =
+      if is_some failure then "External prover failed."
+      else if rc <> 0 then "External prover failed: " ^ proof
+      else "Try this command: " ^
+        produce_answer name (proof, thm_names, conj_pos, ctxt, th, subgoalno)
+    val _ = Output.debug (fn () => "Sledgehammer response (rc = " ^ string_of_int rc ^ "):\n" ^ proof)
+  in (success, message, proof, thm_names, the_filtered_clauses) end;
+
+
+
+(** common provers **)
+
+(* generic TPTP-based provers *)
+
+fun tptp_prover_opts_full max_new theory_const full command timeout ax_clauses fcls name n goal =
+  external_prover
+  (ResAtp.get_relevant max_new theory_const)
+  (ResAtp.prepare_clauses false)
+  (ResHolClause.tptp_write_file (AtpManager.get_full_types()))
+  command
+  ResReconstruct.find_failure
+  (if full then ResReconstruct.structured_proof else ResReconstruct.lemma_list false)
+  timeout ax_clauses fcls name n goal;
+
+(*arbitrary ATP with TPTP input/output and problemfile as last argument*)
+fun tptp_prover_opts max_new theory_const =
+  tptp_prover_opts_full max_new theory_const false;
+
+fun tptp_prover x = tptp_prover_opts 60 true x;
+
+(*for structured proofs: prover must support TSTP*)
+fun full_prover_opts max_new theory_const =
+  tptp_prover_opts_full max_new theory_const true;
+
+fun full_prover x = full_prover_opts 60 true x;
+
+
+(* Vampire *)
+
+(*NB: Vampire does not work without explicit timelimit*)
+
+fun vampire_opts max_new theory_const timeout = tptp_prover_opts
+  max_new theory_const
+  (Path.explode "$VAMPIRE_HOME/vampire",
+    ("--output_syntax tptp --mode casc -t " ^ string_of_int timeout))
+  timeout;
+
+val vampire = vampire_opts 60 false;
+
+fun vampire_opts_full max_new theory_const timeout = full_prover_opts
+  max_new theory_const
+  (Path.explode "$VAMPIRE_HOME/vampire",
+    ("--output_syntax tptp --mode casc -t " ^ string_of_int timeout))
+  timeout;
+
+val vampire_full = vampire_opts_full 60 false;
+
+
+(* E prover *)
+
+fun eprover_opts max_new theory_const timeout = tptp_prover_opts
+  max_new theory_const
+  (Path.explode "$E_HOME/eproof",
+    "--tstp-in --tstp-out -l5 -xAutoDev -tAutoDev --silent --cpu-limit=" ^ string_of_int timeout)
+  timeout;
+
+val eprover = eprover_opts 100 false;
+
+fun eprover_opts_full max_new theory_const timeout = full_prover_opts
+  max_new theory_const
+  (Path.explode "$E_HOME/eproof",
+    "--tstp-in --tstp-out -l5 -xAutoDev -tAutoDev --silent --cpu-limit=" ^ string_of_int timeout)
+  timeout;
+
+val eprover_full = eprover_opts_full 100 false;
+
+
+(* SPASS *)
+
+fun spass_opts max_new theory_const timeout ax_clauses fcls name n goal = external_prover
+  (ResAtp.get_relevant max_new theory_const)
+  (ResAtp.prepare_clauses true)
+  (ResHolClause.dfg_write_file (AtpManager.get_full_types()))
+  (Path.explode "$SPASS_HOME/SPASS",
+    "-Auto -SOS=1 -PGiven=0 -PProblem=0 -Splits=0 -FullRed=0 -DocProof -TimeLimit=" ^
+      string_of_int timeout)
+  ResReconstruct.find_failure
+  (ResReconstruct.lemma_list true)
+  timeout ax_clauses fcls name n goal;
+
+val spass = spass_opts 40 true;
+
+
+(* remote prover invocation via SystemOnTPTP *)
+
+val systems =
+  Synchronized.var "atp_wrapper_systems" ([]: string list);
+
+fun get_systems () =
+  let
+    val (answer, rc) = system_out ("\"$ISABELLE_ATP_MANAGER/SystemOnTPTP\" -w")
+  in
+    if rc <> 0 then error ("Failed to get available systems from SystemOnTPTP:\n" ^ answer)
+    else split_lines answer
+  end;
+
+fun refresh_systems () = Synchronized.change systems (fn _ =>
+  get_systems ());
+
+fun get_system prefix = Synchronized.change_result systems (fn systems =>
+  let val systems = if null systems then get_systems() else systems
+  in (find_first (String.isPrefix prefix) systems, systems) end);
+
+fun remote_prover_opts max_new theory_const args prover_prefix timeout =
+  let val sys =
+    case get_system prover_prefix of
+      NONE => error ("No system like " ^ quote prover_prefix ^ " at SystemOnTPTP")
+    | SOME sys => sys
+  in tptp_prover_opts max_new theory_const
+    (Path.explode "$ISABELLE_ATP_MANAGER/SystemOnTPTP",
+      args ^ " -t " ^ string_of_int timeout ^ " -s " ^ sys) timeout
+  end;
+
+val remote_prover = remote_prover_opts 60 false;
+
+end;
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/ATP_Manager/etc/settings	Mon Aug 10 10:25:00 2009 +0200
@@ -0,0 +1,2 @@
+ISABELLE_ATP_MANAGER="$COMPONENT"
+
--- a/src/HOL/Tools/atp_manager.ML	Mon Aug 10 08:37:37 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,403 +0,0 @@
-(*  Title:      HOL/Tools/atp_manager.ML
-    Author:     Fabian Immler, TU Muenchen
-
-ATP threads are registered here.
-Threads with the same birth-time are seen as one group.
-All threads of a group are killed when one thread of it has been successful,
-or after a certain time,
-or when the maximum number of threads exceeds; then the oldest thread is killed.
-*)
-
-signature ATP_MANAGER =
-sig
-  val get_atps: unit -> string
-  val set_atps: string -> unit
-  val get_max_atps: unit -> int
-  val set_max_atps: int -> unit
-  val get_timeout: unit -> int
-  val set_timeout: int -> unit
-  val get_full_types: unit -> bool
-  val set_full_types: bool -> unit
-  val kill: unit -> unit
-  val info: unit -> unit
-  val messages: int option -> unit
-  type prover = int -> (thm * (string * int)) list option ->
-    (thm * (string * int)) list option -> string -> int ->
-    Proof.context * (thm list * thm) ->
-    bool * string * string * string vector * (thm * (string * int)) list
-  val add_prover: string -> prover -> theory -> theory
-  val print_provers: theory -> unit
-  val get_prover: string -> theory -> prover option
-  val sledgehammer: string list -> Proof.state -> unit
-end;
-
-structure AtpManager: ATP_MANAGER =
-struct
-
-(** preferences **)
-
-val message_store_limit = 20;
-val message_display_limit = 5;
-
-local
-
-val atps = ref "e remote_vampire";
-val max_atps = ref 5;   (* ~1 means infinite number of atps *)
-val timeout = ref 60;
-val full_types = ref false;
-
-in
-
-fun get_atps () = CRITICAL (fn () => ! atps);
-fun set_atps str = CRITICAL (fn () => atps := str);
-
-fun get_max_atps () = CRITICAL (fn () => ! max_atps);
-fun set_max_atps number = CRITICAL (fn () => max_atps := number);
-
-fun get_timeout () = CRITICAL (fn () => ! timeout);
-fun set_timeout time = CRITICAL (fn () => timeout := time);
-
-fun get_full_types () = CRITICAL (fn () => ! full_types);
-fun set_full_types bool = CRITICAL (fn () => full_types := bool);
-
-val _ =
-  ProofGeneralPgip.add_preference Preferences.category_proof
-    (Preferences.string_pref atps
-      "ATP: provers" "Default automatic provers (separated by whitespace)");
-
-val _ =
-  ProofGeneralPgip.add_preference Preferences.category_proof
-    (Preferences.int_pref max_atps
-      "ATP: maximum number" "How many provers may run in parallel");
-
-val _ =
-  ProofGeneralPgip.add_preference Preferences.category_proof
-    (Preferences.int_pref timeout
-      "ATP: timeout" "ATPs will be interrupted after this time (in seconds)");
-
-val _ =
-  ProofGeneralPgip.add_preference Preferences.category_proof
-    (Preferences.bool_pref full_types
-      "ATP: full types" "ATPs will use full type information");
-
-end;
-
-
-
-(** thread management **)
-
-(* data structures over threads *)
-
-structure ThreadHeap = HeapFun
-(
-  type elem = Time.time * Thread.thread;
-  fun ord ((a, _), (b, _)) = Time.compare (a, b);
-);
-
-fun lookup_thread xs = AList.lookup Thread.equal xs;
-fun delete_thread xs = AList.delete Thread.equal xs;
-fun update_thread xs = AList.update Thread.equal xs;
-
-
-(* state of thread manager *)
-
-datatype T = State of
- {managing_thread: Thread.thread option,
-  timeout_heap: ThreadHeap.T,
-  oldest_heap: ThreadHeap.T,
-  active: (Thread.thread * (Time.time * Time.time * string)) list,
-  cancelling: (Thread.thread * (Time.time * Time.time * string)) list,
-  messages: string list,
-  store: string list};
-
-fun make_state managing_thread timeout_heap oldest_heap active cancelling messages store =
-  State {managing_thread = managing_thread, timeout_heap = timeout_heap, oldest_heap = oldest_heap,
-    active = active, cancelling = cancelling, messages = messages, store = store};
-
-val state = Synchronized.var "atp_manager"
-  (make_state NONE ThreadHeap.empty ThreadHeap.empty [] [] [] []);
-
-
-(* unregister thread *)
-
-fun unregister (success, message) thread = Synchronized.change state
-  (fn state as State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
-    (case lookup_thread active thread of
-      SOME (birthtime, _, description) =>
-        let
-          val (group, active') =
-            if success then List.partition (fn (_, (tb, _, _)) => tb = birthtime) active
-            else List.partition (fn (th, _) => Thread.equal (th, thread)) active
-
-          val now = Time.now ()
-          val cancelling' =
-            fold (fn (th, (tb, _, desc)) => update_thread (th, (tb, now, desc))) group cancelling
-
-          val message' = description ^ "\n" ^ message ^
-            (if length group <= 1 then ""
-             else "\nInterrupted " ^ string_of_int (length group - 1) ^ " other group members")
-          val store' = message' ::
-            (if length store <= message_store_limit then store
-             else #1 (chop message_store_limit store))
-        in make_state
-          managing_thread timeout_heap oldest_heap active' cancelling' (message' :: messages) store'
-        end
-    | NONE => state));
-
-
-(* kill excessive atp threads *)
-
-fun excessive_atps active =
-  let val max = get_max_atps ()
-  in length active > max andalso max > ~1 end;
-
-local
-
-fun kill_oldest () =
-  let exception Unchanged in
-    Synchronized.change_result state
-      (fn State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
-        if ThreadHeap.is_empty oldest_heap orelse not (excessive_atps active)
-        then raise Unchanged
-        else
-          let val ((_, oldest_thread), oldest_heap') = ThreadHeap.min_elem oldest_heap
-          in (oldest_thread,
-          make_state managing_thread timeout_heap oldest_heap' active cancelling messages store) end)
-      |> unregister (false, "Interrupted (maximum number of ATPs exceeded)")
-    handle Unchanged => ()
-  end;
-
-in
-
-fun kill_excessive () =
-  let val State {active, ...} = Synchronized.value state
-  in if excessive_atps active then (kill_oldest (); kill_excessive ()) else () end;
-
-end;
-
-fun print_new_messages () =
-  let val to_print = Synchronized.change_result state
-    (fn State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
-      (messages, make_state managing_thread timeout_heap oldest_heap active cancelling [] store))
-  in
-    if null to_print then ()
-    else priority ("Sledgehammer: " ^ space_implode "\n\n" to_print)
-  end;
-
-
-(* start a watching thread -- only one may exist *)
-
-fun check_thread_manager () = Synchronized.change state
-  (fn State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
-    if (case managing_thread of SOME thread => Thread.isActive thread | NONE => false)
-    then make_state managing_thread timeout_heap oldest_heap active cancelling messages store
-    else let val managing_thread = SOME (SimpleThread.fork false (fn () =>
-      let
-        val min_wait_time = Time.fromMilliseconds 300
-        val max_wait_time = Time.fromSeconds 10
-
-        (* wait for next thread to cancel, or maximum*)
-        fun time_limit (State {timeout_heap, ...}) =
-          (case try ThreadHeap.min timeout_heap of
-            NONE => SOME (Time.+ (Time.now (), max_wait_time))
-          | SOME (time, _) => SOME time)
-
-        (* action: find threads whose timeout is reached, and interrupt cancelling threads *)
-        fun action (State {managing_thread, timeout_heap, oldest_heap, active, cancelling,
-                           messages, store}) =
-          let val (timeout_threads, timeout_heap') =
-            ThreadHeap.upto (Time.now (), Thread.self ()) timeout_heap
-          in
-            if null timeout_threads andalso null cancelling andalso not (excessive_atps active)
-            then NONE
-            else
-              let
-                val _ = List.app (SimpleThread.interrupt o #1) cancelling
-                val cancelling' = filter (Thread.isActive o #1) cancelling
-                val state' = make_state
-                  managing_thread timeout_heap' oldest_heap active cancelling' messages store
-              in SOME (map #2 timeout_threads, state') end
-          end
-      in
-        while Synchronized.change_result state
-          (fn st as
-            State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
-            if (null active) andalso (null cancelling) andalso (null messages)
-            then (false, make_state NONE timeout_heap oldest_heap active cancelling messages store)
-            else (true, st))
-        do
-          (Synchronized.timed_access state time_limit action
-            |> these
-            |> List.app (unregister (false, "Interrupted (reached timeout)"));
-            kill_excessive ();
-            print_new_messages ();
-            (*give threads time to respond to interrupt*)
-            OS.Process.sleep min_wait_time)
-      end))
-    in make_state managing_thread timeout_heap oldest_heap active cancelling messages store end);
-
-
-(* thread is registered here by sledgehammer *)
-
-fun register birthtime deadtime (thread, desc) =
- (Synchronized.change state
-    (fn State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
-      let
-        val timeout_heap' = ThreadHeap.insert (deadtime, thread) timeout_heap
-        val oldest_heap' = ThreadHeap.insert (birthtime, thread) oldest_heap
-        val active' = update_thread (thread, (birthtime, deadtime, desc)) active
-      in make_state managing_thread timeout_heap' oldest_heap' active' cancelling messages store end);
-  check_thread_manager ());
-
-
-
-(** user commands **)
-
-(* kill: move all threads to cancelling *)
-
-fun kill () = Synchronized.change state
-  (fn State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
-    let val formerly_active = map (fn (th, (tb, _, desc)) => (th, (tb, Time.now (), desc))) active
-    in make_state
-      managing_thread timeout_heap oldest_heap [] (formerly_active @ cancelling) messages store
-    end);
-
-
-(* ATP info *)
-
-fun info () =
-  let
-    val State {active, cancelling, ...} = Synchronized.value state
-
-    fun running_info (_, (birth_time, dead_time, desc)) = "Running: "
-        ^ (string_of_int o Time.toSeconds) (Time.- (Time.now (), birth_time))
-        ^ " s  --  "
-        ^ (string_of_int o Time.toSeconds) (Time.- (dead_time, Time.now ()))
-        ^ " s to live:\n" ^ desc
-    fun cancelling_info (_, (_, dead_time, desc)) = "Trying to interrupt thread since "
-        ^ (string_of_int o Time.toSeconds) (Time.- (Time.now (), dead_time))
-        ^ " s:\n" ^ desc
-
-    val running =
-      if null active then "No ATPs running."
-      else space_implode "\n\n" ("Running ATPs:" :: map running_info active)
-    val interrupting =
-      if null cancelling then ""
-      else space_implode "\n\n"
-        ("Trying to interrupt the following ATPs:" :: map cancelling_info cancelling)
-
-  in writeln (running ^ "\n" ^ interrupting) end;
-
-fun messages opt_limit =
-  let
-    val limit = the_default message_display_limit opt_limit;
-    val State {store = msgs, ...} = Synchronized.value state
-    val header = "Recent ATP messages" ^
-      (if length msgs <= limit then ":" else " (" ^ string_of_int limit ^ " displayed):");
-  in writeln (space_implode "\n\n" (header :: #1 (chop limit msgs))) end;
-
-
-
-(** The Sledgehammer **)
-
-(* named provers *)
-
-type prover = int -> (thm * (string * int)) list option ->
-  (thm * (string * int)) list option -> string -> int ->
-  Proof.context * (thm list * thm) ->
-  bool * string * string * string vector * (thm * (string * int)) list
-
-fun err_dup_prover name = error ("Duplicate prover: " ^ quote name);
-
-structure Provers = TheoryDataFun
-(
-  type T = (prover * stamp) Symtab.table
-  val empty = Symtab.empty
-  val copy = I
-  val extend = I
-  fun merge _ tabs : T = Symtab.merge (eq_snd op =) tabs
-    handle Symtab.DUP dup => err_dup_prover dup
-);
-
-fun add_prover name prover thy =
-  Provers.map (Symtab.update_new (name, (prover, stamp ()))) thy
-    handle Symtab.DUP dup => err_dup_prover dup;
-
-fun print_provers thy = Pretty.writeln
-  (Pretty.strs ("external provers:" :: sort_strings (Symtab.keys (Provers.get thy))));
-
-fun get_prover name thy = case Symtab.lookup (Provers.get thy) name of
-  NONE => NONE
-| SOME (prover, _) => SOME prover;
-
-(* start prover thread *)
-
-fun start_prover name birthtime deadtime i proof_state =
-  (case get_prover name (Proof.theory_of proof_state) of
-    NONE => warning ("Unknown external prover: " ^ quote name)
-  | SOME prover =>
-      let
-        val (ctxt, (_, goal)) = Proof.get_goal proof_state
-        val desc =
-          "external prover " ^ quote name ^ " for subgoal " ^ string_of_int i ^ ":\n" ^
-            Syntax.string_of_term ctxt (Thm.term_of (Thm.cprem_of goal i))
-        val _ = SimpleThread.fork true (fn () =>
-          let
-            val _ = register birthtime deadtime (Thread.self (), desc)
-            val result =
-              let val (success, message, _, _, _) =
-                prover (get_timeout ()) NONE NONE name i (Proof.get_goal proof_state)
-              in (success, message) end
-              handle ResHolClause.TOO_TRIVIAL
-                => (true, "Empty clause: Try this command: " ^ Markup.markup Markup.sendback "apply metis")
-              | ERROR msg
-                => (false, "Error: " ^ msg)
-            val _ = unregister result (Thread.self ())
-          in () end handle Interrupt => ())
-      in () end);
-
-
-(* sledghammer for first subgoal *)
-
-fun sledgehammer names proof_state =
-  let
-    val provers =
-      if null names then String.tokens (Symbol.is_ascii_blank o String.str) (get_atps ())
-      else names
-    val birthtime = Time.now ()
-    val deadtime = Time.+ (birthtime, Time.fromSeconds (get_timeout ()))
-  in List.app (fn name => start_prover name birthtime deadtime 1 proof_state) provers end;
-
-
-
-(** Isar command syntax **)
-
-local structure K = OuterKeyword and P = OuterParse in
-
-val _ =
-  OuterSyntax.improper_command "atp_kill" "kill all managed provers" K.diag
-    (Scan.succeed (Toplevel.no_timing o Toplevel.imperative kill));
-
-val _ =
-  OuterSyntax.improper_command "atp_info" "print information about managed provers" K.diag
-    (Scan.succeed (Toplevel.no_timing o Toplevel.imperative info));
-
-val _ =
-  OuterSyntax.improper_command "atp_messages" "print recent messages issued by managed provers" K.diag
-    (Scan.option (P.$$$ "(" |-- P.nat --| P.$$$ ")") >>
-      (fn limit => Toplevel.no_timing o Toplevel.imperative (fn () => messages limit)));
-
-val _ =
-  OuterSyntax.improper_command "print_atps" "print external provers" K.diag
-    (Scan.succeed (Toplevel.no_timing o Toplevel.unknown_theory o
-      Toplevel.keep (print_provers o Toplevel.theory_of)));
-
-val _ =
-  OuterSyntax.command "sledgehammer" "call all automatic theorem provers" K.diag
-    (Scan.repeat P.xname >> (fn names => Toplevel.no_timing o Toplevel.unknown_proof o
-      Toplevel.keep (sledgehammer names o Toplevel.proof_of)));
-
-end;
-
-end;
-
--- a/src/HOL/Tools/atp_minimal.ML	Mon Aug 10 08:37:37 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,223 +0,0 @@
-(*  Title:      HOL/Tools/atp_minimal.ML
-    Author:     Philipp Meyer, TU Muenchen
-
-Minimalization of theorem list for metis by using an external automated theorem prover
-*)
-
-structure AtpMinimal: sig end =
-struct
-
-(* output control *)
-
-fun debug str = Output.debug (fn () => str)
-fun debug_fn f = if ! Output.debugging then f () else ()
-fun answer str = Output.writeln str
-fun println str = Output.priority str
-
-fun order_unique name_list = OrdList.make (String.collate Char.compare) name_list
-fun length_string namelist = Int.toString (length namelist)
-
-fun print_names name_thms_pairs =
-  let
-    val names = map fst name_thms_pairs
-    val ordered = order_unique names
-  in
-    app (fn name => (debug ("  " ^ name))) ordered
-  end
-
-
-(* minimalization algorithm *)
-
-local
-  fun isplit (l,r) [] = (l,r)
-    | isplit (l,r) (h::[]) = (h::l, r)
-    | isplit (l,r) (h1::h2::t) = isplit (h1::l, h2::r) t
-in
-  fun split lst = isplit ([],[]) lst
-end
-
-local
-  fun min p sup [] = raise Empty
-    | min p sup [s0] = [s0]
-    | min p sup s =
-      let
-        val (l0, r0) = split s
-      in
-        if p (sup @ l0)
-        then min p sup l0
-        else
-          if p (sup @ r0)
-          then min p sup r0
-          else
-            let
-              val l = min p (sup @ r0) l0
-              val r = min p (sup @ l) r0
-            in
-              l @ r
-            end
-      end
-in
-  (* return a minimal subset v of s that satisfies p
-   @pre p(s) & ~p([]) & monotone(p)
-   @post v subset s & p(v) &
-         forall e in v. ~p(v \ e)
-   *)
-  fun minimal p s = min p [] s
-end
-
-
-(* failure check and producing answer *)
-
-datatype 'a prove_result = Success of 'a | Failure | Timeout | Error
-
-val string_of_result =
-  fn Success _ => "Success"
-   | Failure => "Failure"
-   | Timeout => "Timeout"
-   | Error => "Error"
-
-val failure_strings =
-  [("SPASS beiseite: Ran out of time.", Timeout),
-   ("Timeout", Timeout),
-   ("time limit exceeded", Timeout),
-   ("# Cannot determine problem status within resource limit", Timeout),
-   ("Error", Error)]
-
-fun produce_answer (success, message, result_string, thm_name_vec, filtered) =
-  if success then
-    (Success (Vector.foldr op:: [] thm_name_vec, filtered), result_string)
-  else
-    let
-      val failure = failure_strings |> get_first (fn (s, t) =>
-          if String.isSubstring s result_string then SOME (t, result_string) else NONE)
-    in
-      if is_some failure then
-        the failure
-      else
-        (Failure, result_string)
-    end
-
-
-(* wrapper for calling external prover *)
-
-fun sh_test_thms prover prover_name time_limit subgoalno state filtered name_thms_pairs =
-  let
-    val _ = println ("Testing " ^ (length_string name_thms_pairs) ^ " theorems... ")
-    val name_thm_pairs =
-      flat (map (fn (n, ths) => map_index (fn (i, th) => (n, th)) ths) name_thms_pairs)
-    val _ = debug_fn (fn () => print_names name_thm_pairs)
-    val axclauses = ResAxioms.cnf_rules_pairs (Proof.theory_of state) name_thm_pairs
-    val (result, proof) =
-      produce_answer
-        (prover time_limit (SOME axclauses) filtered prover_name subgoalno (Proof.get_goal state))
-    val _ = println (string_of_result result)
-    val _ = debug proof
-  in
-    (result, proof)
-  end
-
-
-(* minimalization of thms *)
-
-fun minimalize prover prover_name time_limit state name_thms_pairs =
-  let
-    val _ =
-      println ("Minimize called with " ^ length_string name_thms_pairs ^ " theorems, prover: "
-        ^ prover_name ^ ", time limit: " ^ Int.toString time_limit ^ " seconds")
-    val _ = debug_fn (fn () => app (fn (n, tl) =>
-        (debug n; app (fn t =>
-          debug ("  " ^ Display.string_of_thm (Proof.context_of state) t)) tl)) name_thms_pairs)
-    val test_thms_fun = sh_test_thms prover prover_name time_limit 1 state
-    fun test_thms filtered thms =
-      case test_thms_fun filtered thms of (Success _, _) => true | _ => false
-  in
-    (* try prove first to check result and get used theorems *)
-    (case test_thms_fun NONE name_thms_pairs of
-      (Success (used, filtered), _) =>
-        let
-          val ordered_used = order_unique used
-          val to_use =
-            if length ordered_used < length name_thms_pairs then
-              filter (fn (name1, _) => List.exists (equal name1) ordered_used) name_thms_pairs
-            else
-              name_thms_pairs
-          val min_thms = (minimal (test_thms (SOME filtered)) to_use)
-          val min_names = order_unique (map fst min_thms)
-          val _ = println ("Minimal " ^ (length_string min_thms) ^ " theorems")
-          val _ = debug_fn (fn () => print_names min_thms)
-        in
-          answer ("Try this command: " ^
-            Markup.markup Markup.sendback ("apply (metis " ^ space_implode " " min_names ^ ")"))
-        end
-    | (Timeout, _) =>
-        answer ("Timeout: You may need to increase the time limit of " ^
-          Int.toString time_limit ^ " seconds. Call atp_minimize [time=...] ")
-    | (Error, msg) =>
-        answer ("Error in prover: " ^ msg)
-    | (Failure, _) =>
-        answer "Failure: No proof with the theorems supplied")
-    handle ResHolClause.TOO_TRIVIAL =>
-        answer ("Trivial: Try this command: " ^ Markup.markup Markup.sendback "apply metis")
-    | ERROR msg =>
-        answer ("Error: " ^ msg)
-  end
-
-
-(* Isar command and parsing input *)
-
-local structure K = OuterKeyword and P = OuterParse and T = OuterLex in
-
-fun get_thms context =
-  map (fn (name, interval) =>
-    let
-      val thmref = Facts.Named ((name, Position.none), interval)
-      val ths = ProofContext.get_fact context thmref
-      val name' = Facts.string_of_ref thmref
-    in
-      (name', ths)
-    end)
-
-val default_prover = "remote_vampire"
-val default_time_limit = 5
-
-fun get_time_limit_arg time_string =
-  (case Int.fromString time_string of
-    SOME t => t
-  | NONE => error ("Invalid time limit: " ^ quote time_string))
-
-val get_options =
-  let
-    val def = (default_prover, default_time_limit)
-  in
-    foldl (fn ((name, a), (p, t)) =>
-      (case name of
-        "time" => (p, (get_time_limit_arg a))
-      | "atp" => (a, t)
-      | n => error ("Invalid argument: " ^ n))) def
-  end
-
-fun sh_min_command args thm_names state =
-  let
-    val (prover_name, time_limit) = get_options args
-    val prover =
-      case AtpManager.get_prover prover_name (Proof.theory_of state) of
-        SOME prover => prover
-      | NONE => error ("Unknown prover: " ^ quote prover_name)
-    val name_thms_pairs = get_thms (Proof.context_of state) thm_names
-  in
-    minimalize prover prover_name time_limit state name_thms_pairs
-  end
-
-val parse_args = Scan.optional (Args.bracks (P.list (P.xname --| P.$$$ "=" -- P.xname))) []
-val parse_thm_names = Scan.repeat (P.xname -- Scan.option Attrib.thm_sel)
-
-val _ =
-  OuterSyntax.command "atp_minimize" "minimize theorem list with external prover" K.diag
-    (parse_args -- parse_thm_names >> (fn (args, thm_names) =>
-      Toplevel.no_timing o Toplevel.unknown_proof o
-        Toplevel.keep (sh_min_command args thm_names o Toplevel.proof_of)))
-
-end
-
-end
-
--- a/src/HOL/Tools/atp_wrapper.ML	Mon Aug 10 08:37:37 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,218 +0,0 @@
-(*  Title:      HOL/Tools/atp_wrapper.ML
-    Author:     Fabian Immler, TU Muenchen
-
-Wrapper functions for external ATPs.
-*)
-
-signature ATP_WRAPPER =
-sig
-  val destdir: string ref
-  val problem_name: string ref
-  val tptp_prover_opts_full: int -> bool -> bool -> Path.T * string -> AtpManager.prover
-  val tptp_prover_opts: int -> bool -> Path.T * string -> AtpManager.prover
-  val tptp_prover: Path.T * string -> AtpManager.prover
-  val full_prover_opts: int -> bool -> Path.T * string -> AtpManager.prover
-  val full_prover: Path.T * string  -> AtpManager.prover
-  val vampire_opts: int -> bool -> AtpManager.prover
-  val vampire: AtpManager.prover
-  val vampire_opts_full: int -> bool -> AtpManager.prover
-  val vampire_full: AtpManager.prover
-  val eprover_opts: int -> bool  -> AtpManager.prover
-  val eprover: AtpManager.prover
-  val eprover_opts_full: int -> bool -> AtpManager.prover
-  val eprover_full: AtpManager.prover
-  val spass_opts: int -> bool  -> AtpManager.prover
-  val spass: AtpManager.prover
-  val remote_prover_opts: int -> bool -> string -> string -> AtpManager.prover
-  val remote_prover: string -> string -> AtpManager.prover
-  val refresh_systems: unit -> unit
-end;
-
-structure AtpWrapper: ATP_WRAPPER =
-struct
-
-(** generic ATP wrapper **)
-
-(* global hooks for writing problemfiles *)
-
-val destdir = ref "";   (*Empty means write files to /tmp*)
-val problem_name = ref "prob";
-
-
-(* basic template *)
-
-fun external_prover relevance_filter preparer writer (cmd, args) find_failure produce_answer
-  timeout axiom_clauses filtered_clauses name subgoalno goal =
-  let
-    (* path to unique problem file *)
-    val destdir' = ! destdir
-    val problem_name' = ! problem_name
-    fun prob_pathname nr =
-      let val probfile = Path.basic (problem_name' ^ serial_string () ^ "_" ^ string_of_int nr)
-      in if destdir' = "" then File.tmp_path probfile
-        else if File.exists (Path.explode (destdir'))
-        then Path.append  (Path.explode (destdir')) probfile
-        else error ("No such directory: " ^ destdir')
-      end
-
-    (* get clauses and prepare them for writing *)
-    val (ctxt, (chain_ths, th)) = goal
-    val thy = ProofContext.theory_of ctxt
-    val chain_ths = map (Thm.put_name_hint ResReconstruct.chained_hint) chain_ths
-    val goal_cls = #1 (ResAxioms.neg_conjecture_clauses ctxt th subgoalno)
-    val _ = app (fn th => Output.debug (fn _ => Display.string_of_thm ctxt th)) goal_cls
-    val the_filtered_clauses =
-      case filtered_clauses of
-          NONE => relevance_filter goal goal_cls
-        | SOME fcls => fcls
-    val the_axiom_clauses =
-      case axiom_clauses of
-          NONE => the_filtered_clauses
-        | SOME axcls => axcls
-    val (thm_names, clauses) =
-      preparer goal_cls chain_ths the_axiom_clauses the_filtered_clauses thy
-
-    (* write out problem file and call prover *)
-    val probfile = prob_pathname subgoalno
-    val conj_pos = writer probfile clauses
-    val (proof, rc) = system_out (
-      if File.exists cmd then
-        space_implode " " ["exec", File.shell_path cmd, args, File.platform_path probfile]
-      else error ("Bad executable: " ^ Path.implode cmd))
-
-    (* if problemfile has not been exported, delete problemfile; otherwise export proof, too *)
-    val _ =
-      if destdir' = "" then File.rm probfile
-      else File.write (Path.explode (Path.implode probfile ^ "_proof")) proof
-
-    (* check for success and print out some information on failure *)
-    val failure = find_failure proof
-    val success = rc = 0 andalso is_none failure
-    val message =
-      if is_some failure then "External prover failed."
-      else if rc <> 0 then "External prover failed: " ^ proof
-      else "Try this command: " ^
-        produce_answer name (proof, thm_names, conj_pos, ctxt, th, subgoalno)
-    val _ = Output.debug (fn () => "Sledgehammer response (rc = " ^ string_of_int rc ^ "):\n" ^ proof)
-  in (success, message, proof, thm_names, the_filtered_clauses) end;
-
-
-
-(** common provers **)
-
-(* generic TPTP-based provers *)
-
-fun tptp_prover_opts_full max_new theory_const full command timeout ax_clauses fcls name n goal =
-  external_prover
-  (ResAtp.get_relevant max_new theory_const)
-  (ResAtp.prepare_clauses false)
-  (ResHolClause.tptp_write_file (AtpManager.get_full_types()))
-  command
-  ResReconstruct.find_failure
-  (if full then ResReconstruct.structured_proof else ResReconstruct.lemma_list false)
-  timeout ax_clauses fcls name n goal;
-
-(*arbitrary ATP with TPTP input/output and problemfile as last argument*)
-fun tptp_prover_opts max_new theory_const =
-  tptp_prover_opts_full max_new theory_const false;
-
-fun tptp_prover x = tptp_prover_opts 60 true x;
-
-(*for structured proofs: prover must support TSTP*)
-fun full_prover_opts max_new theory_const =
-  tptp_prover_opts_full max_new theory_const true;
-
-fun full_prover x = full_prover_opts 60 true x;
-
-
-(* Vampire *)
-
-(*NB: Vampire does not work without explicit timelimit*)
-
-fun vampire_opts max_new theory_const timeout = tptp_prover_opts
-  max_new theory_const
-  (Path.explode "$VAMPIRE_HOME/vampire",
-    ("--output_syntax tptp --mode casc -t " ^ string_of_int timeout))
-  timeout;
-
-val vampire = vampire_opts 60 false;
-
-fun vampire_opts_full max_new theory_const timeout = full_prover_opts
-  max_new theory_const
-  (Path.explode "$VAMPIRE_HOME/vampire",
-    ("--output_syntax tptp --mode casc -t " ^ string_of_int timeout))
-  timeout;
-
-val vampire_full = vampire_opts_full 60 false;
-
-
-(* E prover *)
-
-fun eprover_opts max_new theory_const timeout = tptp_prover_opts
-  max_new theory_const
-  (Path.explode "$E_HOME/eproof",
-    "--tstp-in --tstp-out -l5 -xAutoDev -tAutoDev --silent --cpu-limit=" ^ string_of_int timeout)
-  timeout;
-
-val eprover = eprover_opts 100 false;
-
-fun eprover_opts_full max_new theory_const timeout = full_prover_opts
-  max_new theory_const
-  (Path.explode "$E_HOME/eproof",
-    "--tstp-in --tstp-out -l5 -xAutoDev -tAutoDev --silent --cpu-limit=" ^ string_of_int timeout)
-  timeout;
-
-val eprover_full = eprover_opts_full 100 false;
-
-
-(* SPASS *)
-
-fun spass_opts max_new theory_const timeout ax_clauses fcls name n goal = external_prover
-  (ResAtp.get_relevant max_new theory_const)
-  (ResAtp.prepare_clauses true)
-  (ResHolClause.dfg_write_file (AtpManager.get_full_types()))
-  (Path.explode "$SPASS_HOME/SPASS",
-    "-Auto -SOS=1 -PGiven=0 -PProblem=0 -Splits=0 -FullRed=0 -DocProof -TimeLimit=" ^
-      string_of_int timeout)
-  ResReconstruct.find_failure
-  (ResReconstruct.lemma_list true)
-  timeout ax_clauses fcls name n goal;
-
-val spass = spass_opts 40 true;
-
-
-(* remote prover invocation via SystemOnTPTP *)
-
-val systems =
-  Synchronized.var "atp_wrapper_systems" ([]: string list);
-
-fun get_systems () =
-  let
-    val (answer, rc) = system_out (("$ISABELLE_HOME/lib/scripts/SystemOnTPTP" |>
-      Path.explode |> File.shell_path) ^ " -w")
-  in
-    if rc <> 0 then error ("Get available systems from SystemOnTPTP:\n" ^ answer)
-    else split_lines answer
-  end;
-
-fun refresh_systems () = Synchronized.change systems (fn _ =>
-  get_systems ());
-
-fun get_system prefix = Synchronized.change_result systems (fn systems =>
-  let val systems = if null systems then get_systems() else systems
-  in (find_first (String.isPrefix prefix) systems, systems) end);
-
-fun remote_prover_opts max_new theory_const args prover_prefix timeout =
-  let val sys =
-    case get_system prover_prefix of
-      NONE => error ("No system like " ^ quote prover_prefix ^ " at SystemOnTPTP")
-    | SOME sys => sys
-  in tptp_prover_opts max_new theory_const
-    (Path.explode "$ISABELLE_HOME/lib/scripts/SystemOnTPTP",
-      args ^ " -t " ^ string_of_int timeout ^ " -s " ^ sys) timeout
-  end;
-
-val remote_prover = remote_prover_opts 60 false;
-
-end;
-
--- a/src/HOL/Tools/inductive_set.ML	Mon Aug 10 08:37:37 2009 +0200
+++ b/src/HOL/Tools/inductive_set.ML	Mon Aug 10 10:25:00 2009 +0200
@@ -9,6 +9,7 @@
 sig
   val to_set_att: thm list -> attribute
   val to_pred_att: thm list -> attribute
+  val to_pred : thm list -> Context.generic -> thm -> thm
   val pred_set_conv_att: attribute
   val add_inductive_i:
     Inductive.inductive_flags ->
--- a/src/HOL/Tools/record.ML	Mon Aug 10 08:37:37 2009 +0200
+++ b/src/HOL/Tools/record.ML	Mon Aug 10 10:25:00 2009 +0200
@@ -4,7 +4,6 @@
 Extensible records with structural subtyping in HOL.
 *)
 
-
 signature BASIC_RECORD =
 sig
   val record_simproc: simproc
@@ -88,6 +87,8 @@
 val RepN = "Rep_";
 val AbsN = "Abs_";
 
+
+
 (*** utilities ***)
 
 fun but_last xs = fst (split_last xs);
@@ -102,6 +103,7 @@
 fun range_type' T =
     range_type T handle Match => T;
 
+
 (* messages *)
 
 fun trace_thm str thm =
@@ -113,12 +115,14 @@
 fun trace_term str t =
     tracing (str ^ Syntax.string_of_term_global Pure.thy t);
 
+
 (* timing *)
 
 val timing = ref false;
 fun timeit_msg s x = if !timing then (warning s; timeit x) else x ();
 fun timing_msg s = if !timing then warning s else ();
 
+
 (* syntax *)
 
 fun prune n xs = Library.drop (n, xs);
@@ -136,6 +140,7 @@
 val (op ===) = Trueprop o HOLogic.mk_eq;
 val (op ==>) = Logic.mk_implies;
 
+
 (* morphisms *)
 
 fun mk_RepN name = suffix ext_typeN (prefix_base RepN name);
@@ -147,6 +152,7 @@
 fun mk_Abs name repT absT =
   Const (mk_AbsN name,repT --> absT);
 
+
 (* constructor *)
 
 fun mk_extC (name,T) Ts  = (suffix extN name, Ts ---> T);
@@ -155,6 +161,7 @@
   let val Ts = map fastype_of ts
   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)
@@ -163,6 +170,7 @@
   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);
@@ -171,6 +179,7 @@
   let val sT = fastype_of s
   in Const (mk_selC sT (c,T)) $ s end;
 
+
 (* updates *)
 
 fun mk_updC sfx sT (c,T) = (suffix sfx c, (T --> T) --> sT --> sT);
@@ -181,6 +190,7 @@
 
 fun mk_upd sfx c v s = mk_upd' sfx c v (fastype_of s) $ s
 
+
 (* types *)
 
 fun dest_recT (typ as Type (c_ext_type, Ts as (T::_))) =
@@ -209,6 +219,8 @@
       val rTs' = if i < 0 then rTs else Library.take (i,rTs)
   in Library.foldl (fn (s,(c,T)) => s ^ c) ("",rTs') end;
 
+
+
 (*** extend theory by record definition ***)
 
 (** record info **)
@@ -766,20 +778,20 @@
       fun match rT T = (Sign.typ_match thy (varifyT rT,T)
                                                 Vartab.empty);
 
-   in if !print_record_type_abbr
-      then (case last_extT T of
-             SOME (name,_)
-              => if name = lastExt
-                 then
-                  (let
-                     val subst = match schemeT T
-                   in
-                    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)
-                 else raise Match (* give print translation of specialised record a chance *)
-            | _ => raise Match)
+   in
+     if !print_record_type_abbr then
+       (case last_extT T of
+         SOME (name, _) =>
+          if name = lastExt then
+            (let
+               val subst = match schemeT T
+             in
+              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)
+           else raise Match (* give print translation of specialised record a chance *)
+        | _ => raise Match)
        else default_tr' ctxt tm
   end
 
@@ -848,6 +860,8 @@
                                (list_comb (Syntax.const name_sfx,ts))
   in (name_sfx, tr') end;
 
+
+
 (** record simprocs **)
 
 val record_quick_and_dirty_sensitive = ref false;
@@ -1279,8 +1293,6 @@
          end)
 
 
-
-
 local
 val inductive_atomize = thms "induct_atomize";
 val inductive_rulify = thms "induct_rulify";
@@ -1363,6 +1375,7 @@
      else Seq.empty
   end handle Subscript => Seq.empty;
 
+
 (* wrapper *)
 
 val record_split_name = "record_split_tac";
@@ -1400,6 +1413,7 @@
 fun induct_type_global name = [case_names_fields, Induct.induct_type name];
 fun cases_type_global name = [case_names_fields, Induct.cases_type name];
 
+
 (* tactics *)
 
 fun simp_all_tac ss simps = ALLGOALS (Simplifier.asm_full_simp_tac (ss addsimps simps));
@@ -1469,7 +1483,9 @@
   end;
 
 fun mixit convs refls =
-  let fun f ((res,lhs,rhs),refl) = ((refl,List.revAppend (lhs,refl::tl rhs))::res,hd rhs::lhs,tl rhs);
+  let
+    fun f ((res,lhs,rhs),refl) =
+      ((refl,List.revAppend (lhs,refl::tl rhs))::res,hd rhs::lhs,tl rhs);
   in #1 (Library.foldl f (([],[],convs),refls)) end;
 
 
@@ -2166,8 +2182,10 @@
       end);
      val equality = timeit_msg "record equality proof:" equality_prf;
 
-    val ((([sel_convs',upd_convs',sel_defs',upd_defs',[split_meta',split_object',split_ex'],derived_defs'],
-            [surjective',equality']),[induct_scheme',induct',cases_scheme',cases']), thms_thy) =
+    val ((([sel_convs', upd_convs', sel_defs', upd_defs',
+          [split_meta', split_object', split_ex'], derived_defs'],
+          [surjective', equality']),
+          [induct_scheme', induct', cases_scheme', cases']), thms_thy) =
       defs_thy
       |> (PureThy.add_thmss o map (Thm.no_attributes o apfst Binding.name))
          [("select_convs", sel_convs_standard),
@@ -2296,6 +2314,7 @@
 val add_record = gen_add_record read_typ read_raw_parent;
 val add_record_i = gen_add_record cert_typ (K I);
 
+
 (* setup theory *)
 
 val setup =
@@ -2304,6 +2323,7 @@
   Simplifier.map_simpset (fn ss =>
     ss addsimprocs [record_simproc, record_upd_simproc, record_eq_simproc]);
 
+
 (* outer syntax *)
 
 local structure P = OuterParse and K = OuterKeyword in
@@ -2320,6 +2340,5 @@
 
 end;
 
-
-structure BasicRecord: BASIC_RECORD = Record;
-open BasicRecord;
+structure Basic_Record: BASIC_RECORD = Record;
+open Basic_Record;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/Mirabelle.thy	Mon Aug 10 10:25:00 2009 +0200
@@ -0,0 +1,14 @@
+(* Title: Mirabelle.thy
+   Author: Jasmin Blanchette and Sascha Boehme
+*)
+
+theory Mirabelle
+imports Main
+uses "mirabelle.ML"
+begin
+
+(* FIXME: use a logfile for each theory file *)
+
+setup Mirabelle.setup
+
+end
--- a/src/HOL/ex/Predicate_Compile.thy	Mon Aug 10 08:37:37 2009 +0200
+++ b/src/HOL/ex/Predicate_Compile.thy	Mon Aug 10 10:25:00 2009 +0200
@@ -1,5 +1,5 @@
 theory Predicate_Compile
-imports Complex_Main
+imports Complex_Main RPred
 uses "predicate_compile.ML"
 begin
 
--- a/src/HOL/ex/Predicate_Compile_ex.thy	Mon Aug 10 08:37:37 2009 +0200
+++ b/src/HOL/ex/Predicate_Compile_ex.thy	Mon Aug 10 10:25:00 2009 +0200
@@ -1,5 +1,5 @@
 theory Predicate_Compile_ex
-imports Complex_Main Predicate_Compile
+imports Main Predicate_Compile
 begin
 
 inductive even :: "nat \<Rightarrow> bool" and odd :: "nat \<Rightarrow> bool" where
@@ -46,27 +46,28 @@
   | "f x \<Longrightarrow> partition f xs ys zs \<Longrightarrow> partition f (x # xs) (x # ys) zs"
   | "\<not> f x \<Longrightarrow> partition f xs ys zs \<Longrightarrow> partition f (x # xs) ys (x # zs)"
 
-(* FIXME: correct handling of parameters *)
-(*
-ML {* reset Predicate_Compile.do_proofs *}
 code_pred partition .
 
 thm partition.equation
-ML {* set Predicate_Compile.do_proofs *}
-*)
+
+inductive is_even :: "nat \<Rightarrow> bool"
+where
+  "n mod 2 = 0 \<Longrightarrow> is_even n"
+
+code_pred is_even .
 
 (* TODO: requires to handle abstractions in parameter positions correctly *)
-(*FIXME values 10 "{(ys, zs). partition (\<lambda>n. n mod 2 = 0)
-  [0, Suc 0, 2, 3, 4, 5, 6, 7] ys zs}" *)
+values 10 "{(ys, zs). partition is_even
+  [0, Suc 0, 2, 3, 4, 5, 6, 7] ys zs}"
 
+values 10 "{zs. partition is_even zs [0, 2] [3, 5]}"
+values 10 "{zs. partition is_even zs [0, 7] [3, 5]}"
 
 lemma [code_pred_intros]:
   "r a b \<Longrightarrow> tranclp r a b"
   "r a b \<Longrightarrow> tranclp r b c \<Longrightarrow> tranclp r a c"
   by auto
 
-(* Setup requires quick and dirty proof *)
-(*
 code_pred tranclp
 proof -
   case tranclp
@@ -74,6 +75,11 @@
 qed
 
 thm tranclp.equation
+(*
+setup {* Predicate_Compile.add_sizelim_equations [@{const_name tranclp}] *}
+setup {* fn thy => exception_trace (fn () => Predicate_Compile.add_quickcheck_equations [@{const_name tranclp}] thy)  *}
+
+thm tranclp.rpred_equation
 *)
 
 inductive succ :: "nat \<Rightarrow> nat \<Rightarrow> bool" where
@@ -83,12 +89,16 @@
 code_pred succ .
 
 thm succ.equation
+<<<<<<< local
 
 values 10 "{(m, n). succ n m}"
 values "{m. succ 0 m}"
 values "{m. succ m 0}"
 
 (* FIXME: why does this not terminate? *)
+=======
+(* FIXME: why does this not terminate? -- value chooses mode [] --> [1] and then starts enumerating all successors *)
+>>>>>>> other
 (*
 values 20 "{n. tranclp succ 10 n}"
 values "{n. tranclp succ n 10}"
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/RPred.thy	Mon Aug 10 10:25:00 2009 +0200
@@ -0,0 +1,48 @@
+theory RPred
+imports Quickcheck Random Predicate
+begin
+
+types 'a rpred = "Random.seed \<Rightarrow> ('a Predicate.pred \<times> Random.seed)"
+
+section {* The RandomPred Monad *}
+
+text {* A monad to combine the random state monad and the predicate monad *}
+
+definition bot :: "'a rpred"
+  where "bot = Pair (bot_class.bot)"
+
+definition return :: "'a => 'a rpred"
+  where "return x = Pair (Predicate.single x)"
+
+definition bind :: "'a rpred \<Rightarrow> ('a \<Rightarrow> 'b rpred) \<Rightarrow> 'b rpred" (infixl "\<guillemotright>=" 60)
+  where "bind RP f =
+  (\<lambda>s. let (P, s') = RP s;
+        (s1, s2) = Random.split_seed s'
+    in (Predicate.bind P (%a. fst (f a s1)), s2))"
+
+definition supp :: "'a rpred \<Rightarrow> 'a rpred \<Rightarrow> 'a rpred" (infixl "\<squnion>" 80)
+where
+  "supp RP1 RP2 = (\<lambda>s. let (P1, s') = RP1 s; (P2, s'') = RP2 s'
+  in (upper_semilattice_class.sup P1 P2, s''))"
+
+definition if_rpred :: "bool \<Rightarrow> unit rpred"
+where
+  "if_rpred b = (if b then return () else bot)"
+
+(* Missing a good definition for negation: not_rpred *)
+
+definition not_rpred :: "unit Predicate.pred \<Rightarrow> unit rpred"
+where
+  "not_rpred = Pair o Predicate.not_pred"
+
+definition lift_pred :: "'a Predicate.pred \<Rightarrow> 'a rpred"
+  where
+  "lift_pred = Pair"
+
+definition lift_random :: "(Random.seed \<Rightarrow> ('a \<times> (unit \<Rightarrow> term)) \<times> Random.seed) \<Rightarrow> 'a rpred"
+  where "lift_random g = scomp g (Pair o (Predicate.single o fst))"
+
+definition map_rpred :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a rpred \<Rightarrow> 'b rpred)"
+where "map_rpred f P = P \<guillemotright>= (return o f)"
+  
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/mirabelle.ML	Mon Aug 10 10:25:00 2009 +0200
@@ -0,0 +1,318 @@
+(* Title:  mirabelle.ML
+   Author: Jasmin Blanchette and Sascha Boehme
+*)
+
+signature MIRABELLE =
+sig
+  type action
+  type settings
+  val register : string -> action -> theory -> theory
+  val invoke : string -> settings -> theory -> theory
+
+  val timeout : int Config.T
+  val verbose : bool Config.T
+  val set_logfile : string -> theory -> theory
+
+  val setup : theory -> theory
+
+  val step_hook : Toplevel.transition -> Toplevel.state -> Toplevel.state ->
+    unit
+
+  val goal_thm_of : Proof.state -> thm
+  val can_apply : (Proof.context -> int -> tactic) -> Proof.state -> bool
+  val theorems_in_proof_term : Thm.thm -> Thm.thm list
+  val theorems_of_sucessful_proof : Toplevel.state -> Thm.thm list
+  val get_setting : settings -> string * string -> string
+  val get_int_setting : settings -> string * int -> int
+
+(* FIXME  val refute_action : action *)
+  val quickcheck_action : action
+  val arith_action : action
+  val sledgehammer_action : action
+  val metis_action : action
+end
+
+
+
+structure Mirabelle (*: MIRABELLE*) =
+struct
+
+(* Mirabelle core *)
+
+type settings = (string * string) list
+type invoked = {pre: Proof.state, post: Toplevel.state option} -> string option
+type action = settings -> invoked
+
+structure Registered = TheoryDataFun
+(
+  type T = action Symtab.table
+  val empty = Symtab.empty
+  val copy = I
+  val extend = I
+  fun merge _ = Symtab.merge (K true)
+)
+
+fun register name act = Registered.map (Symtab.update_new (name, act))
+
+
+structure Invoked = TheoryDataFun
+(
+  type T = (string * invoked) list
+  val empty = []
+  val copy = I
+  val extend = I
+  fun merge _ = Library.merge (K true)
+)
+
+fun invoke name sts thy = 
+  let 
+    val act = 
+      (case Symtab.lookup (Registered.get thy) name of
+        SOME act => act
+      | NONE => error ("The invoked action " ^ quote name ^ 
+          " is not registered."))
+  in Invoked.map (cons (name, act sts)) thy end
+
+val (logfile, setup1) = Attrib.config_string "mirabelle_logfile" ""
+val (timeout, setup2) = Attrib.config_int "mirabelle_timeout" 30
+val (verbose, setup3) = Attrib.config_bool "mirabelle_verbose" true
+val (start_line, setup4) = Attrib.config_int "mirabelle_start_line" 0
+val (end_line, setup5) = Attrib.config_int "mirabelle_end_line" ~1
+
+val setup_config = setup1 #> setup2 #> setup3 #> setup4 #> setup5
+
+fun set_logfile name =
+  let val _ = File.write (Path.explode name) ""   (* erase file content *)
+  in Config.put_thy logfile name end
+
+local
+
+fun log thy s =
+  let fun append_to n = if n = "" then K () else File.append (Path.explode n)
+  in append_to (Config.get_thy thy logfile) (s ^ "\n") end
+  (* FIXME: with multithreading and parallel proofs enabled, we might need to
+     encapsulate this inside a critical section *)
+
+fun verbose_msg verbose msg = if verbose then SOME msg else NONE
+
+fun with_time_limit (verb, secs) f x = TimeLimit.timeLimit secs f x
+  handle TimeLimit.TimeOut => verbose_msg verb "time out"
+       | ERROR msg => verbose_msg verb ("error: " ^ msg)
+
+fun capture_exns verb f x =
+  (case try f x of NONE => verbose_msg verb "exception" | SOME msg => msg)
+
+fun apply_action (c as (verb, _)) st (name, invoked) =
+  Option.map (pair name) (capture_exns verb (with_time_limit c invoked) st)
+
+fun in_range _ _ NONE = true
+  | in_range l r (SOME i) = (l <= i andalso (r < 0 orelse i <= r))
+
+fun only_within_range thy pos f x =
+  let val l = Config.get_thy thy start_line and r = Config.get_thy thy end_line
+  in if in_range l r (Position.line_of pos) then f x else [] end
+
+fun pretty_print verbose pos name msgs =
+  let
+    val file = the_default "unknown file" (Position.file_of pos)
+
+    val str0 = string_of_int o the_default 0
+    val loc = str0 (Position.line_of pos) ^ ":" ^ str0 (Position.column_of pos)
+
+    val full_loc = if verbose then file ^ ":" ^ loc else "at " ^ loc
+    val head = full_loc ^ " (" ^ name ^ "):"
+
+    fun pretty_msg (name, msg) = Pretty.block (map Pretty.str [name, ": ", msg])
+  in
+    Pretty.string_of (Pretty.big_list head (map pretty_msg msgs))
+  end
+
+in
+
+fun basic_hook tr pre post =
+  let
+    val thy = Proof.theory_of pre
+    val pos = Toplevel.pos_of tr
+    val name = Toplevel.name_of tr
+    val verb = Config.get_thy thy verbose
+    val secs = Time.fromSeconds (Config.get_thy thy timeout)
+    val st = {pre=pre, post=post}
+  in
+    Invoked.get thy
+    |> only_within_range thy pos (map_filter (apply_action (verb, secs) st))
+    |> (fn [] => () | msgs => log thy (pretty_print verb pos name msgs))
+  end
+
+end
+
+fun step_hook tr pre post =
+ (* FIXME: might require wrapping into "interruptible" *)
+  if can (Proof.assert_backward o Toplevel.proof_of) pre andalso
+     not (member (op =) ["disable_pr", "enable_pr"] (Toplevel.name_of tr))
+  then basic_hook tr (Toplevel.proof_of pre) (SOME post)
+  else ()   (* FIXME: add theory_hook here *)
+
+
+
+(* Mirabelle utility functions *)
+
+val goal_thm_of = snd o snd o Proof.get_goal
+
+fun can_apply tac st =
+  let val (ctxt, (facts, goal)) = Proof.get_goal st
+  in
+    (case Seq.pull (HEADGOAL (Method.insert_tac facts THEN' tac ctxt) goal) of
+      SOME (thm, _) => true
+    | NONE => false)
+  end
+
+local
+
+fun fold_body_thms f =
+  let
+    fun app n (PBody {thms, ...}) = thms |> fold (fn (i, (name, prop, body)) =>
+      fn (x, seen) =>
+        if Inttab.defined seen i then (x, seen)
+        else
+          let
+            val body' = Future.join body
+            val (x', seen') = app (n + (if name = "" then 0 else 1)) body'
+              (x, Inttab.update (i, ()) seen)
+        in (x' |> n = 0 ? f (name, prop, body'), seen') end)
+  in fn bodies => fn x => #1 (fold (app 0) bodies (x, Inttab.empty)) end
+
+in
+
+fun theorems_in_proof_term thm =
+  let
+    val all_thms = PureThy.all_thms_of (Thm.theory_of_thm thm)
+    fun collect (s, _, _) = if s <> "" then insert (op =) s else I
+    fun member_of xs (x, y) = if member (op =) xs x then SOME y else NONE
+    fun resolve_thms names = map_filter (member_of names) all_thms
+  in
+    resolve_thms (fold_body_thms collect [Thm.proof_body_of thm] [])
+  end
+
+end
+
+fun theorems_of_sucessful_proof state =
+  (case state of
+    NONE => []
+  | SOME st =>
+      if not (Toplevel.is_proof st) then []
+      else theorems_in_proof_term (goal_thm_of (Toplevel.proof_of st)))
+
+fun get_setting settings (key, default) =
+  the_default default (AList.lookup (op =) settings key)
+
+fun get_int_setting settings (key, default) =
+  (case Option.map Int.fromString (AList.lookup (op =) settings key) of
+    SOME (SOME i) => i
+  | SOME NONE => error ("bad option: " ^ key)
+  | NONE => default)
+
+
+
+(* Mirabelle actions *)
+
+(* FIXME
+fun refute_action settings {pre=st, ...} = 
+  let
+    val params   = [("minsize", "2") (*"maxsize", "2"*)]
+    val subgoal = 0
+    val thy     = Proof.theory_of st
+    val thm = goal_thm_of st
+
+    val _ = Refute.refute_subgoal thy parms thm subgoal
+  in
+    val writ_log = Substring.full (the (Symtab.lookup tab "writeln"))
+    val warn_log = Substring.full (the (Symtab.lookup tab "warning"))
+
+    val r =
+      if Substring.isSubstring "model found" writ_log
+      then
+        if Substring.isSubstring "spurious" warn_log
+        then SOME "potential counterexample"
+        else SOME "real counterexample (bug?)"
+      else
+        if Substring.isSubstring "time limit" writ_log
+        then SOME "no counterexample (time out)"
+        else if Substring.isSubstring "Search terminated" writ_log
+        then SOME "no counterexample (normal termination)"
+        else SOME "no counterexample (unknown)"
+  in r end
+*)
+
+fun quickcheck_action settings {pre=st, ...} =
+  let
+    val has_valid_key = member (op =) ["iterations", "size", "generator"] o fst
+    val args = filter has_valid_key settings
+  in
+    (case Quickcheck.quickcheck args 1 st of
+      NONE => SOME "no counterexample"
+    | SOME _ => SOME "counterexample found")
+  end
+
+
+fun arith_action _ {pre=st, ...} = 
+  if can_apply Arith_Data.arith_tac st
+  then SOME "succeeded"
+  else NONE
+
+
+fun sledgehammer_action settings {pre=st, ...} =
+  let
+    val prover_name = hd (space_explode " " (AtpManager.get_atps ()))
+    val thy = Proof.theory_of st
+ 
+    val prover = the (AtpManager.get_prover prover_name thy)
+    val timeout = AtpManager.get_timeout () 
+
+    val (success, message) =
+      let
+        val (success, message, _, _, _) =
+          prover timeout NONE NONE prover_name 1 (Proof.get_goal st)
+      in (success, message) end
+      handle ResHolClause.TOO_TRIVIAL => (true, "trivial")
+           | ERROR msg => (false, "error: " ^ msg)
+  in
+    if success
+    then SOME ("success (" ^ prover_name ^ ": " ^ message ^ ")")
+    else NONE
+  end
+
+
+fun metis_action settings {pre, post} =
+  let
+    val thms = theorems_of_sucessful_proof post
+    val names = map Thm.get_name thms
+
+    val facts = Facts.props (ProofContext.facts_of (Proof.context_of pre))
+
+    fun metis ctxt = MetisTools.metis_tac ctxt (thms @ facts)
+  in
+    (if can_apply metis pre then "succeeded" else "failed")
+    |> suffix (" (" ^ commas names ^ ")")
+    |> SOME
+  end
+
+
+
+(* Mirabelle setup *)
+
+val setup =
+  setup_config #>
+(* FIXME  register "refute" refute_action #> *)
+  register "quickcheck" quickcheck_action #>
+  register "arith" arith_action #>
+  register "sledgehammer" sledgehammer_action #>
+  register "metis" metis_action (* #> FIXME:
+  Context.theory_map (Specification.add_theorem_hook theorem_hook) *)
+
+end
+
+val _ = Toplevel.add_hook Mirabelle.step_hook
+
+(* no multithreading, no parallel proofs *)
+val _ = Multithreading.max_threads := 1
+val _ = Goal.parallel_proofs := 0
--- a/src/HOL/ex/predicate_compile.ML	Mon Aug 10 08:37:37 2009 +0200
+++ b/src/HOL/ex/predicate_compile.ML	Mon Aug 10 10:25:00 2009 +0200
@@ -7,16 +7,17 @@
 signature PREDICATE_COMPILE =
 sig
   type mode = int list option list * int list
-  val add_equations_of: string list -> theory -> theory
+  (*val add_equations_of: bool -> string list -> theory -> theory *)
   val register_predicate : (thm list * thm * int) -> theory -> theory
   val is_registered : theory -> string -> bool
-  val fetch_pred_data : theory -> string -> (thm list * thm * int)  
+ (* val fetch_pred_data : theory -> string -> (thm list * thm * int)  *)
   val predfun_intro_of: theory -> string -> mode -> thm
   val predfun_elim_of: theory -> string -> mode -> thm
   val strip_intro_concl: int -> term -> term * (term list * term list)
   val predfun_name_of: theory -> string -> mode -> string
   val all_preds_of : theory -> string list
   val modes_of: theory -> string -> mode list
+  val string_of_mode : mode -> string
   val intros_of: theory -> string -> thm list
   val nparams_of: theory -> string -> int
   val add_intro: thm -> theory -> theory
@@ -25,12 +26,77 @@
   val code_pred: string -> Proof.context -> Proof.state
   val code_pred_cmd: string -> Proof.context -> Proof.state
   val print_stored_rules: theory -> unit
+  val print_all_modes: theory -> unit
   val do_proofs: bool ref
   val mk_casesrule : Proof.context -> int -> thm list -> term
   val analyze_compr: theory -> term -> term
   val eval_ref: (unit -> term Predicate.pred) option ref
-  val add_equations : string -> theory -> theory
+  val add_equations : string list -> theory -> theory
   val code_pred_intros_attrib : attribute
+  (* used by Quickcheck_Generator *) 
+  (*val funT_of : mode -> typ -> typ
+  val mk_if_pred : term -> term
+  val mk_Eval : term * term -> term*)
+  val mk_tupleT : typ list -> typ
+(*  val mk_predT :  typ -> typ *)
+  (* temporary for testing of the compilation *)
+  datatype indprem = Prem of term list * term | Negprem of term list * term | Sidecond of term |
+    GeneratorPrem of term list * term | Generator of (string * typ);
+  val prepare_intrs: theory -> string list ->
+    (string * typ) list * int * string list * string list * (string * mode list) list *
+    (string * (term list * indprem list) list) list * (string * (int option list * int)) list
+  datatype compilation_funs = CompilationFuns of {
+    mk_predT : typ -> typ,
+    dest_predT : typ -> typ,
+    mk_bot : typ -> term,
+    mk_single : term -> term,
+    mk_bind : term * term -> term,
+    mk_sup : term * term -> term,
+    mk_if : term -> term,
+    mk_not : term -> term,
+    mk_map : typ -> typ -> term -> term -> term,
+    lift_pred : term -> term
+  };  
+  datatype tmode = Mode of mode * int list * tmode option list;
+  type moded_clause = term list * (indprem * tmode) list
+  type 'a pred_mode_table = (string * (mode * 'a) list) list
+  val infer_modes : bool -> theory -> (string * (int list option list * int list) list) list
+    -> (string * (int option list * int)) list -> string list
+    -> (string * (term list * indprem list) list) list
+    -> (moded_clause list) pred_mode_table
+  val infer_modes_with_generator : theory -> (string * (int list option list * int list) list) list
+    -> (string * (int option list * int)) list -> string list
+    -> (string * (term list * indprem list) list) list
+    -> (moded_clause list) pred_mode_table  
+  (*val compile_preds : theory -> compilation_funs -> string list -> string list
+    -> (string * typ) list -> (moded_clause list) pred_mode_table -> term pred_mode_table
+  val rpred_create_definitions :(string * typ) list -> string * mode list
+    -> theory -> theory 
+  val split_smode : int list -> term list -> (term list * term list) *)
+  val print_moded_clauses :
+    theory -> (moded_clause list) pred_mode_table -> unit
+  val print_compiled_terms : theory -> term pred_mode_table -> unit
+  (*val rpred_prove_preds : theory -> term pred_mode_table -> thm pred_mode_table*)
+  val rpred_compfuns : compilation_funs
+  val dest_funT : typ -> typ * typ
+ (* val depending_preds_of : theory -> thm list -> string list *)
+  val add_quickcheck_equations : string list -> theory -> theory
+  val add_sizelim_equations : string list -> theory -> theory
+  val is_inductive_predicate : theory -> string -> bool
+  val terms_vs : term list -> string list
+  val subsets : int -> int -> int list list
+  val check_mode_clause : bool -> theory -> string list ->
+    (string * mode list) list -> (string * mode list) list -> mode -> (term list * indprem list)
+      -> (term list * (indprem * tmode) list) option
+  val string_of_moded_prem : theory -> (indprem * tmode) -> string
+  val all_modes_of : theory -> (string * mode list) list
+  val all_generator_modes_of : theory -> (string * mode list) list
+  val compile_clause : compilation_funs -> term option -> (term list -> term) ->
+    theory -> string list -> string list -> mode -> term -> moded_clause -> term
+  val preprocess_intro : theory -> thm -> thm
+  val is_constrt : theory -> term -> bool
+  val is_predT : typ -> bool
+  val guess_nparams : typ -> int
 end;
 
 structure Predicate_Compile : PREDICATE_COMPILE =
@@ -42,9 +108,8 @@
 
 fun tracing s = (if ! Toplevel.debug then Output.tracing s else ());
 
-fun print_tac s = (if ! Toplevel.debug then Tactical.print_tac s else Seq.single);
-fun new_print_tac s = Tactical.print_tac s
-fun debug_tac msg = (fn st => (Output.tracing msg; Seq.single st));
+fun print_tac s = Seq.single; (* (if ! Toplevel.debug then Tactical.print_tac s else Seq.single); *)
+fun debug_tac msg = Seq.single; (* (fn st => (Output.tracing msg; Seq.single st)); *)
 
 val do_proofs = ref true;
 
@@ -68,46 +133,44 @@
             HOLogic.mk_eq (Free (a, fastype_of b), b) :: mk_eqs a cs
   in mk_eqs x xs end;
 
-fun mk_pred_enumT T = Type (@{type_name Predicate.pred}, [T])
+fun mk_tupleT [] = HOLogic.unitT
+  | mk_tupleT Ts = foldr1 HOLogic.mk_prodT Ts;
 
-fun dest_pred_enumT (Type (@{type_name Predicate.pred}, [T])) = T
-  | dest_pred_enumT T = raise TYPE ("dest_pred_enumT", [T], []);
+fun dest_tupleT (Type (@{type_name Product_Type.unit}, [])) = []
+  | dest_tupleT (Type (@{type_name "*"}, [T1, T2])) = T1 :: (dest_tupleT T2)
+  | dest_tupleT t = [t]
+
+fun mk_tuple [] = HOLogic.unit
+  | mk_tuple ts = foldr1 HOLogic.mk_prod ts;
 
-fun mk_Enum f =
-  let val T as Type ("fun", [T', _]) = fastype_of f
-  in
-    Const (@{const_name Predicate.Pred}, T --> mk_pred_enumT T') $ f    
-  end;
+fun dest_tuple (Const (@{const_name Product_Type.Unity}, _)) = []
+  | dest_tuple (Const (@{const_name Pair}, _) $ t1 $ t2) = t1 :: (dest_tuple t2)
+  | dest_tuple t = [t]
 
-fun mk_Eval (f, x) =
-  let val T = fastype_of x
-  in
-    Const (@{const_name Predicate.eval}, mk_pred_enumT T --> T --> HOLogic.boolT) $ f $ x
+fun mk_scomp (t, u) =
+  let
+    val T = fastype_of t
+    val U = fastype_of u
+    val [A] = binder_types T
+    val D = body_type U 
+  in 
+    Const (@{const_name "scomp"}, T --> U --> A --> D) $ t $ u
   end;
 
-fun mk_empty T = Const (@{const_name Orderings.bot}, mk_pred_enumT T);
-
-fun mk_single t =
-  let val T = fastype_of t
-  in Const(@{const_name Predicate.single}, T --> mk_pred_enumT T) $ t end;
-
-fun mk_bind (x, f) =
-  let val T as Type ("fun", [_, U]) = fastype_of f
+fun dest_funT (Type ("fun",[S, T])) = (S, T)
+  | dest_funT T = raise TYPE ("dest_funT", [T], [])
+ 
+fun mk_fun_comp (t, u) =
+  let
+    val (_, B) = dest_funT (fastype_of t)
+    val (C, A) = dest_funT (fastype_of u)
   in
-    Const (@{const_name Predicate.bind}, fastype_of x --> T --> U) $ x $ f
+    Const(@{const_name "Fun.comp"}, (A --> B) --> (C --> A) --> C --> B) $ t $ u
   end;
 
-val mk_sup = HOLogic.mk_binop @{const_name sup};
-
-fun mk_if_predenum cond = Const (@{const_name Predicate.if_pred},
-  HOLogic.boolT --> mk_pred_enumT HOLogic.unitT) $ cond;
-
-fun mk_not_pred t = let val T = mk_pred_enumT HOLogic.unitT
-  in Const (@{const_name Predicate.not_pred}, T --> T) $ t end
-
-fun mk_pred_map T1 T2 tf tp = Const (@{const_name Predicate.map},
-  (T1 --> T2) --> mk_pred_enumT T1 --> mk_pred_enumT T2) $ tf $ tp;
-
+fun dest_randomT (Type ("fun", [@{typ Random.seed},
+  Type ("*", [Type ("*", [T, @{typ "unit => Code_Eval.term"}]) ,@{typ Random.seed}])])) = T
+  | dest_randomT T = raise TYPE ("dest_randomT", [T], [])
 
 (* destruction of intro rules *)
 
@@ -118,20 +181,40 @@
   val (params, args) = chop nparams all_args
 in (pred, (params, args)) end
 
-(* data structures *)
+(** data structures **)
+
+type smode = int list;
+type mode = smode option list * smode;
+datatype tmode = Mode of mode * int list * tmode option list;
 
-type mode = int list option list * int list; (*pmode FIMXE*)
+fun split_smode is ts =
+  let
+    fun split_smode' _ _ [] = ([], [])
+      | split_smode' is i (t::ts) = (if i mem is then apfst else apsnd) (cons t)
+          (split_smode' is (i+1) ts)
+  in split_smode' is 1 ts end
+
+fun split_mode (iss, is) ts =
+  let
+    val (t1, t2) = chop (length iss) ts 
+  in (t1, split_smode is t2) end
 
 fun string_of_mode (iss, is) = space_implode " -> " (map
   (fn NONE => "X"
     | SOME js => enclose "[" "]" (commas (map string_of_int js)))
        (iss @ [SOME is]));
 
-fun print_modes modes = Output.tracing ("Inferred modes:\n" ^
-  cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map
-    string_of_mode ms)) modes));
+fun string_of_tmode (Mode (predmode, termmode, param_modes)) =
+  "predmode: " ^ (string_of_mode predmode) ^ 
+  (if null param_modes then "" else
+    "; " ^ "params: " ^ commas (map (the_default "NONE" o Option.map string_of_tmode) param_modes))
+    
+datatype indprem = Prem of term list * term | Negprem of term list * term | Sidecond of term |
+  GeneratorPrem of term list * term | Generator of (string * typ);
 
-    
+type moded_clause = term list * (indprem * tmode) list
+type 'a pred_mode_table = (string * (mode * 'a) list) list
+
 datatype predfun_data = PredfunData of {
   name : string,
   definition : thm,
@@ -143,18 +226,30 @@
 fun mk_predfun_data (name, definition, intro, elim) =
   PredfunData {name = name, definition = definition, intro = intro, elim = elim}
 
+datatype function_data = FunctionData of {
+  name : string,
+  equation : thm option (* is not used at all? *)
+};
+
+fun rep_function_data (FunctionData data) = data;
+fun mk_function_data (name, equation) =
+  FunctionData {name = name, equation = equation}
+
 datatype pred_data = PredData of {
   intros : thm list,
   elim : thm option,
   nparams : int,
-  functions : (mode * predfun_data) list
+  functions : (mode * predfun_data) list,
+  generators : (mode * function_data) list,
+  sizelim_functions : (mode * function_data) list 
 };
 
 fun rep_pred_data (PredData data) = data;
-fun mk_pred_data ((intros, elim, nparams), functions) =
-  PredData {intros = intros, elim = elim, nparams = nparams, functions = functions}
-fun map_pred_data f (PredData {intros, elim, nparams, functions}) =
-  mk_pred_data (f ((intros, elim, nparams), functions))
+fun mk_pred_data ((intros, elim, nparams), (functions, generators, sizelim_functions)) =
+  PredData {intros = intros, elim = elim, nparams = nparams,
+    functions = functions, generators = generators, sizelim_functions = sizelim_functions}
+fun map_pred_data f (PredData {intros, elim, nparams, functions, generators, sizelim_functions}) =
+  mk_pred_data (f ((intros, elim, nparams), (functions, generators, sizelim_functions)))
   
 fun eq_option eq (NONE, NONE) = true
   | eq_option eq (SOME x, SOME y) = eq (x, y)
@@ -208,8 +303,8 @@
   (#functions (the_pred_data thy name)) mode)
 
 fun the_predfun_data thy name mode = case lookup_predfun_data thy name mode
- of NONE => error ("No such mode" ^ string_of_mode mode)
-  | SOME data => data;
+  of NONE => error ("No function defined for mode " ^ string_of_mode mode ^ " of predicate " ^ name)
+   | SOME data => data;
 
 val predfun_name_of = #name ooo the_predfun_data
 
@@ -219,6 +314,76 @@
 
 val predfun_elim_of = #elim ooo the_predfun_data
 
+fun lookup_generator_data thy name mode = 
+  Option.map rep_function_data (AList.lookup (op =)
+  (#generators (the_pred_data thy name)) mode)
+  
+fun the_generator_data thy name mode = case lookup_generator_data thy name mode
+  of NONE => error ("No generator defined for mode " ^ string_of_mode mode ^ " of predicate " ^ name)
+   | SOME data => data
+
+val generator_name_of = #name ooo the_generator_data
+
+val generator_modes_of = (map fst) o #generators oo the_pred_data
+
+fun all_generator_modes_of thy =
+  map (fn name => (name, generator_modes_of thy name)) (all_preds_of thy) 
+
+fun lookup_sizelim_function_data thy name mode =
+  Option.map rep_function_data (AList.lookup (op =)
+  (#sizelim_functions (the_pred_data thy name)) mode)
+
+fun the_sizelim_function_data thy name mode = case lookup_sizelim_function_data thy name mode
+  of NONE => error ("No size-limited function defined for mode " ^ string_of_mode mode
+    ^ " of predicate " ^ name)
+   | SOME data => data
+
+val sizelim_function_name_of = #name ooo the_sizelim_function_data
+
+(*val generator_modes_of = (map fst) o #generators oo the_pred_data*)
+     
+(* diagnostic display functions *)
+
+fun print_modes modes = Output.tracing ("Inferred modes:\n" ^
+  cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map
+    string_of_mode ms)) modes));
+
+fun print_pred_mode_table string_of_entry thy pred_mode_table =
+  let
+    fun print_mode pred (mode, entry) =  "mode : " ^ (string_of_mode mode)
+      ^ (string_of_entry pred mode entry)  
+    fun print_pred (pred, modes) =
+      "predicate " ^ pred ^ ": " ^ cat_lines (map (print_mode pred) modes)
+    val _ = Output.tracing (cat_lines (map print_pred pred_mode_table))
+  in () end;
+
+fun string_of_moded_prem thy (Prem (ts, p), tmode) =
+    (Syntax.string_of_term_global thy (list_comb (p, ts))) ^
+    "(" ^ (string_of_tmode tmode) ^ ")"
+  | string_of_moded_prem thy (GeneratorPrem (ts, p), Mode (predmode, is, _)) =
+    (Syntax.string_of_term_global thy (list_comb (p, ts))) ^
+    "(generator_mode: " ^ (string_of_mode predmode) ^ ")"
+  | string_of_moded_prem thy (Generator (v, T), _) =
+    "Generator for " ^ v ^ " of Type " ^ (Syntax.string_of_typ_global thy T)
+  | string_of_moded_prem thy (Negprem (ts, p), Mode (_, is, _)) =
+    (Syntax.string_of_term_global thy (list_comb (p, ts))) ^
+    "(negative mode: " ^ (space_implode ", " (map string_of_int is)) ^ ")"
+  | string_of_moded_prem thy (Sidecond t, Mode (_, is, _)) =
+    (Syntax.string_of_term_global thy t) ^
+    "(sidecond mode: " ^ (space_implode ", " (map string_of_int is)) ^ ")"    
+  | string_of_moded_prem _ _ = error "string_of_moded_prem: unimplemented"
+     
+fun print_moded_clauses thy =
+  let        
+    fun string_of_clause pred mode clauses =
+      cat_lines (map (fn (ts, prems) => (space_implode " --> "
+        (map (string_of_moded_prem thy) prems)) ^ " --> " ^ pred ^ " "
+        ^ (space_implode " " (map (Syntax.string_of_term_global thy) ts))) clauses)
+  in print_pred_mode_table string_of_clause thy end;
+
+fun print_compiled_terms thy =
+  print_pred_mode_table (fn _ => fn _ => Syntax.string_of_term_global thy) thy
+    
 fun print_stored_rules thy =
   let
     val preds = (Graph.keys o PredData.get) thy
@@ -238,6 +403,18 @@
     fold print preds ()
   end;
 
+fun print_all_modes thy =
+  let
+    val _ = writeln ("Inferred modes:")
+    fun print (pred, modes) u =
+      let
+        val _ = writeln ("predicate: " ^ pred)
+        val _ = writeln ("modes: " ^ (commas (map string_of_mode modes)))
+      in u end  
+  in
+    fold print (all_modes_of thy) ()
+  end
+  
 (** preprocessing rules **)  
 
 fun imp_prems_conv cv ct =
@@ -256,24 +433,48 @@
       (Trueprop_conv (Conv.try_conv (Conv.rewr_conv (Thm.symmetric @{thm Predicate.eq_is_eq})))))
     (Thm.transfer thy rule)
 
-fun preprocess_elim thy nargs elimrule = let
-   fun replace_eqs (Const ("Trueprop", _) $ (Const ("op =", T) $ lhs $ rhs)) =
-      HOLogic.mk_Trueprop (Const (@{const_name Predicate.eq}, T) $ lhs $ rhs)
-    | replace_eqs t = t
-   fun preprocess_case t = let
-     val params = Logic.strip_params t
-     val (assums1, assums2) = chop nargs (Logic.strip_assums_hyp t)
-     val assums_hyp' = assums1 @ (map replace_eqs assums2)
-     in list_all (params, Logic.list_implies (assums_hyp', Logic.strip_assums_concl t)) end
-   val prems = Thm.prems_of elimrule
-   val cases' = map preprocess_case (tl prems)
-   val elimrule' = Logic.list_implies ((hd prems) :: cases', Thm.concl_of elimrule)
- in
-   Thm.equal_elim
-     (Thm.symmetric (Conv.implies_concl_conv (MetaSimplifier.rewrite true [@{thm eq_is_eq}])
-        (cterm_of thy elimrule')))
-     elimrule
- end;
+fun preprocess_elim thy nparams elimrule =
+  let
+    fun replace_eqs (Const ("Trueprop", _) $ (Const ("op =", T) $ lhs $ rhs)) =
+       HOLogic.mk_Trueprop (Const (@{const_name Predicate.eq}, T) $ lhs $ rhs)
+     | replace_eqs t = t
+    val prems = Thm.prems_of elimrule
+    val nargs = length (snd (strip_comb (HOLogic.dest_Trueprop (hd prems)))) - nparams
+    fun preprocess_case t =
+     let
+       val params = Logic.strip_params t
+       val (assums1, assums2) = chop nargs (Logic.strip_assums_hyp t)
+       val assums_hyp' = assums1 @ (map replace_eqs assums2)
+     in
+       list_all (params, Logic.list_implies (assums_hyp', Logic.strip_assums_concl t))
+     end 
+    val cases' = map preprocess_case (tl prems)
+    val elimrule' = Logic.list_implies ((hd prems) :: cases', Thm.concl_of elimrule)
+  in
+    Thm.equal_elim
+      (Thm.symmetric (Conv.implies_concl_conv (MetaSimplifier.rewrite true [@{thm eq_is_eq}])
+         (cterm_of thy elimrule')))
+      elimrule
+  end;
+
+(* special case: predicate with no introduction rule *)
+fun noclause thy predname elim = let
+  val T = (Logic.unvarifyT o Sign.the_const_type thy) predname
+  val Ts = binder_types T
+  val names = Name.variant_list []
+        (map (fn i => "x" ^ (string_of_int i)) (1 upto (length Ts)))
+  val vs = map2 (curry Free) names Ts
+  val clausehd = HOLogic.mk_Trueprop (list_comb (Const (predname, T), vs))
+  val intro_t = Logic.mk_implies (@{prop False}, clausehd)
+  val P = HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT))
+  val elim_t = Logic.list_implies ([clausehd, Logic.mk_implies (@{prop False}, P)], P)
+  val intro = Goal.prove (ProofContext.init thy) names [] intro_t
+        (fn {...} => etac @{thm FalseE} 1)
+  val elim = Goal.prove (ProofContext.init thy) ("P" :: names) [] elim_t
+        (fn {...} => etac elim 1) 
+in
+  ([intro], elim)
+end
 
 fun fetch_pred_data thy name =
   case try (Inductive.the_inductive (ProofContext.init thy)) name of
@@ -282,46 +483,79 @@
         fun is_intro_of intro =
           let
             val (const, _) = strip_comb (HOLogic.dest_Trueprop (concl_of intro))
-          in (fst (dest_Const const) = name) end;
-        val intros = map (preprocess_intro thy) (filter is_intro_of (#intrs result)) 
-        val elim = nth (#elims result) (find_index (fn s => s = name) (#names (fst info)))
+          in (fst (dest_Const const) = name) end;      
+        val intros = ind_set_codegen_preproc thy ((map (preprocess_intro thy))
+          (filter is_intro_of (#intrs result)))
+        val pre_elim = nth (#elims result) (find_index (fn s => s = name) (#names (fst info)))
         val nparams = length (Inductive.params_of (#raw_induct result))
-      in (intros, elim, nparams) end
+        val elim = singleton (ind_set_codegen_preproc thy) (preprocess_elim thy nparams pre_elim)
+        val (intros, elim) = if null intros then noclause thy name elim else (intros, elim)
+      in
+        mk_pred_data ((intros, SOME elim, nparams), ([], [], []))
+      end                                                                    
   | NONE => error ("No such predicate: " ^ quote name)
   
 (* updaters *)
 
-fun add_predfun name mode data = let
-    val add = apsnd (cons (mode, mk_predfun_data data))
+fun apfst3 f (x, y, z) =  (f x, y, z)
+fun apsnd3 f (x, y, z) =  (x, f y, z)
+fun aptrd3 f (x, y, z) =  (x, y, f z)
+
+fun add_predfun name mode data =
+  let
+    val add = (apsnd o apfst3 o cons) (mode, mk_predfun_data data)
   in PredData.map (Graph.map_node name (map_pred_data add)) end
 
 fun is_inductive_predicate thy name =
   is_some (try (Inductive.the_inductive (ProofContext.init thy)) name)
 
-fun depending_preds_of thy intros = fold Term.add_consts (map Thm.prop_of intros) [] |> map fst
-    |> filter (fn c => is_inductive_predicate thy c orelse is_registered thy c)
-
+fun depending_preds_of thy (key, value) =
+  let
+    val intros = (#intros o rep_pred_data) value
+  in
+    fold Term.add_const_names (map Thm.prop_of intros) []
+      |> filter (fn c => (not (c = key)) andalso (is_inductive_predicate thy c orelse is_registered thy c))
+  end;
+    
+    
 (* code dependency graph *)    
+(*
 fun dependencies_of thy name =
   let
     val (intros, elim, nparams) = fetch_pred_data thy name 
-    val data = mk_pred_data ((intros, SOME elim, nparams), [])
+    val data = mk_pred_data ((intros, SOME elim, nparams), ([], [], []))
     val keys = depending_preds_of thy intros
   in
     (data, keys)
   end;
+*)
+(* guessing number of parameters *)
+fun find_indexes pred xs =
+  let
+    fun find is n [] = is
+      | find is n (x :: xs) = find (if pred x then (n :: is) else is) (n + 1) xs;
+  in rev (find [] 0 xs) end;
 
-(* TODO: add_edges - by analysing dependencies *)
+fun is_predT (T as Type("fun", [_, _])) = (snd (strip_type T) = HOLogic.boolT)
+  | is_predT _ = false
+  
+fun guess_nparams T =
+  let
+    val argTs = binder_types T
+    val nparams = fold (curry Int.max)
+      (map (fn x => x + 1) (find_indexes is_predT argTs)) 0
+  in nparams end;
+
 fun add_intro thm thy = let
-   val (name, _) = dest_Const (fst (strip_intro_concl 0 (prop_of thm)))
+   val (name, T) = dest_Const (fst (strip_intro_concl 0 (prop_of thm)))
    fun cons_intro gr =
      case try (Graph.get_node gr) name of
        SOME pred_data => Graph.map_node name (map_pred_data
          (apfst (fn (intro, elim, nparams) => (thm::intro, elim, nparams)))) gr
      | NONE =>
        let
-         val nparams = the_default 0 (try (#3 o fetch_pred_data thy) name)
-       in Graph.new_node (name, mk_pred_data (([thm], NONE, nparams), [])) gr end;
+         val nparams = the_default (guess_nparams T)  (try (#nparams o rep_pred_data o (fetch_pred_data thy)) name)
+       in Graph.new_node (name, mk_pred_data (([thm], NONE, nparams), ([], [], []))) gr end;
   in PredData.map cons_intro thy end
 
 fun set_elim thm = let
@@ -333,16 +567,221 @@
 fun set_nparams name nparams = let
     fun set (intros, elim, _ ) = (intros, elim, nparams) 
   in PredData.map (Graph.map_node name (map_pred_data (apfst set))) end
+    
+fun register_predicate (pre_intros, pre_elim, nparams) thy = let
+    val (name, _) = dest_Const (fst (strip_intro_concl nparams (prop_of (hd pre_intros))))
+    (* preprocessing *)
+    val intros = ind_set_codegen_preproc thy (map (preprocess_intro thy) pre_intros)
+    val elim = singleton (ind_set_codegen_preproc thy) (preprocess_elim thy nparams pre_elim)
+  in
+    PredData.map
+      (Graph.new_node (name, mk_pred_data ((intros, SOME elim, nparams), ([], [], [])))) thy
+  end
 
-fun register_predicate (intros, elim, nparams) thy = let
-    val (name, _) = dest_Const (fst (strip_intro_concl nparams (prop_of (hd intros))))
-    fun set _ = (intros, SOME elim, nparams)
+fun set_generator_name pred mode name = 
+  let
+    val set = (apsnd o apsnd3 o cons) (mode, mk_function_data (name, NONE))
   in
-    PredData.map (Graph.new_node (name, mk_pred_data ((intros, SOME elim, nparams), []))
-      #> fold Graph.add_edge (map (pair name) (depending_preds_of thy intros))) thy
+    PredData.map (Graph.map_node pred (map_pred_data set))
+  end
+
+fun set_sizelim_function_name pred mode name = 
+  let
+    val set = (apsnd o aptrd3 o cons) (mode, mk_function_data (name, NONE))
+  in
+    PredData.map (Graph.map_node pred (map_pred_data set))
   end
 
+(** data structures for generic compilation for different monads **)
+
+(* maybe rename functions more generic:
+  mk_predT -> mk_monadT; dest_predT -> dest_monadT
+  mk_single -> mk_return (?)
+*)
+datatype compilation_funs = CompilationFuns of {
+  mk_predT : typ -> typ,
+  dest_predT : typ -> typ,
+  mk_bot : typ -> term,
+  mk_single : term -> term,
+  mk_bind : term * term -> term,
+  mk_sup : term * term -> term,
+  mk_if : term -> term,
+  mk_not : term -> term,
+(*  funT_of : mode -> typ -> typ, *)
+(*  mk_fun_of : theory -> (string * typ) -> mode -> term, *) 
+  mk_map : typ -> typ -> term -> term -> term,
+  lift_pred : term -> term
+};
+
+fun mk_predT (CompilationFuns funs) = #mk_predT funs
+fun dest_predT (CompilationFuns funs) = #dest_predT funs
+fun mk_bot (CompilationFuns funs) = #mk_bot funs
+fun mk_single (CompilationFuns funs) = #mk_single funs
+fun mk_bind (CompilationFuns funs) = #mk_bind funs
+fun mk_sup (CompilationFuns funs) = #mk_sup funs
+fun mk_if (CompilationFuns funs) = #mk_if funs
+fun mk_not (CompilationFuns funs) = #mk_not funs
+(*fun funT_of (CompilationFuns funs) = #funT_of funs*)
+(*fun mk_fun_of (CompilationFuns funs) = #mk_fun_of funs*)
+fun mk_map (CompilationFuns funs) = #mk_map funs
+fun lift_pred (CompilationFuns funs) = #lift_pred funs
+
+fun funT_of compfuns (iss, is) T =
+  let
+    val Ts = binder_types T
+    val (paramTs, (inargTs, outargTs)) = split_mode (iss, is) Ts
+    val paramTs' = map2 (fn NONE => I | SOME is => funT_of compfuns ([], is)) iss paramTs 
+  in
+    (paramTs' @ inargTs) ---> (mk_predT compfuns (mk_tupleT outargTs))
+  end;
+
+fun sizelim_funT_of compfuns (iss, is) T =
+  let
+    val Ts = binder_types T
+    val (paramTs, (inargTs, outargTs)) = split_mode (iss, is) Ts
+    val paramTs' = map2 (fn SOME is => sizelim_funT_of compfuns ([], is) | NONE => I) iss paramTs 
+  in
+    (paramTs' @ inargTs @ [@{typ "code_numeral"}]) ---> (mk_predT compfuns (mk_tupleT outargTs))
+  end;  
+
+fun mk_fun_of compfuns thy (name, T) mode = 
+  Const (predfun_name_of thy name mode, funT_of compfuns mode T)
+
+fun mk_sizelim_fun_of compfuns thy (name, T) mode =
+  Const (sizelim_function_name_of thy name mode, sizelim_funT_of compfuns mode T)
   
+fun mk_generator_of compfuns thy (name, T) mode = 
+  Const (generator_name_of thy name mode, sizelim_funT_of compfuns mode T)
+
+
+structure PredicateCompFuns =
+struct
+
+fun mk_predT T = Type (@{type_name "Predicate.pred"}, [T])
+
+fun dest_predT (Type (@{type_name "Predicate.pred"}, [T])) = T
+  | dest_predT T = raise TYPE ("dest_predT", [T], []);
+
+fun mk_bot T = Const (@{const_name Orderings.bot}, mk_predT T);
+
+fun mk_single t =
+  let val T = fastype_of t
+  in Const(@{const_name Predicate.single}, T --> mk_predT T) $ t end;
+
+fun mk_bind (x, f) =
+  let val T as Type ("fun", [_, U]) = fastype_of f
+  in
+    Const (@{const_name Predicate.bind}, fastype_of x --> T --> U) $ x $ f
+  end;
+
+val mk_sup = HOLogic.mk_binop @{const_name sup};
+
+fun mk_if cond = Const (@{const_name Predicate.if_pred},
+  HOLogic.boolT --> mk_predT HOLogic.unitT) $ cond;
+
+fun mk_not t = let val T = mk_predT HOLogic.unitT
+  in Const (@{const_name Predicate.not_pred}, T --> T) $ t end
+
+fun mk_Enum f =
+  let val T as Type ("fun", [T', _]) = fastype_of f
+  in
+    Const (@{const_name Predicate.Pred}, T --> mk_predT T') $ f    
+  end;
+
+fun mk_Eval (f, x) =
+  let
+    val T = fastype_of x
+  in
+    Const (@{const_name Predicate.eval}, mk_predT T --> T --> HOLogic.boolT) $ f $ x
+  end;
+
+fun mk_map T1 T2 tf tp = Const (@{const_name Predicate.map},
+  (T1 --> T2) --> mk_predT T1 --> mk_predT T2) $ tf $ tp;
+
+val lift_pred = I
+
+val compfuns = CompilationFuns {mk_predT = mk_predT, dest_predT = dest_predT, mk_bot = mk_bot,
+  mk_single = mk_single, mk_bind = mk_bind, mk_sup = mk_sup, mk_if = mk_if, mk_not = mk_not,
+  mk_map = mk_map, lift_pred = lift_pred};
+
+end;
+
+(* termify_code:
+val termT = Type ("Code_Eval.term", []);
+fun termifyT T = HOLogic.mk_prodT (T, HOLogic.unitT --> termT)
+*)
+(*
+fun lift_random random =
+  let
+    val T = dest_randomT (fastype_of random)
+  in
+    mk_scomp (random,
+      mk_fun_comp (HOLogic.pair_const (PredicateCompFuns.mk_predT T) @{typ Random.seed},
+        mk_fun_comp (Const (@{const_name Predicate.single}, T --> (PredicateCompFuns.mk_predT T)),
+          Const (@{const_name "fst"}, HOLogic.mk_prodT (T, @{typ "unit => term"}) --> T)))) 
+  end;
+*)
+ 
+structure RPredCompFuns =
+struct
+
+fun mk_rpredT T =
+  @{typ "Random.seed"} --> HOLogic.mk_prodT (PredicateCompFuns.mk_predT T, @{typ "Random.seed"})
+
+fun dest_rpredT (Type ("fun", [_,
+  Type (@{type_name "*"}, [Type (@{type_name "Predicate.pred"}, [T]), _])])) = T
+  | dest_rpredT T = raise TYPE ("dest_rpredT", [T], []); 
+
+fun mk_bot T = Const(@{const_name RPred.bot}, mk_rpredT T)
+
+fun mk_single t =
+  let
+    val T = fastype_of t
+  in
+    Const (@{const_name RPred.return}, T --> mk_rpredT T) $ t
+  end;
+
+fun mk_bind (x, f) =
+  let
+    val T as (Type ("fun", [_, U])) = fastype_of f
+  in
+    Const (@{const_name RPred.bind}, fastype_of x --> T --> U) $ x $ f
+  end
+
+val mk_sup = HOLogic.mk_binop @{const_name RPred.supp}
+
+fun mk_if cond = Const (@{const_name RPred.if_rpred},
+  HOLogic.boolT --> mk_rpredT HOLogic.unitT) $ cond;
+
+fun mk_not t = error "Negation is not defined for RPred"
+
+fun mk_map t = error "FIXME" (*FIXME*)
+
+fun lift_pred t =
+  let
+    val T = PredicateCompFuns.dest_predT (fastype_of t)
+    val lift_predT = PredicateCompFuns.mk_predT T --> mk_rpredT T 
+  in
+    Const (@{const_name "RPred.lift_pred"}, lift_predT) $ t  
+  end;
+
+val compfuns = CompilationFuns {mk_predT = mk_rpredT, dest_predT = dest_rpredT, mk_bot = mk_bot,
+    mk_single = mk_single, mk_bind = mk_bind, mk_sup = mk_sup, mk_if = mk_if, mk_not = mk_not,
+    mk_map = mk_map, lift_pred = lift_pred};
+
+end;
+(* for external use with interactive mode *)
+val rpred_compfuns = RPredCompFuns.compfuns;
+
+fun lift_random random =
+  let
+    val T = dest_randomT (fastype_of random)
+  in
+    Const (@{const_name lift_random}, (@{typ Random.seed} -->
+      HOLogic.mk_prodT (HOLogic.mk_prodT (T, @{typ "unit => term"}), @{typ Random.seed})) --> 
+      RPredCompFuns.mk_rpredT T) $ random
+  end;
+ 
 (* Mode analysis *)
 
 (*** check if a term contains only constructor functions ***)
@@ -371,12 +810,6 @@
 fun term_vTs tm =
   fold_aterms (fn Free xT => cons xT | _ => I) tm [];
 
-fun get_args is ts = let
-  fun get_args' _ _ [] = ([], [])
-    | get_args' is i (t::ts) = (if member (op =) is i then apfst else apsnd) (cons t)
-        (get_args' is (i+1) ts)
-in get_args' is 1 ts end
-
 (*FIXME this function should not be named merge... make it local instead*)
 fun merge xs [] = xs
   | merge [] ys = ys
@@ -394,11 +827,10 @@
 
 fun cprods xss = foldr (map op :: o cprod) [[]] xss;
 
-datatype hmode = Mode of mode * int list * hmode option list; (*FIXME don't understand
-  why there is another mode type tmode !?*)
 
   
 (*TODO: cleanup function and put together with modes_of_term *)
+(*
 fun modes_of_param default modes t = let
     val (vs, t') = strip_abs t
     val b = length vs
@@ -409,8 +841,8 @@
               error ("Too few arguments for inductive predicate " ^ name)
             else chop (length iss) args;
           val k = length args2;
-          val perm = map (fn i => (find_index (fn t => t = Bound (b - i)) args2) + 1)
-            (1 upto b)
+          val perm = map (fn i => (find_index_eq (Bound (b - i)) args2) + 1)
+            (1 upto b)  
           val partial_mode = (1 upto k) \\ perm
         in
           if not (partial_mode subset is) then [] else
@@ -432,7 +864,9 @@
     | (Free (name, _), args) => the (mk_modes name args)
     | _ => default end
   
-and modes_of_term modes t =
+and
+*)
+fun modes_of_term modes t =
   let
     val ks = 1 upto length (binder_types (fastype_of t));
     val default = [Mode (([], ks), ks, [])];
@@ -455,21 +889,20 @@
           end
         end)) (AList.lookup op = modes name)
 
-  in (case strip_comb t of
+  in
+    case strip_comb (Envir.eta_contract t) of
       (Const (name, _), args) => the_default default (mk_modes name args)
     | (Var ((name, _), _), args) => the (mk_modes name args)
     | (Free (name, _), args) => the (mk_modes name args)
-    | (Abs _, []) => modes_of_param default modes t 
-    | _ => default)
+    | (Abs _, []) => error "Abs at param position" (* modes_of_param default modes t *)
+    | _ => default
   end
-
-datatype indprem = Prem of term list * term | Negprem of term list * term | Sidecond of term;
-
+  
 fun select_mode_prem thy modes vs ps =
   find_first (is_some o snd) (ps ~~ map
     (fn Prem (us, t) => find_first (fn Mode (_, is, _) =>
           let
-            val (in_ts, out_ts) = get_args is us;
+            val (in_ts, out_ts) = split_smode is us;
             val (out_ts', in_ts') = List.partition (is_constrt thy) out_ts;
             val vTs = maps term_vTs out_ts';
             val dupTs = map snd (duplicates (op =) vTs) @
@@ -492,69 +925,139 @@
           else NONE
       ) ps);
 
-fun check_mode_clause thy param_vs modes (iss, is) (ts, ps) =
+fun fold_prem f (Prem (args, _)) = fold f args
+  | fold_prem f (Negprem (args, _)) = fold f args
+  | fold_prem f (Sidecond t) = f t
+
+fun all_subsets [] = [[]]
+  | all_subsets (x::xs) = let val xss' = all_subsets xs in xss' @ (map (cons x) xss') end
+
+fun generator vTs v = 
+  let
+    val T = the (AList.lookup (op =) vTs v)
+  in
+    (Generator (v, T), Mode (([], []), [], []))
+  end;
+
+fun gen_prem (Prem (us, t)) = GeneratorPrem (us, t) 
+  | gen_prem _ = error "gen_prem : invalid input for gen_prem"
+
+fun param_gen_prem param_vs (p as Prem (us, t as Free (v, _))) =
+  if member (op =) param_vs v then
+    GeneratorPrem (us, t)
+  else p  
+  | param_gen_prem param_vs p = p
+  
+fun check_mode_clause with_generator thy param_vs modes gen_modes (iss, is) (ts, ps) =
   let
     val modes' = modes @ List.mapPartial
       (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))
-        (param_vs ~~ iss); 
-    fun check_mode_prems vs [] = SOME vs
-      | check_mode_prems vs ps = (case select_mode_prem thy modes' vs ps of
-          NONE => NONE
-        | SOME (x, _) => check_mode_prems
-            (case x of Prem (us, _) => vs union terms_vs us | _ => vs)
-            (filter_out (equal x) ps))
-    val (in_ts, in_ts') = List.partition (is_constrt thy) (fst (get_args is ts));
+        (param_vs ~~ iss);
+    val gen_modes' = gen_modes @ List.mapPartial
+      (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))
+        (param_vs ~~ iss);  
+    val vTs = distinct (op =) ((fold o fold_prem) Term.add_frees ps (fold Term.add_frees ts []))
+    val prem_vs = distinct (op =) ((fold o fold_prem) Term.add_free_names ps [])
+    fun check_mode_prems acc_ps vs [] = SOME (acc_ps, vs)
+      | check_mode_prems acc_ps vs ps = (case select_mode_prem thy modes' vs ps of
+          NONE =>
+            (if with_generator then
+              (case select_mode_prem thy gen_modes' vs ps of
+                  SOME (p, SOME mode) => check_mode_prems ((gen_prem p, mode) :: acc_ps) 
+                  (case p of Prem (us, _) => vs union terms_vs us | _ => vs)
+                  (filter_out (equal p) ps)
+                | NONE =>
+                  let 
+                    val all_generator_vs = all_subsets (prem_vs \\ vs) |> sort (int_ord o (pairself length))
+                  in
+                    case (find_first (fn generator_vs => is_some
+                      (select_mode_prem thy modes' (vs union generator_vs) ps)) all_generator_vs) of
+                      SOME generator_vs => check_mode_prems ((map (generator vTs) generator_vs) @ acc_ps)
+                        (vs union generator_vs) ps
+                    | NONE => NONE
+                  end)
+            else
+              NONE)
+        | SOME (p, SOME mode) => check_mode_prems ((if with_generator then param_gen_prem param_vs p else p, mode) :: acc_ps) 
+            (case p of Prem (us, _) => vs union terms_vs us | _ => vs)
+            (filter_out (equal p) ps))
+    val (in_ts, in_ts') = List.partition (is_constrt thy) (fst (split_smode is ts));
     val in_vs = terms_vs in_ts;
     val concl_vs = terms_vs ts
   in
-    forall is_eqT (map snd (duplicates (op =) (maps term_vTs in_ts))) andalso
-    forall (is_eqT o fastype_of) in_ts' andalso
-    (case check_mode_prems (param_vs union in_vs) ps of
-       NONE => false
-     | SOME vs => concl_vs subset vs)
+    if forall is_eqT (map snd (duplicates (op =) (maps term_vTs in_ts))) andalso
+    forall (is_eqT o fastype_of) in_ts' then
+      case check_mode_prems [] (param_vs union in_vs) ps of
+         NONE => NONE
+       | SOME (acc_ps, vs) =>
+         if with_generator then
+           SOME (ts, (rev acc_ps) @ (map (generator vTs) (concl_vs \\ vs))) 
+         else
+           if concl_vs subset vs then SOME (ts, rev acc_ps) else NONE
+    else NONE
   end;
 
-fun check_modes_pred thy param_vs preds modes (p, ms) =
+fun check_modes_pred with_generator thy param_vs preds modes gen_modes (p, ms) =
   let val SOME rs = AList.lookup (op =) preds p
   in (p, List.filter (fn m => case find_index
-    (not o check_mode_clause thy param_vs modes m) rs of
+    (is_none o check_mode_clause with_generator thy param_vs modes gen_modes m) rs of
       ~1 => true
-    | i => (tracing ("Clause " ^ string_of_int (i+1) ^ " of " ^
+    | i => (Output.tracing ("Clause " ^ string_of_int (i + 1) ^ " of " ^
       p ^ " violates mode " ^ string_of_mode m); false)) ms)
   end;
 
+fun get_modes_pred with_generator thy param_vs preds modes gen_modes (p, ms) =
+  let
+    val SOME rs = AList.lookup (op =) preds p 
+  in
+    (p, map (fn m =>
+      (m, map (the o check_mode_clause with_generator thy param_vs modes gen_modes m) rs)) ms)
+  end;
+  
 fun fixp f (x : (string * mode list) list) =
   let val y = f x
   in if x = y then x else fixp f y end;
 
-fun infer_modes thy extra_modes arities param_vs preds = fixp (fn modes =>
-  map (check_modes_pred thy param_vs preds (modes @ extra_modes)) modes)
-    (map (fn (s, (ks, k)) => (s, cprod (cprods (map
-      (fn NONE => [NONE]
-        | SOME k' => map SOME (subsets 1 k')) ks),
-      subsets 1 k))) arities);
+fun modes_of_arities arities =
+  (map (fn (s, (ks, k)) => (s, cprod (cprods (map
+            (fn NONE => [NONE]
+              | SOME k' => map SOME (subsets 1 k')) ks),
+            subsets 1 k))) arities)
+  
+fun infer_modes with_generator thy extra_modes arities param_vs preds =
+  let
+    val modes =
+      fixp (fn modes =>
+        map (check_modes_pred with_generator thy param_vs preds (modes @ extra_modes) []) modes)
+          (modes_of_arities arities)
+  in
+    map (get_modes_pred with_generator thy param_vs preds (modes @ extra_modes) []) modes
+  end;
 
+fun remove_from rem [] = []
+  | remove_from rem ((k, vs) :: xs) =
+    (case AList.lookup (op =) rem k of
+      NONE => (k, vs)
+    | SOME vs' => (k, vs \\ vs'))
+    :: remove_from rem xs
+    
+fun infer_modes_with_generator thy extra_modes arities param_vs preds =
+  let
+    val prednames = map fst preds
+    val extra_modes = all_modes_of thy
+    val gen_modes = all_generator_modes_of thy
+      |> filter_out (fn (name, _) => member (op =) prednames name)
+    val starting_modes = remove_from extra_modes (modes_of_arities arities) 
+    val modes =
+      fixp (fn modes =>
+        map (check_modes_pred true thy param_vs preds extra_modes (gen_modes @ modes)) modes)
+         starting_modes 
+  in
+    map (get_modes_pred true thy param_vs preds extra_modes (gen_modes @ modes)) modes
+  end;
 
 (* term construction *)
 
-(* for simple modes (e.g. parameters) only: better call it param_funT *)
-(* or even better: remove it and only use funT'_of - some modifications to funT'_of necessary *) 
-fun funT_of T NONE = T
-  | funT_of T (SOME mode) = let
-     val Ts = binder_types T;
-     val (Us1, Us2) = get_args mode Ts
-   in Us1 ---> (mk_pred_enumT (HOLogic.mk_tupleT Us2)) end;
-
-fun funT'_of (iss, is) T = let
-    val Ts = binder_types T
-    val (paramTs, argTs) = chop (length iss) Ts
-    val paramTs' = map2 (fn SOME is => funT'_of ([], is) | NONE => I) iss paramTs 
-    val (inargTs, outargTs) = get_args is argTs
-  in
-    (paramTs' @ inargTs) ---> (mk_pred_enumT (HOLogic.mk_tupleT outargTs))
-  end; 
-
-
 fun mk_v (names, vs) s T = (case AList.lookup (op =) vs s of
       NONE => (Free (s, T), (names, (s, [])::vs))
     | SOME xs =>
@@ -573,104 +1076,135 @@
       in (t' $ u', nvs'') end
   | distinct_v x nvs = (x, nvs);
 
-fun compile_match thy eqs eqs' out_ts success_t =
-  let 
+fun compile_match thy compfuns eqs eqs' out_ts success_t =
+  let
     val eqs'' = maps mk_eq eqs @ eqs'
     val names = fold Term.add_free_names (success_t :: eqs'' @ out_ts) [];
     val name = Name.variant names "x";
     val name' = Name.variant (name :: names) "y";
-    val T = HOLogic.mk_tupleT (map fastype_of out_ts);
+    val T = mk_tupleT (map fastype_of out_ts);
     val U = fastype_of success_t;
-    val U' = dest_pred_enumT U;
+    val U' = dest_predT compfuns U;
     val v = Free (name, T);
     val v' = Free (name', T);
   in
     lambda v (fst (Datatype.make_case
       (ProofContext.init thy) false [] v
-      [(HOLogic.mk_tuple out_ts,
+      [(mk_tuple out_ts,
         if null eqs'' then success_t
         else Const (@{const_name HOL.If}, HOLogic.boolT --> U --> U --> U) $
           foldr1 HOLogic.mk_conj eqs'' $ success_t $
-            mk_empty U'),
-       (v', mk_empty U')]))
+            mk_bot compfuns U'),
+       (v', mk_bot compfuns U')]))
   end;
 
-fun compile_param_ext thy modes (NONE, t) = t
-  | compile_param_ext thy modes (m as SOME (Mode ((iss, is'), is, ms)), t) =
+(*FIXME function can be removed*)
+fun mk_funcomp f t =
+  let
+    val names = Term.add_free_names t [];
+    val Ts = binder_types (fastype_of t);
+    val vs = map Free
+      (Name.variant_list names (replicate (length Ts) "x") ~~ Ts)
+  in
+    fold_rev lambda vs (f (list_comb (t, vs)))
+  end;
+(*
+fun compile_param_ext thy compfuns modes (NONE, t) = t
+  | compile_param_ext thy compfuns modes (m as SOME (Mode ((iss, is'), is, ms)), t) =
       let
         val (vs, u) = strip_abs t
-        val (ivs, ovs) = get_args is vs    
+        val (ivs, ovs) = split_mode is vs    
         val (f, args) = strip_comb u
         val (params, args') = chop (length ms) args
-        val (inargs, outargs) = get_args is' args'
+        val (inargs, outargs) = split_mode is' args'
         val b = length vs
-        val perm = map (fn i => find_index (fn t => t = Bound (b - i)) args' + 1) (1 upto b)
+        val perm = map (fn i => (find_index_eq (Bound (b - i)) args') + 1) (1 upto b)
         val outp_perm =
-          snd (get_args is perm)
+          snd (split_mode is perm)
           |> map (fn i => i - length (filter (fn x => x < i) is'))
-        val names = [] (* TODO *)
+        val names = [] -- TODO
         val out_names = Name.variant_list names (replicate (length outargs) "x")
         val f' = case f of
             Const (name, T) =>
               if AList.defined op = modes name then
-                Const (predfun_name_of thy name (iss, is'), funT'_of (iss, is') T)
+                mk_predfun_of thy compfuns (name, T) (iss, is')
               else error "compile param: Not an inductive predicate with correct mode"
-          | Free (name, T) => Free (name, funT_of T (SOME is'))
-        val outTs = HOLogic.strip_tupleT (dest_pred_enumT (body_type (fastype_of f')))
+          | Free (name, T) => Free (name, param_funT_of compfuns T (SOME is'))
+        val outTs = dest_tupleT (dest_predT compfuns (body_type (fastype_of f')))
         val out_vs = map Free (out_names ~~ outTs)
         val params' = map (compile_param thy modes) (ms ~~ params)
         val f_app = list_comb (f', params' @ inargs)
-        val single_t = (mk_single (HOLogic.mk_tuple (map (fn i => nth out_vs (i - 1)) outp_perm)))
-        val match_t = compile_match thy [] [] out_vs single_t
+        val single_t = (mk_single compfuns (mk_tuple (map (fn i => nth out_vs (i - 1)) outp_perm)))
+        val match_t = compile_match thy compfuns [] [] out_vs single_t
       in list_abs (ivs,
-        mk_bind (f_app, match_t))
+        mk_bind compfuns (f_app, match_t))
       end
-  | compile_param_ext _ _ _ = error "compile params"
+  | compile_param_ext _ _ _ _ = error "compile params"
+*)
 
-and compile_param thy modes (NONE, t) = t
- | compile_param thy modes (m as SOME (Mode ((iss, is'), is, ms)), t) =
-   (* (case t of
-     Abs _ => error "compile_param: Invalid term" *) (* compile_param_ext thy modes (m, t) *)
-   (*  |  _ => let *)
-   let  
+fun compile_param size thy compfuns (NONE, t) = t
+  | compile_param size thy compfuns (m as SOME (Mode ((iss, is'), is, ms)), t) =
+   let
      val (f, args) = strip_comb (Envir.eta_contract t)
      val (params, args') = chop (length ms) args
-     val params' = map (compile_param thy modes) (ms ~~ params)
-     val f' = case f of
-        Const (name, T) =>
-          if AList.defined op = modes name then
-             Const (predfun_name_of thy name (iss, is'), funT'_of (iss, is') T)
-          else error "compile param: Not an inductive predicate with correct mode"
-      | Free (name, T) => Free (name, funT_of T (SOME is'))
+     val params' = map (compile_param size thy compfuns) (ms ~~ params)
+     val mk_fun_of = case size of NONE => mk_fun_of | SOME _ => mk_sizelim_fun_of
+     val funT_of = case size of NONE => funT_of | SOME _ => sizelim_funT_of
+     val f' =
+       case f of
+         Const (name, T) =>
+           mk_fun_of compfuns thy (name, T) (iss, is')
+       | Free (name, T) => Free (name, funT_of compfuns (iss, is') T)
+       | _ => error ("PredicateCompiler: illegal parameter term")
    in list_comb (f', params' @ args') end
- | compile_param _ _ _ = error "compile params"
-  
-  
-fun compile_expr thy modes (SOME (Mode (mode, is, ms)), t) =
-      (case strip_comb t of
-         (Const (name, T), params) =>
-           if AList.defined op = modes name then
-             let
-               val (Ts, Us) = get_args is
-                 (curry Library.drop (length ms) (fst (strip_type T)))
-               val params' = map (compile_param thy modes) (ms ~~ params)
-             in list_comb (Const (predfun_name_of thy name mode, ((map fastype_of params') @ Ts) --->
-               mk_pred_enumT (HOLogic.mk_tupleT Us)), params')
-             end
-           else error "not a valid inductive expression"
-       | (Free (name, T), args) =>
-         (*if name mem param_vs then *)
-         (* Higher order mode call *)
-         let val r = Free (name, funT_of T (SOME is))
-         in list_comb (r, args) end)
-  | compile_expr _ _ _ = error "not a valid inductive expression"
+   
+fun compile_expr size thy ((Mode (mode, is, ms)), t) =
+  case strip_comb t of
+    (Const (name, T), params) =>
+       let
+         val params' = map (compile_param size thy PredicateCompFuns.compfuns) (ms ~~ params)
+         val mk_fun_of = case size of NONE => mk_fun_of | SOME _ => mk_sizelim_fun_of
+       in
+         list_comb (mk_fun_of PredicateCompFuns.compfuns thy (name, T) mode, params')
+       end
+  | (Free (name, T), args) =>
+       let 
+         val funT_of = case size of NONE => funT_of | SOME _ => sizelim_funT_of 
+       in
+         list_comb (Free (name, funT_of PredicateCompFuns.compfuns ([], is) T), args)
+       end;
+       
+fun compile_gen_expr size thy compfuns ((Mode (mode, is, ms)), t) =
+  case strip_comb t of
+    (Const (name, T), params) =>
+      let
+        val params' = map (compile_param size thy compfuns) (ms ~~ params)
+      in
+        list_comb (mk_generator_of compfuns thy (name, T) mode, params')
+      end
+    | (Free (name, T), args) =>
+      list_comb (Free (name, sizelim_funT_of RPredCompFuns.compfuns ([], is) T), args)
+          
+(** specific rpred functions -- move them to the correct place in this file *)
 
+(* uncommented termify code; causes more trouble than expected at first *) 
+(*
+fun mk_valtermify_term (t as Const (c, T)) = HOLogic.mk_prod (t, Abs ("u", HOLogic.unitT, HOLogic.reflect_term t))
+  | mk_valtermify_term (Free (x, T)) = Free (x, termifyT T) 
+  | mk_valtermify_term (t1 $ t2) =
+    let
+      val T = fastype_of t1
+      val (T1, T2) = dest_funT T
+      val t1' = mk_valtermify_term t1
+      val t2' = mk_valtermify_term t2
+    in
+      Const ("Code_Eval.valapp", termifyT T --> termifyT T1 --> termifyT T2) $ t1' $ t2'
+    end
+  | mk_valtermify_term _ = error "Not a valid term for mk_valtermify_term"
+*)
 
-fun compile_clause thy all_vs param_vs modes (iss, is) (ts, ps) inp =
+fun compile_clause compfuns size final_term thy all_vs param_vs (iss, is) inp (ts, moded_ps) =
   let
-    val modes' = modes @ List.mapPartial
-      (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))
-        (param_vs ~~ iss);
     fun check_constrt t (names, eqs) =
       if is_constrt thy t then (t, (names, eqs)) else
         let
@@ -678,7 +1212,7 @@
           val v = Free (s, fastype_of t)
         in (v, (s::names, HOLogic.mk_eq (v, t)::eqs)) end;
 
-    val (in_ts, out_ts) = get_args is ts;
+    val (in_ts, out_ts) = split_smode is ts;
     val (in_ts', (all_vs', eqs)) =
       fold_map check_constrt in_ts (all_vs, []);
 
@@ -689,15 +1223,16 @@
             val (out_ts''', (names'', constr_vs)) = fold_map distinct_v
               out_ts'' (names', map (rpair []) vs);
           in
-            compile_match thy constr_vs (eqs @ eqs') out_ts'''
-              (mk_single (HOLogic.mk_tuple out_ts))
+          (* termify code:
+            compile_match thy compfuns constr_vs (eqs @ eqs') out_ts'''
+              (mk_single compfuns (mk_tuple (map mk_valtermify_term out_ts)))
+           *)
+            compile_match thy compfuns constr_vs (eqs @ eqs') out_ts'''
+              (final_term out_ts)
           end
-      | compile_prems out_ts vs names ps =
+      | compile_prems out_ts vs names ((p, mode as Mode ((_, is), _, _)) :: ps) =
           let
             val vs' = distinct (op =) (flat (vs :: map term_vs out_ts));
-            val SOME (p, mode as SOME (Mode (_, js, _))) =
-              select_mode_prem thy modes' vs' ps
-            val ps' = filter_out (equal p) ps
             val (out_ts', (names', eqs)) =
               fold_map check_constrt out_ts (names, [])
             val (out_ts'', (names'', constr_vs')) = fold_map distinct_v
@@ -705,67 +1240,97 @@
             val (compiled_clause, rest) = case p of
                Prem (us, t) =>
                  let
-                   val (in_ts, out_ts''') = get_args js us;
-                   val u = list_comb (compile_expr thy modes (mode, t), in_ts)
-                   val rest = compile_prems out_ts''' vs' names'' ps'
+                   val (in_ts, out_ts''') = split_smode is us;
+                   val args = case size of
+                     NONE => in_ts
+                   | SOME size_t => in_ts @ [size_t]
+                   val u = lift_pred compfuns
+                     (list_comb (compile_expr size thy (mode, t), args))                     
+                   val rest = compile_prems out_ts''' vs' names'' ps
                  in
                    (u, rest)
                  end
              | Negprem (us, t) =>
                  let
-                   val (in_ts, out_ts''') = get_args js us
-                   val u = list_comb (compile_expr thy modes (mode, t), in_ts)
-                   val rest = compile_prems out_ts''' vs' names'' ps'
+                   val (in_ts, out_ts''') = split_smode is us
+                   val u = lift_pred compfuns
+                     (mk_not PredicateCompFuns.compfuns (list_comb (compile_expr NONE thy (mode, t), in_ts)))
+                   val rest = compile_prems out_ts''' vs' names'' ps
                  in
-                   (mk_not_pred u, rest)
+                   (u, rest)
                  end
              | Sidecond t =>
                  let
-                   val rest = compile_prems [] vs' names'' ps';
+                   val rest = compile_prems [] vs' names'' ps;
                  in
-                   (mk_if_predenum t, rest)
+                   (mk_if compfuns t, rest)
+                 end
+             | GeneratorPrem (us, t) =>
+                 let
+                   val (in_ts, out_ts''') = split_smode is us;
+                   val args = case size of
+                     NONE => in_ts
+                   | SOME size_t => in_ts @ [size_t]
+                   val u = list_comb (compile_gen_expr size thy compfuns (mode, t), args)
+                   val rest = compile_prems out_ts''' vs' names'' ps
+                 in
+                   (u, rest)
+                 end
+             | Generator (v, T) =>
+                 let
+                   val u = lift_random (HOLogic.mk_random T @{term "1::code_numeral"})
+                   val rest = compile_prems [Free (v, T)]  vs' names'' ps;
+                 in
+                   (u, rest)
                  end
           in
-            compile_match thy constr_vs' eqs out_ts'' 
-              (mk_bind (compiled_clause, rest))
+            compile_match thy compfuns constr_vs' eqs out_ts'' 
+              (mk_bind compfuns (compiled_clause, rest))
           end
-    val prem_t = compile_prems in_ts' param_vs all_vs' ps;
+    val prem_t = compile_prems in_ts' param_vs all_vs' moded_ps;
   in
-    mk_bind (mk_single inp, prem_t)
+    mk_bind compfuns (mk_single compfuns inp, prem_t)
   end
 
-fun compile_pred thy all_vs param_vs modes s T cls mode =
+fun compile_pred compfuns mk_fun_of use_size thy all_vs param_vs s T mode moded_cls =
   let
-    val Ts = binder_types T;
-    val (Ts1, Ts2) = chop (length param_vs) Ts;
-    val Ts1' = map2 funT_of Ts1 (fst mode)
-    val (Us1, Us2) = get_args (snd mode) Ts2;
-    val xnames = Name.variant_list param_vs
+    val (Ts1, (Us1, Us2)) = split_mode mode (binder_types T)
+    val funT_of = if use_size then sizelim_funT_of else funT_of 
+    val Ts1' = map2 (fn NONE => I | SOME is => funT_of compfuns ([], is)) (fst mode) Ts1
+    val xnames = Name.variant_list (all_vs @ param_vs)
       (map (fn i => "x" ^ string_of_int i) (snd mode));
+    val size_name = Name.variant (all_vs @ param_vs @ xnames) "size"
+    (* termify code: val xs = map2 (fn s => fn T => Free (s, termifyT T)) xnames Us1; *)
     val xs = map2 (fn s => fn T => Free (s, T)) xnames Us1;
+    val params = map2 (fn s => fn T => Free (s, T)) param_vs Ts1'
+    val size = Free (size_name, @{typ "code_numeral"})
+    val decr_size =
+      if use_size then
+        SOME (Const ("HOL.minus_class.minus", @{typ "code_numeral => code_numeral => code_numeral"})
+          $ size $ Const ("HOL.one_class.one", @{typ "Code_Numeral.code_numeral"}))
+      else
+        NONE
     val cl_ts =
-      map (fn cl => compile_clause thy
-        all_vs param_vs modes mode cl (HOLogic.mk_tuple xs)) cls;
-    val mode_id = predfun_name_of thy s mode
+      map (compile_clause compfuns decr_size (fn out_ts => mk_single compfuns (mk_tuple out_ts))
+        thy all_vs param_vs mode (mk_tuple xs)) moded_cls;
+    val t = foldr1 (mk_sup compfuns) cl_ts
+    val T' = mk_predT compfuns (mk_tupleT Us2)
+    val size_t = Const (@{const_name "If"}, @{typ bool} --> T' --> T' --> T')
+      $ HOLogic.mk_eq (size, @{term "0 :: code_numeral"})
+      $ mk_bot compfuns (dest_predT compfuns T') $ t
+    val fun_const = mk_fun_of compfuns thy (s, T) mode
+    val eq = if use_size then
+      (list_comb (fun_const, params @ xs @ [size]), size_t)
+    else
+      (list_comb (fun_const, params @ xs), t)
   in
-    HOLogic.mk_Trueprop (HOLogic.mk_eq
-      (list_comb (Const (mode_id, (Ts1' @ Us1) --->
-           mk_pred_enumT (HOLogic.mk_tupleT Us2)),
-         map2 (fn s => fn T => Free (s, T)) param_vs Ts1' @ xs),
-       foldr1 mk_sup cl_ts))
+    HOLogic.mk_Trueprop (HOLogic.mk_eq eq)
   end;
-
-fun compile_preds thy all_vs param_vs modes preds =
-  map (fn (s, (T, cls)) =>
-    map (compile_pred thy all_vs param_vs modes s T cls)
-      ((the o AList.lookup (op =) modes) s)) preds;
-
-
+  
 (* special setup for simpset *)                  
 val HOL_basic_ss' = HOL_basic_ss setSolver 
   (mk_solver "all_tac_solver" (fn _ => fn _ => all_tac))
 
-
 (* Definition of executable functions and their intro and elim rules *)
 
 fun print_arities arities = tracing ("Arities:\n" ^
@@ -780,41 +1345,40 @@
   val argnames = Name.variant_list names
         (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts)));
   val args = map Free (argnames ~~ Ts)
-  val (inargs, outargs) = get_args mode args
-  val r = mk_Eval (list_comb (x, inargs), HOLogic.mk_tuple outargs)
+  val (inargs, outargs) = split_smode mode args
+  val r = PredicateCompFuns.mk_Eval (list_comb (x, inargs), mk_tuple outargs)
   val t = fold_rev lambda args r 
 in
   (t, argnames @ names)
 end;
 
-fun create_intro_elim_rule nparams mode defthm mode_id funT pred thy =
+fun create_intro_elim_rule (mode as (iss, is)) defthm mode_id funT pred thy =
 let
   val Ts = binder_types (fastype_of pred)
   val funtrm = Const (mode_id, funT)
   val argnames = Name.variant_list []
         (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts)));
-  val (Ts1, Ts2) = chop nparams Ts;
-  val Ts1' = map2 funT_of Ts1 (fst mode)
+  val (Ts1, Ts2) = chop (length iss) Ts;
+  val Ts1' = map2 (fn NONE => I | SOME is => funT_of (PredicateCompFuns.compfuns) ([], is)) iss Ts1
   val args = map Free (argnames ~~ (Ts1' @ Ts2))
-  val (params, io_args) = chop nparams args
-  val (inargs, outargs) = get_args (snd mode) io_args
+  val (params, ioargs) = chop (length iss) args
+  val (inargs, outargs) = split_smode is ioargs
   val param_names = Name.variant_list argnames
-    (map (fn i => "p" ^ string_of_int i) (1 upto nparams))
+    (map (fn i => "p" ^ string_of_int i) (1 upto (length iss)))
   val param_vs = map Free (param_names ~~ Ts1)
-  val (params', names) = fold_map mk_Eval_of ((params ~~ Ts1) ~~ (fst mode)) []
-  val predpropI = HOLogic.mk_Trueprop (list_comb (pred, param_vs @ io_args))
-  val predpropE = HOLogic.mk_Trueprop (list_comb (pred, params' @ io_args))
+  val (params', names) = fold_map mk_Eval_of ((params ~~ Ts1) ~~ iss) []
+  val predpropI = HOLogic.mk_Trueprop (list_comb (pred, param_vs @ ioargs))
+  val predpropE = HOLogic.mk_Trueprop (list_comb (pred, params' @ ioargs))
   val param_eqs = map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (param_vs ~~ params')
   val funargs = params @ inargs
-  val funpropE = HOLogic.mk_Trueprop (mk_Eval (list_comb (funtrm, funargs),
-                  if null outargs then Free("y", HOLogic.unitT) else HOLogic.mk_tuple outargs))
-  val funpropI = HOLogic.mk_Trueprop (mk_Eval (list_comb (funtrm, funargs),
-                   HOLogic.mk_tuple outargs))
+  val funpropE = HOLogic.mk_Trueprop (PredicateCompFuns.mk_Eval (list_comb (funtrm, funargs),
+                  if null outargs then Free("y", HOLogic.unitT) else mk_tuple outargs))
+  val funpropI = HOLogic.mk_Trueprop (PredicateCompFuns.mk_Eval (list_comb (funtrm, funargs),
+                   mk_tuple outargs))
   val introtrm = Logic.list_implies (predpropI :: param_eqs, funpropI)
-  val _ = Output.tracing (Syntax.string_of_term_global thy introtrm) 
   val simprules = [defthm, @{thm eval_pred},
                    @{thm "split_beta"}, @{thm "fst_conv"}, @{thm "snd_conv"}]
-  val unfolddef_tac = (Simplifier.asm_full_simp_tac (HOL_basic_ss addsimps simprules) 1)
+  val unfolddef_tac = Simplifier.asm_full_simp_tac (HOL_basic_ss addsimps simprules) 1
   val introthm = Goal.prove (ProofContext.init thy) (argnames @ param_names @ ["y"]) [] introtrm (fn {...} => unfolddef_tac)
   val P = HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT));
   val elimtrm = Logic.list_implies ([funpropE, Logic.mk_implies (predpropE, P)], P)
@@ -823,53 +1387,92 @@
   (introthm, elimthm)
 end;
 
-fun create_definitions preds nparams (name, modes) thy =
+fun create_constname_of_mode thy prefix name mode = 
   let
-    val _ = tracing "create definitions"
+    fun string_of_mode mode = if null mode then "0"
+      else space_implode "_" (map string_of_int mode)
+    val HOmode = space_implode "_and_"
+      (fold (fn NONE => I | SOME mode => cons (string_of_mode mode)) (fst mode) [])
+  in
+    (Sign.full_bname thy (prefix ^ (Long_Name.base_name name))) ^
+      (if HOmode = "" then "_" else "_for_" ^ HOmode ^ "_yields_") ^ (string_of_mode (snd mode))
+  end;
+  
+fun create_definitions preds (name, modes) thy =
+  let
+    val compfuns = PredicateCompFuns.compfuns
     val T = AList.lookup (op =) preds name |> the
-    fun create_definition mode thy = let
-      fun string_of_mode mode = if null mode then "0"
-        else space_implode "_" (map string_of_int mode)
-      val HOmode = let
-        fun string_of_HOmode m s = case m of NONE => s | SOME mode => s ^ "__" ^ (string_of_mode mode)    
-        in (fold string_of_HOmode (fst mode) "") end;
-      val mode_id = name ^ (if HOmode = "" then "_" else HOmode ^ "___")
-        ^ (string_of_mode (snd mode))
-      val Ts = binder_types T;
-      val (Ts1, Ts2) = chop nparams Ts;
-      val Ts1' = map2 funT_of Ts1 (fst mode)
-      val (Us1, Us2) = get_args (snd mode) Ts2;
+    fun create_definition (mode as (iss, is)) thy = let
+      val mode_cname = create_constname_of_mode thy "" name mode
+      val mode_cbasename = Long_Name.base_name mode_cname
+      val Ts = binder_types T
+      val (Ts1, Ts2) = chop (length iss) Ts
+      val (Us1, Us2) =  split_smode is Ts2
+      val Ts1' = map2 (fn NONE => I | SOME is => funT_of compfuns ([], is)) iss Ts1
+      val funT = (Ts1' @ Us1) ---> (mk_predT compfuns (mk_tupleT Us2))
       val names = Name.variant_list []
         (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts)));
-      val xs = map Free (names ~~ (Ts1' @ Ts2));
-      val (xparams, xargs) = chop nparams xs;
-      val (xparams', names') = fold_map mk_Eval_of ((xparams ~~ Ts1) ~~ (fst mode)) names
-      val (xins, xouts) = get_args (snd mode) xargs;
+      val xs = map Free (names ~~ (Ts1' @ Ts2));                   
+      val (xparams, xargs) = chop (length iss) xs;
+      val (xins, xouts) = split_smode is xargs 
+      val (xparams', names') = fold_map mk_Eval_of ((xparams ~~ Ts1) ~~ iss) names
       fun mk_split_lambda [] t = lambda (Free (Name.variant names' "x", HOLogic.unitT)) t
-       | mk_split_lambda [x] t = lambda x t
-       | mk_split_lambda xs t = let
-         fun mk_split_lambda' (x::y::[]) t = HOLogic.mk_split (lambda x (lambda y t))
-           | mk_split_lambda' (x::xs) t = HOLogic.mk_split (lambda x (mk_split_lambda' xs t))
-         in mk_split_lambda' xs t end;
-      val predterm = mk_Enum (mk_split_lambda xouts (list_comb (Const (name, T), xparams' @ xargs)))
-      val funT = (Ts1' @ Us1) ---> (mk_pred_enumT (HOLogic.mk_tupleT Us2))
-      val mode_id = Sign.full_bname thy (Long_Name.base_name mode_id)
-      val lhs = list_comb (Const (mode_id, funT), xparams @ xins)
+        | mk_split_lambda [x] t = lambda x t
+        | mk_split_lambda xs t =
+        let
+          fun mk_split_lambda' (x::y::[]) t = HOLogic.mk_split (lambda x (lambda y t))
+            | mk_split_lambda' (x::xs) t = HOLogic.mk_split (lambda x (mk_split_lambda' xs t))
+        in
+          mk_split_lambda' xs t
+        end;
+      val predterm = PredicateCompFuns.mk_Enum (mk_split_lambda xouts
+        (list_comb (Const (name, T), xparams' @ xargs)))
+      val lhs = list_comb (Const (mode_cname, funT), xparams @ xins)
       val def = Logic.mk_equals (lhs, predterm)
       val ([definition], thy') = thy |>
-        Sign.add_consts_i [(Binding.name (Long_Name.base_name mode_id), funT, NoSyn)] |>
-        PureThy.add_defs false [((Binding.name (Long_Name.base_name mode_id ^ "_def"), def), [])]
-      val (intro, elim) = create_intro_elim_rule nparams mode definition mode_id funT (Const (name, T)) thy'
-      in thy' |> add_predfun name mode (mode_id, definition, intro, elim)
-        |> PureThy.store_thm (Binding.name (Long_Name.base_name mode_id ^ "I"), intro) |> snd
-        |> PureThy.store_thm (Binding.name (Long_Name.base_name mode_id ^ "E"), elim)  |> snd
+        Sign.add_consts_i [(Binding.name mode_cbasename, funT, NoSyn)] |>
+        PureThy.add_defs false [((Binding.name (mode_cbasename ^ "_def"), def), [])]
+      val (intro, elim) =
+        create_intro_elim_rule mode definition mode_cname funT (Const (name, T)) thy'
+      in thy' |> add_predfun name mode (mode_cname, definition, intro, elim)
+        |> PureThy.store_thm (Binding.name (mode_cbasename ^ "I"), intro) |> snd
+        |> PureThy.store_thm (Binding.name (mode_cbasename ^ "E"), elim)  |> snd
         |> Theory.checkpoint
       end;
   in
     fold create_definition modes thy
   end;
 
-(**************************************************************************************)
+fun sizelim_create_definitions preds (name, modes) thy =
+  let
+    val T = AList.lookup (op =) preds name |> the
+    fun create_definition mode thy =
+      let
+        val mode_cname = create_constname_of_mode thy "sizelim_" name mode
+        val funT = sizelim_funT_of PredicateCompFuns.compfuns mode T
+      in
+        thy |> Sign.add_consts_i [(Binding.name (Long_Name.base_name mode_cname), funT, NoSyn)]
+        |> set_sizelim_function_name name mode mode_cname 
+      end;
+  in
+    fold create_definition modes thy
+  end;
+    
+fun rpred_create_definitions preds (name, modes) thy =
+  let
+    val T = AList.lookup (op =) preds name |> the
+    fun create_definition mode thy =
+      let
+        val mode_cname = create_constname_of_mode thy "gen_" name mode
+        val funT = sizelim_funT_of RPredCompFuns.compfuns mode T
+      in
+        thy |> Sign.add_consts_i [(Binding.name (Long_Name.base_name mode_cname), funT, NoSyn)]
+        |> set_generator_name name mode mode_cname 
+      end;
+  in
+    fold create_definition modes thy
+  end;
+  
 (* Proving equivalence of term *)
 
 fun is_Type (Type _) = true
@@ -892,65 +1495,48 @@
 
 (* MAJOR FIXME:  prove_params should be simple
  - different form of introrule for parameters ? *)
-fun prove_param thy modes (NONE, t) =
-  all_tac 
-| prove_param thy modes (m as SOME (Mode (mode, is, ms)), t) =
-  REPEAT_DETERM (etac @{thm thin_rl} 1)
-  THEN REPEAT_DETERM (rtac @{thm ext} 1)
-  THEN (rtac @{thm iffI} 1)
-  THEN new_print_tac "prove_param"
-  (* proof in one direction *)
-  THEN (atac 1)
-  (* proof in the other direction *)
-  THEN (atac 1)
-  THEN new_print_tac "after prove_param"
-(*  let
-    val  (f, args) = strip_comb t
+fun prove_param thy (NONE, t) = TRY (rtac @{thm refl} 1)
+  | prove_param thy (m as SOME (Mode (mode, is, ms)), t) =
+  let
+    val  (f, args) = strip_comb (Envir.eta_contract t)
     val (params, _) = chop (length ms) args
     val f_tac = case f of
-        Const (name, T) => simp_tac (HOL_basic_ss addsimps 
-           (@{thm eval_pred}::(predfun_definition_of thy name mode)::
-           @{thm "Product_Type.split_conv"}::[])) 1
-      | Free _ => all_tac
-      | Abs _ => error "TODO: implement here"
-  in  
-    print_tac "before simplification in prove_args:"
+      Const (name, T) => simp_tac (HOL_basic_ss addsimps 
+         (@{thm eval_pred}::(predfun_definition_of thy name mode)::
+         @{thm "Product_Type.split_conv"}::[])) 1
+    | Free _ => TRY (rtac @{thm refl} 1)
+    | Abs _ => error "prove_param: No valid parameter term"
+  in
+    REPEAT_DETERM (etac @{thm thin_rl} 1)
+    THEN REPEAT_DETERM (rtac @{thm ext} 1)
+    THEN print_tac "prove_param"
     THEN f_tac
     THEN print_tac "after simplification in prove_args"
-    THEN (EVERY (map (prove_param thy modes) (ms ~~ params)))
+    THEN (EVERY (map (prove_param thy) (ms ~~ params)))
     THEN (REPEAT_DETERM (atac 1))
   end
-*)
-fun prove_expr thy modes (SOME (Mode (mode, is, ms)), t, us) (premposition : int) =
-  (case strip_comb t of
-    (Const (name, T), args) =>
-      if AList.defined op = modes name then (let
-          val introrule = predfun_intro_of thy name mode
-          (*val (in_args, out_args) = get_args is us
-          val (pred, rargs) = strip_comb (HOLogic.dest_Trueprop
-            (hd (Logic.strip_imp_prems (prop_of introrule))))
-          val nparams = length ms (* get_nparams thy (fst (dest_Const pred)) *)
-          val (_, args) = chop nparams rargs
-          val subst = map (pairself (cterm_of thy)) (args ~~ us)
-          val inst_introrule = Drule.cterm_instantiate subst introrule*)
-         (* the next line is old and probably wrong *)
-          val (args1, args2) = chop (length ms) args
-        in
+
+fun prove_expr thy (Mode (mode, is, ms), t, us) (premposition : int) =
+  case strip_comb t of
+    (Const (name, T), args) =>  
+      let
+        val introrule = predfun_intro_of thy name mode
+        val (args1, args2) = chop (length ms) args
+      in
         rtac @{thm bindI} 1
         THEN print_tac "before intro rule:"
         (* for the right assumption in first position *)
         THEN rotate_tac premposition 1
+        THEN debug_tac (Display.string_of_thm (ProofContext.init thy) introrule)
         THEN rtac introrule 1
-        THEN new_print_tac "after intro rule"
+        THEN print_tac "after intro rule"
         (* work with parameter arguments *)
         THEN (atac 1)
-        THEN (new_print_tac "parameter goal")
-        THEN (EVERY (map (prove_param thy modes) (ms ~~ args1)))
-        THEN (REPEAT_DETERM (atac 1)) end)
-      else error "Prove expr if case not implemented"
-    | _ => rtac @{thm bindI} 1
-           THEN atac 1)
-  | prove_expr _ _ _ _ =  error "Prove expr not implemented"
+        THEN (print_tac "parameter goal")
+        THEN (EVERY (map (prove_param thy) (ms ~~ args1)))
+        THEN (REPEAT_DETERM (atac 1))
+      end
+  | _ => rtac @{thm bindI} 1 THEN atac 1
 
 fun SOLVED tac st = FILTER (fn st' => nprems_of st' = nprems_of st - 1) tac st; 
 
@@ -994,105 +1580,86 @@
     (* need better control here! *)
   end
 
-fun prove_clause thy nargs all_vs param_vs modes (iss, is) (ts, ps) = let
-  val modes' = modes @ List.mapPartial
-   (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))
-     (param_vs ~~ iss);
-  fun check_constrt ((names, eqs), t) =
-      if is_constrt thy t then ((names, eqs), t) else
-        let
-          val s = Name.variant names "x";
-          val v = Free (s, fastype_of t)
-        in ((s::names, HOLogic.mk_eq (v, t)::eqs), v) end;
-  
-  val (in_ts, clause_out_ts) = get_args is ts;
-  val ((all_vs', eqs), in_ts') =
-      (*FIXME*) Library.foldl_map check_constrt ((all_vs, []), in_ts);
-  fun prove_prems out_ts vs [] =
-    (prove_match thy out_ts)
-    THEN asm_simp_tac HOL_basic_ss' 1
-    THEN print_tac "before the last rule of singleI:"
-    THEN (rtac (if null clause_out_ts then @{thm singleI_unit} else @{thm singleI}) 1)
-  | prove_prems out_ts vs rps =
-    let
-      val vs' = distinct (op =) (flat (vs :: map term_vs out_ts));
-      val SOME (p, mode as SOME (Mode ((iss, js), _, param_modes))) =
-        select_mode_prem thy modes' vs' rps;
-      val premposition = (find_index (equal p) ps) + nargs
-      val rps' = filter_out (equal p) rps;
-      val rest_tac = (case p of Prem (us, t) =>
-          let
-            val (in_ts, out_ts''') = get_args js us
-            val rec_tac = prove_prems out_ts''' vs' rps'
-          in
-            print_tac "before clause:"
-            THEN asm_simp_tac HOL_basic_ss 1
-            THEN print_tac "before prove_expr:"
-            THEN prove_expr thy modes (mode, t, us) premposition
-            THEN print_tac "after prove_expr:"
-            THEN rec_tac
-          end
-        | Negprem (us, t) =>
-          let
-            val (in_ts, out_ts''') = get_args js us
-            val rec_tac = prove_prems out_ts''' vs' rps'
-            val name = (case strip_comb t of (Const (c, _), _) => SOME c | _ => NONE)
-            val (_, params) = strip_comb t
-          in
-            rtac @{thm bindI} 1
-            THEN (if (is_some name) then
-                simp_tac (HOL_basic_ss addsimps [predfun_definition_of thy (the name) (iss, js)]) 1
-                THEN rtac @{thm not_predI} 1
-                (* FIXME: work with parameter arguments *)
-                THEN (EVERY (map (prove_param thy modes) (param_modes ~~ params)))
-              else
-                rtac @{thm not_predI'} 1)
-            THEN (REPEAT_DETERM (atac 1))
-            THEN rec_tac
-          end
-        | Sidecond t =>
-         rtac @{thm bindI} 1
-         THEN rtac @{thm if_predI} 1
-         THEN print_tac "before sidecond:"
-         THEN prove_sidecond thy modes t
-         THEN print_tac "after sidecond:"
-         THEN prove_prems [] vs' rps')
-    in (prove_match thy out_ts)
-        THEN rest_tac
-    end;
-  val prems_tac = prove_prems in_ts' param_vs ps
-in
-  rtac @{thm bindI} 1
-  THEN rtac @{thm singleI} 1
-  THEN prems_tac
-end;
+fun prove_clause thy nargs modes (iss, is) (_, clauses) (ts, moded_ps) =
+  let
+    val (in_ts, clause_out_ts) = split_smode is ts;
+    fun prove_prems out_ts [] =
+      (prove_match thy out_ts)
+      THEN asm_simp_tac HOL_basic_ss' 1
+      THEN (rtac (if null clause_out_ts then @{thm singleI_unit} else @{thm singleI}) 1)
+    | prove_prems out_ts ((p, mode as Mode ((iss, is), _, param_modes)) :: ps) =
+      let
+        val premposition = (find_index (equal p) clauses) + nargs
+        val rest_tac = (case p of Prem (us, t) =>
+            let
+              val (_, out_ts''') = split_smode is us
+              val rec_tac = prove_prems out_ts''' ps
+            in
+              print_tac "before clause:"
+              THEN asm_simp_tac HOL_basic_ss 1
+              THEN print_tac "before prove_expr:"
+              THEN prove_expr thy (mode, t, us) premposition
+              THEN print_tac "after prove_expr:"
+              THEN rec_tac
+            end
+          | Negprem (us, t) =>
+            let
+              val (_, out_ts''') = split_smode is us
+              val rec_tac = prove_prems out_ts''' ps
+              val name = (case strip_comb t of (Const (c, _), _) => SOME c | _ => NONE)
+              val (_, params) = strip_comb t
+            in
+              rtac @{thm bindI} 1
+              THEN (if (is_some name) then
+                  simp_tac (HOL_basic_ss addsimps [predfun_definition_of thy (the name) (iss, is)]) 1
+                  THEN rtac @{thm not_predI} 1
+                  THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1
+                  THEN (REPEAT_DETERM (atac 1))
+                  (* FIXME: work with parameter arguments *)
+                  THEN (EVERY (map (prove_param thy) (param_modes ~~ params)))
+                else
+                  rtac @{thm not_predI'} 1)
+                  THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1
+              THEN rec_tac
+            end
+          | Sidecond t =>
+           rtac @{thm bindI} 1
+           THEN rtac @{thm if_predI} 1
+           THEN print_tac "before sidecond:"
+           THEN prove_sidecond thy modes t
+           THEN print_tac "after sidecond:"
+           THEN prove_prems [] ps)
+      in (prove_match thy out_ts)
+          THEN rest_tac
+      end;
+    val prems_tac = prove_prems in_ts moded_ps
+  in
+    rtac @{thm bindI} 1
+    THEN rtac @{thm singleI} 1
+    THEN prems_tac
+  end;
 
 fun select_sup 1 1 = []
   | select_sup _ 1 = [rtac @{thm supI1}]
   | select_sup n i = (rtac @{thm supI2})::(select_sup (n - 1) (i - 1));
 
-fun prove_one_direction thy all_vs param_vs modes clauses ((pred, T), mode) = let
-(*  val ind_result = Inductive.the_inductive (ProofContext.init thy) pred
-  val index = find_index (fn s => s = pred) (#names (fst ind_result))
-  val (_, T) = dest_Const (nth (#preds (snd ind_result)) index) *)
-  val nargs = length (binder_types T) - nparams_of thy pred
-  val pred_case_rule = singleton (ind_set_codegen_preproc thy)
-    (preprocess_elim thy nargs (the_elim_of thy pred))
-  (* FIXME preprocessor |> Simplifier.full_simplify (HOL_basic_ss addsimps [@{thm Predicate.memb_code}])*)
-in
-  REPEAT_DETERM (CHANGED (rewtac @{thm "split_paired_all"}))
-  THEN etac (predfun_elim_of thy pred mode) 1
-  THEN etac pred_case_rule 1
-  THEN (EVERY (map
-         (fn i => EVERY' (select_sup (length clauses) i) i) 
-           (1 upto (length clauses))))
-  THEN (EVERY (map (prove_clause thy nargs all_vs param_vs modes mode) clauses))
-  THEN new_print_tac "proved one direction"
-end;
+fun prove_one_direction thy clauses preds modes pred mode moded_clauses =
+  let
+    val T = the (AList.lookup (op =) preds pred)
+    val nargs = length (binder_types T) - nparams_of thy pred
+    val pred_case_rule = the_elim_of thy pred
+  in
+    REPEAT_DETERM (CHANGED (rewtac @{thm "split_paired_all"}))
+    THEN etac (predfun_elim_of thy pred mode) 1
+    THEN etac pred_case_rule 1
+    THEN (EVERY (map
+           (fn i => EVERY' (select_sup (length moded_clauses) i) i) 
+             (1 upto (length moded_clauses))))
+    THEN (EVERY (map2 (prove_clause thy nargs modes mode) clauses moded_clauses))
+    THEN print_tac "proved one direction"
+  end;
 
-(*******************************************************************************************************)
-(* Proof in the other direction ************************************************************************)
-(*******************************************************************************************************)
+(** Proof in the other direction **)
 
 fun prove_match2 thy out_ts = let
   fun split_term_tac (Free _) = all_tac
@@ -1114,51 +1681,50 @@
       end
     else all_tac
   in
-    split_term_tac (HOLogic.mk_tuple out_ts)
+    split_term_tac (mk_tuple out_ts)
     THEN (DETERM (TRY ((Splitter.split_asm_tac [@{thm "split_if_asm"}] 1) THEN (etac @{thm botE} 2))))
   end
 
 (* VERY LARGE SIMILIRATIY to function prove_param 
 -- join both functions
 *)
+(* TODO: remove function *)
 
-fun prove_param2 thy modes (NONE, t) = all_tac 
-  | prove_param2 thy modes (m as SOME (Mode (mode, is, ms)), t) = let
-    val  (f, args) = strip_comb t
+fun prove_param2 thy (NONE, t) = all_tac 
+  | prove_param2 thy (m as SOME (Mode (mode, is, ms)), t) = let
+    val  (f, args) = strip_comb (Envir.eta_contract t)
     val (params, _) = chop (length ms) args
     val f_tac = case f of
         Const (name, T) => full_simp_tac (HOL_basic_ss addsimps 
            (@{thm eval_pred}::(predfun_definition_of thy name mode)
            :: @{thm "Product_Type.split_conv"}::[])) 1
       | Free _ => all_tac
+      | _ => error "prove_param2: illegal parameter term"
   in  
     print_tac "before simplification in prove_args:"
     THEN f_tac
     THEN print_tac "after simplification in prove_args"
-    THEN (EVERY (map (prove_param2 thy modes) (ms ~~ params)))
+    THEN (EVERY (map (prove_param2 thy) (ms ~~ params)))
   end
 
 
-fun prove_expr2 thy modes (SOME (Mode (mode, is, ms)), t) = 
+fun prove_expr2 thy (Mode (mode, is, ms), t) = 
   (case strip_comb t of
     (Const (name, T), args) =>
-      if AList.defined op = modes name then
-        etac @{thm bindE} 1
-        THEN (REPEAT_DETERM (CHANGED (rewtac @{thm "split_paired_all"})))
-        THEN new_print_tac "prove_expr2-before"
-        THEN (debug_tac (Syntax.string_of_term_global thy
-          (prop_of (predfun_elim_of thy name mode))))
-        THEN (etac (predfun_elim_of thy name mode) 1)
-        THEN new_print_tac "prove_expr2"
-        (* TODO -- FIXME: replace remove_last_goal*)
-        (* THEN (EVERY (replicate (length args) (remove_last_goal thy))) *)
-        THEN (EVERY (map (prove_param thy modes) (ms ~~ args)))
-        THEN new_print_tac "finished prove_expr2"
-      
-      else error "Prove expr2 if case not implemented"
+      etac @{thm bindE} 1
+      THEN (REPEAT_DETERM (CHANGED (rewtac @{thm "split_paired_all"})))
+      THEN print_tac "prove_expr2-before"
+      THEN (debug_tac (Syntax.string_of_term_global thy
+        (prop_of (predfun_elim_of thy name mode))))
+      THEN (etac (predfun_elim_of thy name mode) 1)
+      THEN print_tac "prove_expr2"
+      THEN (EVERY (map (prove_param2 thy) (ms ~~ args)))
+      THEN print_tac "finished prove_expr2"      
     | _ => etac @{thm bindE} 1)
-  | prove_expr2 _ _ _ = error "Prove expr2 not implemented"
-
+    
+(* FIXME: what is this for? *)
+(* replace defined by has_mode thy pred *)
+(* TODO: rewrite function *)
 fun prove_sidecond2 thy modes t = let
   fun preds_of t nameTs = case strip_comb t of 
     (f as Const (name, T), args) =>
@@ -1176,147 +1742,140 @@
    THEN print_tac "after sidecond2 simplification"
    end
   
-fun prove_clause2 thy all_vs param_vs modes (iss, is) (ts, ps) pred i = let
-  val modes' = modes @ List.mapPartial
-   (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))
-     (param_vs ~~ iss);
-  fun check_constrt ((names, eqs), t) =
-      if is_constrt thy t then ((names, eqs), t) else
-        let
-          val s = Name.variant names "x";
-          val v = Free (s, fastype_of t)
-        in ((s::names, HOLogic.mk_eq (v, t)::eqs), v) end;
-  val pred_intro_rule = nth (intros_of thy pred) (i - 1)
-    |> preprocess_intro thy
-    |> (fn thm => hd (ind_set_codegen_preproc thy [thm]))
-    (* FIXME preprocess |> Simplifier.full_simplify (HOL_basic_ss addsimps [@ {thm Predicate.memb_code}]) *)
-  val (in_ts, clause_out_ts) = get_args is ts;
-  val ((all_vs', eqs), in_ts') =
-      (*FIXME*) Library.foldl_map check_constrt ((all_vs, []), in_ts);
-  fun prove_prems2 out_ts vs [] =
-    print_tac "before prove_match2 - last call:"
-    THEN prove_match2 thy out_ts
-    THEN print_tac "after prove_match2 - last call:"
-    THEN (etac @{thm singleE} 1)
-    THEN (REPEAT_DETERM (etac @{thm Pair_inject} 1))
-    THEN (asm_full_simp_tac HOL_basic_ss' 1)
-    THEN (REPEAT_DETERM (etac @{thm Pair_inject} 1))
-    THEN (asm_full_simp_tac HOL_basic_ss' 1)
-    THEN SOLVED (print_tac "state before applying intro rule:"
+fun prove_clause2 thy modes pred (iss, is) (ts, ps) i =
+  let
+    val pred_intro_rule = nth (intros_of thy pred) (i - 1)
+    val (in_ts, clause_out_ts) = split_smode is ts;
+    fun prove_prems2 out_ts [] =
+      print_tac "before prove_match2 - last call:"
+      THEN prove_match2 thy out_ts
+      THEN print_tac "after prove_match2 - last call:"
+      THEN (etac @{thm singleE} 1)
+      THEN (REPEAT_DETERM (etac @{thm Pair_inject} 1))
+      THEN (asm_full_simp_tac HOL_basic_ss' 1)
+      THEN (REPEAT_DETERM (etac @{thm Pair_inject} 1))
+      THEN (asm_full_simp_tac HOL_basic_ss' 1)
+      THEN SOLVED (print_tac "state before applying intro rule:"
       THEN (rtac pred_intro_rule 1)
       (* How to handle equality correctly? *)
       THEN (print_tac "state before assumption matching")
       THEN (REPEAT (atac 1 ORELSE 
          (CHANGED (asm_full_simp_tac HOL_basic_ss' 1)
           THEN print_tac "state after simp_tac:"))))
-  | prove_prems2 out_ts vs ps = let
-      val vs' = distinct (op =) (flat (vs :: map term_vs out_ts));
-      val SOME (p, mode as SOME (Mode ((iss, js), _, param_modes))) =
-        select_mode_prem thy modes' vs' ps;
-      val ps' = filter_out (equal p) ps;
-      val rest_tac = (case p of Prem (us, t) =>
+    | prove_prems2 out_ts ((p, mode as Mode ((iss, is), _, param_modes)) :: ps) =
+      let
+        val rest_tac = (case p of
+          Prem (us, t) =>
           let
-            val (in_ts, out_ts''') = get_args js us
-            val rec_tac = prove_prems2 out_ts''' vs' ps'
+            val (_, out_ts''') = split_smode is us
+            val rec_tac = prove_prems2 out_ts''' ps
           in
-            (prove_expr2 thy modes (mode, t)) THEN rec_tac
+            (prove_expr2 thy (mode, t)) THEN rec_tac
           end
         | Negprem (us, t) =>
           let
-            val (in_ts, out_ts''') = get_args js us
-            val rec_tac = prove_prems2 out_ts''' vs' ps'
+            val (_, out_ts''') = split_smode is us
+            val rec_tac = prove_prems2 out_ts''' ps
             val name = (case strip_comb t of (Const (c, _), _) => SOME c | _ => NONE)
             val (_, params) = strip_comb t
           in
             print_tac "before neg prem 2"
             THEN etac @{thm bindE} 1
             THEN (if is_some name then
-                full_simp_tac (HOL_basic_ss addsimps [predfun_definition_of thy (the name) (iss, js)]) 1 
+                full_simp_tac (HOL_basic_ss addsimps [predfun_definition_of thy (the name) (iss, is)]) 1 
                 THEN etac @{thm not_predE} 1
-                THEN (EVERY (map (prove_param thy modes) (param_modes ~~ params)))
+                THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1
+                THEN (EVERY (map (prove_param2 thy) (param_modes ~~ params)))
               else
                 etac @{thm not_predE'} 1)
             THEN rec_tac
           end 
         | Sidecond t =>
-            etac @{thm bindE} 1
-            THEN etac @{thm if_predE} 1
-            THEN prove_sidecond2 thy modes t 
-            THEN prove_prems2 [] vs' ps')
-    in print_tac "before prove_match2:"
-       THEN prove_match2 thy out_ts
-       THEN print_tac "after prove_match2:"
-       THEN rest_tac
-    end;
-  val prems_tac = prove_prems2 in_ts' param_vs ps 
-in
-  new_print_tac "starting prove_clause2"
-  THEN etac @{thm bindE} 1
-  THEN (etac @{thm singleE'} 1)
-  THEN (TRY (etac @{thm Pair_inject} 1))
-  THEN print_tac "after singleE':"
-  THEN prems_tac
-end;
+          etac @{thm bindE} 1
+          THEN etac @{thm if_predE} 1
+          THEN prove_sidecond2 thy modes t 
+          THEN prove_prems2 [] ps)
+      in print_tac "before prove_match2:"
+         THEN prove_match2 thy out_ts
+         THEN print_tac "after prove_match2:"
+         THEN rest_tac
+      end;
+    val prems_tac = prove_prems2 in_ts ps 
+  in
+    print_tac "starting prove_clause2"
+    THEN etac @{thm bindE} 1
+    THEN (etac @{thm singleE'} 1)
+    THEN (TRY (etac @{thm Pair_inject} 1))
+    THEN print_tac "after singleE':"
+    THEN prems_tac
+  end;
  
-fun prove_other_direction thy all_vs param_vs modes clauses (pred, mode) = let
-  fun prove_clause (clause, i) =
-    (if i < length clauses then etac @{thm supE} 1 else all_tac)
-    THEN (prove_clause2 thy all_vs param_vs modes mode clause pred i)
-in
-  (DETERM (TRY (rtac @{thm unit.induct} 1)))
-   THEN (REPEAT_DETERM (CHANGED (rewtac @{thm split_paired_all})))
-   THEN (rtac (predfun_intro_of thy pred mode) 1)
-   THEN (REPEAT_DETERM (rtac @{thm refl} 2))
-   THEN (EVERY (map prove_clause (clauses ~~ (1 upto (length clauses)))))
-end;
+fun prove_other_direction thy modes pred mode moded_clauses =
+  let
+    fun prove_clause clause i =
+      (if i < length moded_clauses then etac @{thm supE} 1 else all_tac)
+      THEN (prove_clause2 thy modes pred mode clause i)
+  in
+    (DETERM (TRY (rtac @{thm unit.induct} 1)))
+     THEN (REPEAT_DETERM (CHANGED (rewtac @{thm split_paired_all})))
+     THEN (rtac (predfun_intro_of thy pred mode) 1)
+     THEN (REPEAT_DETERM (rtac @{thm refl} 2))
+     THEN (EVERY (map2 prove_clause moded_clauses (1 upto (length moded_clauses))))
+  end;
+
+(** proof procedure **)
 
-fun prove_pred thy all_vs param_vs modes clauses (((pred, T), mode), t) = let
-  val ctxt = ProofContext.init thy
-  val clauses' = the (AList.lookup (op =) clauses pred)
-in
-  Goal.prove ctxt (Term.fold_aterms (fn Free (x, _) => insert (op =) x | _ => I) t []) [] t
-    (if !do_proofs then
-      (fn _ =>
-      rtac @{thm pred_iffI} 1
-      THEN prove_one_direction thy all_vs param_vs modes clauses' ((pred, T), mode)
-      THEN print_tac "proved one direction"
-      THEN prove_other_direction thy all_vs param_vs modes clauses' (pred, mode)
-      THEN print_tac "proved other direction")
-     else (fn _ => mycheat_tac thy 1))
-end;
+fun prove_pred thy clauses preds modes pred mode (moded_clauses, compiled_term) =
+  let
+    val ctxt = ProofContext.init thy
+    val clauses = the (AList.lookup (op =) clauses pred)
+  in
+    Goal.prove ctxt (Term.add_free_names compiled_term []) [] compiled_term
+      (if !do_proofs then
+        (fn _ =>
+        rtac @{thm pred_iffI} 1
+        THEN prove_one_direction thy clauses preds modes pred mode moded_clauses
+        THEN print_tac "proved one direction"
+        THEN prove_other_direction thy modes pred mode moded_clauses
+        THEN print_tac "proved other direction")
+       else (fn _ => mycheat_tac thy 1))
+  end;
 
-fun prove_preds thy all_vs param_vs modes clauses pmts =
-  map (prove_pred thy all_vs param_vs modes clauses) pmts
+(* composition of mode inference, definition, compilation and proof *)
+
+(** auxillary combinators for table of preds and modes **)
 
-(* special case: inductive predicate with no clauses *)
-fun noclause (predname, T) thy = let
-  val Ts = binder_types T
-  val names = Name.variant_list []
-        (map (fn i => "x" ^ (string_of_int i)) (1 upto (length Ts)))
-  val vs = map2 (curry Free) names Ts
-  val clausehd =  HOLogic.mk_Trueprop (list_comb(Const (predname, T), vs))
-  val intro_t = Logic.mk_implies (@{prop False}, clausehd)
-  val P = HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT))
-  val elim_t = Logic.list_implies ([clausehd, Logic.mk_implies (@{prop False}, P)], P)
-  val intro = Goal.prove (ProofContext.init thy) names [] intro_t
-        (fn {...} => etac @{thm FalseE} 1)
-  val elim = Goal.prove (ProofContext.init thy) ("P" :: names) [] elim_t
-        (fn {...} => etac (the_elim_of thy predname) 1) 
-in
-  add_intro intro thy
-  |> set_elim elim
-end
+fun map_preds_modes f preds_modes_table =
+  map (fn (pred, modes) =>
+    (pred, map (fn (mode, value) => (mode, f pred mode value)) modes)) preds_modes_table
 
+fun join_preds_modes table1 table2 =
+  map_preds_modes (fn pred => fn mode => fn value =>
+    (value, the (AList.lookup (op =) (the (AList.lookup (op =) table2 pred)) mode))) table1
+    
+fun maps_modes preds_modes_table =
+  map (fn (pred, modes) =>
+    (pred, map (fn (mode, value) => value) modes)) preds_modes_table  
+    
+fun compile_preds compfuns mk_fun_of use_size thy all_vs param_vs preds moded_clauses =
+  map_preds_modes (fn pred => compile_pred compfuns mk_fun_of use_size thy all_vs param_vs pred
+      (the (AList.lookup (op =) preds pred))) moded_clauses  
+  
+fun prove thy clauses preds modes moded_clauses compiled_terms =
+  map_preds_modes (prove_pred thy clauses preds modes)
+    (join_preds_modes moded_clauses compiled_terms)
+
+fun prove_by_skip thy _ _ _ _ compiled_terms =
+  map_preds_modes (fn pred => fn mode => fn t => Drule.standard (SkipProof.make_thm thy t))
+    compiled_terms
+    
 fun prepare_intrs thy prednames =
   let
-    (* FIXME: preprocessing moved to fetch_pred_data *)
-    val intrs = map (preprocess_intro thy) (maps (intros_of thy) prednames)
-      |> ind_set_codegen_preproc thy (*FIXME preprocessor
-      |> map (Simplifier.full_simplify (HOL_basic_ss addsimps [@ {thm Predicate.memb_code}]))*)
+    val intrs = maps (intros_of thy) prednames
       |> map (Logic.unvarify o prop_of)
     val nparams = nparams_of thy (hd prednames)
+    val extra_modes = all_modes_of thy |> filter_out (fn (name, _) => member (op =) prednames name)
     val preds = distinct (op =) (map (dest_Const o fst o (strip_intro_concl nparams)) intrs)
-    val extra_modes = all_modes_of thy |> filter_out (fn (name, _) => member (op =) prednames name)
     val _ $ u = Logic.strip_imp_concl (hd intrs);
     val params = List.take (snd (strip_comb u), nparams);
     val param_vs = maps term_vs params
@@ -1324,7 +1883,7 @@
     fun dest_prem t =
       (case strip_comb t of
         (v as Free _, ts) => if v mem params then Prem (ts, v) else Sidecond t
-      | (c as Const (@{const_name Not}, _), [t]) => (case dest_prem t of
+      | (c as Const (@{const_name Not}, _), [t]) => (case dest_prem t of          
           Prem (ts, t) => Negprem (ts, t)
         | Negprem _ => error ("Double negation not allowed in premise: " ^ (Syntax.string_of_term_global thy (c $ t))) 
         | Sidecond t => Sidecond (c $ t))
@@ -1352,46 +1911,95 @@
     val (clauses, arities) = fold add_clause intrs ([], []);
   in (preds, nparams, all_vs, param_vs, extra_modes, clauses, arities) end;
 
-fun arrange kvs =
+(** main function of predicate compiler **)
+
+fun add_equations_of steps prednames thy =
   let
-    fun add (key, value) table =
-      AList.update op = (key, these (AList.lookup op = table key) @ [value]) table
-  in fold add kvs [] end;
-        
-(* main function *)
+    val _ = Output.tracing ("Starting predicate compiler for predicates " ^ commas prednames ^ "...")
+    val (preds, nparams, all_vs, param_vs, extra_modes, clauses, arities) =
+      prepare_intrs thy prednames
+    val _ = Output.tracing "Infering modes..."
+    val moded_clauses = #infer_modes steps thy extra_modes arities param_vs clauses 
+    val modes = map (fn (p, mps) => (p, map fst mps)) moded_clauses
+    val _ = print_modes modes
+    val _ = print_moded_clauses thy moded_clauses
+    val _ = Output.tracing "Defining executable functions..."
+    val thy' = fold (#create_definitions steps preds) modes thy
+      |> Theory.checkpoint
+    val _ = Output.tracing "Compiling equations..."
+    val compiled_terms =
+      (#compile_preds steps) thy' all_vs param_vs preds moded_clauses
+    val _ = print_compiled_terms thy' compiled_terms
+    val _ = Output.tracing "Proving equations..."
+    val result_thms = #prove steps thy' clauses preds (extra_modes @ modes)
+      moded_clauses compiled_terms
+    val qname = #qname steps
+    (* val attrib = gn thy => Attrib.attribute_i thy Code.add_eqn_attrib *)
+    val attrib = fn thy => Attrib.attribute_i thy (Attrib.internal (K (Thm.declaration_attribute
+      (fn thm => Context.mapping (Code.add_eqn thm) I))))
+    val thy'' = fold (fn (name, result_thms) => fn thy => snd (PureThy.add_thmss
+      [((Binding.qualify true (Long_Name.base_name name) (Binding.name qname), result_thms),
+        [attrib thy ])] thy))
+      (maps_modes result_thms) thy'
+      |> Theory.checkpoint
+  in
+    thy''
+  end
 
-fun add_equations_of prednames thy =
-let
-  val _ = tracing ("starting add_equations with " ^ commas prednames ^ "...")
-  (* null clause handling *)
-  (*
-  val thy' = fold (fn pred as (predname, T) => fn thy =>
-    if null (intros_of thy predname) then noclause pred thy else thy) preds thy
-    *)
-  val (preds, nparams, all_vs, param_vs, extra_modes, clauses, arities) =
-    prepare_intrs thy prednames
-  val _ = tracing "Infering modes..."
-  val modes = infer_modes thy extra_modes arities param_vs clauses
-  val _ = print_modes modes
-  val _ = tracing "Defining executable functions..."
-  val thy' = fold (create_definitions preds nparams) modes thy |> Theory.checkpoint
-  val clauses' = map (fn (s, cls) => (s, (the (AList.lookup (op =) preds s), cls))) clauses
-  val _ = tracing "Compiling equations..."
-  val ts = compile_preds thy' all_vs param_vs (extra_modes @ modes) clauses'
-  val _ = map (Output.tracing o (Syntax.string_of_term_global thy')) (flat ts)
-  val pred_mode =
-    maps (fn (s, (T, _)) => map (pair (s, T)) ((the o AList.lookup (op =) modes) s)) clauses' 
-  val _ = Output.tracing "Proving equations..."
-  val result_thms =
-    prove_preds thy' all_vs param_vs (extra_modes @ modes) clauses (pred_mode ~~ (flat ts))
-  val thy'' = fold (fn (name, result_thms) => fn thy => snd (PureThy.add_thmss
-    [((Binding.qualify true (Long_Name.base_name name) (Binding.name "equation"), result_thms),
-      [Attrib.attribute_i thy Code.add_default_eqn_attrib])] thy))
-    (arrange ((map (fn ((name, _), _) => name) pred_mode) ~~ result_thms)) thy'
-    |> Theory.checkpoint
-in
-  thy''
-end
+fun extend' value_of edges_of key (G, visited) =
+  let
+    val (G', v) = case try (Graph.get_node G) key of
+        SOME v => (G, v)
+      | NONE => (Graph.new_node (key, value_of key) G, value_of key)
+    val (G'', visited') = fold (extend' value_of edges_of) (edges_of (key, v) \\ visited)
+      (G', key :: visited) 
+  in
+    (fold (Graph.add_edge o (pair key)) (edges_of (key, v)) G'', visited')
+  end;
+
+fun extend value_of edges_of key G = fst (extend' value_of edges_of key (G, [])) 
+  
+fun gen_add_equations steps names thy =
+  let
+    val thy' = PredData.map (fold (extend (fetch_pred_data thy) (depending_preds_of thy)) names) thy
+      |> Theory.checkpoint;
+    fun strong_conn_of gr keys =
+      Graph.strong_conn (Graph.subgraph (member (op =) (Graph.all_succs gr keys)) gr)
+    val scc = strong_conn_of (PredData.get thy') names
+    val thy'' = fold_rev
+      (fn preds => fn thy =>
+        if #are_not_defined steps thy preds then add_equations_of steps preds thy else thy)
+      scc thy' |> Theory.checkpoint
+  in thy'' end
+
+(* different instantiantions of the predicate compiler *)
+
+val add_equations = gen_add_equations
+  {infer_modes = infer_modes false,
+  create_definitions = create_definitions,
+  compile_preds = compile_preds PredicateCompFuns.compfuns mk_fun_of false,
+  prove = prove,
+  are_not_defined = (fn thy => forall (null o modes_of thy)),
+  qname = "equation"}
+
+val add_sizelim_equations = gen_add_equations
+  {infer_modes = infer_modes false,
+  create_definitions = sizelim_create_definitions,
+  compile_preds = compile_preds PredicateCompFuns.compfuns mk_sizelim_fun_of true,
+  prove = prove_by_skip,
+  are_not_defined = (fn thy => fn preds => true), (* TODO *)
+  qname = "sizelim_equation"
+  }
+  
+val add_quickcheck_equations = gen_add_equations
+  {infer_modes = infer_modes_with_generator,
+  create_definitions = rpred_create_definitions,
+  compile_preds = compile_preds RPredCompFuns.compfuns mk_generator_of true,
+  prove = prove_by_skip,
+  are_not_defined = (fn thy => fn preds => true), (* TODO *)
+  qname = "rpred_equation"}
+
+(** user interface **)
 
 (* generation of case rules from user-given introduction rules *)
 
@@ -1404,7 +2012,8 @@
     val (argnames, ctxt2) = Variable.variant_fixes
       (map (fn i => "a" ^ string_of_int i) (1 upto (length args))) ctxt1
     val argvs = map2 (curry Free) argnames (map fastype_of args)
-    fun mk_case intro = let
+    fun mk_case intro =
+      let
         val (_, (_, args)) = strip_intro_concl nparams intro
         val prems = Logic.strip_imp_prems intro
         val eqprems = map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (argvs ~~ args)
@@ -1412,46 +2021,30 @@
           (fn t as Free _ =>
               if member (op aconv) params t then I else insert (op aconv) t
            | _ => I) (args @ prems) []
-        in fold Logic.all frees (Logic.list_implies (eqprems @ prems, prop)) end
+      in fold Logic.all frees (Logic.list_implies (eqprems @ prems, prop)) end
     val assm = HOLogic.mk_Trueprop (list_comb (pred, params @ argvs))
     val cases = map mk_case intros
   in Logic.list_implies (assm :: cases, prop) end;
 
-fun add_equations name thy =
-  let
-    val thy' = PredData.map (Graph.extend (dependencies_of thy) name) thy |> Theory.checkpoint;
-    (*val preds = Graph.all_preds (PredData.get thy') [name] |> filter_out (has_elim thy') *)
-    fun strong_conn_of gr keys =
-      Graph.strong_conn (Graph.subgraph (member (op =) (Graph.all_succs gr keys)) gr)
-    val scc = strong_conn_of (PredData.get thy') [name]
-    val thy'' = fold_rev
-      (fn preds => fn thy =>
-        if forall (null o modes_of thy) preds then add_equations_of preds thy else thy)
-      scc thy' |> Theory.checkpoint
-  in thy'' end
+(* code_pred_intro attribute *)
 
-  
 fun attrib f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I);
 
 val code_pred_intros_attrib = attrib add_intro;
 
-(** user interface **)
-
 local
 
 (* TODO: make TheoryDataFun to GenericDataFun & remove duplication of local theory and theory *)
 (* TODO: must create state to prove multiple cases *)
 fun generic_code_pred prep_const raw_const lthy =
   let
-  
     val thy = ProofContext.theory_of lthy
     val const = prep_const thy raw_const
-    
-    val lthy' = LocalTheory.theory (PredData.map (Graph.extend (dependencies_of thy) const)) lthy
+    val lthy' = LocalTheory.theory (PredData.map
+        (extend (fetch_pred_data thy) (depending_preds_of thy) const)) lthy
       |> LocalTheory.checkpoint
     val thy' = ProofContext.theory_of lthy'
     val preds = Graph.all_preds (PredData.get thy') [const] |> filter_out (has_elim thy')
-    
     fun mk_cases const =
       let
         val nparams = nparams_of thy' const
@@ -1463,11 +2056,16 @@
         assumes = [("", Logic.strip_imp_prems case_rule)],
         binds = [], cases = []}) cases_rules
     val case_env = map2 (fn p => fn c => (Long_Name.base_name p, SOME c)) preds cases
-    val _ = Output.tracing (commas (map fst case_env))
-    val lthy'' = ProofContext.add_cases true case_env lthy'
-    
-    fun after_qed thms =
-      LocalTheory.theory (fold set_elim (map the_single thms) #> add_equations const)
+    val lthy'' = lthy'
+      |> fold Variable.auto_fixes cases_rules 
+      |> ProofContext.add_cases true case_env
+    fun after_qed thms goal_ctxt =
+      let
+        val global_thms = ProofContext.export goal_ctxt
+          (ProofContext.init (ProofContext.theory_of goal_ctxt)) (map the_single thms)
+      in
+        goal_ctxt |> LocalTheory.theory (fold set_elim global_thms #> add_equations [const])
+      end  
   in
     Proof.theorem_i NONE after_qed (map (single o (rpair [])) cases_rules) lthy''
   end;
@@ -1523,9 +2121,8 @@
       | [m] => m
       | m :: _ :: _ => (warning ("Multiple modes possible for comprehension "
                 ^ Syntax.string_of_term_global thy t_compr); m);
-    val (inargs, outargs) = get_args user_mode args;
-    val t_pred = list_comb (compile_expr thy (all_modes_of thy) (SOME m, list_comb (pred, params)),
-      inargs);
+    val (inargs, outargs) = split_smode user_mode args;
+    val t_pred = list_comb (compile_expr NONE thy (m, list_comb (pred, params)), inargs);
     val t_eval = if null outargs then t_pred else let
         val outargs_bounds = map (fn Bound i => i) outargs;
         val outargsTs = map (nth Ts) outargs_bounds;
@@ -1537,14 +2134,14 @@
         val arrange = funpow (length outargs_bounds - 1) HOLogic.mk_split
           (Term.list_abs (map (pair "") outargsTs,
             HOLogic.mk_ptuple fp T_compr (map Bound arrange_bounds)))
-      in mk_pred_map T_pred T_compr arrange t_pred end
+      in mk_map PredicateCompFuns.compfuns T_pred T_compr arrange t_pred end
   in t_eval end;
 
 fun eval thy t_compr =
   let
     val t = analyze_compr thy t_compr;
-    val T = dest_pred_enumT (fastype_of t);
-    val t' = mk_pred_map T HOLogic.termT (HOLogic.term_of_const T) t;
+    val T = dest_predT PredicateCompFuns.compfuns (fastype_of t);
+    val t' = mk_map PredicateCompFuns.compfuns T HOLogic.termT (HOLogic.term_of_const T) t;
   in (T, Code_ML.eval NONE ("Predicate_Compile.eval_ref", eval_ref) Predicate.map thy t' []) end;
 
 fun values ctxt k t_compr =
--- a/src/Pure/Concurrent/future.ML	Mon Aug 10 08:37:37 2009 +0200
+++ b/src/Pure/Concurrent/future.ML	Mon Aug 10 10:25:00 2009 +0200
@@ -120,11 +120,10 @@
 fun SYNCHRONIZED name = SimpleThread.synchronized name lock;
 
 fun wait cond = (*requires SYNCHRONIZED*)
-  Multithreading.sync_wait NONE cond lock;
+  Multithreading.sync_wait NONE NONE cond lock;
 
-fun wait_interruptible timeout cond = (*requires SYNCHRONIZED*)
-  interruptible (fn () =>
-    ignore (Multithreading.sync_wait (SOME (Time.+ (Time.now (), timeout))) cond lock)) ();
+fun wait_timeout timeout cond = (*requires SYNCHRONIZED*)
+  Multithreading.sync_wait NONE (SOME (Time.+ (Time.now (), timeout))) cond lock;
 
 fun signal cond = (*requires SYNCHRONIZED*)
   ConditionVar.signal cond;
@@ -149,11 +148,11 @@
         val res =
           if ok then
             Exn.capture (fn () =>
-             (Thread.testInterrupt ();
-              Multithreading.with_attributes Multithreading.restricted_interrupts
-                (fn _ => fn () => e ())) ()) ()
+              Multithreading.with_attributes Multithreading.private_interrupts (fn _ => e ())) ()
           else Exn.Exn Exn.Interrupt;
-        val _ = Synchronized.change result (K (SOME res));
+        val _ = Synchronized.change result
+          (fn NONE => SOME res
+            | SOME _ => raise Fail "Duplicate assignment of future value");
       in
         (case res of
           Exn.Exn exn => (Task_Queue.cancel_group group exn; false)
@@ -276,20 +275,23 @@
         broadcast_work ());
 
     (*delay loop*)
-    val _ = wait_interruptible next_round scheduler_event
-      handle Exn.Interrupt =>
-        (Multithreading.tracing 1 (fn () => "Interrupt");
-          List.app do_cancel (Task_Queue.cancel_all (! queue)));
+    val _ = Exn.release (wait_timeout next_round scheduler_event);
 
     (*shutdown*)
     val _ = if Task_Queue.is_empty (! queue) then do_shutdown := true else ();
     val continue = not (! do_shutdown andalso null (! workers));
     val _ = if continue then () else scheduler := NONE;
     val _ = broadcast scheduler_event;
-  in continue end;
+  in continue end
+  handle Exn.Interrupt =>
+   (Multithreading.tracing 1 (fn () => "Interrupt");
+    uninterruptible (fn _ => fn () => List.app do_cancel (Task_Queue.cancel_all (! queue))) ();
+    scheduler_next ());
 
 fun scheduler_loop () =
-  while SYNCHRONIZED "scheduler" (fn () => scheduler_next ()) do ();
+  Multithreading.with_attributes
+    (Multithreading.sync_interrupts Multithreading.public_interrupts)
+    (fn _ => while SYNCHRONIZED "scheduler" (fn () => scheduler_next ()) do ());
 
 fun scheduler_active () = (*requires SYNCHRONIZED*)
   (case ! scheduler of NONE => false | SOME thread => Thread.isActive thread);
@@ -393,12 +395,11 @@
 
 fun interruptible_task f x =
   if Multithreading.available then
-   (Thread.testInterrupt ();
     Multithreading.with_attributes
       (if is_worker ()
-       then Multithreading.restricted_interrupts
-       else Multithreading.regular_interrupts)
-      (fn _ => fn x => f x) x)
+       then Multithreading.private_interrupts
+       else Multithreading.public_interrupts)
+      (fn _ => f x)
   else interruptible f x;
 
 (*cancel: present and future group members will be interrupted eventually*)
--- a/src/Pure/Concurrent/simple_thread.ML	Mon Aug 10 08:37:37 2009 +0200
+++ b/src/Pure/Concurrent/simple_thread.ML	Mon Aug 10 10:25:00 2009 +0200
@@ -16,7 +16,7 @@
 
 fun fork interrupts body =
   Thread.fork (fn () => exception_trace (fn () => body ()),
-    if interrupts then Multithreading.regular_interrupts else Multithreading.no_interrupts);
+    if interrupts then Multithreading.public_interrupts else Multithreading.no_interrupts);
 
 fun interrupt thread = Thread.interrupt thread handle Thread _ => ();
 
--- a/src/Pure/Concurrent/synchronized.ML	Mon Aug 10 08:37:37 2009 +0200
+++ b/src/Pure/Concurrent/synchronized.ML	Mon Aug 10 10:25:00 2009 +0200
@@ -48,9 +48,10 @@
           (case f x of
             SOME (y, x') => (var := x'; SOME y)
           | NONE =>
-              if Multithreading.sync_wait (time_limit x) cond lock
-              then try_change ()
-              else NONE)
+              (case Multithreading.sync_wait NONE (time_limit x) cond lock of
+                Exn.Result true => try_change ()
+              | Exn.Result false => NONE
+              | Exn.Exn exn => reraise exn))
         end;
       val res = try_change ();
       val _ = ConditionVar.broadcast cond;
--- a/src/Pure/ML-Systems/multithreading.ML	Mon Aug 10 08:37:37 2009 +0200
+++ b/src/Pure/ML-Systems/multithreading.ML	Mon Aug 10 10:25:00 2009 +0200
@@ -13,20 +13,21 @@
 signature MULTITHREADING =
 sig
   include BASIC_MULTITHREADING
-  val trace: int ref
-  val tracing: int -> (unit -> string) -> unit
-  val tracing_time: bool -> Time.time -> (unit -> string) -> unit
-  val real_time: ('a -> unit) -> 'a -> Time.time
   val available: bool
   val max_threads: int ref
   val max_threads_value: unit -> int
   val enabled: unit -> bool
   val no_interrupts: Thread.threadAttribute list
-  val regular_interrupts: Thread.threadAttribute list
-  val restricted_interrupts: Thread.threadAttribute list
-  val with_attributes: Thread.threadAttribute list ->
-    (Thread.threadAttribute list -> 'a -> 'b) -> 'a -> 'b
-  val sync_wait: Time.time option -> ConditionVar.conditionVar -> Mutex.mutex -> bool
+  val public_interrupts: Thread.threadAttribute list
+  val private_interrupts: Thread.threadAttribute list
+  val sync_interrupts: Thread.threadAttribute list -> Thread.threadAttribute list
+  val with_attributes: Thread.threadAttribute list -> (Thread.threadAttribute list -> 'a) -> 'a
+  val sync_wait: Thread.threadAttribute list option -> Time.time option ->
+    ConditionVar.conditionVar -> Mutex.mutex -> bool Exn.result
+  val trace: int ref
+  val tracing: int -> (unit -> string) -> unit
+  val tracing_time: bool -> Time.time -> (unit -> string) -> unit
+  val real_time: ('a -> unit) -> 'a -> Time.time
   val self_critical: unit -> bool
   val serial: unit -> int
 end;
@@ -34,14 +35,6 @@
 structure Multithreading: MULTITHREADING =
 struct
 
-(* tracing *)
-
-val trace = ref (0: int);
-fun tracing _ _ = ();
-fun tracing_time _ _ _ = ();
-fun real_time f x = (f x; Time.zeroTime);
-
-
 (* options *)
 
 val available = false;
@@ -52,18 +45,22 @@
 
 (* attributes *)
 
-val no_interrupts =
-  [Thread.EnableBroadcastInterrupt false, Thread.InterruptState Thread.InterruptDefer];
+val no_interrupts = [];
+val public_interrupts = [];
+val private_interrupts = [];
+fun sync_interrupts _ = [];
 
-val regular_interrupts =
-  [Thread.EnableBroadcastInterrupt true, Thread.InterruptState Thread.InterruptAsynchOnce];
+fun with_attributes _ e = e [];
 
-val restricted_interrupts =
-  [Thread.EnableBroadcastInterrupt false, Thread.InterruptState Thread.InterruptAsynchOnce];
+fun sync_wait _ _ _ _ = Exn.Result true;
+
+
+(* tracing *)
 
-fun with_attributes _ f x = f [] x;
-
-fun sync_wait _ _ _ = false;
+val trace = ref (0: int);
+fun tracing _ _ = ();
+fun tracing_time _ _ _ = ();
+fun real_time f x = (f x; Time.zeroTime);
 
 
 (* critical section *)
--- a/src/Pure/ML-Systems/multithreading_polyml.ML	Mon Aug 10 08:37:37 2009 +0200
+++ b/src/Pure/ML-Systems/multithreading_polyml.ML	Mon Aug 10 10:25:00 2009 +0200
@@ -27,31 +27,6 @@
 structure Multithreading: MULTITHREADING =
 struct
 
-(* tracing *)
-
-val trace = ref 0;
-
-fun tracing level msg =
-  if level > ! trace then ()
-  else (TextIO.output (TextIO.stdErr, (">>> " ^ msg () ^ "\n")); TextIO.flushOut TextIO.stdErr)
-    handle _ (*sic*) => ();
-
-fun tracing_time detailed time =
-  tracing
-   (if not detailed then 5
-    else if Time.>= (time, Time.fromMilliseconds 1000) then 1
-    else if Time.>= (time, Time.fromMilliseconds 100) then 2
-    else if Time.>= (time, Time.fromMilliseconds 10) then 3
-    else if Time.>= (time, Time.fromMilliseconds 1) then 4 else 5);
-
-fun real_time f x =
-  let
-    val timer = Timer.startRealTimer ();
-    val () = f x;
-    val time = Timer.checkRealTimer timer;
-  in time end;
-
-
 (* options *)
 
 val available = true;
@@ -91,57 +66,76 @@
 val no_interrupts =
   [Thread.EnableBroadcastInterrupt false, Thread.InterruptState Thread.InterruptDefer];
 
-val regular_interrupts =
+val public_interrupts =
   [Thread.EnableBroadcastInterrupt true, Thread.InterruptState Thread.InterruptAsynchOnce];
 
-val restricted_interrupts =
+val private_interrupts =
   [Thread.EnableBroadcastInterrupt false, Thread.InterruptState Thread.InterruptAsynchOnce];
 
+val sync_interrupts = map
+  (fn x as Thread.InterruptState Thread.InterruptDefer => x
+    | Thread.InterruptState _ => Thread.InterruptState Thread.InterruptSynch
+    | x => x);
+
 val safe_interrupts = map
   (fn Thread.InterruptState Thread.InterruptAsynch =>
       Thread.InterruptState Thread.InterruptAsynchOnce
     | x => x);
 
-fun with_attributes new_atts f x =
+fun with_attributes new_atts e =
   let
     val orig_atts = safe_interrupts (Thread.getAttributes ());
     val result = Exn.capture (fn () =>
-      (Thread.setAttributes (safe_interrupts new_atts); f orig_atts x)) ();
+      (Thread.setAttributes (safe_interrupts new_atts); e orig_atts)) ();
     val _ = Thread.setAttributes orig_atts;
   in Exn.release result end;
 
 
-(* regular interruptibility *)
+(* portable wrappers *)
+
+fun interruptible f x = with_attributes public_interrupts (fn _ => f x);
 
-fun interruptible f x =
-  (Thread.testInterrupt (); with_attributes regular_interrupts (fn _ => fn x => f x) x);
-
-fun uninterruptible f =
-  with_attributes no_interrupts (fn atts => fn x =>
-    f (fn g => with_attributes atts (fn _ => fn y => g y)) x);
+fun uninterruptible f x =
+  with_attributes no_interrupts (fn atts =>
+    f (fn g => fn y => with_attributes atts (fn _ => g y)) x);
 
 
 (* synchronous wait *)
 
-fun sync_attributes e =
+fun sync_wait opt_atts time cond lock =
+  with_attributes
+    (sync_interrupts (case opt_atts of SOME atts => atts | NONE => Thread.getAttributes ()))
+    (fn _ =>
+      (case time of
+        SOME t => Exn.Result (ConditionVar.waitUntil (cond, lock, t))
+      | NONE => (ConditionVar.wait (cond, lock); Exn.Result true))
+      handle exn => Exn.Exn exn);
+
+
+(* tracing *)
+
+val trace = ref 0;
+
+fun tracing level msg =
+  if level > ! trace then ()
+  else uninterruptible (fn _ => fn () =>
+    (TextIO.output (TextIO.stdErr, (">>> " ^ msg () ^ "\n")); TextIO.flushOut TextIO.stdErr)
+      handle _ (*sic*) => ()) ();
+
+fun tracing_time detailed time =
+  tracing
+   (if not detailed then 5
+    else if Time.>= (time, Time.fromMilliseconds 1000) then 1
+    else if Time.>= (time, Time.fromMilliseconds 100) then 2
+    else if Time.>= (time, Time.fromMilliseconds 10) then 3
+    else if Time.>= (time, Time.fromMilliseconds 1) then 4 else 5);
+
+fun real_time f x =
   let
-    val orig_atts = Thread.getAttributes ();
-    val broadcast =
-      (case List.find (fn Thread.EnableBroadcastInterrupt _ => true | _ => false) orig_atts of
-        NONE => Thread.EnableBroadcastInterrupt false
-      | SOME att => att);
-    val interrupt_state =
-      (case List.find (fn Thread.InterruptState _ => true | _ => false) orig_atts of
-        NONE => Thread.InterruptState Thread.InterruptDefer
-      | SOME (state as Thread.InterruptState Thread.InterruptDefer) => state
-      | _ => Thread.InterruptState Thread.InterruptSynch);
-  in with_attributes [broadcast, interrupt_state] (fn _ => fn () => e ()) () end;
-
-fun sync_wait time cond lock =
-  sync_attributes (fn () =>
-    (case time of
-      SOME t => ConditionVar.waitUntil (cond, lock, t)
-    | NONE => (ConditionVar.wait (cond, lock); true)));
+    val timer = Timer.startRealTimer ();
+    val () = f x;
+    val time = Timer.checkRealTimer timer;
+  in time end;
 
 
 (* execution with time limit *)
@@ -169,7 +163,7 @@
 
 (* system shell processes, with propagation of interrupts *)
 
-fun system_out script = uninterruptible (fn restore_attributes => fn () =>
+fun system_out script = with_attributes no_interrupts (fn orig_atts =>
   let
     val script_name = OS.FileSys.tmpName ();
     val _ = write_file script_name script;
@@ -180,13 +174,12 @@
     (*result state*)
     datatype result = Wait | Signal | Result of int;
     val result = ref Wait;
-    val result_mutex = Mutex.mutex ();
-    val result_cond = ConditionVar.conditionVar ();
+    val lock = Mutex.mutex ();
+    val cond = ConditionVar.conditionVar ();
     fun set_result res =
-      (Mutex.lock result_mutex; result := res; Mutex.unlock result_mutex;
-        ConditionVar.signal result_cond);
+      (Mutex.lock lock; result := res; ConditionVar.signal cond; Mutex.unlock lock);
 
-    val _ = Mutex.lock result_mutex;
+    val _ = Mutex.lock lock;
 
     (*system thread*)
     val system_thread = Thread.fork (fn () =>
@@ -216,11 +209,12 @@
       handle OS.SysErr _ => () | IO.Io _ =>
         (OS.Process.sleep (Time.fromMilliseconds 100); if n > 0 then kill (n - 1) else ());
 
-    val _ = while ! result = Wait do
-      restore_attributes (fn () =>
-        (ignore (sync_wait (SOME (Time.+ (Time.now (), Time.fromMilliseconds 100)))
-            result_cond result_mutex)
-          handle Exn.Interrupt => kill 10)) ();
+    val _ =
+      while ! result = Wait do
+        let val res =
+          sync_wait (SOME orig_atts)
+            (SOME (Time.+ (Time.now (), Time.fromMilliseconds 100))) cond lock
+        in case res of Exn.Exn Exn.Interrupt => kill 10 | _ => () end;
 
     (*cleanup*)
     val output = read_file output_name handle IO.Io _ => "";
@@ -229,7 +223,7 @@
     val _ = OS.FileSys.remove output_name handle OS.SysErr _ => ();
     val _ = Thread.interrupt system_thread handle Thread _ => ();
     val rc = (case ! result of Signal => raise Exn.Interrupt | Result rc => rc);
-  in (output, rc) end) ();
+  in (output, rc) end);
 
 
 (* critical section -- may be nested within the same thread *)
--- a/src/Pure/System/isabelle_system.scala	Mon Aug 10 08:37:37 2009 +0200
+++ b/src/Pure/System/isabelle_system.scala	Mon Aug 10 10:25:00 2009 +0200
@@ -298,6 +298,12 @@
 
   /** Isabelle resources **/
 
+  /* components */
+
+  def components(): List[String] =
+    getenv("ISABELLE_COMPONENTS").split(":").toList
+
+
   /* find logics */
 
   def find_logics(): List[String] =
--- a/src/Pure/subgoal.ML	Mon Aug 10 08:37:37 2009 +0200
+++ b/src/Pure/subgoal.ML	Mon Aug 10 10:25:00 2009 +0200
@@ -102,7 +102,7 @@
 *)
 fun lift_subgoals params asms th =
   let
-    val lift = fold_rev Thm.all_name params o curry Drule.list_implies asms;
+    fun lift ct = fold_rev Thm.all_name params (Drule.list_implies (asms, ct));
     val unlift =
       fold (Thm.elim_implies o Thm.assume) asms o
       Drule.forall_elim_list (map #2 params) o Thm.assume;
@@ -133,14 +133,14 @@
 fun GEN_FOCUS flags tac ctxt i st =
   if Thm.nprems_of st < i then Seq.empty
   else
-    let val (args as {context, params, asms, ...}, st') = gen_focus flags ctxt i st;
-    in Seq.lifts (retrofit context ctxt params asms i) (tac args st') st end;
+    let val (args as {context = ctxt', params, asms, ...}, st') = gen_focus flags ctxt i st;
+    in Seq.lifts (retrofit ctxt' ctxt params asms i) (tac args st') st end;
 
 val FOCUS_PARAMS = GEN_FOCUS (false, false);
 val FOCUS_PREMS = GEN_FOCUS (true, false);
 val FOCUS = GEN_FOCUS (true, true);
 
-fun SUBPROOF tac = FOCUS (FILTER Thm.no_prems o tac);
+fun SUBPROOF tac ctxt = FOCUS (Seq.map (tap (Goal.check_finished ctxt)) oo tac) ctxt;
 
 end;
 
--- a/src/Tools/quickcheck.ML	Mon Aug 10 08:37:37 2009 +0200
+++ b/src/Tools/quickcheck.ML	Mon Aug 10 10:25:00 2009 +0200
@@ -11,6 +11,7 @@
   val test_term: Proof.context -> bool -> string option -> int -> int -> term ->
     (string * term) list option
   val add_generator: string * (Proof.context -> term -> int -> term list option) -> theory -> theory
+  val quickcheck: (string * string) list -> int -> Proof.state -> (string * term) list option
 end;
 
 structure Quickcheck : QUICKCHECK =
@@ -215,18 +216,21 @@
     |> (Data.map o apsnd o map_test_params) f
   end;
 
-fun quickcheck_cmd args i state =
+fun quickcheck args i state =
   let
-    val prf = Toplevel.proof_of state;
-    val thy = Toplevel.theory_of state;
-    val ctxt = Toplevel.context_of state;
+    val thy = Proof.theory_of state;
+    val ctxt = Proof.context_of state;
     val default_params = (dest_test_params o snd o Data.get) thy;
     val f = fold (parse_test_param_inst ctxt) args;
     val (((size, iterations), default_type), (generator_name, insts)) =
       f (default_params, (NONE, []));
-    val counterex = test_goal false generator_name size iterations
-      default_type insts i [] prf;
-  in (Pretty.writeln o pretty_counterex ctxt) counterex end;
+  in
+    test_goal false generator_name size iterations default_type insts i [] state
+  end;
+
+fun quickcheck_cmd args i state =
+  quickcheck args i (Toplevel.proof_of state)
+  |> Pretty.writeln o pretty_counterex (Toplevel.context_of state);
 
 local structure P = OuterParse and K = OuterKeyword in